NEWS
author wenzelm
Sat, 09 Jan 2016 12:35:07 +0100
changeset 62107 f74a33b14200
parent 62098 b1b2834bb493
child 62108 0046bacc5f5b
permissions -rw-r--r--
discontinued \<struct> syntax;

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

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


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

*** General ***

* 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>

* 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'.


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

* IDE support for the source-level debugger of Poly/ML, to work with
Isabelle/ML and official Standard ML. Configuration option "ML_debugger"
and commands 'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug',
'SML_file_no_debug' control compilation of sources with 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.

* 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.

* 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".

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

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


*** 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
(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.

* 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>).

* 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. 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.
  - 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'.

* 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.

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

* Library/Multiset:
  - Renamed multiset inclusion operators:
      < ~> <#
      \<subset> ~> \<subset>#
      <= ~> <=#
      \<le> ~> \<le>#
      \<subseteq> ~> \<subseteq>#
    INCOMPATIBILITY.
  - "'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-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.

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

* 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.



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
INCOMPATIBILITY, explicit instantiation sometimes needs to refer to
different index.

* Lexical separation of signed and unsigned numerals: categories "num"
and "float" are unsigned. INCOMPATIBILITY: subtle change in precedence
of numeral signs, particularly in expressions involving infix syntax
like "(- 1) ^ n".

* Old inner token category "xnum" has been discontinued.  Potential
INCOMPATIBILITY for exotic syntax: may use mixfix grammar with "num"
token category instead.


*** HOL ***

* New (co)datatype package:
  - The 'datatype_new' command has been renamed 'datatype'. The old
    command of that name is now called 'old_datatype' and is provided
    by "~~/src/HOL/Library/Old_Datatype.thy". See
    'isabelle doc datatypes' for information on porting.
    INCOMPATIBILITY.
  - Renamed theorems:
      disc_corec ~> corec_disc
      disc_corec_iff ~> corec_disc_iff
      disc_exclude ~> distinct_disc
      disc_exhaust ~> exhaust_disc
      disc_map_iff ~> map_disc_iff
      sel_corec ~> corec_sel
      sel_exhaust ~> exhaust_sel
      sel_map ~> map_sel
      sel_set ~> set_sel
      sel_split ~> split_sel
      sel_split_asm ~> split_sel_asm
      strong_coinduct ~> coinduct_strong
      weak_case_cong ~> case_cong_weak
    INCOMPATIBILITY.
  - The "no_code" option to "free_constructors", "datatype_new", and
    "codatatype" has been renamed "plugins del: code".
    INCOMPATIBILITY.
  - The rules "set_empty" have been removed. They are easy
    consequences of other set rules "by auto".
    INCOMPATIBILITY.
  - The rule "set_cases" is now registered with the "[cases set]"
    attribute. This can influence the behavior of the "cases" proof
    method when more than one case rule is applicable (e.g., an
    assumption is of the form "w : set ws" and the method "cases w"
    is invoked). The solution is to specify the case rule explicitly
    (e.g. "cases w rule: widget.exhaust").
    INCOMPATIBILITY.
  - Renamed theories:
      BNF_Comp ~> BNF_Composition
      BNF_FP_Base ~> BNF_Fixpoint_Base
      BNF_GFP ~> BNF_Greatest_Fixpoint
      BNF_LFP ~> BNF_Least_Fixpoint
      BNF_Constructions_on_Wellorders ~> BNF_Wellorder_Constructions
      Cardinals/Constructions_on_Wellorders ~> Cardinals/Wellorder_Constructions
    INCOMPATIBILITY.
  - Lifting and Transfer setup for basic HOL types sum and prod (also
    option) is now performed by the BNF package. Theories Lifting_Sum,
    Lifting_Product and Lifting_Option from Main became obsolete and
    were removed. Changed definitions of the relators rel_prod and
    rel_sum (using inductive).
    INCOMPATIBILITY: use rel_prod.simps and rel_sum.simps instead
    of rel_prod_def and rel_sum_def.
    Minor INCOMPATIBILITY: (rarely used by name) transfer theorem names
    changed (e.g. map_prod_transfer ~> prod.map_transfer).
  - Parametricity theorems for map functions, relators, set functions,
    constructors, case combinators, discriminators, selectors and
    (co)recursors are automatically proved and registered as transfer
    rules.

* Old datatype package:
  - The old 'datatype' command has been renamed 'old_datatype', and
    'rep_datatype' has been renamed 'old_rep_datatype'. They are
    provided by "~~/src/HOL/Library/Old_Datatype.thy". See
    'isabelle doc datatypes' for information on porting.
    INCOMPATIBILITY.
  - Renamed theorems:
      weak_case_cong ~> case_cong_weak
    INCOMPATIBILITY.
  - Renamed theory:
      ~~/src/HOL/Datatype.thy ~> ~~/src/HOL/Library/Old_Datatype.thy
    INCOMPATIBILITY.

* Nitpick:
  - Fixed soundness bug related to the strict and non-strict subset
    operations.

* Sledgehammer:
  - CVC4 is now included with Isabelle instead of CVC3 and run by
    default.
  - Z3 is now always enabled by default, now that it is fully open
    source. The "z3_non_commercial" option is discontinued.
  - Minimization is now always enabled by default.
    Removed sub-command:
      min
  - Proof reconstruction, both one-liners and Isar, has been
    dramatically improved.
  - Improved support for CVC4 and veriT.

* Old and new SMT modules:
  - The old 'smt' method has been renamed 'old_smt' and moved to
    'src/HOL/Library/Old_SMT.thy'. It is provided for compatibility,
    until applications have been ported to use the new 'smt' method. For
    the method to work, an older version of Z3 (e.g. Z3 3.2 or 4.0) must
    be installed, and the environment variable "OLD_Z3_SOLVER" must
    point to it.
    INCOMPATIBILITY.
  - The 'smt2' method has been renamed 'smt'.
    INCOMPATIBILITY.
  - New option 'smt_reconstruction_step_timeout' to limit the
    reconstruction time of Z3 proof steps in the new 'smt' method.
  - New option 'smt_statistics' to display statistics of the new 'smt'
    method, especially runtime statistics of Z3 proof reconstruction.

* Lifting: command 'lift_definition' allows to execute lifted constants
that have as a return type a datatype containing a subtype. This
overcomes long-time limitations in the area of code generation and
lifting, and avoids tedious workarounds.

* Command and antiquotation "value" provide different evaluation slots
(again), where the previous strategy (NBE after ML) serves as default.
Minor INCOMPATIBILITY.

* Add NO_MATCH-simproc, allows to check for syntactic non-equality.

* field_simps: Use NO_MATCH-simproc for distribution rules, to avoid
non-termination in case of distributing a division. With this change
field_simps is in some cases slightly less powerful, if it fails try to
add algebra_simps, or use divide_simps. Minor INCOMPATIBILITY.

* Separate class no_zero_divisors has been given up in favour of fully
algebraic semiring_no_zero_divisors. INCOMPATIBILITY.

* Class linordered_semidom really requires no zero divisors.
INCOMPATIBILITY.

* Classes division_ring, field and linordered_field always demand
"inverse 0 = 0". Given up separate classes division_ring_inverse_zero,
field_inverse_zero and linordered_field_inverse_zero. INCOMPATIBILITY.

* Classes cancel_ab_semigroup_add / cancel_monoid_add specify explicit
additive inverse operation. INCOMPATIBILITY.

* Complex powers and square roots. The functions "ln" and "powr" are now
overloaded for types real and complex, and 0 powr y = 0 by definition.
INCOMPATIBILITY: type constraints may be necessary.

* The functions "sin" and "cos" are now defined for any type of sort
"{real_normed_algebra_1,banach}" type, so in particular on "real" and
"complex" uniformly. Minor INCOMPATIBILITY: type constraints may be
needed.

* New library of properties of the complex transcendental functions sin,
cos, tan, exp, Ln, Arctan, Arcsin, Arccos. Ported from HOL Light.

* The factorial function, "fact", now has type "nat => 'a" (of a sort
that admits numeric types including nat, int, real and complex.
INCOMPATIBILITY: an expression such as "fact 3 = 6" may require a type
constraint, and the combination "real (fact k)" is likely to be
unsatisfactory. If a type conversion is still necessary, then use
"of_nat (fact k)" or "real_of_nat (fact k)".

* Removed functions "natfloor" and "natceiling", use "nat o floor" and
"nat o ceiling" instead. A few of the lemmas have been retained and
adapted: in their names "natfloor"/"natceiling" has been replaced by
"nat_floor"/"nat_ceiling".

* Qualified some duplicated fact names required for boostrapping the
type class hierarchy:
  ab_add_uminus_conv_diff ~> diff_conv_add_uminus
  field_inverse_zero ~> inverse_zero
  field_divide_inverse ~> divide_inverse
  field_inverse ~> left_inverse
Minor INCOMPATIBILITY.

* Eliminated fact duplicates:
  mult_less_imp_less_right ~> mult_right_less_imp_less
  mult_less_imp_less_left ~> mult_left_less_imp_less
Minor INCOMPATIBILITY.

* Fact consolidation: even_less_0_iff is subsumed by
double_add_less_zero_iff_single_add_less_zero (simp by default anyway).

* Generalized and consolidated some theorems concerning divsibility:
  dvd_reduce ~> dvd_add_triv_right_iff
  dvd_plus_eq_right ~> dvd_add_right_iff
  dvd_plus_eq_left ~> dvd_add_left_iff
Minor INCOMPATIBILITY.

* "even" and "odd" are mere abbreviations for "2 dvd _" and "~ 2 dvd _"
and part of theory Main.
  even_def ~> even_iff_mod_2_eq_zero
INCOMPATIBILITY.

* Lemma name consolidation: divide_Numeral1 ~> divide_numeral_1. Minor
INCOMPATIBILITY.

* Bootstrap of listsum as special case of abstract product over lists.
Fact rename:
    listsum_def ~> listsum.eq_foldr
INCOMPATIBILITY.

* Product over lists via constant "listprod".

* Theory List: renamed drop_Suc_conv_tl and nth_drop' to
Cons_nth_drop_Suc.

* New infrastructure for compiling, running, evaluating and testing
generated code in target languages in HOL/Library/Code_Test. See
HOL/Codegenerator_Test/Code_Test* for examples.

* Library/Multiset:
  - Introduced "replicate_mset" operation.
  - Introduced alternative characterizations of the multiset ordering in
    "Library/Multiset_Order".
  - Renamed multiset ordering:
      <# ~> #<#
      <=# ~> #<=#
      \<subset># ~> #\<subset>#
      \<subseteq># ~> #\<subseteq>#
    INCOMPATIBILITY.
  - Introduced abbreviations for ill-named multiset operations:
      <#, \<subset># abbreviate < (strict subset)
      <=#, \<le>#, \<subseteq># abbreviate <= (subset or equal)
    INCOMPATIBILITY.
  - Renamed
      in_multiset_of ~> in_multiset_in_set
      Multiset.fold ~> fold_mset
      Multiset.filter ~> filter_mset
    INCOMPATIBILITY.
  - Removed mcard, is equal to size.
  - Added attributes:
      image_mset.id [simp]
      image_mset_id [simp]
      elem_multiset_of_set [simp, intro]
      comp_fun_commute_plus_mset [simp]
      comp_fun_commute.fold_mset_insert [OF comp_fun_commute_plus_mset, simp]
      in_mset_fold_plus_iff [iff]
      set_of_Union_mset [simp]
      in_Union_mset_iff [iff]
    INCOMPATIBILITY.

* Library/Sum_of_Squares: simplified and improved "sos" method. Always
use local CSDP executable, which is much faster than the NEOS server.
The "sos_cert" functionality is invoked as "sos" with additional
argument. Minor INCOMPATIBILITY.

* HOL-Decision_Procs: New counterexample generator quickcheck
[approximation] for inequalities of transcendental functions. Uses
hardware floating point arithmetic to randomly discover potential
counterexamples. Counterexamples are certified with the "approximation"
method. See HOL/Decision_Procs/ex/Approximation_Quickcheck_Ex.thy for
examples.

* HOL-Probability: Reworked measurability prover
  - applies destructor rules repeatedly
  - removed application splitting (replaced by destructor rule)
  - added congruence rules to rewrite measure spaces under the sets
    projection

* New proof method "rewrite" (in theory ~~/src/HOL/Library/Rewrite) for
single-step rewriting with subterm selection based on patterns.


*** ML ***

* Subtle change of name space policy: undeclared entries are now
considered inaccessible, instead of accessible via the fully-qualified
internal name. This mainly affects Name_Space.intern (and derivatives),
which may produce an unexpected Long_Name.hidden prefix. Note that
contemporary applications use the strict Name_Space.check (and
derivatives) instead, which is not affected by the change. Potential
INCOMPATIBILITY in rare applications of Name_Space.intern.

* Subtle change of error semantics of Toplevel.proof_of: regular user
ERROR instead of internal Toplevel.UNDEF.

* Basic combinators map, fold, fold_map, split_list, apply are available
as parameterized antiquotations, e.g. @{map 4} for lists of quadruples.

* Renamed "pairself" to "apply2", in accordance to @{apply 2}.
INCOMPATIBILITY.

* Former combinators NAMED_CRITICAL and CRITICAL for central critical
sections have been discontinued, in favour of the more elementary
Multithreading.synchronized and its high-level derivative
Synchronized.var (which is usually sufficient in applications). Subtle
INCOMPATIBILITY: synchronized access needs to be atomic and cannot be
nested.

* Synchronized.value (ML) is actually synchronized (as in Scala): subtle
change of semantics with minimal potential for INCOMPATIBILITY.

* The main operations to certify logical entities are Thm.ctyp_of and
Thm.cterm_of with a local context; old-style global theory variants are
available as Thm.global_ctyp_of and Thm.global_cterm_of.
INCOMPATIBILITY.

* Elementary operations in module Thm are no longer pervasive.
INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of,
Thm.term_of etc.

* Proper context for various elementary tactics: assume_tac,
resolve_tac, eresolve_tac, dresolve_tac, forward_tac, match_tac,
compose_tac, Splitter.split_tac etc. INCOMPATIBILITY.

* Tactical PARALLEL_ALLGOALS is the most common way to refer to
PARALLEL_GOALS.

* Goal.prove_multi is superseded by the fully general Goal.prove_common,
which also allows to specify a fork priority.

* Antiquotation @{command_spec "COMMAND"} is superseded by
@{command_keyword COMMAND} (usually without quotes and with PIDE
markup). Minor INCOMPATIBILITY.

* Cartouches within ML sources are turned into values of type
Input.source (with formal position information).


*** System ***

* The Isabelle tool "update_cartouches" changes theory files to use
cartouches instead of old-style {* verbatim *} or `alt_string` tokens.

* The Isabelle tool "build" provides new options -X, -k, -x.

* Discontinued old-fashioned "codegen" tool. Code generation can always
be externally triggered using an appropriate ROOT file plus a
corresponding theory. Parametrization is possible using environment
variables, or ML snippets in the most extreme cases. Minor
INCOMPATIBILITY.

* JVM system property "isabelle.threads" determines size of Scala thread
pool, like Isabelle system option "threads" for ML.

* JVM system property "isabelle.laf" determines the default Swing
look-and-feel, via internal class name or symbolic name as in the jEdit
menu Global Options / Appearance.

* Support for Proof General and Isar TTY loop has been discontinued.
Minor INCOMPATIBILITY, use standard PIDE infrastructure instead.



New in Isabelle2014 (August 2014)
---------------------------------

*** General ***

* Support for official Standard ML within the Isabelle context.
Command 'SML_file' reads and evaluates the given Standard ML file.
Toplevel bindings are stored within the theory context; the initial
environment is restricted to the Standard ML implementation of
Poly/ML, without the add-ons of Isabelle/ML.  Commands 'SML_import'
and 'SML_export' allow to exchange toplevel bindings between the two
separate environments.  See also ~~/src/Tools/SML/Examples.thy for
some examples.

* Standard tactics and proof methods such as "clarsimp", "auto" and
"safe" now preserve equality hypotheses "x = expr" where x is a free
variable.  Locale assumptions and chained facts containing "x"
continue to be useful.  The new method "hypsubst_thin" and the
configuration option "hypsubst_thin" (within the attribute name space)
restore the previous behavior.  INCOMPATIBILITY, especially where
induction is done after these methods or when the names of free and
bound variables clash.  As first approximation, old proofs may be
repaired by "using [[hypsubst_thin = true]]" in the critical spot.

* More static checking of proof methods, which allows the system to
form a closure over the concrete syntax.  Method arguments should be
processed in the original proof context as far as possible, before
operating on the goal state.  In any case, the standard discipline for
subgoal-addressing needs to be observed: no subgoals or a subgoal
number that is out of range produces an empty result sequence, not an
exception.  Potential INCOMPATIBILITY for non-conformant tactical
proof tools.

* Lexical syntax (inner and outer) supports text cartouches with
arbitrary nesting, and without escapes of quotes etc.  The Prover IDE
supports input via ` (backquote).

* The outer syntax categories "text" (for formal comments and document
markup commands) and "altstring" (for literal fact references) allow
cartouches as well, in addition to the traditional mix of quotations.

* Syntax of document antiquotation @{rail} now uses \<newline> instead
of "\\", to avoid the optical illusion of escaped backslash within
string token.  General renovation of its syntax using text cartouches.
Minor INCOMPATIBILITY.

* Discontinued legacy_isub_isup, which was a temporary workaround for
Isabelle/ML in Isabelle2013-1.  The prover process no longer accepts
old identifier syntax with \<^isub> or \<^isup>.  Potential
INCOMPATIBILITY.

* Document antiquotation @{url} produces markup for the given URL,
which results in an active hyperlink within the text.

* Document antiquotation @{file_unchecked} is like @{file}, but does
not check existence within the file-system.

* Updated and extended manuals: codegen, datatypes, implementation,
isar-ref, jedit, system.


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

* Improved Document panel: simplified interaction where every single
mouse click (re)opens document via desktop environment or as jEdit
buffer.

* Support for Navigator plugin (with toolbar buttons), with connection
to PIDE hyperlinks.

* Auxiliary files ('ML_file' etc.) are managed by the Prover IDE.
Open text buffers take precedence over copies within the file-system.

* Improved support for Isabelle/ML, with jEdit mode "isabelle-ml" for
auxiliary ML files.

* Improved syntactic and semantic completion mechanism, with simple
templates, completion language context, name-space completion,
file-name completion, spell-checker completion.

* Refined GUI popup for completion: more robust key/mouse event
handling and propagation to enclosing text area -- avoid loosing
keystrokes with slow / remote graphics displays.

* Completion popup supports both ENTER and TAB (default) to select an
item, depending on Isabelle options.

* Refined insertion of completion items wrt. jEdit text: multiple
selections, rectangular selections, rectangular selection as "tall
caret".

* Integrated spell-checker for document text, comments etc. with
completion popup and context-menu.

* More general "Query" panel supersedes "Find" panel, with GUI access
to commands 'find_theorems' and 'find_consts', as well as print
operations for the context.  Minor incompatibility in keyboard
shortcuts etc.: replace action isabelle-find by isabelle-query.

* Search field for all output panels ("Output", "Query", "Info" etc.)
to highlight text via regular expression.

* Option "jedit_print_mode" (see also "Plugin Options / Isabelle /
General") allows to specify additional print modes for the prover
process, without requiring old-fashioned command-line invocation of
"isabelle jedit -m MODE".

* More support for remote files (e.g. http) using standard Java
networking operations instead of jEdit virtual file-systems.

* Empty editors buffers that are no longer required (e.g.\ via theory
imports) are automatically removed from the document model.

* Improved monitor panel.

* Improved Console/Scala plugin: more uniform scala.Console output,
more robust treatment of threads and interrupts.

* Improved management of dockable windows: clarified keyboard focus
and window placement wrt. main editor view; optional menu item to
"Detach" a copy where this makes sense.

* New Simplifier Trace panel provides an interactive view of the
simplification process, enabled by the "simp_trace_new" attribute
within the context.


*** Pure ***

* Low-level type-class commands 'classes', 'classrel', 'arities' have
been discontinued to avoid the danger of non-trivial axiomatization
that is not immediately visible.  INCOMPATIBILITY, use regular
'instance' command with proof.  The required OFCLASS(...) theorem
might be postulated via 'axiomatization' beforehand, or the proof
finished trivially if the underlying class definition is made vacuous
(without any assumptions).  See also Isabelle/ML operations
Axclass.class_axiomatization, Axclass.classrel_axiomatization,
Axclass.arity_axiomatization.

* Basic constants of Pure use more conventional names and are always
qualified.  Rare INCOMPATIBILITY, but with potentially serious
consequences, notably for tools in Isabelle/ML.  The following
renaming needs to be applied:

  ==             ~>  Pure.eq
  ==>            ~>  Pure.imp
  all            ~>  Pure.all
  TYPE           ~>  Pure.type
  dummy_pattern  ~>  Pure.dummy_pattern

Systematic porting works by using the following theory setup on a
*previous* Isabelle version to introduce the new name accesses for the
old constants:

setup {*
  fn thy => thy
    |> Sign.root_path
    |> Sign.const_alias (Binding.qualify true "Pure" @{binding eq}) "=="
    |> Sign.const_alias (Binding.qualify true "Pure" @{binding imp}) "==>"
    |> Sign.const_alias (Binding.qualify true "Pure" @{binding all}) "all"
    |> Sign.restore_naming thy
*}

Thus ML antiquotations like @{const_name Pure.eq} may be used already.
Later the application is moved to the current Isabelle version, and
the auxiliary aliases are deleted.

* Attributes "where" and "of" allow an optional context of local
variables ('for' declaration): these variables become schematic in the
instantiated theorem.

* Obsolete attribute "standard" has been discontinued (legacy since
Isabelle2012).  Potential INCOMPATIBILITY, use explicit 'for' context
where instantiations with schematic variables are intended (for
declaration commands like 'lemmas' or attributes like "of").  The
following temporary definition may help to port old applications:

  attribute_setup standard =
    "Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))"

* More thorough check of proof context for goal statements and
attributed fact expressions (concerning background theory, declared
hyps).  Potential INCOMPATIBILITY, tools need to observe standard
context discipline.  See also Assumption.add_assumes and the more
primitive Thm.assume_hyps.

* Inner syntax token language allows regular quoted strings "..."
(only makes sense in practice, if outer syntax is delimited
differently, e.g. via cartouches).

* Command 'print_term_bindings' supersedes 'print_binds' for clarity,
but the latter is retained some time as Proof General legacy.

* Code generator preprocessor: explicit control of simp tracing on a
per-constant basis.  See attribute "code_preproc".


*** HOL ***

* Code generator: enforce case of identifiers only for strict target
language requirements.  INCOMPATIBILITY.

* Code generator: explicit proof contexts in many ML interfaces.
INCOMPATIBILITY.

* Code generator: minimize exported identifiers by default.  Minor
INCOMPATIBILITY.

* Code generation for SML and OCaml: dropped arcane "no_signatures"
option.  Minor INCOMPATIBILITY.

* "declare [[code abort: ...]]" replaces "code_abort ...".
INCOMPATIBILITY.

* "declare [[code drop: ...]]" drops all code equations associated
with the given constants.

* Code generations are provided for make, fields, extend and truncate
operations on records.

* Command and antiquotation "value" are now hardcoded against nbe and
ML.  Minor INCOMPATIBILITY.

* Renamed command 'enriched_type' to 'functor'. INCOMPATIBILITY.

* The symbol "\<newline>" may be used within char or string literals
to represent (Char Nibble0 NibbleA), i.e. ASCII newline.

* Qualified String.implode and String.explode.  INCOMPATIBILITY.

* Simplifier: Enhanced solver of preconditions of rewrite rules can
now deal with conjunctions.  For help with converting proofs, the old
behaviour of the simplifier can be restored like this: declare/using
[[simp_legacy_precond]].  This configuration option will disappear
again in the future.  INCOMPATIBILITY.

* Simproc "finite_Collect" is no longer enabled by default, due to
spurious crashes and other surprises.  Potential INCOMPATIBILITY.

* Moved new (co)datatype package and its dependencies from session
  "HOL-BNF" to "HOL".  The commands 'bnf', 'wrap_free_constructors',
  'datatype_new', 'codatatype', 'primcorec', 'primcorecursive' are now
  part of theory "Main".

  Theory renamings:
    FunDef.thy ~> Fun_Def.thy (and Fun_Def_Base.thy)
    Library/Wfrec.thy ~> Wfrec.thy
    Library/Zorn.thy ~> Zorn.thy
    Cardinals/Order_Relation.thy ~> Order_Relation.thy
    Library/Order_Union.thy ~> Cardinals/Order_Union.thy
    Cardinals/Cardinal_Arithmetic_Base.thy ~> BNF_Cardinal_Arithmetic.thy
    Cardinals/Cardinal_Order_Relation_Base.thy ~> BNF_Cardinal_Order_Relation.thy
    Cardinals/Constructions_on_Wellorders_Base.thy ~> BNF_Constructions_on_Wellorders.thy
    Cardinals/Wellorder_Embedding_Base.thy ~> BNF_Wellorder_Embedding.thy
    Cardinals/Wellorder_Relation_Base.thy ~> BNF_Wellorder_Relation.thy
    BNF/Ctr_Sugar.thy ~> Ctr_Sugar.thy
    BNF/Basic_BNFs.thy ~> Basic_BNFs.thy
    BNF/BNF_Comp.thy ~> BNF_Comp.thy
    BNF/BNF_Def.thy ~> BNF_Def.thy
    BNF/BNF_FP_Base.thy ~> BNF_FP_Base.thy
    BNF/BNF_GFP.thy ~> BNF_GFP.thy
    BNF/BNF_LFP.thy ~> BNF_LFP.thy
    BNF/BNF_Util.thy ~> BNF_Util.thy
    BNF/Coinduction.thy ~> Coinduction.thy
    BNF/More_BNFs.thy ~> Library/More_BNFs.thy
    BNF/Countable_Type.thy ~> Library/Countable_Set_Type.thy
    BNF/Examples/* ~> BNF_Examples/*

  New theories:
    Wellorder_Extension.thy (split from Zorn.thy)
    Library/Cardinal_Notations.thy
    Library/BNF_Axomatization.thy
    BNF_Examples/Misc_Primcorec.thy
    BNF_Examples/Stream_Processor.thy

  Discontinued theories:
    BNF/BNF.thy
    BNF/Equiv_Relations_More.thy

INCOMPATIBILITY.

* New (co)datatype package:
  - Command 'primcorec' is fully implemented.
  - Command 'datatype_new' generates size functions ("size_xxx" and
    "size") as required by 'fun'.
  - BNFs are integrated with the Lifting tool and new-style
    (co)datatypes with Transfer.
  - Renamed commands:
      datatype_new_compat ~> datatype_compat
      primrec_new ~> primrec
      wrap_free_constructors ~> free_constructors
    INCOMPATIBILITY.
  - The generated constants "xxx_case" and "xxx_rec" have been renamed
    "case_xxx" and "rec_xxx" (e.g., "prod_case" ~> "case_prod").
    INCOMPATIBILITY.
  - The constant "xxx_(un)fold" and related theorems are no longer
    generated.  Use "xxx_(co)rec" or define "xxx_(un)fold" manually
    using "prim(co)rec".
    INCOMPATIBILITY.
  - No discriminators are generated for nullary constructors by
    default, eliminating the need for the odd "=:" syntax.
    INCOMPATIBILITY.
  - No discriminators or selectors are generated by default by
    "datatype_new", unless custom names are specified or the new
    "discs_sels" option is passed.
    INCOMPATIBILITY.

* Old datatype package:
  - The generated theorems "xxx.cases" and "xxx.recs" have been
    renamed "xxx.case" and "xxx.rec" (e.g., "sum.cases" ->
    "sum.case").  INCOMPATIBILITY.
  - The generated constants "xxx_case", "xxx_rec", and "xxx_size" have
    been renamed "case_xxx", "rec_xxx", and "size_xxx" (e.g.,
    "prod_case" ~> "case_prod").  INCOMPATIBILITY.

* The types "'a list" and "'a option", their set and map functions,
  their relators, and their selectors are now produced using the new
  BNF-based datatype package.

  Renamed constants:
    Option.set ~> set_option
    Option.map ~> map_option
    option_rel ~> rel_option

  Renamed theorems:
    set_def ~> set_rec[abs_def]
    map_def ~> map_rec[abs_def]
    Option.map_def ~> map_option_case[abs_def] (with "case_option" instead of "rec_option")
    option.recs ~> option.rec
    list_all2_def ~> list_all2_iff
    set.simps ~> set_simps (or the slightly different "list.set")
    map.simps ~> list.map
    hd.simps ~> list.sel(1)
    tl.simps ~> list.sel(2-3)
    the.simps ~> option.sel

INCOMPATIBILITY.

* The following map functions and relators have been renamed:
    sum_map ~> map_sum
    map_pair ~> map_prod
    prod_rel ~> rel_prod
    sum_rel ~> rel_sum
    fun_rel ~> rel_fun
    set_rel ~> rel_set
    filter_rel ~> rel_filter
    fset_rel ~> rel_fset (in "src/HOL/Library/FSet.thy")
    cset_rel ~> rel_cset (in "src/HOL/Library/Countable_Set_Type.thy")
    vset ~> rel_vset (in "src/HOL/Library/Quotient_Set.thy")

INCOMPATIBILITY.

* Lifting and Transfer:
  - a type variable as a raw type is supported
  - stronger reflexivity prover
  - rep_eq is always generated by lift_definition
  - setup for Lifting/Transfer is now automated for BNFs
    + holds for BNFs that do not contain a dead variable
    + relator_eq, relator_mono, relator_distr, relator_domain,
      relator_eq_onp, quot_map, transfer rules for bi_unique, bi_total,
      right_unique, right_total, left_unique, left_total are proved
      automatically
    + definition of a predicator is generated automatically
    + simplification rules for a predicator definition are proved
      automatically for datatypes
  - consolidation of the setup of Lifting/Transfer
    + property that a relator preservers reflexivity is not needed any
      more
      Minor INCOMPATIBILITY.
    + left_total and left_unique rules are now transfer rules
      (reflexivity_rule attribute not needed anymore)
      INCOMPATIBILITY.
    + Domainp does not have to be a separate assumption in
      relator_domain theorems (=> more natural statement)
      INCOMPATIBILITY.
  - registration of code equations is more robust
    Potential INCOMPATIBILITY.
  - respectfulness proof obligation is preprocessed to a more readable
    form
    Potential INCOMPATIBILITY.
  - eq_onp is always unfolded in respectfulness proof obligation
    Potential INCOMPATIBILITY.
  - unregister lifting setup for Code_Numeral.integer and
    Code_Numeral.natural
    Potential INCOMPATIBILITY.
  - Lifting.invariant -> eq_onp
    INCOMPATIBILITY.

* New internal SAT solver "cdclite" that produces models and proof
traces.  This solver replaces the internal SAT solvers "enumerate" and
"dpll".  Applications that explicitly used one of these two SAT
solvers should use "cdclite" instead. In addition, "cdclite" is now
the default SAT solver for the "sat" and "satx" proof methods and
corresponding tactics; the old default can be restored using "declare
[[sat_solver = zchaff_with_proofs]]".  Minor INCOMPATIBILITY.

* SMT module: A new version of the SMT module, temporarily called
"SMT2", uses SMT-LIB 2 and supports recent versions of Z3 (e.g.,
4.3). The new proof method is called "smt2". CVC3 and CVC4 are also
supported as oracles. Yices is no longer supported, because no version
of the solver can handle both SMT-LIB 2 and quantifiers.

* Activation of Z3 now works via "z3_non_commercial" system option
(without requiring restart), instead of former settings variable
"Z3_NON_COMMERCIAL".  The option can be edited in Isabelle/jEdit menu
Plugin Options / Isabelle / General.

* Sledgehammer:
  - Z3 can now produce Isar proofs.
  - MaSh overhaul:
    . New SML-based learning algorithms eliminate the dependency on
      Python and increase performance and reliability.
    . MaSh and MeSh are now used by default together with the
      traditional MePo (Meng-Paulson) relevance filter. To disable
      MaSh, set the "MaSh" system option in Isabelle/jEdit Plugin
      Options / Isabelle / General to "none".
  - New option:
      smt_proofs
  - Renamed options:
      isar_compress ~> compress
      isar_try0 ~> try0

INCOMPATIBILITY.

* Removed solvers remote_cvc3 and remote_z3. Use cvc3 and z3 instead.

* Nitpick:
  - Fixed soundness bug whereby mutually recursive datatypes could
    take infinite values.
  - Fixed soundness bug with low-level number functions such as
    "Abs_Integ" and "Rep_Integ".
  - Removed "std" option.
  - Renamed "show_datatypes" to "show_types" and "hide_datatypes" to
    "hide_types".

* Metis: Removed legacy proof method 'metisFT'. Use 'metis
(full_types)' instead. INCOMPATIBILITY.

* Try0: Added 'algebra' and 'meson' to the set of proof methods.

* Adjustion of INF and SUP operations:
  - Elongated constants INFI and SUPR to INFIMUM and SUPREMUM.
  - Consolidated theorem names containing INFI and SUPR: have INF and
    SUP instead uniformly.
  - More aggressive normalization of expressions involving INF and Inf
    or SUP and Sup.
  - INF_image and SUP_image do not unfold composition.
  - Dropped facts INF_comp, SUP_comp.
  - Default congruence rules strong_INF_cong and strong_SUP_cong, with
    simplifier implication in premises.  Generalize and replace former
    INT_cong, SUP_cong

INCOMPATIBILITY.

* SUP and INF generalized to conditionally_complete_lattice.

* Swapped orientation of facts image_comp and vimage_comp:

  image_compose ~> image_comp [symmetric]
  image_comp ~> image_comp [symmetric]
  vimage_compose ~> vimage_comp [symmetric]
  vimage_comp ~> vimage_comp [symmetric]

INCOMPATIBILITY.

* Theory reorganization: split of Big_Operators.thy into
Groups_Big.thy and Lattices_Big.thy.

* Consolidated some facts about big group operators:

    setsum_0' ~> setsum.neutral
    setsum_0 ~> setsum.neutral_const
    setsum_addf ~> setsum.distrib
    setsum_cartesian_product ~> setsum.cartesian_product
    setsum_cases ~> setsum.If_cases
    setsum_commute ~> setsum.commute
    setsum_cong ~> setsum.cong
    setsum_delta ~> setsum.delta
    setsum_delta' ~> setsum.delta'
    setsum_diff1' ~> setsum.remove
    setsum_empty ~> setsum.empty
    setsum_infinite ~> setsum.infinite
    setsum_insert ~> setsum.insert
    setsum_inter_restrict'' ~> setsum.inter_filter
    setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left
    setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right
    setsum_mono_zero_left ~> setsum.mono_neutral_left
    setsum_mono_zero_right ~> setsum.mono_neutral_right
    setsum_reindex ~> setsum.reindex
    setsum_reindex_cong ~> setsum.reindex_cong
    setsum_reindex_nonzero ~> setsum.reindex_nontrivial
    setsum_restrict_set ~> setsum.inter_restrict
    setsum_Plus ~> setsum.Plus
    setsum_setsum_restrict ~> setsum.commute_restrict
    setsum_Sigma ~> setsum.Sigma
    setsum_subset_diff ~> setsum.subset_diff
    setsum_Un_disjoint ~> setsum.union_disjoint
    setsum_UN_disjoint ~> setsum.UNION_disjoint
    setsum_Un_Int ~> setsum.union_inter
    setsum_Union_disjoint ~> setsum.Union_disjoint
    setsum_UNION_zero ~> setsum.Union_comp
    setsum_Un_zero ~> setsum.union_inter_neutral
    strong_setprod_cong ~> setprod.strong_cong
    strong_setsum_cong ~> setsum.strong_cong
    setprod_1' ~> setprod.neutral
    setprod_1 ~> setprod.neutral_const
    setprod_cartesian_product ~> setprod.cartesian_product
    setprod_cong ~> setprod.cong
    setprod_delta ~> setprod.delta
    setprod_delta' ~> setprod.delta'
    setprod_empty ~> setprod.empty
    setprod_infinite ~> setprod.infinite
    setprod_insert ~> setprod.insert
    setprod_mono_one_cong_left ~> setprod.mono_neutral_cong_left
    setprod_mono_one_cong_right ~> setprod.mono_neutral_cong_right
    setprod_mono_one_left ~> setprod.mono_neutral_left
    setprod_mono_one_right ~> setprod.mono_neutral_right
    setprod_reindex ~> setprod.reindex
    setprod_reindex_cong ~> setprod.reindex_cong
    setprod_reindex_nonzero ~> setprod.reindex_nontrivial
    setprod_Sigma ~> setprod.Sigma
    setprod_subset_diff ~> setprod.subset_diff
    setprod_timesf ~> setprod.distrib
    setprod_Un2 ~> setprod.union_diff2
    setprod_Un_disjoint ~> setprod.union_disjoint
    setprod_UN_disjoint ~> setprod.UNION_disjoint
    setprod_Un_Int ~> setprod.union_inter
    setprod_Union_disjoint ~> setprod.Union_disjoint
    setprod_Un_one ~> setprod.union_inter_neutral

  Dropped setsum_cong2 (simple variant of setsum.cong).
  Dropped setsum_inter_restrict' (simple variant of setsum.inter_restrict)
  Dropped setsum_reindex_id, setprod_reindex_id
    (simple variants of setsum.reindex [symmetric], setprod.reindex [symmetric]).

INCOMPATIBILITY.

* Abolished slightly odd global lattice interpretation for min/max.

  Fact consolidations:
    min_max.inf_assoc ~> min.assoc
    min_max.inf_commute ~> min.commute
    min_max.inf_left_commute ~> min.left_commute
    min_max.inf_idem ~> min.idem
    min_max.inf_left_idem ~> min.left_idem
    min_max.inf_right_idem ~> min.right_idem
    min_max.sup_assoc ~> max.assoc
    min_max.sup_commute ~> max.commute
    min_max.sup_left_commute ~> max.left_commute
    min_max.sup_idem ~> max.idem
    min_max.sup_left_idem ~> max.left_idem
    min_max.sup_inf_distrib1 ~> max_min_distrib2
    min_max.sup_inf_distrib2 ~> max_min_distrib1
    min_max.inf_sup_distrib1 ~> min_max_distrib2
    min_max.inf_sup_distrib2 ~> min_max_distrib1
    min_max.distrib ~> min_max_distribs
    min_max.inf_absorb1 ~> min.absorb1
    min_max.inf_absorb2 ~> min.absorb2
    min_max.sup_absorb1 ~> max.absorb1
    min_max.sup_absorb2 ~> max.absorb2
    min_max.le_iff_inf ~> min.absorb_iff1
    min_max.le_iff_sup ~> max.absorb_iff2
    min_max.inf_le1 ~> min.cobounded1
    min_max.inf_le2 ~> min.cobounded2
    le_maxI1, min_max.sup_ge1 ~> max.cobounded1
    le_maxI2, min_max.sup_ge2 ~> max.cobounded2
    min_max.le_infI1 ~> min.coboundedI1
    min_max.le_infI2 ~> min.coboundedI2
    min_max.le_supI1 ~> max.coboundedI1
    min_max.le_supI2 ~> max.coboundedI2
    min_max.less_infI1 ~> min.strict_coboundedI1
    min_max.less_infI2 ~> min.strict_coboundedI2
    min_max.less_supI1 ~> max.strict_coboundedI1
    min_max.less_supI2 ~> max.strict_coboundedI2
    min_max.inf_mono ~> min.mono
    min_max.sup_mono ~> max.mono
    min_max.le_infI, min_max.inf_greatest ~> min.boundedI
    min_max.le_supI, min_max.sup_least ~> max.boundedI
    min_max.le_inf_iff ~> min.bounded_iff
    min_max.le_sup_iff ~> max.bounded_iff

For min_max.inf_sup_aci, prefer (one of) min.commute, min.assoc,
min.left_commute, min.left_idem, max.commute, max.assoc,
max.left_commute, max.left_idem directly.

For min_max.inf_sup_ord, prefer (one of) min.cobounded1,
min.cobounded2, max.cobounded1m max.cobounded2 directly.

For min_ac or max_ac, prefer more general collection ac_simps.

INCOMPATIBILITY.

* Theorem disambiguation Inf_le_Sup (on finite sets) ~>
Inf_fin_le_Sup_fin.  INCOMPATIBILITY.

* Qualified constant names Wellfounded.acc, Wellfounded.accp.
INCOMPATIBILITY.

* Fact generalization and consolidation:
    neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1

INCOMPATIBILITY.

* Purely algebraic definition of even.  Fact generalization and
  consolidation:
    nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd
    even_zero_(nat|int) ~> even_zero

INCOMPATIBILITY.

* Abolished neg_numeral.
  - Canonical representation for minus one is "- 1".
  - Canonical representation for other negative numbers is "- (numeral _)".
  - When devising rule sets for number calculation, consider the
    following canonical cases: 0, 1, numeral _, - 1, - numeral _.
  - HOLogic.dest_number also recognizes numerals in non-canonical forms
    like "numeral One", "- numeral One", "- 0" and even "- ... - _".
  - Syntax for negative numerals is mere input syntax.

INCOMPATIBILITY.

* Reduced name variants for rules on associativity and commutativity:

    add_assoc ~> add.assoc
    add_commute ~> add.commute
    add_left_commute ~> add.left_commute
    mult_assoc ~> mult.assoc
    mult_commute ~> mult.commute
    mult_left_commute ~> mult.left_commute
    nat_add_assoc ~> add.assoc
    nat_add_commute ~> add.commute
    nat_add_left_commute ~> add.left_commute
    nat_mult_assoc ~> mult.assoc
    nat_mult_commute ~> mult.commute
    eq_assoc ~> iff_assoc
    eq_left_commute ~> iff_left_commute

INCOMPATIBILITY.

* Fact collections add_ac and mult_ac are considered old-fashioned.
Prefer ac_simps instead, or specify rules
(add|mult).(assoc|commute|left_commute) individually.

* Elimination of fact duplicates:
    equals_zero_I ~> minus_unique
    diff_eq_0_iff_eq ~> right_minus_eq
    nat_infinite ~> infinite_UNIV_nat
    int_infinite ~> infinite_UNIV_int

INCOMPATIBILITY.

* Fact name consolidation:
    diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
    minus_le_self_iff ~> neg_less_eq_nonneg
    le_minus_self_iff ~> less_eq_neg_nonpos
    neg_less_nonneg ~> neg_less_pos
    less_minus_self_iff ~> less_neg_neg [simp]

INCOMPATIBILITY.

* More simplification rules on unary and binary minus:
add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
minus_add_cancel, uminus_add_conv_diff.  These correspondingly have
been taken away from fact collections algebra_simps and field_simps.
INCOMPATIBILITY.

To restore proofs, the following patterns are helpful:

a) Arbitrary failing proof not involving "diff_def":
Consider simplification with algebra_simps or field_simps.

b) Lifting rules from addition to subtraction:
Try with "using <rule for addition> of [... "- _" ...]" by simp".

c) Simplification with "diff_def": just drop "diff_def".
Consider simplification with algebra_simps or field_simps;
or the brute way with
"simp add: diff_conv_add_uminus del: add_uminus_conv_diff".

* Introduce bdd_above and bdd_below in theory
Conditionally_Complete_Lattices, use them instead of explicitly
stating boundedness of sets.

* ccpo.admissible quantifies only over non-empty chains to allow more
syntax-directed proof rules; the case of the empty chain shows up as
additional case in fixpoint induction proofs.  INCOMPATIBILITY.

* Removed and renamed theorems in Series:
  summable_le         ~>  suminf_le
  suminf_le           ~>  suminf_le_const
  series_pos_le       ~>  setsum_le_suminf
  series_pos_less     ~>  setsum_less_suminf
  suminf_ge_zero      ~>  suminf_nonneg
  suminf_gt_zero      ~>  suminf_pos
  suminf_gt_zero_iff  ~>  suminf_pos_iff
  summable_sumr_LIMSEQ_suminf  ~>  summable_LIMSEQ
  suminf_0_le         ~>  suminf_nonneg [rotate]
  pos_summable        ~>  summableI_nonneg_bounded
  ratio_test          ~>  summable_ratio_test

  removed series_zero, replaced by sums_finite

  removed auxiliary lemmas:

    sumr_offset, sumr_offset2, sumr_offset3, sumr_offset4, sumr_group,
    half, le_Suc_ex_iff, lemma_realpow_diff_sumr,
    real_setsum_nat_ivl_bounded, summable_le2, ratio_test_lemma2,
    sumr_minus_one_realpow_zerom, sumr_one_lb_realpow_zero,
    summable_convergent_sumr_iff, sumr_diff_mult_const

INCOMPATIBILITY.

* Replace (F)DERIV syntax by has_derivative:
  - "(f has_derivative f') (at x within s)" replaces "FDERIV f x : s : f'"

  - "(f has_field_derivative f') (at x within s)" replaces "DERIV f x : s : f'"

  - "f differentiable at x within s" replaces "_ differentiable _ in _" syntax

  - removed constant isDiff

  - "DERIV f x : f'" and "FDERIV f x : f'" syntax is only available as
    input syntax.

  - "DERIV f x : s : f'" and "FDERIV f x : s : f'" syntax removed.

  - Renamed FDERIV_... lemmas to has_derivative_...

  - renamed deriv (the syntax constant used for "DERIV _ _ :> _") to DERIV

  - removed DERIV_intros, has_derivative_eq_intros

  - introduced derivative_intros and deriative_eq_intros which
    includes now rules for DERIV, has_derivative and
    has_vector_derivative.

  - Other renamings:
    differentiable_def        ~>  real_differentiable_def
    differentiableE           ~>  real_differentiableE
    fderiv_def                ~>  has_derivative_at
    field_fderiv_def          ~>  field_has_derivative_at
    isDiff_der                ~>  differentiable_def
    deriv_fderiv              ~>  has_field_derivative_def
    deriv_def                 ~>  DERIV_def

INCOMPATIBILITY.

* Include more theorems in continuous_intros. Remove the
continuous_on_intros, isCont_intros collections, these facts are now
in continuous_intros.

* Theorems about complex numbers are now stated only using Re and Im,
the Complex constructor is not used anymore. It is possible to use
primcorec to defined the behaviour of a complex-valued function.

Removed theorems about the Complex constructor from the simpset, they
are available as the lemma collection legacy_Complex_simps. This
especially removes

    i_complex_of_real: "ii * complex_of_real r = Complex 0 r".

Instead the reverse direction is supported with
    Complex_eq: "Complex a b = a + \<i> * b"

Moved csqrt from Fundamental_Algebra_Theorem to Complex.

  Renamings:
    Re/Im                  ~>  complex.sel
    complex_Re/Im_zero     ~>  zero_complex.sel
    complex_Re/Im_add      ~>  plus_complex.sel
    complex_Re/Im_minus    ~>  uminus_complex.sel
    complex_Re/Im_diff     ~>  minus_complex.sel
    complex_Re/Im_one      ~>  one_complex.sel
    complex_Re/Im_mult     ~>  times_complex.sel
    complex_Re/Im_inverse  ~>  inverse_complex.sel
    complex_Re/Im_scaleR   ~>  scaleR_complex.sel
    complex_Re/Im_i        ~>  ii.sel
    complex_Re/Im_cnj      ~>  cnj.sel
    Re/Im_cis              ~>  cis.sel

    complex_divide_def   ~>  divide_complex_def
    complex_norm_def     ~>  norm_complex_def
    cmod_def             ~>  norm_complex_de

  Removed theorems:
    complex_zero_def
    complex_add_def
    complex_minus_def
    complex_diff_def
    complex_one_def
    complex_mult_def
    complex_inverse_def
    complex_scaleR_def

INCOMPATIBILITY.

* Theory Lubs moved HOL image to HOL-Library. It is replaced by
Conditionally_Complete_Lattices.  INCOMPATIBILITY.

* HOL-Library: new theory src/HOL/Library/Tree.thy.

* HOL-Library: removed theory src/HOL/Library/Kleene_Algebra.thy; it
is subsumed by session Kleene_Algebra in AFP.

* HOL-Library / theory RBT: various constants and facts are hidden;
lifting setup is unregistered.  INCOMPATIBILITY.

* HOL-Cardinals: new theory src/HOL/Cardinals/Ordinal_Arithmetic.thy.

* HOL-Word: bit representations prefer type bool over type bit.
INCOMPATIBILITY.

* HOL-Word:
  - Abandoned fact collection "word_arith_alts", which is a duplicate
    of "word_arith_wis".
  - Dropped first (duplicated) element in fact collections
    "sint_word_ariths", "word_arith_alts", "uint_word_ariths",
    "uint_word_arith_bintrs".

* HOL-Number_Theory:
  - consolidated the proofs of the binomial theorem
  - the function fib is again of type nat => nat and not overloaded
  - no more references to Old_Number_Theory in the HOL libraries
    (except the AFP)

INCOMPATIBILITY.

* HOL-Multivariate_Analysis:
  - Type class ordered_real_vector for ordered vector spaces.
  - New theory Complex_Basic_Analysis defining complex derivatives,
    holomorphic functions, etc., ported from HOL Light's canal.ml.
  - Changed order of ordered_euclidean_space to be compatible with
    pointwise ordering on products. Therefore instance of
    conditionally_complete_lattice and ordered_real_vector.
    INCOMPATIBILITY: use box instead of greaterThanLessThan or
    explicit set-comprehensions with eucl_less for other (half-)open
    intervals.
  - removed dependencies on type class ordered_euclidean_space with
    introduction of "cbox" on euclidean_space
    - renamed theorems:
        interval ~> box
        mem_interval ~> mem_box
        interval_eq_empty ~> box_eq_empty
        interval_ne_empty ~> box_ne_empty
        interval_sing(1) ~> cbox_sing
        interval_sing(2) ~> box_sing
        subset_interval_imp ~> subset_box_imp
        subset_interval ~> subset_box
        open_interval ~> open_box
        closed_interval ~> closed_cbox
        interior_closed_interval ~> interior_cbox
        bounded_closed_interval ~> bounded_cbox
        compact_interval ~> compact_cbox
        bounded_subset_closed_interval_symmetric ~> bounded_subset_cbox_symmetric
        bounded_subset_closed_interval ~> bounded_subset_cbox
        mem_interval_componentwiseI ~> mem_box_componentwiseI
        convex_box ~> convex_prod
        rel_interior_real_interval ~> rel_interior_real_box
        convex_interval ~> convex_box
        convex_hull_eq_real_interval ~> convex_hull_eq_real_cbox
        frechet_derivative_within_closed_interval ~> frechet_derivative_within_cbox
        content_closed_interval' ~> content_cbox'
        elementary_subset_interval ~> elementary_subset_box
        diameter_closed_interval ~> diameter_cbox
        frontier_closed_interval ~> frontier_cbox
        frontier_open_interval ~> frontier_box
        bounded_subset_open_interval_symmetric ~> bounded_subset_box_symmetric
        closure_open_interval ~> closure_box
        open_closed_interval_convex ~> open_cbox_convex
        open_interval_midpoint ~> box_midpoint
        content_image_affinity_interval ~> content_image_affinity_cbox
        is_interval_interval ~> is_interval_cbox + is_interval_box + is_interval_closed_interval
        bounded_interval ~> bounded_closed_interval + bounded_boxes

    - respective theorems for intervals over the reals:
        content_closed_interval + content_cbox
        has_integral + has_integral_real
        fine_division_exists + fine_division_exists_real
        has_integral_null + has_integral_null_real
        tagged_division_union_interval + tagged_division_union_interval_real
        has_integral_const + has_integral_const_real
        integral_const + integral_const_real
        has_integral_bound + has_integral_bound_real
        integrable_continuous + integrable_continuous_real
        integrable_subinterval + integrable_subinterval_real
        has_integral_reflect_lemma + has_integral_reflect_lemma_real
        integrable_reflect + integrable_reflect_real
        integral_reflect + integral_reflect_real
        image_affinity_interval + image_affinity_cbox
        image_smult_interval + image_smult_cbox
        integrable_const + integrable_const_ivl
        integrable_on_subinterval + integrable_on_subcbox

  - renamed theorems:
    derivative_linear         ~>  has_derivative_bounded_linear
    derivative_is_linear      ~>  has_derivative_linear
    bounded_linear_imp_linear ~>  bounded_linear.linear

* HOL-Probability:
  - Renamed positive_integral to nn_integral:

    . Renamed all lemmas "*positive_integral*" to *nn_integral*"
      positive_integral_positive ~> nn_integral_nonneg

    . Renamed abbreviation integral\<^sup>P to integral\<^sup>N.

  - replaced the Lebesgue integral on real numbers by the more general
    Bochner integral for functions into a real-normed vector space.

    integral_zero               ~>  integral_zero / integrable_zero
    integral_minus              ~>  integral_minus / integrable_minus
    integral_add                ~>  integral_add / integrable_add
    integral_diff               ~>  integral_diff / integrable_diff
    integral_setsum             ~>  integral_setsum / integrable_setsum
    integral_multc              ~>  integral_mult_left / integrable_mult_left
    integral_cmult              ~>  integral_mult_right / integrable_mult_right
    integral_triangle_inequality~>  integral_norm_bound
    integrable_nonneg           ~>  integrableI_nonneg
    integral_positive           ~>  integral_nonneg_AE
    integrable_abs_iff          ~>  integrable_abs_cancel
    positive_integral_lim_INF   ~>  nn_integral_liminf
    lebesgue_real_affine        ~>  lborel_real_affine
    borel_integral_has_integral ~>  has_integral_lebesgue_integral
    integral_indicator          ~>
         integral_real_indicator / integrable_real_indicator
    positive_integral_fst       ~>  nn_integral_fst'
    positive_integral_fst_measurable ~> nn_integral_fst
    positive_integral_snd_measurable ~> nn_integral_snd

    integrable_fst_measurable   ~>
         integral_fst / integrable_fst / AE_integrable_fst

    integrable_snd_measurable   ~>
         integral_snd / integrable_snd / AE_integrable_snd

    integral_monotone_convergence  ~>
         integral_monotone_convergence / integrable_monotone_convergence

    integral_monotone_convergence_at_top  ~>
         integral_monotone_convergence_at_top /
         integrable_monotone_convergence_at_top

    has_integral_iff_positive_integral_lebesgue  ~>
         has_integral_iff_has_bochner_integral_lebesgue_nonneg

    lebesgue_integral_has_integral  ~>
         has_integral_integrable_lebesgue_nonneg

    positive_integral_lebesgue_has_integral  ~>
         integral_has_integral_lebesgue_nonneg /
         integrable_has_integral_lebesgue_nonneg

    lebesgue_integral_real_affine  ~>
         nn_integral_real_affine

    has_integral_iff_positive_integral_lborel  ~>
         integral_has_integral_nonneg / integrable_has_integral_nonneg

    The following theorems where removed:

    lebesgue_integral_nonneg
    lebesgue_integral_uminus
    lebesgue_integral_cmult
    lebesgue_integral_multc
    lebesgue_integral_cmult_nonneg
    integral_cmul_indicator
    integral_real

  - Formalized properties about exponentially, Erlang, and normal
    distributed random variables.

* HOL-Decision_Procs: Separate command 'approximate' for approximative
computation in src/HOL/Decision_Procs/Approximation.  Minor
INCOMPATIBILITY.


*** Scala ***

* The signature and semantics of Document.Snapshot.cumulate_markup /
select_markup have been clarified.  Markup is now traversed in the
order of reports given by the prover: later markup is usually more
specific and may override results accumulated so far.  The elements
guard is mandatory and checked precisely.  Subtle INCOMPATIBILITY.

* Substantial reworking of internal PIDE protocol communication
channels.  INCOMPATIBILITY.


*** ML ***

* Subtle change of semantics of Thm.eq_thm: theory stamps are not
compared (according to Thm.thm_ord), but assumed to be covered by the
current background theory.  Thus equivalent data produced in different
branches of the theory graph usually coincides (e.g. relevant for
theory merge).  Note that the softer Thm.eq_thm_prop is often more
appropriate than Thm.eq_thm.

* Proper context for basic Simplifier operations: rewrite_rule,
rewrite_goals_rule, rewrite_goals_tac etc. INCOMPATIBILITY, need to
pass runtime Proof.context (and ensure that the simplified entity
actually belongs to it).

* Proper context discipline for read_instantiate and instantiate_tac:
variables that are meant to become schematic need to be given as
fixed, and are generalized by the explicit context of local variables.
This corresponds to Isar attributes "where" and "of" with 'for'
declaration.  INCOMPATIBILITY, also due to potential change of indices
of schematic variables.

* Moved ML_Compiler.exn_trace and other operations on exceptions to
structure Runtime.  Minor INCOMPATIBILITY.

* Discontinued old Toplevel.debug in favour of system option
"ML_exception_trace", which may be also declared within the context
via "declare [[ML_exception_trace = true]]".  Minor INCOMPATIBILITY.

* Renamed configuration option "ML_trace" to "ML_source_trace". Minor
INCOMPATIBILITY.

* Configuration option "ML_print_depth" controls the pretty-printing
depth of the ML compiler within the context.  The old print_depth in
ML is still available as default_print_depth, but rarely used.  Minor
INCOMPATIBILITY.

* Toplevel function "use" refers to raw ML bootstrap environment,
without Isar context nor antiquotations.  Potential INCOMPATIBILITY.
Note that 'ML_file' is the canonical command to load ML files into the
formal context.

* Simplified programming interface to define ML antiquotations, see
structure ML_Antiquotation.  Minor INCOMPATIBILITY.

* ML antiquotation @{here} refers to its source position, which is
occasionally useful for experimentation and diagnostic purposes.

* ML antiquotation @{path} produces a Path.T value, similarly to
Path.explode, but with compile-time check against the file-system and
some PIDE markup.  Note that unlike theory source, ML does not have a
well-defined master directory, so an absolute symbolic path
specification is usually required, e.g. "~~/src/HOL".

* ML antiquotation @{print} inlines a function to print an arbitrary
ML value, which is occasionally useful for diagnostic or demonstration
purposes.


*** System ***

* Proof General with its traditional helper scripts is now an optional
Isabelle component, e.g. see ProofGeneral-4.2-2 from the Isabelle
component repository http://isabelle.in.tum.de/components/.  Note that
the "system" manual provides general explanations about add-on
components, especially those that are not bundled with the release.

* The raw Isabelle process executable has been renamed from
"isabelle-process" to "isabelle_process", which conforms to common
shell naming conventions, and allows to define a shell function within
the Isabelle environment to avoid dynamic path lookup.  Rare
incompatibility for old tools that do not use the ISABELLE_PROCESS
settings variable.

* Former "isabelle tty" has been superseded by "isabelle console",
with implicit build like "isabelle jedit", and without the mostly
obsolete Isar TTY loop.

* Simplified "isabelle display" tool.  Settings variables DVI_VIEWER
and PDF_VIEWER now refer to the actual programs, not shell
command-lines.  Discontinued option -c: invocation may be asynchronous
via desktop environment, without any special precautions.  Potential
INCOMPATIBILITY with ambitious private settings.

* Removed obsolete "isabelle unsymbolize".  Note that the usual format
for email communication is the Unicode rendering of Isabelle symbols,
as produced by Isabelle/jEdit, for example.

* Removed obsolete tool "wwwfind". Similar functionality may be
integrated into Isabelle/jEdit eventually.

* Improved 'display_drafts' concerning desktop integration and
repeated invocation in PIDE front-end: re-use single file
$ISABELLE_HOME_USER/tmp/drafts.pdf and corresponding views.

* Session ROOT specifications require explicit 'document_files' for
robust dependencies on LaTeX sources.  Only these explicitly given
files are copied to the document output directory, before document
processing is started.

* Windows: support for regular TeX installation (e.g. MiKTeX) instead
of TeX Live from Cygwin.



New in Isabelle2013-2 (December 2013)
-------------------------------------

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

* More robust editing of running commands with internal forks,
e.g. non-terminating 'by' steps.

* More relaxed Sledgehammer panel: avoid repeated application of query
after edits surrounding the command location.

* More status information about commands that are interrupted
accidentally (via physical event or Poly/ML runtime system signal,
e.g. out-of-memory).


*** System ***

* More robust termination of external processes managed by
Isabelle/ML: support cancellation of tasks within the range of
milliseconds, as required for PIDE document editing with automatically
tried tools (e.g. Sledgehammer).

* Reactivated Isabelle/Scala kill command for external processes on
Mac OS X, which was accidentally broken in Isabelle2013-1 due to a
workaround for some Debian/Ubuntu Linux versions from 2013.



New in Isabelle2013-1 (November 2013)
-------------------------------------

*** General ***

* Discontinued obsolete 'uses' within theory header.  Note that
commands like 'ML_file' work without separate declaration of file
dependencies.  Minor INCOMPATIBILITY.

* Discontinued redundant 'use' command, which was superseded by
'ML_file' in Isabelle2013.  Minor INCOMPATIBILITY.

* Simplified subscripts within identifiers, using plain \<^sub>
instead of the second copy \<^isub> and \<^isup>.  Superscripts are
only for literal tokens within notation; explicit mixfix annotations
for consts or fixed variables may be used as fall-back for unusual
names.  Obsolete \<twosuperior> has been expanded to \<^sup>2 in
Isabelle/HOL.  INCOMPATIBILITY, use "isabelle update_sub_sup" to
standardize symbols as a starting point for further manual cleanup.
The ML reference variable "legacy_isub_isup" may be set as temporary
workaround, to make the prover accept a subset of the old identifier
syntax.

* Document antiquotations: term style "isub" has been renamed to
"sub".  Minor INCOMPATIBILITY.

* Uniform management of "quick_and_dirty" as system option (see also
"isabelle options"), configuration option within the context (see also
Config.get in Isabelle/ML), and attribute in Isabelle/Isar.  Minor
INCOMPATIBILITY, need to use more official Isabelle means to access
quick_and_dirty, instead of historical poking into mutable reference.

* Renamed command 'print_configs' to 'print_options'.  Minor
INCOMPATIBILITY.

* Proper diagnostic command 'print_state'.  Old 'pr' (with its
implicit change of some global references) is retained for now as
control command, e.g. for ProofGeneral 3.7.x.

* Discontinued 'print_drafts' command with its old-fashioned PS output
and Unix command-line print spooling.  Minor INCOMPATIBILITY: use
'display_drafts' instead and print via the regular document viewer.

* Updated and extended "isar-ref" and "implementation" manual,
eliminated old "ref" manual.


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

* New manual "jedit" for Isabelle/jEdit, see isabelle doc or
Documentation panel.

* Dockable window "Documentation" provides access to Isabelle
documentation.

* Dockable window "Find" provides query operations for formal entities
(GUI front-end to 'find_theorems' command).

* Dockable window "Sledgehammer" manages asynchronous / parallel
sledgehammer runs over existing document sources, independently of
normal editing and checking process.

* Dockable window "Timing" provides an overview of relevant command
timing information, depending on option jedit_timing_threshold.  The
same timing information is shown in the extended tooltip of the
command keyword, when hovering the mouse over it while the CONTROL or
COMMAND modifier is pressed.

* Improved dockable window "Theories": Continuous checking of proof
document (visible and required parts) may be controlled explicitly,
using check box or shortcut "C+e ENTER".  Individual theory nodes may
be marked explicitly as required and checked in full, using check box
or shortcut "C+e SPACE".

* Improved completion mechanism, which is now managed by the
Isabelle/jEdit plugin instead of SideKick.  Refined table of Isabelle
symbol abbreviations (see $ISABELLE_HOME/etc/symbols).

* Standard jEdit keyboard shortcut C+b complete-word is remapped to
isabelle.complete for explicit completion in Isabelle sources.
INCOMPATIBILITY wrt. jEdit defaults, may have to invent new shortcuts
to resolve conflict.

* Improved support of various "minor modes" for Isabelle NEWS,
options, session ROOT etc., with completion and SideKick tree view.

* Strictly monotonic document update, without premature cancellation of
running transactions that are still needed: avoid reset/restart of
such command executions while editing.

* Support for asynchronous print functions, as overlay to existing
document content.

* Support for automatic tools in HOL, which try to prove or disprove
toplevel theorem statements.

* Action isabelle.reset-font-size resets main text area font size
according to Isabelle/Scala plugin option "jedit_font_reset_size" (see
also "Plugin Options / Isabelle / General").  It can be bound to some
keyboard shortcut by the user (e.g. C+0 and/or C+NUMPAD0).

* File specifications in jEdit (e.g. file browser) may refer to
$ISABELLE_HOME and $ISABELLE_HOME_USER on all platforms.  Discontinued
obsolete $ISABELLE_HOME_WINDOWS variable.

* Improved support for Linux look-and-feel "GTK+", see also "Utilities
/ Global Options / Appearance".

* Improved support of native Mac OS X functionality via "MacOSX"
plugin, which is now enabled by default.


*** Pure ***

* Commands 'interpretation' and 'sublocale' are now target-sensitive.
In particular, 'interpretation' allows for non-persistent
interpretation within "context ... begin ... end" blocks offering a
light-weight alternative to 'sublocale'.  See "isar-ref" manual for
details.

* Improved locales diagnostic command 'print_dependencies'.

* Discontinued obsolete 'axioms' command, which has been marked as
legacy since Isabelle2009-2.  INCOMPATIBILITY, use 'axiomatization'
instead, while observing its uniform scope for polymorphism.

* Discontinued empty name bindings in 'axiomatization'.
INCOMPATIBILITY.

* System option "proofs" has been discontinued.  Instead the global
state of Proofterm.proofs is persistently compiled into logic images
as required, notably HOL-Proofs.  Users no longer need to change
Proofterm.proofs dynamically.  Minor INCOMPATIBILITY.

* Syntax translation functions (print_translation etc.) always depend
on Proof.context.  Discontinued former "(advanced)" option -- this is
now the default.  Minor INCOMPATIBILITY.

* Former global reference trace_unify_fail is now available as
configuration option "unify_trace_failure" (global context only).

* SELECT_GOAL now retains the syntactic context of the overall goal
state (schematic variables etc.).  Potential INCOMPATIBILITY in rare
situations.


*** HOL ***

* Stronger precedence of syntax for big intersection and union on
sets, in accordance with corresponding lattice operations.
INCOMPATIBILITY.

* Notation "{p:A. P}" now allows tuple patterns as well.

* Nested case expressions are now translated in a separate check phase
rather than during parsing. The data for case combinators is separated
from the datatype package. The declaration attribute
"case_translation" can be used to register new case combinators:

  declare [[case_translation case_combinator constructor1 ... constructorN]]

* Code generator:
  - 'code_printing' unifies 'code_const' / 'code_type' / 'code_class' /
    'code_instance'.
  - 'code_identifier' declares name hints for arbitrary identifiers in
    generated code, subsuming 'code_modulename'.

See the isar-ref manual for syntax diagrams, and the HOL theories for
examples.

* Attibute 'code': 'code' now declares concrete and abstract code
equations uniformly.  Use explicit 'code equation' and 'code abstract'
to distinguish both when desired.

* Discontinued theories Code_Integer and Efficient_Nat by a more
fine-grain stack of theories Code_Target_Int, Code_Binary_Nat,
Code_Target_Nat and Code_Target_Numeral.  See the tutorial on code
generation for details.  INCOMPATIBILITY.

* Numeric types are mapped by default to target language numerals:
natural (replaces former code_numeral) and integer (replaces former
code_int).  Conversions are available as integer_of_natural /
natural_of_integer / integer_of_nat / nat_of_integer (in HOL) and
Code_Numeral.integer_of_natural / Code_Numeral.natural_of_integer (in
ML).  INCOMPATIBILITY.

* Function package: For mutually recursive functions f and g, separate
cases rules f.cases and g.cases are generated instead of unusable
f_g.cases which exposed internal sum types. Potential INCOMPATIBILITY,
in the case that the unusable rule was used nevertheless.

* Function package: For each function f, new rules f.elims are
generated, which eliminate equalities of the form "f x = t".

* New command 'fun_cases' derives ad-hoc elimination rules for
function equations as simplified instances of f.elims, analogous to
inductive_cases.  See ~~/src/HOL/ex/Fundefs.thy for some examples.

* Lifting:
  - parametrized correspondence relations are now supported:
    + parametricity theorems for the raw term can be specified in
      the command lift_definition, which allow us to generate stronger
      transfer rules
    + setup_lifting generates stronger transfer rules if parametric
      correspondence relation can be generated
    + various new properties of the relator must be specified to support
      parametricity
    + parametricity theorem for the Quotient relation can be specified
  - setup_lifting generates domain rules for the Transfer package
  - stronger reflexivity prover of respectfulness theorems for type
    copies
  - ===> and --> are now local. The symbols can be introduced
    by interpreting the locale lifting_syntax (typically in an
    anonymous context)
  - Lifting/Transfer relevant parts of Library/Quotient_* are now in
    Main. Potential INCOMPATIBILITY
  - new commands for restoring and deleting Lifting/Transfer context:
    lifting_forget, lifting_update
  - the command print_quotmaps was renamed to print_quot_maps.
    INCOMPATIBILITY

* Transfer:
  - better support for domains in Transfer: replace Domainp T
    by the actual invariant in a transferred goal
  - transfer rules can have as assumptions other transfer rules
  - Experimental support for transferring from the raw level to the
    abstract level: Transfer.transferred attribute
  - Attribute version of the transfer method: untransferred attribute

* Reification and reflection:
  - Reification is now directly available in HOL-Main in structure
    "Reification".
  - Reflection now handles multiple lists with variables also.
  - The whole reflection stack has been decomposed into conversions.
INCOMPATIBILITY.

* Revised devices for recursive definitions over finite sets:
  - Only one fundamental fold combinator on finite set remains:
    Finite_Set.fold :: ('a => 'b => 'b) => 'b => 'a set => 'b
    This is now identity on infinite sets.
  - Locales ("mini packages") for fundamental definitions with
    Finite_Set.fold: folding, folding_idem.
  - Locales comm_monoid_set, semilattice_order_set and
    semilattice_neutr_order_set for big operators on sets.
    See theory Big_Operators for canonical examples.
    Note that foundational constants comm_monoid_set.F and
    semilattice_set.F correspond to former combinators fold_image
    and fold1 respectively.  These are now gone.  You may use
    those foundational constants as substitutes, but it is
    preferable to interpret the above locales accordingly.
  - Dropped class ab_semigroup_idem_mult (special case of lattice,
    no longer needed in connection with Finite_Set.fold etc.)
  - Fact renames:
      card.union_inter ~> card_Un_Int [symmetric]
      card.union_disjoint ~> card_Un_disjoint
INCOMPATIBILITY.

* Locale hierarchy for abstract orderings and (semi)lattices.

* Complete_Partial_Order.admissible is defined outside the type class
ccpo, but with mandatory prefix ccpo. Admissibility theorems lose the
class predicate assumption or sort constraint when possible.
INCOMPATIBILITY.

* Introduce type class "conditionally_complete_lattice": Like a
complete lattice but does not assume the existence of the top and
bottom elements.  Allows to generalize some lemmas about reals and
extended reals.  Removed SupInf and replaced it by the instantiation
of conditionally_complete_lattice for real. Renamed lemmas about
conditionally-complete lattice from Sup_... to cSup_... and from
Inf_...  to cInf_... to avoid hidding of similar complete lattice
lemmas.

* Introduce type class linear_continuum as combination of
conditionally-complete lattices and inner dense linorders which have
more than one element.  INCOMPATIBILITY.

* Introduced type classes order_top and order_bot. The old classes top
and bot only contain the syntax without assumptions.  INCOMPATIBILITY:
Rename bot -> order_bot, top -> order_top

* Introduce type classes "no_top" and "no_bot" for orderings without
top and bottom elements.

* Split dense_linorder into inner_dense_order and no_top, no_bot.

* Complex_Main: Unify and move various concepts from
HOL-Multivariate_Analysis to HOL-Complex_Main.

 - Introduce type class (lin)order_topology and
   linear_continuum_topology.  Allows to generalize theorems about
   limits and order.  Instances are reals and extended reals.

 - continuous and continuos_on from Multivariate_Analysis:
   "continuous" is the continuity of a function at a filter.  "isCont"
   is now an abbrevitation: "isCont x f == continuous (at _) f".

   Generalized continuity lemmas from isCont to continuous on an
   arbitrary filter.

 - compact from Multivariate_Analysis. Use Bolzano's lemma to prove
   compactness of closed intervals on reals. Continuous functions
   attain infimum and supremum on compact sets. The inverse of a
   continuous function is continuous, when the function is continuous
   on a compact set.

 - connected from Multivariate_Analysis. Use it to prove the
   intermediate value theorem. Show connectedness of intervals on
   linear_continuum_topology).

 - first_countable_topology from Multivariate_Analysis. Is used to
   show equivalence of properties on the neighbourhood filter of x and
   on all sequences converging to x.

 - FDERIV: Definition of has_derivative moved to Deriv.thy. Moved
   theorems from Library/FDERIV.thy to Deriv.thy and base the
   definition of DERIV on FDERIV. Add variants of DERIV and FDERIV
   which are restricted to sets, i.e. to represent derivatives from
   left or right.

 - Removed the within-filter. It is replaced by the principal filter:

     F within X = inf F (principal X)

 - Introduce "at x within U" as a single constant, "at x" is now an
   abbreviation for "at x within UNIV"

 - Introduce named theorem collections tendsto_intros,
   continuous_intros, continuous_on_intros and FDERIV_intros. Theorems
   in tendsto_intros (or FDERIV_intros) are also available as
   tendsto_eq_intros (or FDERIV_eq_intros) where the right-hand side
   is replaced by a congruence rule. This allows to apply them as
   intro rules and then proving equivalence by the simplifier.

 - Restructured theories in HOL-Complex_Main:

   + Moved RealDef and RComplete into Real

   + Introduced Topological_Spaces and moved theorems about
     topological spaces, filters, limits and continuity to it

   + Renamed RealVector to Real_Vector_Spaces

   + Split Lim, SEQ, Series into Topological_Spaces,
     Real_Vector_Spaces, and Limits

   + Moved Ln and Log to Transcendental

   + Moved theorems about continuity from Deriv to Topological_Spaces

 - Remove various auxiliary lemmas.

INCOMPATIBILITY.

* Nitpick:
  - Added option "spy".
  - Reduce incidence of "too high arity" errors.

* Sledgehammer:
  - Renamed option:
      isar_shrink ~> isar_compress
    INCOMPATIBILITY.
  - Added options "isar_try0", "spy".
  - Better support for "isar_proofs".
  - MaSh has been fined-tuned and now runs as a local server.

* Improved support for ad hoc overloading of constants (see also
isar-ref manual and ~~/src/HOL/ex/Adhoc_Overloading_Examples.thy).

* Library/Polynomial.thy:
  - Use lifting for primitive definitions.
  - Explicit conversions from and to lists of coefficients, used for
    generated code.
  - Replaced recursion operator poly_rec by fold_coeffs.
  - Prefer pre-existing gcd operation for gcd.
  - Fact renames:
    poly_eq_iff ~> poly_eq_poly_eq_iff
    poly_ext ~> poly_eqI
    expand_poly_eq ~> poly_eq_iff
IMCOMPATIBILITY.

* New Library/Simps_Case_Conv.thy: Provides commands simps_of_case and
case_of_simps to convert function definitions between a list of
equations with patterns on the lhs and a single equation with case
expressions on the rhs. See also Ex/Simps_Case_Conv_Examples.thy.

* New Library/FSet.thy: type of finite sets defined as a subtype of
sets defined by Lifting/Transfer.

* Discontinued theory src/HOL/Library/Eval_Witness.  INCOMPATIBILITY.

* Consolidation of library theories on product orders:

    Product_Lattice ~> Product_Order -- pointwise order on products
    Product_ord ~> Product_Lexorder -- lexicographic order on products

INCOMPATIBILITY.

* Imperative-HOL: The MREC combinator is considered legacy and no
longer included by default. INCOMPATIBILITY, use partial_function
instead, or import theory Legacy_Mrec as a fallback.

* HOL-Algebra: Discontinued theories ~~/src/HOL/Algebra/abstract and
~~/src/HOL/Algebra/poly.  Existing theories should be based on
~~/src/HOL/Library/Polynomial instead.  The latter provides
integration with HOL's type classes for rings.  INCOMPATIBILITY.

* HOL-BNF:
  - Various improvements to BNF-based (co)datatype package, including
    new commands "primrec_new", "primcorec", and
    "datatype_new_compat", as well as documentation. See
    "datatypes.pdf" for details.
  - New "coinduction" method to avoid some boilerplate (compared to
    coinduct).
  - Renamed keywords:
    data ~> datatype_new
    codata ~> codatatype
    bnf_def ~> bnf
  - Renamed many generated theorems, including
    discs ~> disc
    map_comp' ~> map_comp
    map_id' ~> map_id
    sels ~> sel
    set_map' ~> set_map
    sets ~> set
IMCOMPATIBILITY.


*** ML ***

* Spec_Check is a Quickcheck tool for Isabelle/ML.  The ML function
"check_property" allows to check specifications of the form "ALL x y
z. prop x y z".  See also ~~/src/Tools/Spec_Check/ with its
Examples.thy in particular.

* Improved printing of exception trace in Poly/ML 5.5.1, with regular
tracing output in the command transaction context instead of physical
stdout.  See also Toplevel.debug, Toplevel.debugging and
ML_Compiler.exn_trace.

* ML type "theory" is now immutable, without any special treatment of
drafts or linear updates (which could lead to "stale theory" errors in
the past).  Discontinued obsolete operations like Theory.copy,
Theory.checkpoint, and the auxiliary type theory_ref.  Minor
INCOMPATIBILITY.

* More uniform naming of goal functions for skipped proofs:

    Skip_Proof.prove  ~>  Goal.prove_sorry
    Skip_Proof.prove_global  ~>  Goal.prove_sorry_global

Minor INCOMPATIBILITY.

* Simplifier tactics and tools use proper Proof.context instead of
historic type simpset.  Old-style declarations like addsimps,
addsimprocs etc. operate directly on Proof.context.  Raw type simpset
retains its use as snapshot of the main Simplifier context, using
simpset_of and put_simpset on Proof.context.  INCOMPATIBILITY -- port
old tools by making them depend on (ctxt : Proof.context) instead of
(ss : simpset), then turn (simpset_of ctxt) into ctxt.

* Modifiers for classical wrappers (e.g. addWrapper, delWrapper)
operate on Proof.context instead of claset, for uniformity with addIs,
addEs, addDs etc. Note that claset_of and put_claset allow to manage
clasets separately from the context.

* Discontinued obsolete ML antiquotations @{claset} and @{simpset}.
INCOMPATIBILITY, use @{context} instead.

* Antiquotation @{theory_context A} is similar to @{theory A}, but
presents the result as initial Proof.context.


*** System ***

* Discontinued obsolete isabelle usedir, mkdir, make -- superseded by
"isabelle build" in Isabelle2013.  INCOMPATIBILITY.

* Discontinued obsolete isabelle-process options -f and -u (former
administrative aliases of option -e).  Minor INCOMPATIBILITY.

* Discontinued obsolete isabelle print tool, and PRINT_COMMAND
settings variable.

* Discontinued ISABELLE_DOC_FORMAT settings variable and historic
document formats: dvi.gz, ps, ps.gz -- the default document format is
always pdf.

* Isabelle settings variable ISABELLE_BUILD_JAVA_OPTIONS allows to
specify global resources of the JVM process run by isabelle build.

* Toplevel executable $ISABELLE_HOME/bin/isabelle_scala_script allows
to run Isabelle/Scala source files as standalone programs.

* Improved "isabelle keywords" tool (for old-style ProofGeneral
keyword tables): use Isabelle/Scala operations, which inspect outer
syntax without requiring to build sessions first.

* Sessions may be organized via 'chapter' specifications in the ROOT
file, which determines a two-level hierarchy of browser info.  The old
tree-like organization via implicit sub-session relation (with its
tendency towards erratic fluctuation of URLs) has been discontinued.
The default chapter is called "Unsorted".  Potential INCOMPATIBILITY
for HTML presentation of theories.



New in Isabelle2013 (February 2013)
-----------------------------------

*** General ***

* Theorem status about oracles and unfinished/failed future proofs is
no longer printed by default, since it is incompatible with
incremental / parallel checking of the persistent document model.  ML
function Thm.peek_status may be used to inspect a snapshot of the
ongoing evaluation process.  Note that in batch mode --- notably
isabelle build --- the system ensures that future proofs of all
accessible theorems in the theory context are finished (as before).

* Configuration option show_markup controls direct inlining of markup
into the printed representation of formal entities --- notably type
and sort constraints.  This enables Prover IDE users to retrieve that
information via tooltips in the output window, for example.

* Command 'ML_file' evaluates ML text from a file directly within the
theory, without any predeclaration via 'uses' in the theory header.

* Old command 'use' command and corresponding keyword 'uses' in the
theory header are legacy features and will be discontinued soon.
Tools that load their additional source files may imitate the
'ML_file' implementation, such that the system can take care of
dependencies properly.

* Discontinued obsolete method fastsimp / tactic fast_simp_tac, which
is called fastforce / fast_force_tac already since Isabelle2011-1.

* Updated and extended "isar-ref" and "implementation" manual, reduced
remaining material in old "ref" manual.

* Improved support for auxiliary contexts that indicate block structure
for specifications.  Nesting of "context fixes ... context assumes ..."
and "class ... context ...".

* Attribute "consumes" allows a negative value as well, which is
interpreted relatively to the total number of premises of the rule in
the target context.  This form of declaration is stable when exported
from a nested 'context' with additional assumptions.  It is the
preferred form for definitional packages, notably cases/rules produced
in HOL/inductive and HOL/function.

* More informative error messages for Isar proof commands involving
lazy enumerations (method applications etc.).

* Refined 'help' command to retrieve outer syntax commands according
to name patterns (with clickable results).


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

* Parallel terminal proofs ('by') are enabled by default, likewise
proofs that are built into packages like 'datatype', 'function'.  This
allows to "run ahead" checking the theory specifications on the
surface, while the prover is still crunching on internal
justifications.  Unfinished / cancelled proofs are restarted as
required to complete full proof checking eventually.

* Improved output panel with tooltips, hyperlinks etc. based on the
same Rich_Text_Area as regular Isabelle/jEdit buffers.  Activation of
tooltips leads to some window that supports the same recursively,
which can lead to stacks of tooltips as the semantic document content
is explored.  ESCAPE closes the whole stack, individual windows may be
closed separately, or detached to become independent jEdit dockables.

* Improved support for commands that produce graph output: the text
message contains a clickable area to open a new instance of the graph
browser on demand.

* More robust incremental parsing of outer syntax (partial comments,
malformed symbols).  Changing the balance of open/close quotes and
comment delimiters works more conveniently with unfinished situations
that frequently occur in user interaction.

* More efficient painting and improved reactivity when editing large
files.  More scalable management of formal document content.

* Smarter handling of tracing messages: prover process pauses after
certain number of messages per command transaction, with some user
dialog to stop or continue.  This avoids swamping the front-end with
potentially infinite message streams.

* More plugin options and preferences, based on Isabelle/Scala.  The
jEdit plugin option panel provides access to some Isabelle/Scala
options, including tuning parameters for editor reactivity and color
schemes.

* Dockable window "Symbols" provides some editing support for Isabelle
symbols.

* Dockable window "Monitor" shows ML runtime statistics.  Note that
continuous display of the chart slows down the system.

* Improved editing support for control styles: subscript, superscript,
bold, reset of style -- operating on single symbols or text
selections.  Cf. keyboard shortcuts C+e DOWN/UP/RIGHT/LEFT.

* Actions isabelle.increase-font-size and isabelle.decrease-font-size
adjust the main text area font size, and its derivatives for output,
tooltips etc.  Cf. keyboard shortcuts C-PLUS and C-MINUS, which often
need to be adapted to local keyboard layouts.

* More reactive completion popup by default: use \t (TAB) instead of
\n (NEWLINE) to minimize intrusion into regular flow of editing.  See
also "Plugin Options / SideKick / General / Code Completion Options".

* Implicit check and build dialog of the specified logic session
image.  For example, HOL, HOLCF, HOL-Nominal can be produced on
demand, without bundling big platform-dependent heap images in the
Isabelle distribution.

* Uniform Java 7 platform on Linux, Mac OS X, Windows: recent updates
from Oracle provide better multi-platform experience.  This version is
now bundled exclusively with Isabelle.


*** Pure ***

* Code generation for Haskell: restrict unqualified imports from
Haskell Prelude to a small set of fundamental operations.

* Command 'export_code': relative file names are interpreted
relatively to master directory of current theory rather than the
rather arbitrary current working directory.  INCOMPATIBILITY.

* Discontinued obsolete attribute "COMP".  Potential INCOMPATIBILITY,
use regular rule composition via "OF" / "THEN", or explicit proof
structure instead.  Note that Isabelle/ML provides a variety of
operators like COMP, INCR_COMP, COMP_INCR, which need to be applied
with some care where this is really required.

* Command 'typ' supports an additional variant with explicit sort
constraint, to infer and check the most general type conforming to a
given sort.  Example (in HOL):

  typ "_ * _ * bool * unit" :: finite

* Command 'locale_deps' visualizes all locales and their relations as
a Hasse diagram.


*** HOL ***

* Sledgehammer:

  - Added MaSh relevance filter based on machine-learning; see the
    Sledgehammer manual for details.
  - Polished Isar proofs generated with "isar_proofs" option.
  - Rationalized type encodings ("type_enc" option).
  - Renamed "kill_provers" subcommand to "kill_all".
  - Renamed options:
      isar_proof ~> isar_proofs
      isar_shrink_factor ~> isar_shrink
      max_relevant ~> max_facts
      relevance_thresholds ~> fact_thresholds

* Quickcheck: added an optimisation for equality premises.  It is
switched on by default, and can be switched off by setting the
configuration quickcheck_optimise_equality to false.

* Quotient: only one quotient can be defined by quotient_type
INCOMPATIBILITY.

* Lifting:
  - generation of an abstraction function equation in lift_definition
  - quot_del attribute
  - renamed no_abs_code -> no_code (INCOMPATIBILITY.)

* Simproc "finite_Collect" rewrites set comprehensions into pointfree
expressions.

* Preprocessing of the code generator rewrites set comprehensions into
pointfree expressions.

* The SMT solver Z3 has now by default a restricted set of directly
supported features. For the full set of features (div/mod, nonlinear
arithmetic, datatypes/records) with potential proof reconstruction
failures, enable the configuration option "z3_with_extensions".  Minor
INCOMPATIBILITY.

* Simplified 'typedef' specifications: historical options for implicit
set definition and alternative name have been discontinued.  The
former behavior of "typedef (open) t = A" is now the default, but
written just "typedef t = A".  INCOMPATIBILITY, need to adapt theories
accordingly.

* Removed constant "chars"; prefer "Enum.enum" on type "char"
directly.  INCOMPATIBILITY.

* Moved operation product, sublists and n_lists from theory Enum to
List.  INCOMPATIBILITY.

* Theorem UN_o generalized to SUP_comp.  INCOMPATIBILITY.

* Class "comm_monoid_diff" formalises properties of bounded
subtraction, with natural numbers and multisets as typical instances.

* Added combinator "Option.these" with type "'a option set => 'a set".

* Theory "Transitive_Closure": renamed lemmas

  reflcl_tranclp -> reflclp_tranclp
  rtranclp_reflcl -> rtranclp_reflclp

INCOMPATIBILITY.

* Theory "Rings": renamed lemmas (in class semiring)

  left_distrib ~> distrib_right
  right_distrib ~> distrib_left

INCOMPATIBILITY.

* Generalized the definition of limits:

  - Introduced the predicate filterlim (LIM x F. f x :> G) which
    expresses that when the input values x converge to F then the
    output f x converges to G.

  - Added filters for convergence to positive (at_top) and negative
    infinity (at_bot).

  - Moved infinity in the norm (at_infinity) from
    Multivariate_Analysis to Complex_Main.

  - Removed real_tendsto_inf, it is superseded by "LIM x F. f x :>
    at_top".

INCOMPATIBILITY.

* Theory "Library/Option_ord" provides instantiation of option type to
lattice type classes.

* Theory "Library/Multiset": renamed

    constant fold_mset ~> Multiset.fold
    fact fold_mset_commute ~> fold_mset_comm

INCOMPATIBILITY.

* Renamed theory Library/List_Prefix to Library/Sublist, with related
changes as follows.

  - Renamed constants (and related lemmas)

      prefix ~> prefixeq
      strict_prefix ~> prefix

  - Replaced constant "postfix" by "suffixeq" with swapped argument
    order (i.e., "postfix xs ys" is now "suffixeq ys xs") and dropped
    old infix syntax "xs >>= ys"; use "suffixeq ys xs" instead.
    Renamed lemmas accordingly.

  - Added constant "list_hembeq" for homeomorphic embedding on
    lists. Added abbreviation "sublisteq" for special case
    "list_hembeq (op =)".

  - Theory Library/Sublist no longer provides "order" and "bot" type
    class instances for the prefix order (merely corresponding locale
    interpretations). The type class instances are now in theory
    Library/Prefix_Order.

  - The sublist relation of theory Library/Sublist_Order is now based
    on "Sublist.sublisteq".  Renamed lemmas accordingly:

      le_list_append_le_same_iff ~> Sublist.sublisteq_append_le_same_iff
      le_list_append_mono ~> Sublist.list_hembeq_append_mono
      le_list_below_empty ~> Sublist.list_hembeq_Nil, Sublist.list_hembeq_Nil2
      le_list_Cons_EX ~> Sublist.list_hembeq_ConsD
      le_list_drop_Cons2 ~> Sublist.sublisteq_Cons2'
      le_list_drop_Cons_neq ~> Sublist.sublisteq_Cons2_neq
      le_list_drop_Cons ~> Sublist.sublisteq_Cons'
      le_list_drop_many ~> Sublist.sublisteq_drop_many
      le_list_filter_left ~> Sublist.sublisteq_filter_left
      le_list_rev_drop_many ~> Sublist.sublisteq_rev_drop_many
      le_list_rev_take_iff ~> Sublist.sublisteq_append
      le_list_same_length ~> Sublist.sublisteq_same_length
      le_list_take_many_iff ~> Sublist.sublisteq_append'
      less_eq_list.drop ~> less_eq_list_drop
      less_eq_list.induct ~> less_eq_list_induct
      not_le_list_length ~> Sublist.not_sublisteq_length

INCOMPATIBILITY.

* New theory Library/Countable_Set.

* Theory Library/Debug and Library/Parallel provide debugging and
parallel execution for code generated towards Isabelle/ML.

* Theory Library/FuncSet: Extended support for Pi and extensional and
introduce the extensional dependent function space "PiE". Replaced
extensional_funcset by an abbreviation, and renamed lemmas from
extensional_funcset to PiE as follows:

  extensional_empty  ~>  PiE_empty
  extensional_funcset_empty_domain  ~>  PiE_empty_domain
  extensional_funcset_empty_range  ~>  PiE_empty_range
  extensional_funcset_arb  ~>  PiE_arb
  extensional_funcset_mem  ~>  PiE_mem
  extensional_funcset_extend_domainI  ~>  PiE_fun_upd
  extensional_funcset_restrict_domain  ~>  fun_upd_in_PiE
  extensional_funcset_extend_domain_eq  ~>  PiE_insert_eq
  card_extensional_funcset  ~>  card_PiE
  finite_extensional_funcset  ~>  finite_PiE

INCOMPATIBILITY.

* Theory Library/FinFun: theory of almost everywhere constant
functions (supersedes the AFP entry "Code Generation for Functions as
Data").

* Theory Library/Phantom: generic phantom type to make a type
parameter appear in a constant's type.  This alternative to adding
TYPE('a) as another parameter avoids unnecessary closures in generated
code.

* Theory Library/RBT_Impl: efficient construction of red-black trees
from sorted associative lists. Merging two trees with rbt_union may
return a structurally different tree than before.  Potential
INCOMPATIBILITY.

* Theory Library/IArray: immutable arrays with code generation.

* Theory Library/Finite_Lattice: theory of finite lattices.

* HOL/Multivariate_Analysis: replaced

  "basis :: 'a::euclidean_space => nat => real"
  "\<Chi>\<Chi> :: (nat => real) => 'a::euclidean_space"

on euclidean spaces by using the inner product "_ \<bullet> _" with
vectors from the Basis set: "\<Chi>\<Chi> i. f i" is superseded by
"SUM i : Basis. f i * r i".

  With this change the following constants are also changed or removed:

    DIM('a) :: nat  ~>  card (Basis :: 'a set)   (is an abbreviation)
    a $$ i  ~>  inner a i  (where i : Basis)
    cart_base i  removed
    \<pi>, \<pi>'  removed

  Theorems about these constants where removed.

  Renamed lemmas:

    component_le_norm  ~>  Basis_le_norm
    euclidean_eq  ~>  euclidean_eq_iff
    differential_zero_maxmin_component  ~>  differential_zero_maxmin_cart
    euclidean_simps  ~>  inner_simps
    independent_basis  ~>  independent_Basis
    span_basis  ~>  span_Basis
    in_span_basis  ~>  in_span_Basis
    norm_bound_component_le  ~>  norm_boound_Basis_le
    norm_bound_component_lt  ~>  norm_boound_Basis_lt
    component_le_infnorm  ~>  Basis_le_infnorm

INCOMPATIBILITY.

* HOL/Probability:

  - Added simproc "measurable" to automatically prove measurability.

  - Added induction rules for sigma sets with disjoint union
    (sigma_sets_induct_disjoint) and for Borel-measurable functions
    (borel_measurable_induct).

  - Added the Daniell-Kolmogorov theorem (the existence the limit of a
    projective family).

* HOL/Cardinals: Theories of ordinals and cardinals (supersedes the
AFP entry "Ordinals_and_Cardinals").

* HOL/BNF: New (co)datatype package based on bounded natural functors
with support for mixed, nested recursion and interesting non-free
datatypes.

* HOL/Finite_Set and Relation: added new set and relation operations
expressed by Finite_Set.fold.

* New theory HOL/Library/RBT_Set: implementation of sets by red-black
trees for the code generator.

* HOL/Library/RBT and HOL/Library/Mapping have been converted to
Lifting/Transfer.
possible INCOMPATIBILITY.

* HOL/Set: renamed Set.project -> Set.filter
INCOMPATIBILITY.


*** Document preparation ***

* Dropped legacy antiquotations "term_style" and "thm_style", since
styles may be given as arguments to "term" and "thm" already.
Discontinued legacy styles "prem1" .. "prem19".

* Default LaTeX rendering for \<euro> is now based on eurosym package,
instead of slightly exotic babel/greek.

* Document variant NAME may use different LaTeX entry point
document/root_NAME.tex if that file exists, instead of the common
document/root.tex.

* Simplified custom document/build script, instead of old-style
document/IsaMakefile.  Minor INCOMPATIBILITY.


*** ML ***

* The default limit for maximum number of worker threads is now 8,
instead of 4, in correspondence to capabilities of contemporary
hardware and Poly/ML runtime system.

* Type Seq.results and related operations support embedded error
messages within lazy enumerations, and thus allow to provide
informative errors in the absence of any usable results.

* Renamed Position.str_of to Position.here to emphasize that this is a
formal device to inline positions into message text, but not
necessarily printing visible text.


*** System ***

* Advanced support for Isabelle sessions and build management, see
"system" manual for the chapter of that name, especially the "isabelle
build" tool and its examples.  The "isabelle mkroot" tool prepares
session root directories for use with "isabelle build", similar to
former "isabelle mkdir" for "isabelle usedir".  Note that this affects
document preparation as well.  INCOMPATIBILITY, isabelle usedir /
mkdir / make are rendered obsolete.

* Discontinued obsolete Isabelle/build script, it is superseded by the
regular isabelle build tool.  For example:

  isabelle build -s -b HOL

* Discontinued obsolete "isabelle makeall".

* Discontinued obsolete IsaMakefile and ROOT.ML files from the
Isabelle distribution, except for rudimentary src/HOL/IsaMakefile that
provides some traditional targets that invoke "isabelle build".  Note
that this is inefficient!  Applications of Isabelle/HOL involving
"isabelle make" should be upgraded to use "isabelle build" directly.

* The "isabelle options" tool prints Isabelle system options, as
required for "isabelle build", for example.

* The "isabelle logo" tool produces EPS and PDF format simultaneously.
Minor INCOMPATIBILITY in command-line options.

* The "isabelle install" tool has now a simpler command-line.  Minor
INCOMPATIBILITY.

* The "isabelle components" tool helps to resolve add-on components
that are not bundled, or referenced from a bare-bones repository
version of Isabelle.

* Settings variable ISABELLE_PLATFORM_FAMILY refers to the general
platform family: "linux", "macos", "windows".

* The ML system is configured as regular component, and no longer
picked up from some surrounding directory.  Potential INCOMPATIBILITY
for home-made settings.

* Improved ML runtime statistics (heap, threads, future tasks etc.).

* Discontinued support for Poly/ML 5.2.1, which was the last version
without exception positions and advanced ML compiler/toplevel
configuration.

* Discontinued special treatment of Proof General -- no longer guess
PROOFGENERAL_HOME based on accidental file-system layout.  Minor
INCOMPATIBILITY: provide PROOFGENERAL_HOME and PROOFGENERAL_OPTIONS
settings manually, or use a Proof General version that has been
bundled as Isabelle component.



New in Isabelle2012 (May 2012)
------------------------------

*** General ***

* Prover IDE (PIDE) improvements:

  - more robust Sledgehammer integration (as before the sledgehammer
    command-line needs to be typed into the source buffer)
  - markup for bound variables
  - markup for types of term variables (displayed as tooltips)
  - support for user-defined Isar commands within the running session
  - improved support for Unicode outside original 16bit range
    e.g. glyph for \<A> (thanks to jEdit 4.5.1)

* Forward declaration of outer syntax keywords within the theory
header -- minor INCOMPATIBILITY for user-defined commands.  Allow new
commands to be used in the same theory where defined.

* Auxiliary contexts indicate block structure for specifications with
additional parameters and assumptions.  Such unnamed contexts may be
nested within other targets, like 'theory', 'locale', 'class',
'instantiation' etc.  Results from the local context are generalized
accordingly and applied to the enclosing target context.  Example:

  context
    fixes x y z :: 'a
    assumes xy: "x = y" and yz: "y = z"
  begin

  lemma my_trans: "x = z" using xy yz by simp

  end

  thm my_trans

The most basic application is to factor-out context elements of
several fixes/assumes/shows theorem statements, e.g. see
~~/src/HOL/Isar_Examples/Group_Context.thy

Any other local theory specification element works within the "context
... begin ... end" block as well.

* Bundled declarations associate attributed fact expressions with a
given name in the context.  These may be later included in other
contexts.  This allows to manage context extensions casually, without
the logical dependencies of locales and locale interpretation.  See
commands 'bundle', 'include', 'including' etc. in the isar-ref manual.

* Commands 'lemmas' and 'theorems' allow local variables using 'for'
declaration, and results are standardized before being stored.  Thus
old-style "standard" after instantiation or composition of facts
becomes obsolete.  Minor INCOMPATIBILITY, due to potential change of
indices of schematic variables.

* Rule attributes in local theory declarations (e.g. locale or class)
are now statically evaluated: the resulting theorem is stored instead
of the original expression.  INCOMPATIBILITY in rare situations, where
the historic accident of dynamic re-evaluation in interpretations
etc. was exploited.

* New tutorial "Programming and Proving in Isabelle/HOL"
("prog-prove").  It completely supersedes "A Tutorial Introduction to
Structured Isar Proofs" ("isar-overview"), which has been removed.  It
also supersedes "Isabelle/HOL, A Proof Assistant for Higher-Order
Logic" as the recommended beginners tutorial, but does not cover all
of the material of that old tutorial.

* Updated and extended reference manuals: "isar-ref",
"implementation", "system"; reduced remaining material in old "ref"
manual.


*** Pure ***

* Command 'definition' no longer exports the foundational "raw_def"
into the user context.  Minor INCOMPATIBILITY, may use the regular
"def" result with attribute "abs_def" to imitate the old version.

* Attribute "abs_def" turns an equation of the form "f x y == t" into
"f == %x y. t", which ensures that "simp" or "unfold" steps always
expand it.  This also works for object-logic equality.  (Formerly
undocumented feature.)

* Sort constraints are now propagated in simultaneous statements, just
like type constraints.  INCOMPATIBILITY in rare situations, where
distinct sorts used to be assigned accidentally.  For example:

  lemma "P (x::'a::foo)" and "Q (y::'a::bar)"  -- "now illegal"

  lemma "P (x::'a)" and "Q (y::'a::bar)"
    -- "now uniform 'a::bar instead of default sort for first occurrence (!)"

* Rule composition via attribute "OF" (or ML functions OF/MRS) is more
tolerant against multiple unifiers, as long as the final result is
unique.  (As before, rules are composed in canonical right-to-left
order to accommodate newly introduced premises.)

* Renamed some inner syntax categories:

    num ~> num_token
    xnum ~> xnum_token
    xstr ~> str_token

Minor INCOMPATIBILITY.  Note that in practice "num_const" or
"num_position" etc. are mainly used instead (which also include
position information via constraints).

* Simplified configuration options for syntax ambiguity: see
"syntax_ambiguity_warning" and "syntax_ambiguity_limit" in isar-ref
manual.  Minor INCOMPATIBILITY.

* Discontinued configuration option "syntax_positions": atomic terms
in parse trees are always annotated by position constraints.

* Old code generator for SML and its commands 'code_module',
'code_library', 'consts_code', 'types_code' have been discontinued.
Use commands of the generic code generator instead.  INCOMPATIBILITY.

* Redundant attribute "code_inline" has been discontinued. Use
"code_unfold" instead.  INCOMPATIBILITY.

* Dropped attribute "code_unfold_post" in favor of the its dual
"code_abbrev", which yields a common pattern in definitions like

  definition [code_abbrev]: "f = t"

INCOMPATIBILITY.

* Obsolete 'types' command has been discontinued.  Use 'type_synonym'
instead.  INCOMPATIBILITY.

* Discontinued old "prems" fact, which used to refer to the accidental
collection of foundational premises in the context (already marked as
legacy since Isabelle2011).


*** HOL ***

* Type 'a set is now a proper type constructor (just as before
Isabelle2008).  Definitions mem_def and Collect_def have disappeared.
Non-trivial INCOMPATIBILITY.  For developments keeping predicates and
sets separate, it is often sufficient to rephrase some set S that has
been accidentally used as predicates by "%x. x : S", and some
predicate P that has been accidentally used as set by "{x. P x}".
Corresponding proofs in a first step should be pruned from any
tinkering with former theorems mem_def and Collect_def as far as
possible.

For developments which deliberately mix predicates and sets, a
planning step is necessary to determine what should become a predicate
and what a set.  It can be helpful to carry out that step in
Isabelle2011-1 before jumping right into the current release.

* Code generation by default implements sets as container type rather
than predicates.  INCOMPATIBILITY.

* New type synonym 'a rel = ('a * 'a) set

* The representation of numerals has changed.  Datatype "num"
represents strictly positive binary numerals, along with functions
"numeral :: num => 'a" and "neg_numeral :: num => 'a" to represent
positive and negated numeric literals, respectively.  See also
definitions in ~~/src/HOL/Num.thy.  Potential INCOMPATIBILITY, some
user theories may require adaptations as follows:

  - Theorems with number_ring or number_semiring constraints: These
    classes are gone; use comm_ring_1 or comm_semiring_1 instead.

  - Theories defining numeric types: Remove number, number_semiring,
    and number_ring instances. Defer all theorems about numerals until
    after classes one and semigroup_add have been instantiated.

  - Numeral-only simp rules: Replace each rule having a "number_of v"
    pattern with two copies, one for numeral and one for neg_numeral.

  - Theorems about subclasses of semiring_1 or ring_1: These classes
    automatically support numerals now, so more simp rules and
    simprocs may now apply within the proof.

  - Definitions and theorems using old constructors Pls/Min/Bit0/Bit1:
    Redefine using other integer operations.

* Transfer: New package intended to generalize the existing
"descending" method and related theorem attributes from the Quotient
package.  (Not all functionality is implemented yet, but future
development will focus on Transfer as an eventual replacement for the
corresponding parts of the Quotient package.)

  - transfer_rule attribute: Maintains a collection of transfer rules,
    which relate constants at two different types. Transfer rules may
    relate different type instances of the same polymorphic constant,
    or they may relate an operation on a raw type to a corresponding
    operation on an abstract type (quotient or subtype). For example:

    ((A ===> B) ===> list_all2 A ===> list_all2 B) map map
    (cr_int ===> cr_int ===> cr_int) (%(x,y) (u,v). (x+u, y+v)) plus_int

  - transfer method: Replaces a subgoal on abstract types with an
    equivalent subgoal on the corresponding raw types. Constants are
    replaced with corresponding ones according to the transfer rules.
    Goals are generalized over all free variables by default; this is
    necessary for variables whose types change, but can be overridden
    for specific variables with e.g. "transfer fixing: x y z".  The
    variant transfer' method allows replacing a subgoal with one that
    is logically stronger (rather than equivalent).

  - relator_eq attribute: Collects identity laws for relators of
    various type constructors, e.g. "list_all2 (op =) = (op =)".  The
    transfer method uses these lemmas to infer transfer rules for
    non-polymorphic constants on the fly.

  - transfer_prover method: Assists with proving a transfer rule for a
    new constant, provided the constant is defined in terms of other
    constants that already have transfer rules. It should be applied
    after unfolding the constant definitions.

  - HOL/ex/Transfer_Int_Nat.thy: Example theory demonstrating transfer
    from type nat to type int.

* Lifting: New package intended to generalize the quotient_definition
facility of the Quotient package; designed to work with Transfer.

  - lift_definition command: Defines operations on an abstract type in
    terms of a corresponding operation on a representation
    type.  Example syntax:

    lift_definition dlist_insert :: "'a => 'a dlist => 'a dlist"
      is List.insert

    Users must discharge a respectfulness proof obligation when each
    constant is defined. (For a type copy, i.e. a typedef with UNIV,
    the proof is discharged automatically.) The obligation is
    presented in a user-friendly, readable form; a respectfulness
    theorem in the standard format and a transfer rule are generated
    by the package.

  - Integration with code_abstype: For typedefs (e.g. subtypes
    corresponding to a datatype invariant, such as dlist),
    lift_definition generates a code certificate theorem and sets up
    code generation for each constant.

  - setup_lifting command: Sets up the Lifting package to work with a
    user-defined type. The user must provide either a quotient theorem
    or a type_definition theorem.  The package configures transfer
    rules for equality and quantifiers on the type, and sets up the
    lift_definition command to work with the type.

  - Usage examples: See Quotient_Examples/Lift_DList.thy,
    Quotient_Examples/Lift_RBT.thy, Quotient_Examples/Lift_FSet.thy,
    Word/Word.thy and Library/Float.thy.

* Quotient package:

  - The 'quotient_type' command now supports a 'morphisms' option with
    rep and abs functions, similar to typedef.

  - 'quotient_type' sets up new types to work with the Lifting and
    Transfer packages, as with 'setup_lifting'.

  - The 'quotient_definition' command now requires the user to prove a
    respectfulness property at the point where the constant is
    defined, similar to lift_definition; INCOMPATIBILITY.

  - Renamed predicate 'Quotient' to 'Quotient3', and renamed theorems
    accordingly, INCOMPATIBILITY.

* New diagnostic command 'find_unused_assms' to find potentially
superfluous assumptions in theorems using Quickcheck.

* Quickcheck:

  - Quickcheck returns variable assignments as counterexamples, which
    allows to reveal the underspecification of functions under test.
    For example, refuting "hd xs = x", it presents the variable
    assignment xs = [] and x = a1 as a counterexample, assuming that
    any property is false whenever "hd []" occurs in it.

    These counterexample are marked as potentially spurious, as
    Quickcheck also returns "xs = []" as a counterexample to the
    obvious theorem "hd xs = hd xs".

    After finding a potentially spurious counterexample, Quickcheck
    continues searching for genuine ones.

    By default, Quickcheck shows potentially spurious and genuine
    counterexamples. The option "genuine_only" sets quickcheck to only
    show genuine counterexamples.

  - The command 'quickcheck_generator' creates random and exhaustive
    value generators for a given type and operations.

    It generates values by using the operations as if they were
    constructors of that type.

  - Support for multisets.

  - Added "use_subtype" options.

  - Added "quickcheck_locale" configuration to specify how to process
    conjectures in a locale context.

* Nitpick: Fixed infinite loop caused by the 'peephole_optim' option
and affecting 'rat' and 'real'.

* Sledgehammer:
  - Integrated more tightly with SPASS, as described in the ITP 2012
    paper "More SPASS with Isabelle".
  - Made it try "smt" as a fallback if "metis" fails or times out.
  - Added support for the following provers: Alt-Ergo (via Why3 and
    TFF1), iProver, iProver-Eq.
  - Sped up the minimizer.
  - Added "lam_trans", "uncurry_aliases", and "minimize" options.
  - Renamed "slicing" ("no_slicing") option to "slice" ("dont_slice").
  - Renamed "sound" option to "strict".

* Metis: Added possibility to specify lambda translations scheme as a
parenthesized argument (e.g., "by (metis (lifting) ...)").

* SMT: Renamed "smt_fixed" option to "smt_read_only_certificates".

* Command 'try0': Renamed from 'try_methods'. INCOMPATIBILITY.

* New "case_product" attribute to generate a case rule doing multiple
case distinctions at the same time.  E.g.

  list.exhaust [case_product nat.exhaust]

produces a rule which can be used to perform case distinction on both
a list and a nat.

* New "eventually_elim" method as a generalized variant of the
eventually_elim* rules.  Supports structured proofs.

* Typedef with implicit set definition is considered legacy.  Use
"typedef (open)" form instead, which will eventually become the
default.

* Record: code generation can be switched off manually with

  declare [[record_coden = false]]  -- "default true"

* Datatype: type parameters allow explicit sort constraints.

* Concrete syntax for case expressions includes constraints for source
positions, and thus produces Prover IDE markup for its bindings.
INCOMPATIBILITY for old-style syntax translations that augment the
pattern notation; e.g. see src/HOL/HOLCF/One.thy for translations of
one_case.

* Clarified attribute "mono_set": pure declaration without modifying
the result of the fact expression.

* More default pred/set conversions on a couple of relation operations
and predicates.  Added powers of predicate relations.  Consolidation
of some relation theorems:

  converse_def ~> converse_unfold
  rel_comp_def ~> relcomp_unfold
  symp_def ~> (modified, use symp_def and sym_def instead)
  transp_def ~> transp_trans
  Domain_def ~> Domain_unfold
  Range_def ~> Domain_converse [symmetric]

Generalized theorems INF_INT_eq, INF_INT_eq2, SUP_UN_eq, SUP_UN_eq2.

See theory "Relation" for examples for making use of pred/set
conversions by means of attributes "to_set" and "to_pred".

INCOMPATIBILITY.

* Renamed facts about the power operation on relations, i.e., relpow
to match the constant's name:

  rel_pow_1 ~> relpow_1
  rel_pow_0_I ~> relpow_0_I
  rel_pow_Suc_I ~> relpow_Suc_I
  rel_pow_Suc_I2 ~> relpow_Suc_I2
  rel_pow_0_E ~> relpow_0_E
  rel_pow_Suc_E ~> relpow_Suc_E
  rel_pow_E ~> relpow_E
  rel_pow_Suc_D2 ~> relpow_Suc_D2
  rel_pow_Suc_E2 ~> relpow_Suc_E2
  rel_pow_Suc_D2' ~> relpow_Suc_D2'
  rel_pow_E2 ~> relpow_E2
  rel_pow_add ~> relpow_add
  rel_pow_commute ~> relpow
  rel_pow_empty ~> relpow_empty:
  rtrancl_imp_UN_rel_pow ~> rtrancl_imp_UN_relpow
  rel_pow_imp_rtrancl ~> relpow_imp_rtrancl
  rtrancl_is_UN_rel_pow ~> rtrancl_is_UN_relpow
  rtrancl_imp_rel_pow ~> rtrancl_imp_relpow
  rel_pow_fun_conv ~> relpow_fun_conv
  rel_pow_finite_bounded1 ~> relpow_finite_bounded1
  rel_pow_finite_bounded ~> relpow_finite_bounded
  rtrancl_finite_eq_rel_pow ~> rtrancl_finite_eq_relpow
  trancl_finite_eq_rel_pow ~> trancl_finite_eq_relpow
  single_valued_rel_pow ~> single_valued_relpow

INCOMPATIBILITY.

* Theory Relation: Consolidated constant name for relation composition
and corresponding theorem names:

  - Renamed constant rel_comp to relcomp.

  - Dropped abbreviation pred_comp. Use relcompp instead.

  - Renamed theorems:

    rel_compI ~> relcompI
    rel_compEpair ~> relcompEpair
    rel_compE ~> relcompE
    pred_comp_rel_comp_eq ~> relcompp_relcomp_eq
    rel_comp_empty1 ~> relcomp_empty1
    rel_comp_mono ~> relcomp_mono
    rel_comp_subset_Sigma ~> relcomp_subset_Sigma
    rel_comp_distrib ~> relcomp_distrib
    rel_comp_distrib2 ~> relcomp_distrib2
    rel_comp_UNION_distrib ~> relcomp_UNION_distrib
    rel_comp_UNION_distrib2 ~> relcomp_UNION_distrib2
    single_valued_rel_comp ~> single_valued_relcomp
    rel_comp_def ~> relcomp_unfold
    converse_rel_comp ~> converse_relcomp
    pred_compI ~> relcomppI
    pred_compE ~> relcomppE
    pred_comp_bot1 ~> relcompp_bot1
    pred_comp_bot2 ~> relcompp_bot2
    transp_pred_comp_less_eq ~> transp_relcompp_less_eq
    pred_comp_mono ~> relcompp_mono
    pred_comp_distrib ~> relcompp_distrib
    pred_comp_distrib2 ~> relcompp_distrib2
    converse_pred_comp ~> converse_relcompp

    finite_rel_comp ~> finite_relcomp

    set_rel_comp ~> set_relcomp

INCOMPATIBILITY.

* Theory Divides: Discontinued redundant theorems about div and mod.
INCOMPATIBILITY, use the corresponding generic theorems instead.

  DIVISION_BY_ZERO ~> div_by_0, mod_by_0
  zdiv_self ~> div_self
  zmod_self ~> mod_self
  zdiv_zero ~> div_0
  zmod_zero ~> mod_0
  zdiv_zmod_equality ~> div_mod_equality2
  zdiv_zmod_equality2 ~> div_mod_equality
  zmod_zdiv_trivial ~> mod_div_trivial
  zdiv_zminus_zminus ~> div_minus_minus
  zmod_zminus_zminus ~> mod_minus_minus
  zdiv_zminus2 ~> div_minus_right
  zmod_zminus2 ~> mod_minus_right
  zdiv_minus1_right ~> div_minus1_right
  zmod_minus1_right ~> mod_minus1_right
  zdvd_mult_div_cancel ~> dvd_mult_div_cancel
  zmod_zmult1_eq ~> mod_mult_right_eq
  zpower_zmod ~> power_mod
  zdvd_zmod ~> dvd_mod
  zdvd_zmod_imp_zdvd ~> dvd_mod_imp_dvd
  mod_mult_distrib ~> mult_mod_left
  mod_mult_distrib2 ~> mult_mod_right

* Removed redundant theorems nat_mult_2 and nat_mult_2_right; use
generic mult_2 and mult_2_right instead. INCOMPATIBILITY.

* Finite_Set.fold now qualified.  INCOMPATIBILITY.

* Consolidated theorem names concerning fold combinators:

  inf_INFI_fold_inf ~> inf_INF_fold_inf
  sup_SUPR_fold_sup ~> sup_SUP_fold_sup
  INFI_fold_inf ~> INF_fold_inf
  SUPR_fold_sup ~> SUP_fold_sup
  union_set ~> union_set_fold
  minus_set ~> minus_set_fold
  INFI_set_fold ~> INF_set_fold
  SUPR_set_fold ~> SUP_set_fold
  INF_code ~> INF_set_foldr
  SUP_code ~> SUP_set_foldr
  foldr.simps ~> foldr.simps (in point-free formulation)
  foldr_fold_rev ~> foldr_conv_fold
  foldl_fold ~> foldl_conv_fold
  foldr_foldr ~> foldr_conv_foldl
  foldl_foldr ~> foldl_conv_foldr
  fold_set_remdups ~> fold_set_fold_remdups
  fold_set ~> fold_set_fold
  fold1_set ~> fold1_set_fold

INCOMPATIBILITY.

* Dropped rarely useful theorems concerning fold combinators:
foldl_apply, foldl_fun_comm, foldl_rev, fold_weak_invariant,
rev_foldl_cons, fold_set_remdups, fold_set, fold_set1,
concat_conv_foldl, foldl_weak_invariant, foldl_invariant,
foldr_invariant, foldl_absorb0, foldl_foldr1_lemma, foldl_foldr1,
listsum_conv_fold, listsum_foldl, sort_foldl_insort, foldl_assoc,
foldr_conv_foldl, start_le_sum, elem_le_sum, sum_eq_0_conv.
INCOMPATIBILITY.  For the common phrases "%xs. List.foldr plus xs 0"
and "List.foldl plus 0", prefer "List.listsum".  Otherwise it can be
useful to boil down "List.foldr" and "List.foldl" to "List.fold" by
unfolding "foldr_conv_fold" and "foldl_conv_fold".

* Dropped lemmas minus_set_foldr, union_set_foldr, union_coset_foldr,
inter_coset_foldr, Inf_fin_set_foldr, Sup_fin_set_foldr,
Min_fin_set_foldr, Max_fin_set_foldr, Inf_set_foldr, Sup_set_foldr,
INF_set_foldr, SUP_set_foldr.  INCOMPATIBILITY.  Prefer corresponding
lemmas over fold rather than foldr, or make use of lemmas
fold_conv_foldr and fold_rev.

* Congruence rules Option.map_cong and Option.bind_cong for recursion
through option types.

* "Transitive_Closure.ntrancl": bounded transitive closure on
relations.

* Constant "Set.not_member" now qualified.  INCOMPATIBILITY.

* Theory Int: Discontinued many legacy theorems specific to type int.
INCOMPATIBILITY, use the corresponding generic theorems instead.

  zminus_zminus ~> minus_minus
  zminus_0 ~> minus_zero
  zminus_zadd_distrib ~> minus_add_distrib
  zadd_commute ~> add_commute
  zadd_assoc ~> add_assoc
  zadd_left_commute ~> add_left_commute
  zadd_ac ~> add_ac
  zmult_ac ~> mult_ac
  zadd_0 ~> add_0_left
  zadd_0_right ~> add_0_right
  zadd_zminus_inverse2 ~> left_minus
  zmult_zminus ~> mult_minus_left
  zmult_commute ~> mult_commute
  zmult_assoc ~> mult_assoc
  zadd_zmult_distrib ~> left_distrib
  zadd_zmult_distrib2 ~> right_distrib
  zdiff_zmult_distrib ~> left_diff_distrib
  zdiff_zmult_distrib2 ~> right_diff_distrib
  zmult_1 ~> mult_1_left
  zmult_1_right ~> mult_1_right
  zle_refl ~> order_refl
  zle_trans ~> order_trans
  zle_antisym ~> order_antisym
  zle_linear ~> linorder_linear
  zless_linear ~> linorder_less_linear
  zadd_left_mono ~> add_left_mono
  zadd_strict_right_mono ~> add_strict_right_mono
  zadd_zless_mono ~> add_less_le_mono
  int_0_less_1 ~> zero_less_one
  int_0_neq_1 ~> zero_neq_one
  zless_le ~> less_le
  zpower_zadd_distrib ~> power_add
  zero_less_zpower_abs_iff ~> zero_less_power_abs_iff
  zero_le_zpower_abs ~> zero_le_power_abs

* Theory Deriv: Renamed

  DERIV_nonneg_imp_nonincreasing ~> DERIV_nonneg_imp_nondecreasing

* Theory Library/Multiset: Improved code generation of multisets.

* Theory HOL/Library/Set_Algebras: Addition and multiplication on sets
are expressed via type classes again. The special syntax
\<oplus>/\<otimes> has been replaced by plain +/*. Removed constant
setsum_set, which is now subsumed by Big_Operators.setsum.
INCOMPATIBILITY.

* Theory HOL/Library/Diagonalize has been removed. INCOMPATIBILITY,
use theory HOL/Library/Nat_Bijection instead.

* Theory HOL/Library/RBT_Impl: Backing implementation of red-black
trees is now inside a type class context.  Names of affected
operations and lemmas have been prefixed by rbt_.  INCOMPATIBILITY for
theories working directly with raw red-black trees, adapt the names as
follows:

  Operations:
  bulkload -> rbt_bulkload
  del_from_left -> rbt_del_from_left
  del_from_right -> rbt_del_from_right
  del -> rbt_del
  delete -> rbt_delete
  ins -> rbt_ins
  insert -> rbt_insert
  insertw -> rbt_insert_with
  insert_with_key -> rbt_insert_with_key
  map_entry -> rbt_map_entry
  lookup -> rbt_lookup
  sorted -> rbt_sorted
  tree_greater -> rbt_greater
  tree_less -> rbt_less
  tree_less_symbol -> rbt_less_symbol
  union -> rbt_union
  union_with -> rbt_union_with
  union_with_key -> rbt_union_with_key

  Lemmas:
  balance_left_sorted -> balance_left_rbt_sorted
  balance_left_tree_greater -> balance_left_rbt_greater
  balance_left_tree_less -> balance_left_rbt_less
  balance_right_sorted -> balance_right_rbt_sorted
  balance_right_tree_greater -> balance_right_rbt_greater
  balance_right_tree_less -> balance_right_rbt_less
  balance_sorted -> balance_rbt_sorted
  balance_tree_greater -> balance_rbt_greater
  balance_tree_less -> balance_rbt_less
  bulkload_is_rbt -> rbt_bulkload_is_rbt
  combine_sorted -> combine_rbt_sorted
  combine_tree_greater -> combine_rbt_greater
  combine_tree_less -> combine_rbt_less
  delete_in_tree -> rbt_delete_in_tree
  delete_is_rbt -> rbt_delete_is_rbt
  del_from_left_tree_greater -> rbt_del_from_left_rbt_greater
  del_from_left_tree_less -> rbt_del_from_left_rbt_less
  del_from_right_tree_greater -> rbt_del_from_right_rbt_greater
  del_from_right_tree_less -> rbt_del_from_right_rbt_less
  del_in_tree -> rbt_del_in_tree
  del_inv1_inv2 -> rbt_del_inv1_inv2
  del_sorted -> rbt_del_rbt_sorted
  del_tree_greater -> rbt_del_rbt_greater
  del_tree_less -> rbt_del_rbt_less
  dom_lookup_Branch -> dom_rbt_lookup_Branch
  entries_lookup -> entries_rbt_lookup
  finite_dom_lookup -> finite_dom_rbt_lookup
  insert_sorted -> rbt_insert_rbt_sorted
  insertw_is_rbt -> rbt_insertw_is_rbt
  insertwk_is_rbt -> rbt_insertwk_is_rbt
  insertwk_sorted -> rbt_insertwk_rbt_sorted
  insertw_sorted -> rbt_insertw_rbt_sorted
  ins_sorted -> ins_rbt_sorted
  ins_tree_greater -> ins_rbt_greater
  ins_tree_less -> ins_rbt_less
  is_rbt_sorted -> is_rbt_rbt_sorted
  lookup_balance -> rbt_lookup_balance
  lookup_bulkload -> rbt_lookup_rbt_bulkload
  lookup_delete -> rbt_lookup_rbt_delete
  lookup_Empty -> rbt_lookup_Empty
  lookup_from_in_tree -> rbt_lookup_from_in_tree
  lookup_in_tree -> rbt_lookup_in_tree
  lookup_ins -> rbt_lookup_ins
  lookup_insert -> rbt_lookup_rbt_insert
  lookup_insertw -> rbt_lookup_rbt_insertw
  lookup_insertwk -> rbt_lookup_rbt_insertwk
  lookup_keys -> rbt_lookup_keys
  lookup_map -> rbt_lookup_map
  lookup_map_entry -> rbt_lookup_rbt_map_entry
  lookup_tree_greater -> rbt_lookup_rbt_greater
  lookup_tree_less -> rbt_lookup_rbt_less
  lookup_union -> rbt_lookup_rbt_union
  map_entry_color_of -> rbt_map_entry_color_of
  map_entry_inv1 -> rbt_map_entry_inv1
  map_entry_inv2 -> rbt_map_entry_inv2
  map_entry_is_rbt -> rbt_map_entry_is_rbt
  map_entry_sorted -> rbt_map_entry_rbt_sorted
  map_entry_tree_greater -> rbt_map_entry_rbt_greater
  map_entry_tree_less -> rbt_map_entry_rbt_less
  map_tree_greater -> map_rbt_greater
  map_tree_less -> map_rbt_less
  map_sorted -> map_rbt_sorted
  paint_sorted -> paint_rbt_sorted
  paint_lookup -> paint_rbt_lookup
  paint_tree_greater -> paint_rbt_greater
  paint_tree_less -> paint_rbt_less
  sorted_entries -> rbt_sorted_entries
  tree_greater_eq_trans -> rbt_greater_eq_trans
  tree_greater_nit -> rbt_greater_nit
  tree_greater_prop -> rbt_greater_prop
  tree_greater_simps -> rbt_greater_simps
  tree_greater_trans -> rbt_greater_trans
  tree_less_eq_trans -> rbt_less_eq_trans
  tree_less_nit -> rbt_less_nit
  tree_less_prop -> rbt_less_prop
  tree_less_simps -> rbt_less_simps
  tree_less_trans -> rbt_less_trans
  tree_ord_props -> rbt_ord_props
  union_Branch -> rbt_union_Branch
  union_is_rbt -> rbt_union_is_rbt
  unionw_is_rbt -> rbt_unionw_is_rbt
  unionwk_is_rbt -> rbt_unionwk_is_rbt
  unionwk_sorted -> rbt_unionwk_rbt_sorted

* Theory HOL/Library/Float: Floating point numbers are now defined as
a subset of the real numbers.  All operations are defined using the
lifing-framework and proofs use the transfer method.  INCOMPATIBILITY.

  Changed Operations:
  float_abs -> abs
  float_nprt -> nprt
  float_pprt -> pprt
  pow2 -> use powr
  round_down -> float_round_down
  round_up -> float_round_up
  scale -> exponent

  Removed Operations:
  ceiling_fl, lb_mult, lb_mod, ub_mult, ub_mod

  Renamed Lemmas:
  abs_float_def -> Float.compute_float_abs
  bitlen_ge0 -> bitlen_nonneg
  bitlen.simps -> Float.compute_bitlen
  float_components -> Float_mantissa_exponent
  float_divl.simps -> Float.compute_float_divl
  float_divr.simps -> Float.compute_float_divr
  float_eq_odd -> mult_powr_eq_mult_powr_iff
  float_power -> real_of_float_power
  lapprox_posrat_def -> Float.compute_lapprox_posrat
  lapprox_rat.simps -> Float.compute_lapprox_rat
  le_float_def' -> Float.compute_float_le
  le_float_def -> less_eq_float.rep_eq
  less_float_def' -> Float.compute_float_less
  less_float_def -> less_float.rep_eq
  normfloat_def -> Float.compute_normfloat
  normfloat_imp_odd_or_zero -> mantissa_not_dvd and mantissa_noteq_0
  normfloat -> normfloat_def
  normfloat_unique -> use normfloat_def
  number_of_float_Float -> Float.compute_float_numeral, Float.compute_float_neg_numeral
  one_float_def -> Float.compute_float_one
  plus_float_def -> Float.compute_float_plus
  rapprox_posrat_def -> Float.compute_rapprox_posrat
  rapprox_rat.simps -> Float.compute_rapprox_rat
  real_of_float_0 -> zero_float.rep_eq
  real_of_float_1 -> one_float.rep_eq
  real_of_float_abs -> abs_float.rep_eq
  real_of_float_add -> plus_float.rep_eq
  real_of_float_minus -> uminus_float.rep_eq
  real_of_float_mult -> times_float.rep_eq
  real_of_float_simp -> Float.rep_eq
  real_of_float_sub -> minus_float.rep_eq
  round_down.simps -> Float.compute_float_round_down
  round_up.simps -> Float.compute_float_round_up
  times_float_def -> Float.compute_float_times
  uminus_float_def -> Float.compute_float_uminus
  zero_float_def -> Float.compute_float_zero

  Lemmas not necessary anymore, use the transfer method:
  bitlen_B0, bitlen_B1, bitlen_ge1, bitlen_Min, bitlen_Pls, float_divl,
  float_divr, float_le_simp, float_less1_mantissa_bound,
  float_less_simp, float_less_zero, float_le_zero,
  float_pos_less1_e_neg, float_pos_m_pos, float_split, float_split2,
  floor_pos_exp, lapprox_posrat, lapprox_posrat_bottom, lapprox_rat,
  lapprox_rat_bottom, normalized_float, rapprox_posrat,
  rapprox_posrat_le1, rapprox_rat, real_of_float_ge0_exp,
  real_of_float_neg_exp, real_of_float_nge0_exp, round_down floor_fl,
  round_up, zero_le_float, zero_less_float

* New theory HOL/Library/DAList provides an abstract type for
association lists with distinct keys.

* Session HOL/IMP: Added new theory of abstract interpretation of
annotated commands.

* Session HOL-Import: Re-implementation from scratch is faster,
simpler, and more scalable.  Requires a proof bundle, which is
available as an external component.  Discontinued old (and mostly
dead) Importer for HOL4 and HOL Light.  INCOMPATIBILITY.

* Session HOL-Word: Discontinued many redundant theorems specific to
type 'a word. INCOMPATIBILITY, use the corresponding generic theorems
instead.

  word_sub_alt ~> word_sub_wi
  word_add_alt ~> word_add_def
  word_mult_alt ~> word_mult_def
  word_minus_alt ~> word_minus_def
  word_0_alt ~> word_0_wi
  word_1_alt ~> word_1_wi
  word_add_0 ~> add_0_left
  word_add_0_right ~> add_0_right
  word_mult_1 ~> mult_1_left
  word_mult_1_right ~> mult_1_right
  word_add_commute ~> add_commute
  word_add_assoc ~> add_assoc
  word_add_left_commute ~> add_left_commute
  word_mult_commute ~> mult_commute
  word_mult_assoc ~> mult_assoc
  word_mult_left_commute ~> mult_left_commute
  word_left_distrib ~> left_distrib
  word_right_distrib ~> right_distrib
  word_left_minus ~> left_minus
  word_diff_0_right ~> diff_0_right
  word_diff_self ~> diff_self
  word_sub_def ~> diff_minus
  word_diff_minus ~> diff_minus
  word_add_ac ~> add_ac
  word_mult_ac ~> mult_ac
  word_plus_ac0 ~> add_0_left add_0_right add_ac
  word_times_ac1 ~> mult_1_left mult_1_right mult_ac
  word_order_trans ~> order_trans
  word_order_refl ~> order_refl
  word_order_antisym ~> order_antisym
  word_order_linear ~> linorder_linear
  lenw1_zero_neq_one ~> zero_neq_one
  word_number_of_eq ~> number_of_eq
  word_of_int_add_hom ~> wi_hom_add
  word_of_int_sub_hom ~> wi_hom_sub
  word_of_int_mult_hom ~> wi_hom_mult
  word_of_int_minus_hom ~> wi_hom_neg
  word_of_int_succ_hom ~> wi_hom_succ
  word_of_int_pred_hom ~> wi_hom_pred
  word_of_int_0_hom ~> word_0_wi
  word_of_int_1_hom ~> word_1_wi

* Session HOL-Word: New proof method "word_bitwise" for splitting
machine word equalities and inequalities into logical circuits,
defined in HOL/Word/WordBitwise.thy.  Supports addition, subtraction,
multiplication, shifting by constants, bitwise operators and numeric
constants.  Requires fixed-length word types, not 'a word.  Solves
many standard word identities outright and converts more into first
order problems amenable to blast or similar.  See also examples in
HOL/Word/Examples/WordExamples.thy.

* Session HOL-Probability: Introduced the type "'a measure" to
represent measures, this replaces the records 'a algebra and 'a
measure_space.  The locales based on subset_class now have two
locale-parameters the space \<Omega> and the set of measurable sets M.
The product of probability spaces uses now the same constant as the
finite product of sigma-finite measure spaces "PiM :: ('i => 'a)
measure".  Most constants are defined now outside of locales and gain
an additional parameter, like null_sets, almost_eventually or \<mu>'.
Measure space constructions for distributions and densities now got
their own constants distr and density.  Instead of using locales to
describe measure spaces with a finite space, the measure count_space
and point_measure is introduced.  INCOMPATIBILITY.

  Renamed constants:
  measure -> emeasure
  finite_measure.\<mu>' -> measure
  product_algebra_generator -> prod_algebra
  product_prob_space.emb -> prod_emb
  product_prob_space.infprod_algebra -> PiM

  Removed locales:
  completeable_measure_space
  finite_measure_space
  finite_prob_space
  finite_product_finite_prob_space
  finite_product_sigma_algebra
  finite_sigma_algebra
  measure_space
  pair_finite_prob_space
  pair_finite_sigma_algebra
  pair_finite_space
  pair_sigma_algebra
  product_sigma_algebra

  Removed constants:
  conditional_space
  distribution -> use distr measure, or distributed predicate
  image_space
  joint_distribution -> use distr measure, or distributed predicate
  pair_measure_generator
  product_prob_space.infprod_algebra -> use PiM
  subvimage

  Replacement theorems:
  finite_additivity_sufficient -> ring_of_sets.countably_additiveI_finite
  finite_measure.empty_measure -> measure_empty
  finite_measure.finite_continuity_from_above -> finite_measure.finite_Lim_measure_decseq
  finite_measure.finite_continuity_from_below -> finite_measure.finite_Lim_measure_incseq
  finite_measure.finite_measure_countably_subadditive -> finite_measure.finite_measure_subadditive_countably
  finite_measure.finite_measure_eq -> finite_measure.emeasure_eq_measure
  finite_measure.finite_measure -> finite_measure.emeasure_finite
  finite_measure.finite_measure_finite_singleton -> finite_measure.finite_measure_eq_setsum_singleton
  finite_measure.positive_measure' -> measure_nonneg
  finite_measure.real_measure -> finite_measure.emeasure_real
  finite_product_prob_space.finite_measure_times -> finite_product_prob_space.finite_measure_PiM_emb
  finite_product_sigma_algebra.in_P -> sets_PiM_I_finite
  finite_product_sigma_algebra.P_empty -> space_PiM_empty, sets_PiM_empty
  information_space.conditional_entropy_eq -> information_space.conditional_entropy_simple_distributed
  information_space.conditional_entropy_positive -> information_space.conditional_entropy_nonneg_simple
  information_space.conditional_mutual_information_eq_mutual_information -> information_space.conditional_mutual_information_eq_mutual_information_simple
  information_space.conditional_mutual_information_generic_positive -> information_space.conditional_mutual_information_nonneg_simple
  information_space.conditional_mutual_information_positive -> information_space.conditional_mutual_information_nonneg_simple
  information_space.entropy_commute -> information_space.entropy_commute_simple
  information_space.entropy_eq -> information_space.entropy_simple_distributed
  information_space.entropy_generic_eq -> information_space.entropy_simple_distributed
  information_space.entropy_positive -> information_space.entropy_nonneg_simple
  information_space.entropy_uniform_max -> information_space.entropy_uniform
  information_space.KL_eq_0_imp -> information_space.KL_eq_0_iff_eq
  information_space.KL_eq_0 -> information_space.KL_same_eq_0
  information_space.KL_ge_0 -> information_space.KL_nonneg
  information_space.mutual_information_eq -> information_space.mutual_information_simple_distributed
  information_space.mutual_information_positive -> information_space.mutual_information_nonneg_simple
  Int_stable_cuboids -> Int_stable_atLeastAtMost
  Int_stable_product_algebra_generator -> positive_integral
  measure_preserving -> equality "distr M N f = N" "f : measurable M N"
  measure_space.additive -> emeasure_additive
  measure_space.AE_iff_null_set -> AE_iff_null
  measure_space.almost_everywhere_def -> eventually_ae_filter
  measure_space.almost_everywhere_vimage -> AE_distrD
  measure_space.continuity_from_above -> INF_emeasure_decseq
  measure_space.continuity_from_above_Lim -> Lim_emeasure_decseq
  measure_space.continuity_from_below_Lim -> Lim_emeasure_incseq
  measure_space.continuity_from_below -> SUP_emeasure_incseq
  measure_space_density -> emeasure_density
  measure_space.density_is_absolutely_continuous -> absolutely_continuousI_density
  measure_space.integrable_vimage -> integrable_distr
  measure_space.integral_translated_density -> integral_density
  measure_space.integral_vimage -> integral_distr
  measure_space.measure_additive -> plus_emeasure
  measure_space.measure_compl -> emeasure_compl
  measure_space.measure_countable_increasing -> emeasure_countable_increasing
  measure_space.measure_countably_subadditive -> emeasure_subadditive_countably
  measure_space.measure_decseq -> decseq_emeasure
  measure_space.measure_Diff -> emeasure_Diff
  measure_space.measure_Diff_null_set -> emeasure_Diff_null_set
  measure_space.measure_eq_0 -> emeasure_eq_0
  measure_space.measure_finitely_subadditive -> emeasure_subadditive_finite
  measure_space.measure_finite_singleton -> emeasure_eq_setsum_singleton
  measure_space.measure_incseq -> incseq_emeasure
  measure_space.measure_insert -> emeasure_insert
  measure_space.measure_mono -> emeasure_mono
  measure_space.measure_not_negative -> emeasure_not_MInf
  measure_space.measure_preserving_Int_stable -> measure_eqI_generator_eq
  measure_space.measure_setsum -> setsum_emeasure
  measure_space.measure_setsum_split -> setsum_emeasure_cover
  measure_space.measure_space_vimage -> emeasure_distr
  measure_space.measure_subadditive_finite -> emeasure_subadditive_finite
  measure_space.measure_subadditive -> subadditive
  measure_space.measure_top -> emeasure_space
  measure_space.measure_UN_eq_0 -> emeasure_UN_eq_0
  measure_space.measure_Un_null_set -> emeasure_Un_null_set
  measure_space.positive_integral_translated_density -> positive_integral_density
  measure_space.positive_integral_vimage -> positive_integral_distr
  measure_space.real_continuity_from_above -> Lim_measure_decseq
  measure_space.real_continuity_from_below -> Lim_measure_incseq
  measure_space.real_measure_countably_subadditive -> measure_subadditive_countably
  measure_space.real_measure_Diff -> measure_Diff
  measure_space.real_measure_finite_Union -> measure_finite_Union
  measure_space.real_measure_setsum_singleton -> measure_eq_setsum_singleton
  measure_space.real_measure_subadditive -> measure_subadditive
  measure_space.real_measure_Union -> measure_Union
  measure_space.real_measure_UNION -> measure_UNION
  measure_space.simple_function_vimage -> simple_function_comp
  measure_space.simple_integral_vimage -> simple_integral_distr
  measure_space.simple_integral_vimage -> simple_integral_distr
  measure_unique_Int_stable -> measure_eqI_generator_eq
  measure_unique_Int_stable_vimage -> measure_eqI_generator_eq
  pair_sigma_algebra.measurable_cut_fst -> sets_Pair1
  pair_sigma_algebra.measurable_cut_snd -> sets_Pair2
  pair_sigma_algebra.measurable_pair_image_fst -> measurable_Pair1
  pair_sigma_algebra.measurable_pair_image_snd -> measurable_Pair2
  pair_sigma_algebra.measurable_product_swap -> measurable_pair_swap_iff
  pair_sigma_algebra.pair_sigma_algebra_measurable -> measurable_pair_swap
  pair_sigma_algebra.pair_sigma_algebra_swap_measurable -> measurable_pair_swap'
  pair_sigma_algebra.sets_swap -> sets_pair_swap
  pair_sigma_finite.measure_cut_measurable_fst -> pair_sigma_finite.measurable_emeasure_Pair1
  pair_sigma_finite.measure_cut_measurable_snd -> pair_sigma_finite.measurable_emeasure_Pair2
  pair_sigma_finite.measure_preserving_swap -> pair_sigma_finite.distr_pair_swap
  pair_sigma_finite.pair_measure_alt2 -> pair_sigma_finite.emeasure_pair_measure_alt2
  pair_sigma_finite.pair_measure_alt -> pair_sigma_finite.emeasure_pair_measure_alt
  pair_sigma_finite.pair_measure_times -> pair_sigma_finite.emeasure_pair_measure_Times
  prob_space.indep_distribution_eq_measure -> prob_space.indep_vars_iff_distr_eq_PiM
  prob_space.indep_var_distributionD -> prob_space.indep_var_distribution_eq
  prob_space.measure_space_1 -> prob_space.emeasure_space_1
  prob_space.prob_space_vimage -> prob_space_distr
  prob_space.random_variable_restrict -> measurable_restrict
  prob_space_unique_Int_stable -> measure_eqI_prob_space
  product_algebraE -> prod_algebraE_all
  product_algebra_generator_der -> prod_algebra_eq_finite
  product_algebra_generator_into_space -> prod_algebra_sets_into_space
  product_algebraI -> sets_PiM_I_finite
  product_measure_exists -> product_sigma_finite.sigma_finite
  product_prob_space.finite_index_eq_finite_product -> product_prob_space.sets_PiM_generator
  product_prob_space.finite_measure_infprod_emb_Pi -> product_prob_space.measure_PiM_emb
  product_prob_space.infprod_spec -> product_prob_space.emeasure_PiM_emb_not_empty
  product_prob_space.measurable_component -> measurable_component_singleton
  product_prob_space.measurable_emb -> measurable_prod_emb
  product_prob_space.measurable_into_infprod_algebra -> measurable_PiM_single
  product_prob_space.measurable_singleton_infprod -> measurable_component_singleton
  product_prob_space.measure_emb -> emeasure_prod_emb
  product_prob_space.measure_preserving_restrict -> product_prob_space.distr_restrict
  product_sigma_algebra.product_algebra_into_space -> space_closed
  product_sigma_finite.measure_fold -> product_sigma_finite.distr_merge
  product_sigma_finite.measure_preserving_component_singelton -> product_sigma_finite.distr_singleton
  product_sigma_finite.measure_preserving_merge -> product_sigma_finite.distr_merge
  sequence_space.measure_infprod -> sequence_space.measure_PiM_countable
  sets_product_algebra -> sets_PiM
  sigma_algebra.measurable_sigma -> measurable_measure_of
  sigma_finite_measure.disjoint_sigma_finite -> sigma_finite_disjoint
  sigma_finite_measure.RN_deriv_vimage -> sigma_finite_measure.RN_deriv_distr
  sigma_product_algebra_sigma_eq -> sigma_prod_algebra_sigma_eq
  space_product_algebra -> space_PiM

* Session HOL-TPTP: support to parse and import TPTP problems (all
languages) into Isabelle/HOL.


*** FOL ***

* New "case_product" attribute (see HOL).


*** ZF ***

* Greater support for structured proofs involving induction or case
analysis.

* Much greater use of mathematical symbols.

* Removal of many ML theorem bindings.  INCOMPATIBILITY.


*** ML ***

* Antiquotation @{keyword "name"} produces a parser for outer syntax
from a minor keyword introduced via theory header declaration.

* Antiquotation @{command_spec "name"} produces the
Outer_Syntax.command_spec from a major keyword introduced via theory
header declaration; it can be passed to Outer_Syntax.command etc.

* Local_Theory.define no longer hard-wires default theorem name
"foo_def", but retains the binding as given.  If that is Binding.empty
/ Attrib.empty_binding, the result is not registered as user-level
fact.  The Local_Theory.define_internal variant allows to specify a
non-empty name (used for the foundation in the background theory),
while omitting the fact binding in the user-context.  Potential
INCOMPATIBILITY for derived definitional packages: need to specify
naming policy for primitive definitions more explicitly.

* Renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in
conformance with similar operations in structure Term and Logic.

* Antiquotation @{attributes [...]} embeds attribute source
representation into the ML text, which is particularly useful with
declarations like Local_Theory.note.

* Structure Proof_Context follows standard naming scheme.  Old
ProofContext has been discontinued.  INCOMPATIBILITY.

* Refined Local_Theory.declaration {syntax, pervasive}, with subtle
change of semantics: update is applied to auxiliary local theory
context as well.

* Modernized some old-style infix operations:

  addeqcongs    ~> Simplifier.add_eqcong
  deleqcongs    ~> Simplifier.del_eqcong
  addcongs      ~> Simplifier.add_cong
  delcongs      ~> Simplifier.del_cong
  setmksimps    ~> Simplifier.set_mksimps
  setmkcong     ~> Simplifier.set_mkcong
  setmksym      ~> Simplifier.set_mksym
  setmkeqTrue   ~> Simplifier.set_mkeqTrue
  settermless   ~> Simplifier.set_termless
  setsubgoaler  ~> Simplifier.set_subgoaler
  addsplits     ~> Splitter.add_split
  delsplits     ~> Splitter.del_split


*** System ***

* USER_HOME settings variable points to cross-platform user home
directory, which coincides with HOME on POSIX systems only.  Likewise,
the Isabelle path specification "~" now expands to $USER_HOME, instead
of former $HOME.  A different default for USER_HOME may be set
explicitly in shell environment, before Isabelle settings are
evaluated.  Minor INCOMPATIBILITY: need to adapt Isabelle path where
the generic user home was intended.

* ISABELLE_HOME_WINDOWS refers to ISABELLE_HOME in windows file name
notation, which is useful for the jEdit file browser, for example.

* ISABELLE_JDK_HOME settings variable points to JDK with javac and jar
(not just JRE).



New in Isabelle2011-1 (October 2011)
------------------------------------

*** General ***

* Improved Isabelle/jEdit Prover IDE (PIDE), which can be invoked as
"isabelle jedit" or "ISABELLE_HOME/Isabelle" on the command line.

  - Management of multiple theory files directly from the editor
    buffer store -- bypassing the file-system (no requirement to save
    files for checking).

  - Markup of formal entities within the text buffer, with semantic
    highlighting, tooltips and hyperlinks to jump to defining source
    positions.

  - Improved text rendering, with sub/superscripts in the source
    buffer (including support for copy/paste wrt. output panel, HTML
    theory output and other non-Isabelle text boxes).

  - Refined scheduling of proof checking and printing of results,
    based on interactive editor view.  (Note: jEdit folding and
    narrowing allows to restrict buffer perspectives explicitly.)

  - Reduced CPU performance requirements, usable on machines with few
    cores.

  - Reduced memory requirements due to pruning of unused document
    versions (garbage collection).

See also ~~/src/Tools/jEdit/README.html for further information,
including some remaining limitations.

* Theory loader: source files are exclusively located via the master
directory of each theory node (where the .thy file itself resides).
The global load path (such as src/HOL/Library) has been discontinued.
Note that the path element ~~ may be used to reference theories in the
Isabelle home folder -- for instance, "~~/src/HOL/Library/FuncSet".
INCOMPATIBILITY.

* Theory loader: source files are identified by content via SHA1
digests.  Discontinued former path/modtime identification and optional
ISABELLE_FILE_IDENT plugin scripts.

* Parallelization of nested Isar proofs is subject to
Goal.parallel_proofs_threshold (default 100).  See also isabelle
usedir option -Q.

* Name space: former unsynchronized references are now proper
configuration options, with more conventional names:

  long_names   ~> names_long
  short_names  ~> names_short
  unique_names ~> names_unique

Minor INCOMPATIBILITY, need to declare options in context like this:

  declare [[names_unique = false]]

* Literal facts `prop` may contain dummy patterns, e.g. `_ = _`.  Note
that the result needs to be unique, which means fact specifications
may have to be refined after enriching a proof context.

* Attribute "case_names" has been refined: the assumptions in each case
can be named now by following the case name with [name1 name2 ...].

* Isabelle/Isar reference manual has been updated and extended:
  - "Synopsis" provides a catalog of main Isar language concepts.
  - Formal references in syntax diagrams, via @{rail} antiquotation.
  - Updated material from classic "ref" manual, notably about
    "Classical Reasoner".


*** HOL ***

* Class bot and top require underlying partial order rather than
preorder: uniqueness of bot and top is guaranteed.  INCOMPATIBILITY.

* Class complete_lattice: generalized a couple of lemmas from sets;
generalized theorems INF_cong and SUP_cong.  New type classes for
complete boolean algebras and complete linear orders.  Lemmas
Inf_less_iff, less_Sup_iff, INF_less_iff, less_SUP_iff now reside in
class complete_linorder.

Changed proposition of lemmas Inf_bool_def, Sup_bool_def, Inf_fun_def,
Sup_fun_def, Inf_apply, Sup_apply.

Removed redundant lemmas (the right hand side gives hints how to
replace them for (metis ...), or (simp only: ...) proofs):

  Inf_singleton ~> Inf_insert [where A="{}", unfolded Inf_empty inf_top_right]
  Sup_singleton ~> Sup_insert [where A="{}", unfolded Sup_empty sup_bot_right]
  Inf_binary ~> Inf_insert, Inf_empty, and inf_top_right
  Sup_binary ~> Sup_insert, Sup_empty, and sup_bot_right
  Int_eq_Inter ~> Inf_insert, Inf_empty, and inf_top_right
  Un_eq_Union ~> Sup_insert, Sup_empty, and sup_bot_right
  Inter_def ~> INF_def, image_def
  Union_def ~> SUP_def, image_def
  INT_eq ~> INF_def, and image_def
  UN_eq ~> SUP_def, and image_def
  INF_subset ~> INF_superset_mono [OF _ order_refl]

More consistent and comprehensive names:

  INTER_eq_Inter_image ~> INF_def
  UNION_eq_Union_image ~> SUP_def
  INFI_def ~> INF_def
  SUPR_def ~> SUP_def
  INF_leI ~> INF_lower
  INF_leI2 ~> INF_lower2
  le_INFI ~> INF_greatest
  le_SUPI ~> SUP_upper
  le_SUPI2 ~> SUP_upper2
  SUP_leI ~> SUP_least
  INFI_bool_eq ~> INF_bool_eq
  SUPR_bool_eq ~> SUP_bool_eq
  INFI_apply ~> INF_apply
  SUPR_apply ~> SUP_apply
  INTER_def ~> INTER_eq
  UNION_def ~> UNION_eq

INCOMPATIBILITY.

* Renamed theory Complete_Lattice to Complete_Lattices.
INCOMPATIBILITY.

* Theory Complete_Lattices: lemmas Inf_eq_top_iff, INF_eq_top_iff,
INF_image, Inf_insert, INF_top, Inf_top_conv, INF_top_conv, SUP_bot,
Sup_bot_conv, SUP_bot_conv, Sup_eq_top_iff, SUP_eq_top_iff, SUP_image,
Sup_insert are now declared as [simp].  INCOMPATIBILITY.

* Theory Lattice: lemmas compl_inf_bot, compl_le_comp_iff,
compl_sup_top, inf_idem, inf_left_idem, inf_sup_absorb, sup_idem,
sup_inf_absob, sup_left_idem are now declared as [simp].  Minor
INCOMPATIBILITY.

* Added syntactic classes "inf" and "sup" for the respective
constants.  INCOMPATIBILITY: Changes in the argument order of the
(mostly internal) locale predicates for some derived classes.

* Theorem collections ball_simps and bex_simps do not contain theorems
referring to UNION any longer; these have been moved to collection
UN_ball_bex_simps.  INCOMPATIBILITY.

* Theory Archimedean_Field: floor now is defined as parameter of a
separate type class floor_ceiling.

* Theory Finite_Set: more coherent development of fold_set locales:

    locale fun_left_comm ~> locale comp_fun_commute
    locale fun_left_comm_idem ~> locale comp_fun_idem

Both use point-free characterization; interpretation proofs may need
adjustment.  INCOMPATIBILITY.

* Theory Limits: Type "'a net" has been renamed to "'a filter", in
accordance with standard mathematical terminology. INCOMPATIBILITY.

* Theory Complex_Main: The locale interpretations for the
bounded_linear and bounded_bilinear locales have been removed, in
order to reduce the number of duplicate lemmas. Users must use the
original names for distributivity theorems, potential INCOMPATIBILITY.

  divide.add ~> add_divide_distrib
  divide.diff ~> diff_divide_distrib
  divide.setsum ~> setsum_divide_distrib
  mult.add_right ~> right_distrib
  mult.diff_right ~> right_diff_distrib
  mult_right.setsum ~> setsum_right_distrib
  mult_left.diff ~> left_diff_distrib

* Theory Complex_Main: Several redundant theorems have been removed or
replaced by more general versions. INCOMPATIBILITY.

  real_diff_def ~> minus_real_def
  real_divide_def ~> divide_real_def
  real_less_def ~> less_le
  real_abs_def ~> abs_real_def
  real_sgn_def ~> sgn_real_def
  real_mult_commute ~> mult_commute
  real_mult_assoc ~> mult_assoc
  real_mult_1 ~> mult_1_left
  real_add_mult_distrib ~> left_distrib
  real_zero_not_eq_one ~> zero_neq_one
  real_mult_inverse_left ~> left_inverse
  INVERSE_ZERO ~> inverse_zero
  real_le_refl ~> order_refl
  real_le_antisym ~> order_antisym
  real_le_trans ~> order_trans
  real_le_linear ~> linear
  real_le_eq_diff ~> le_iff_diff_le_0
  real_add_left_mono ~> add_left_mono
  real_mult_order ~> mult_pos_pos
  real_mult_less_mono2 ~> mult_strict_left_mono
  real_of_int_real_of_nat ~> real_of_int_of_nat_eq
  real_0_le_divide_iff ~> zero_le_divide_iff
  realpow_two_disj ~> power2_eq_iff
  real_squared_diff_one_factored ~> square_diff_one_factored
  realpow_two_diff ~> square_diff_square_factored
  reals_complete2 ~> complete_real
  real_sum_squared_expand ~> power2_sum
  exp_ln_eq ~> ln_unique
  expi_add ~> exp_add
  expi_zero ~> exp_zero
  lemma_DERIV_subst ~> DERIV_cong
  LIMSEQ_Zfun_iff ~> tendsto_Zfun_iff
  LIMSEQ_const ~> tendsto_const
  LIMSEQ_norm ~> tendsto_norm
  LIMSEQ_add ~> tendsto_add
  LIMSEQ_minus ~> tendsto_minus
  LIMSEQ_minus_cancel ~> tendsto_minus_cancel
  LIMSEQ_diff ~> tendsto_diff
  bounded_linear.LIMSEQ ~> bounded_linear.tendsto
  bounded_bilinear.LIMSEQ ~> bounded_bilinear.tendsto
  LIMSEQ_mult ~> tendsto_mult
  LIMSEQ_inverse ~> tendsto_inverse
  LIMSEQ_divide ~> tendsto_divide
  LIMSEQ_pow ~> tendsto_power
  LIMSEQ_setsum ~> tendsto_setsum
  LIMSEQ_setprod ~> tendsto_setprod
  LIMSEQ_norm_zero ~> tendsto_norm_zero_iff
  LIMSEQ_rabs_zero ~> tendsto_rabs_zero_iff
  LIMSEQ_imp_rabs ~> tendsto_rabs
  LIMSEQ_add_minus ~> tendsto_add [OF _ tendsto_minus]
  LIMSEQ_add_const ~> tendsto_add [OF _ tendsto_const]
  LIMSEQ_diff_const ~> tendsto_diff [OF _ tendsto_const]
  LIMSEQ_Complex ~> tendsto_Complex
  LIM_ident ~> tendsto_ident_at
  LIM_const ~> tendsto_const
  LIM_add ~> tendsto_add
  LIM_add_zero ~> tendsto_add_zero
  LIM_minus ~> tendsto_minus
  LIM_diff ~> tendsto_diff
  LIM_norm ~> tendsto_norm
  LIM_norm_zero ~> tendsto_norm_zero
  LIM_norm_zero_cancel ~> tendsto_norm_zero_cancel
  LIM_norm_zero_iff ~> tendsto_norm_zero_iff
  LIM_rabs ~> tendsto_rabs
  LIM_rabs_zero ~> tendsto_rabs_zero
  LIM_rabs_zero_cancel ~> tendsto_rabs_zero_cancel
  LIM_rabs_zero_iff ~> tendsto_rabs_zero_iff
  LIM_compose ~> tendsto_compose
  LIM_mult ~> tendsto_mult
  LIM_scaleR ~> tendsto_scaleR
  LIM_of_real ~> tendsto_of_real
  LIM_power ~> tendsto_power
  LIM_inverse ~> tendsto_inverse
  LIM_sgn ~> tendsto_sgn
  isCont_LIM_compose ~> isCont_tendsto_compose
  bounded_linear.LIM ~> bounded_linear.tendsto
  bounded_linear.LIM_zero ~> bounded_linear.tendsto_zero
  bounded_bilinear.LIM ~> bounded_bilinear.tendsto
  bounded_bilinear.LIM_prod_zero ~> bounded_bilinear.tendsto_zero
  bounded_bilinear.LIM_left_zero ~> bounded_bilinear.tendsto_left_zero
  bounded_bilinear.LIM_right_zero ~> bounded_bilinear.tendsto_right_zero
  LIM_inverse_fun ~> tendsto_inverse [OF tendsto_ident_at]

* Theory Complex_Main: The definition of infinite series was
generalized.  Now it is defined on the type class {topological_space,
comm_monoid_add}.  Hence it is useable also for extended real numbers.

* Theory Complex_Main: The complex exponential function "expi" is now
a type-constrained abbreviation for "exp :: complex => complex"; thus
several polymorphic lemmas about "exp" are now applicable to "expi".

* Code generation:

  - Theory Library/Code_Char_ord provides native ordering of
    characters in the target language.

  - Commands code_module and code_library are legacy, use export_code
    instead.

  - Method "evaluation" is legacy, use method "eval" instead.

  - Legacy evaluator "SML" is deactivated by default.  May be
    reactivated by the following theory command:

      setup {* Value.add_evaluator ("SML", Codegen.eval_term) *}

* Declare ext [intro] by default.  Rare INCOMPATIBILITY.

* New proof method "induction" that gives induction hypotheses the
name "IH", thus distinguishing them from further hypotheses that come
from rule induction.  The latter are still called "hyps".  Method
"induction" is a thin wrapper around "induct" and follows the same
syntax.

* Method "fastsimp" has been renamed to "fastforce", but "fastsimp" is
still available as a legacy feature for some time.

* Nitpick:
  - Added "need" and "total_consts" options.
  - Reintroduced "show_skolems" option by popular demand.
  - Renamed attribute: nitpick_def ~> nitpick_unfold.
    INCOMPATIBILITY.

* Sledgehammer:
  - Use quasi-sound (and efficient) translations by default.
  - Added support for the following provers: E-ToFoF, LEO-II,
    Satallax, SNARK, Waldmeister, and Z3 with TPTP syntax.
  - Automatically preplay and minimize proofs before showing them if
    this can be done within reasonable time.
  - sledgehammer available_provers ~> sledgehammer supported_provers.
    INCOMPATIBILITY.
  - Added "preplay_timeout", "slicing", "type_enc", "sound",
    "max_mono_iters", and "max_new_mono_instances" options.
  - Removed "explicit_apply" and "full_types" options as well as "Full
    Types" Proof General menu item. INCOMPATIBILITY.

* Metis:
  - Removed "metisF" -- use "metis" instead. INCOMPATIBILITY.
  - Obsoleted "metisFT" -- use "metis (full_types)" instead.
    INCOMPATIBILITY.

* Command 'try':
  - Renamed 'try_methods' and added "simp:", "intro:", "dest:", and
    "elim:" options. INCOMPATIBILITY.
  - Introduced 'try' that not only runs 'try_methods' but also
    'solve_direct', 'sledgehammer', 'quickcheck', and 'nitpick'.

* Quickcheck:
  - Added "eval" option to evaluate terms for the found counterexample
    (currently only supported by the default (exhaustive) tester).
  - Added post-processing of terms to obtain readable counterexamples
    (currently only supported by the default (exhaustive) tester).
  - New counterexample generator quickcheck[narrowing] enables
    narrowing-based testing.  Requires the Glasgow Haskell compiler
    with its installation location defined in the Isabelle settings
    environment as ISABELLE_GHC.
  - Removed quickcheck tester "SML" based on the SML code generator
    (formly in HOL/Library).

* Function package: discontinued option "tailrec".  INCOMPATIBILITY,
use 'partial_function' instead.

* Theory Library/Extended_Reals replaces now the positive extended
reals found in probability theory. This file is extended by
Multivariate_Analysis/Extended_Real_Limits.

* Theory Library/Old_Recdef: old 'recdef' package has been moved here,
from where it must be imported explicitly if it is really required.
INCOMPATIBILITY.

* Theory Library/Wfrec: well-founded recursion combinator "wfrec" has
been moved here.  INCOMPATIBILITY.

* Theory Library/Saturated provides type of numbers with saturated
arithmetic.

* Theory Library/Product_Lattice defines a pointwise ordering for the
product type 'a * 'b, and provides instance proofs for various order
and lattice type classes.

* Theory Library/Countable now provides the "countable_datatype" proof
method for proving "countable" class instances for datatypes.

* Theory Library/Cset_Monad allows do notation for computable sets
(cset) via the generic monad ad-hoc overloading facility.

* Library: Theories of common data structures are split into theories
for implementation, an invariant-ensuring type, and connection to an
abstract type. INCOMPATIBILITY.

  - RBT is split into RBT and RBT_Mapping.
  - AssocList is split and renamed into AList and AList_Mapping.
  - DList is split into DList_Impl, DList, and DList_Cset.
  - Cset is split into Cset and List_Cset.

* Theory Library/Nat_Infinity has been renamed to
Library/Extended_Nat, with name changes of the following types and
constants:

  type inat   ~> type enat
  Fin         ~> enat
  Infty       ~> infinity (overloaded)
  iSuc        ~> eSuc
  the_Fin     ~> the_enat

Every theorem name containing "inat", "Fin", "Infty", or "iSuc" has
been renamed accordingly. INCOMPATIBILITY.

* Session Multivariate_Analysis: The euclidean_space type class now
fixes a constant "Basis :: 'a set" consisting of the standard
orthonormal basis for the type. Users now have the option of
quantifying over this set instead of using the "basis" function, e.g.
"ALL x:Basis. P x" vs "ALL i<DIM('a). P (basis i)".

* Session Multivariate_Analysis: Type "('a, 'b) cart" has been renamed
to "('a, 'b) vec" (the syntax "'a ^ 'b" remains unaffected). Constants
"Cart_nth" and "Cart_lambda" have been respectively renamed to
"vec_nth" and "vec_lambda"; theorems mentioning those names have
changed to match. Definition theorems for overloaded constants now use
the standard "foo_vec_def" naming scheme. A few other theorems have
been renamed as follows (INCOMPATIBILITY):

  Cart_eq          ~> vec_eq_iff
  dist_nth_le_cart ~> dist_vec_nth_le
  tendsto_vector   ~> vec_tendstoI
  Cauchy_vector    ~> vec_CauchyI

* Session Multivariate_Analysis: Several duplicate theorems have been
removed, and other theorems have been renamed or replaced with more
general versions. INCOMPATIBILITY.

  finite_choice ~> finite_set_choice
  eventually_conjI ~> eventually_conj
  eventually_and ~> eventually_conj_iff
  eventually_false ~> eventually_False
  setsum_norm ~> norm_setsum
  Lim_sequentially ~> LIMSEQ_def
  Lim_ident_at ~> LIM_ident
  Lim_const ~> tendsto_const
  Lim_cmul ~> tendsto_scaleR [OF tendsto_const]
  Lim_neg ~> tendsto_minus
  Lim_add ~> tendsto_add
  Lim_sub ~> tendsto_diff
  Lim_mul ~> tendsto_scaleR
  Lim_vmul ~> tendsto_scaleR [OF _ tendsto_const]
  Lim_null_norm ~> tendsto_norm_zero_iff [symmetric]
  Lim_linear ~> bounded_linear.tendsto
  Lim_component ~> tendsto_euclidean_component
  Lim_component_cart ~> tendsto_vec_nth
  Lim_inner ~> tendsto_inner [OF tendsto_const]
  dot_lsum ~> inner_setsum_left
  dot_rsum ~> inner_setsum_right
  continuous_cmul ~> continuous_scaleR [OF continuous_const]
  continuous_neg ~> continuous_minus
  continuous_sub ~> continuous_diff
  continuous_vmul ~> continuous_scaleR [OF _ continuous_const]
  continuous_mul ~> continuous_scaleR
  continuous_inv ~> continuous_inverse
  continuous_at_within_inv ~> continuous_at_within_inverse
  continuous_at_inv ~> continuous_at_inverse
  continuous_at_norm ~> continuous_norm [OF continuous_at_id]
  continuous_at_infnorm ~> continuous_infnorm [OF continuous_at_id]
  continuous_at_component ~> continuous_component [OF continuous_at_id]
  continuous_on_neg ~> continuous_on_minus
  continuous_on_sub ~> continuous_on_diff
  continuous_on_cmul ~> continuous_on_scaleR [OF continuous_on_const]
  continuous_on_vmul ~> continuous_on_scaleR [OF _ continuous_on_const]
  continuous_on_mul ~> continuous_on_scaleR
  continuous_on_mul_real ~> continuous_on_mult
  continuous_on_inner ~> continuous_on_inner [OF continuous_on_const]
  continuous_on_norm ~> continuous_on_norm [OF continuous_on_id]
  continuous_on_inverse ~> continuous_on_inv
  uniformly_continuous_on_neg ~> uniformly_continuous_on_minus
  uniformly_continuous_on_sub ~> uniformly_continuous_on_diff
  subset_interior ~> interior_mono
  subset_closure ~> closure_mono
  closure_univ ~> closure_UNIV
  real_arch_lt ~> reals_Archimedean2
  real_arch ~> reals_Archimedean3
  real_abs_norm ~> abs_norm_cancel
  real_abs_sub_norm ~> norm_triangle_ineq3
  norm_cauchy_schwarz_abs ~> Cauchy_Schwarz_ineq2

* Session HOL-Probability:
  - Caratheodory's extension lemma is now proved for ring_of_sets.
  - Infinite products of probability measures are now available.
  - Sigma closure is independent, if the generator is independent
  - Use extended reals instead of positive extended
    reals. INCOMPATIBILITY.

* Session HOLCF: Discontinued legacy theorem names, INCOMPATIBILITY.

  expand_fun_below ~> fun_below_iff
  below_fun_ext ~> fun_belowI
  expand_cfun_eq ~> cfun_eq_iff
  ext_cfun ~> cfun_eqI
  expand_cfun_below ~> cfun_below_iff
  below_cfun_ext ~> cfun_belowI
  monofun_fun_fun ~> fun_belowD
  monofun_fun_arg ~> monofunE
  monofun_lub_fun ~> adm_monofun [THEN admD]
  cont_lub_fun ~> adm_cont [THEN admD]
  cont2cont_Rep_CFun ~> cont2cont_APP
  cont_Rep_CFun_app ~> cont_APP_app
  cont_Rep_CFun_app_app ~> cont_APP_app_app
  cont_cfun_fun ~> cont_Rep_cfun1 [THEN contE]
  cont_cfun_arg ~> cont_Rep_cfun2 [THEN contE]
  contlub_cfun ~> lub_APP [symmetric]
  contlub_LAM ~> lub_LAM [symmetric]
  thelubI ~> lub_eqI
  UU_I ~> bottomI
  lift_distinct1 ~> lift.distinct(1)
  lift_distinct2 ~> lift.distinct(2)
  Def_not_UU ~> lift.distinct(2)
  Def_inject ~> lift.inject
  below_UU_iff ~> below_bottom_iff
  eq_UU_iff ~> eq_bottom_iff


*** Document preparation ***

* Antiquotation @{rail} layouts railroad syntax diagrams, see also
isar-ref manual, both for description and actual application of the
same.

* Antiquotation @{value} evaluates the given term and presents its
result.

* Antiquotations: term style "isub" provides ad-hoc conversion of
variables x1, y23 into subscripted form x\<^isub>1,
y\<^isub>2\<^isub>3.

* Predefined LaTeX macros for Isabelle symbols \<bind> and \<then>
(e.g. see ~~/src/HOL/Library/Monad_Syntax.thy).

* Localized \isabellestyle switch can be used within blocks or groups
like this:

  \isabellestyle{it}  %preferred default
  {\isabellestylett @{text "typewriter stuff"}}

* Discontinued special treatment of hard tabulators.  Implicit
tab-width is now defined as 1.  Potential INCOMPATIBILITY for visual
layouts.


*** ML ***

* The inner syntax of sort/type/term/prop supports inlined YXML
representations within quoted string tokens.  By encoding logical
entities via Term_XML (in ML or Scala) concrete syntax can be
bypassed, which is particularly useful for producing bits of text
under external program control.

* Antiquotations for ML and document preparation are managed as theory
data, which requires explicit setup.

* Isabelle_Process.is_active allows tools to check if the official
process wrapper is running (Isabelle/Scala/jEdit) or the old TTY loop
(better known as Proof General).

* Structure Proof_Context follows standard naming scheme.  Old
ProofContext is still available for some time as legacy alias.

* Structure Timing provides various operations for timing; supersedes
former start_timing/end_timing etc.

* Path.print is the official way to show file-system paths to users
(including quotes etc.).

* Inner syntax: identifiers in parse trees of generic categories
"logic", "aprop", "idt" etc. carry position information (disguised as
type constraints).  Occasional INCOMPATIBILITY with non-compliant
translations that choke on unexpected type constraints.  Positions can
be stripped in ML translations via Syntax.strip_positions /
Syntax.strip_positions_ast, or via the syntax constant
"_strip_positions" within parse trees.  As last resort, positions can
be disabled via the configuration option Syntax.positions, which is
called "syntax_positions" in Isar attribute syntax.

* Discontinued special status of various ML structures that contribute
to structure Syntax (Ast, Lexicon, Mixfix, Parser, Printer etc.): less
pervasive content, no inclusion in structure Syntax.  INCOMPATIBILITY,
refer directly to Ast.Constant, Lexicon.is_identifier,
Syntax_Trans.mk_binder_tr etc.

* Typed print translation: discontinued show_sorts argument, which is
already available via context of "advanced" translation.

* Refined PARALLEL_GOALS tactical: degrades gracefully for schematic
goal states; body tactic needs to address all subgoals uniformly.

* Slightly more special eq_list/eq_set, with shortcut involving
pointer equality (assumes that eq relation is reflexive).

* Classical tactics use proper Proof.context instead of historic types
claset/clasimpset.  Old-style declarations like addIs, addEs, addDs
operate directly on Proof.context.  Raw type claset retains its use as
snapshot of the classical context, which can be recovered via
(put_claset HOL_cs) etc.  Type clasimpset has been discontinued.
INCOMPATIBILITY, classical tactics and derived proof methods require
proper Proof.context.


*** System ***

* Discontinued support for Poly/ML 5.2, which was the last version
without proper multithreading and TimeLimit implementation.

* Discontinued old lib/scripts/polyml-platform, which has been
obsolete since Isabelle2009-2.

* Various optional external tools are referenced more robustly and
uniformly by explicit Isabelle settings as follows:

  ISABELLE_CSDP   (formerly CSDP_EXE)
  ISABELLE_GHC    (formerly EXEC_GHC or GHC_PATH)
  ISABELLE_OCAML  (formerly EXEC_OCAML)
  ISABELLE_SWIPL  (formerly EXEC_SWIPL)
  ISABELLE_YAP    (formerly EXEC_YAP)

Note that automated detection from the file-system or search path has
been discontinued.  INCOMPATIBILITY.

* Scala layer provides JVM method invocation service for static
methods of type (String)String, see Invoke_Scala.method in ML.  For
example:

  Invoke_Scala.method "java.lang.System.getProperty" "java.home"

Together with YXML.string_of_body/parse_body and XML.Encode/Decode
this allows to pass structured values between ML and Scala.

* The IsabelleText fonts includes some further glyphs to support the
Prover IDE.  Potential INCOMPATIBILITY: users who happen to have
installed a local copy (which is normally *not* required) need to
delete or update it from ~~/lib/fonts/.



New in Isabelle2011 (January 2011)
----------------------------------

*** General ***

* Experimental Prover IDE based on Isabelle/Scala and jEdit (see
src/Tools/jEdit).  This also serves as IDE for Isabelle/ML, with
useful tooltips and hyperlinks produced from its static analysis.  The
bundled component provides an executable Isabelle tool that can be run
like this:

  Isabelle2011/bin/isabelle jedit

* Significantly improved Isabelle/Isar implementation manual.

* System settings: ISABELLE_HOME_USER now includes ISABELLE_IDENTIFIER
(and thus refers to something like $HOME/.isabelle/Isabelle2011),
while the default heap location within that directory lacks that extra
suffix.  This isolates multiple Isabelle installations from each
other, avoiding problems with old settings in new versions.
INCOMPATIBILITY, need to copy/upgrade old user settings manually.

* Source files are always encoded as UTF-8, instead of old-fashioned
ISO-Latin-1.  INCOMPATIBILITY.  Isabelle LaTeX documents might require
the following package declarations:

  \usepackage[utf8]{inputenc}
  \usepackage{textcomp}

* Explicit treatment of UTF-8 sequences as Isabelle symbols, such that
a Unicode character is treated as a single symbol, not a sequence of
non-ASCII bytes as before.  Since Isabelle/ML string literals may
contain symbols without further backslash escapes, Unicode can now be
used here as well.  Recall that Symbol.explode in ML provides a
consistent view on symbols, while raw explode (or String.explode)
merely give a byte-oriented representation.

* Theory loader: source files are primarily located via the master
directory of each theory node (where the .thy file itself resides).
The global load path is still partially available as legacy feature.
Minor INCOMPATIBILITY due to subtle change in file lookup: use
explicit paths, relatively to the theory.

* Special treatment of ML file names has been discontinued.
Historically, optional extensions .ML or .sml were added on demand --
at the cost of clarity of file dependencies.  Recall that Isabelle/ML
files exclusively use the .ML extension.  Minor INCOMPATIBILITY.

* Various options that affect pretty printing etc. are now properly
handled within the context via configuration options, instead of
unsynchronized references or print modes.  There are both ML Config.T
entities and Isar declaration attributes to access these.

  ML (Config.T)                 Isar (attribute)

  eta_contract                  eta_contract
  show_brackets                 show_brackets
  show_sorts                    show_sorts
  show_types                    show_types
  show_question_marks           show_question_marks
  show_consts                   show_consts
  show_abbrevs                  show_abbrevs

  Syntax.ast_trace              syntax_ast_trace
  Syntax.ast_stat               syntax_ast_stat
  Syntax.ambiguity_level        syntax_ambiguity_level

  Goal_Display.goals_limit      goals_limit
  Goal_Display.show_main_goal   show_main_goal

  Method.rule_trace             rule_trace

  Thy_Output.display            thy_output_display
  Thy_Output.quotes             thy_output_quotes
  Thy_Output.indent             thy_output_indent
  Thy_Output.source             thy_output_source
  Thy_Output.break              thy_output_break

Note that corresponding "..._default" references in ML may only be
changed globally at the ROOT session setup, but *not* within a theory.
The option "show_abbrevs" supersedes the former print mode
"no_abbrevs" with inverted meaning.

* More systematic naming of some configuration options.
INCOMPATIBILITY.

  trace_simp  ~>  simp_trace
  debug_simp  ~>  simp_debug

* Support for real valued configuration options, using simplistic
floating-point notation that coincides with the inner syntax for
float_token.

* Support for real valued preferences (with approximative PGIP type):
front-ends need to accept "pgint" values in float notation.
INCOMPATIBILITY.

* The IsabelleText font now includes Cyrillic, Hebrew, Arabic from
DejaVu Sans.

* Discontinued support for Poly/ML 5.0 and 5.1 versions.


*** Pure ***

* Command 'type_synonym' (with single argument) replaces somewhat
outdated 'types', which is still available as legacy feature for some
time.

* Command 'nonterminal' (with 'and' separated list of arguments)
replaces somewhat outdated 'nonterminals'.  INCOMPATIBILITY.

* Command 'notepad' replaces former 'example_proof' for
experimentation in Isar without any result.  INCOMPATIBILITY.

* Locale interpretation commands 'interpret' and 'sublocale' accept
lists of equations to map definitions in a locale to appropriate
entities in the context of the interpretation.  The 'interpretation'
command already provided this functionality.

* Diagnostic command 'print_dependencies' prints the locale instances
that would be activated if the specified expression was interpreted in
the current context.  Variant "print_dependencies!" assumes a context
without interpretations.

* Diagnostic command 'print_interps' prints interpretations in proofs
in addition to interpretations in theories.

* Discontinued obsolete 'global' and 'local' commands to manipulate
the theory name space.  Rare INCOMPATIBILITY.  The ML functions
Sign.root_path and Sign.local_path may be applied directly where this
feature is still required for historical reasons.

* Discontinued obsolete 'constdefs' command.  INCOMPATIBILITY, use