diff --git a/NEWS b/NEWS new file mode 100644 --- /dev/null +++ b/NEWS @@ -0,0 +1,17186 @@ +Isabelle NEWS -- history of user-relevant changes +================================================= + +(Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.) + + +New in Isabelle2022 (October 2022) +---------------------------------- + +*** General *** + +* The instantiation of schematic goals is now displayed explicitly as a +list of variable assignments. This works for proof state output, and at +the end of a toplevel proof. In rare situations, a proof command or +proof method may violate the intended goal discipline, by not producing +an instance of the original goal, but there is only a warning, no hard +error. + +* Session ROOT files support 'chapter_definition' entries (optional). +This allows to associate additional information as follows: + + - "chapter_definition NAME (GROUPS)" to make all sessions that belong + to this chapter members of the given groups + + - "chapter_definition NAME description TEXT" to provide a description + for presentation purposes + +* Old-style {* verbatim *} tokens have been discontinued (legacy feature +since Isabelle2019). INCOMPATIBILITY, use \cartouche\ syntax instead. + + +*** Isabelle/jEdit Prover IDE *** + +* Command 'print_state' outputs a plain message, i.e. "writeln" instead +of "state". Thus it is displayed in the Output panel, even if the option +"editor_output_state" is disabled. + + +*** Isabelle/VSCode Prover IDE *** + +* VSCodium, an open-source distribution of VSCode without MS telemetry, +has been bundled with Isabelle as add-on component. The command-line +tool "isabelle vscode" automatically configures it as Isabelle/VSCode +and starts the application. This includes special support for the +UTF-8-Isabelle encoding and the corresponding Isabelle fonts. + +* Command-line tools "isabelle electron" and "isabelle node" provide +access to the underlying technologies of VSCodium, for use in other +applications. This essentially provides a freely programmable Chromium +browser engine that works uniformly on all platforms. + +Example: + + URL="https://isabelle.in.tum.de" isabelle electron \ + --app="$(isabelle getenv -b ISABELLE_HOME)"/src/Tools/Electron/test + + +*** HTML/PDF presentation *** + +* Management of dependencies has become more robust and accurate, +following the session build hierarchy, and the up-to-date notion of +"isabelle build". Changed sessions and updated builds will cause new +HTML presentation, when that is enabled eventually. Unchanged sessions +retain their HTML output that is already present. Thus HTML presentation +for basic sessions like "HOL" and "HOL-Analysis" is produced at most +once, as required by user sessions. + +* HTML presentation no longer supports README.html, which was meant as +add-on to the index.html of a session. Rare INCOMPATIBILITY, consider +using a separate theory "README" with Isabelle document markup/markdown. + +* ML files (and other auxiliary files) are presented with detailed +hyperlinks, just like regular theory sources. + +* Support for external hyperlinks (URLs). + +* Support for internal hyperlinks to files that belong formally to the +presented session. + + +*** HOL *** + +* HOL record types: new simproc that sorts update expressions, guarded +by configuration option "record_sort_updates" (default: false). Some +examples are in theory "HOL-Examples.Records". + +* Meson: added support for polymorphic "using" facts. Minor +INCOMPATIBILITY. + +* Moved auxiliary computation constant "divmod_nat" to theory +"HOL.Euclidean_Division". Minor INCOMPATIBILITY. + +* Renamed attribute "arith_split" to "linarith_split". Minor +INCOMPATIBILITY. + +* Theory "HOL.Rings": rule split_of_bool_asm is not split any longer, +analogously to split_if_asm. INCOMPATIBILITY. + +* Theory "HOL.Bit_Operations": rule bit_0 is not default [simp] any +longer. INCOMPATIBILITY. + +* Streamlined primitive definitions of division and modulus on integers. +INCOMPATIBILITY. + +* Theory "HOL.Fun": + - Added predicate monotone_on and redefined monotone to be an + abbreviation. Lemma monotone_def is explicitly provided for backward + compatibility but its usage is discouraged. Minor INCOMPATIBILITY. + - Changed argument order of mono_on and strict_mono_on to uniformize + with monotone_on and the general "characterizing set at the beginning + of predicates" preference. Also change them to be abbreviations + of monotone_of. Lemmas mono_on_def and strict_mono_on_def are + explicitly provided for backward compatibility but their usage is + discouraged. INCOMPATIBILITY. + - Move mono, strict_mono, antimono, and relevant lemmas to Fun theory. + Also change them to be abbreviations of mono_on, strict_mono_on, + and monotone, respectively. Lemmas mono_def, strict_mono_def, and + antimono_def are explicitly provided for backward compatibility but + their usage is discouraged. Minor INCOMPATIBILITY. + - Added lemmas. + monotone_onD + monotone_onI + monotone_on_empty[simp] + monotone_on_o + monotone_on_subset + +* Theory "HOL.Relation": + - Added predicate reflp_on and redefined reflp to be an abbreviation. + Lemma reflp_def is explicitly provided for backward compatibility + but its usage is discouraged. Minor INCOMPATIBILITY. + - Added predicate totalp_on and abbreviation totalp. + - Replaced HOL.implies by Pure.imp in lemma reflp_mono for consistency + with other lemmas. Minor INCOMPATIBILITY. + - Added lemmas. + preorder.asymp_greater + preorder.asymp_less + reflp_onD + reflp_onI + reflp_on_Inf + reflp_on_Sup + reflp_on_empty[simp] + reflp_on_inf + reflp_on_mono + reflp_on_subset + reflp_on_sup + total_on_subset + totalpD + totalpI + totalp_onD + totalp_onI + totalp_on_empty[simp] + totalp_on_subset + totalp_on_total_on_eq[pred_set_conv] + +* Theory "HOL.Transitive_Closure": + - Added lemmas. + total_on_trancl + totalp_on_tranclp + +* New theory "HOL-Library.NList" of fixed length lists. + +* New Theory "HOL-Library.Code_Abstract_Char" implements characters by +target language integers, sacrificing pattern patching in exchange for +dramatically increased performance for comparisons. + +* Theory "HOL-Library.Char_ord": streamlined logical specifications. +Minor INCOMPATIBILITY. + +* Theory "HOL-Library.Multiset": + - Consolidated operation and fact names. + multp ~> multp_code + multeqp ~> multeqp_code + multp_cancel_add_mset ~> multp_cancel_add_mset0 + multp_cancel_add_mset0[simplified] ~> multp_cancel_add_mset + multp_code_iff ~> multp_code_iff_mult + multeqp_code_iff ~> multeqp_code_iff_reflcl_mult + Minor INCOMPATIBILITY. + - Moved mult1_lessE out of preorder type class and add explicit + assumption. Minor INCOMPATIBILITY. + - Added predicate multp equivalent to set mult. Reuse name previously + used for what is now called multp_code. Minor INCOMPATIBILITY. + - Lifted multiple lemmas from mult to multp. + - Redefined less_multiset to be based on multp. INCOMPATIBILITY. + - Added lemmas. + Multiset.bex_greatest_element + Multiset.bex_least_element + filter_mset_cong + filter_mset_cong0 + image_mset_eq_image_mset_plusD + image_mset_eq_plusD + image_mset_eq_plus_image_msetD + image_mset_filter_mset_swap + monotone_multp_multp_image_mset + monotone_on_multp_multp_image_mset + multp_image_mset_image_msetD + +* Theory "HOL-Library.Sublist": added lemma map_mono_strict_suffix. + +* Theory "HOL-ex.Sum_of_Powers" has been deleted. The same material is +in the AFP as Bernoulli. + +* Session HOL-Algebra: some facts have been renamed to avoid fact name +clashes on interpretation: + + is_ring ~> ring_axioms + cring ~> cring_axioms + R_def ~> R_m_def + +INCOMPATIBILITY. + +* Nitpick: To avoid technical issues, prefer non-JNI solvers to JNI +solvers by default. Minor INCOMPATIBILITY. + +* Sledgehammer: + - Redesigned multithreading to provide more fine grained prover + schedules. The binary option 'slice' has been replaced by a numeric + value 'slices' indicating the number of desired slices. Stronger + provers can now be run by more than one thread simultaneously. The + new option 'max_proofs' controls the number of proofs shown. + INCOMPATIBILITY. + - Introduced sledgehammer_outcome data type and changed return type of + ML function Sledgehammer.run_sledgehammer from "bool * (string * + string list)" to "bool * (sledgehammer_outcome * string)". The + former value can be recomputed with "apsnd (ATP_Util.map_prod + Sledgehammer.short_string_of_sledgehammer_outcome single)". + INCOMPATIBILITY. + - Added support for TX0 and TX1 TPTP formats and $ite/$let expressions + in TH0 and TH1. + - Added support for cvc5. + - Generate Isar proofs by default when and only when the one-liner + proof fails to replay and the Isar proof succeeds. + - Replaced option "sledgehammer_atp_dest_dir" by + "sledgehammer_atp_problem_dest_dir", for problem files, and + "sledgehammer_atp_proof_dest_dir", for proof files. Minor + INCOMPATIBILITY. + - Removed support for experimental prover 'z3_tptp'. + - The fastest successfully preplayed proof is always suggested. + - All SMT solvers but Z3 now resort to suggest (smt (verit)) when no + proof could be preplayed. + - Added new "some_preplayed" value to "expect" option to specify that + some successfully preplayed proof is expected. This is in contrast + to the "some" value which doesn't specify whether preplay succeeds + or not. + +* Mirabelle: + - Replaced sledgehammer option "keep" by "keep_probs", for problems + files, and "keep_proofs" for proof files. Minor INCOMPATIBILITY. + - Added option "-r INT" to randomize the goals with a given 64-bit + seed before selection. + - Added option "-y" for a dry run. + - Renamed run_action to run in Mirabelle.action record. Minor + INCOMPATIBILITY. + - Run the actions on goals before commands "unfolding" and "using". + +* (Co)datatype package: + - BNFs now require a strict cardinality bound (o). + Minor INCOMPATIBILITY for existing manual BNF declarations. + - Lemma map_ident_strong is now generated for all BNFs. + +* More ambitious minimization of case expressions in generated code. + +* Code generation for Scala: type annotations in pattern bindings are +printed in a way suitable for Scala 3. + + +*** ML *** + +* Type Bytes.T supports scalable byte strings, beyond the limit of +String.maxSize (approx. 64 MB on 64_32 architecture). + +* Operations for XZ compression (via Isabelle/Scala): + + XZ.compress: Bytes.T -> Bytes.T + XZ.uncompress: Bytes.T -> Bytes.T + + +*** System *** + +* Isabelle/Scala is now based on Scala 3. This is a completely different +compiler ("dotty") and a quite different source language (we are using +the classic Java-style syntax, not the new Python-style syntax). +Occasional INCOMPATIBILITY, see also the official Scala documentation +https://docs.scala-lang.org/scala3/guides/migration/compatibility-intro.html + +* External Isabelle tools implemented as .scala scripts are no longer +supported. INCOMPATIBILITY, instead provide a proper Isabelle/Scala +module with etc/build.props and "services" for a suitable class instance +of isabelle.Isabelle_Scala_Tools. For example, see +$ISABELLE_HOME/etc/build.props and its isabelle.Tools, which defines the +standard Isabelle tools. + +* The session build database now maintains an additional "uuid" column +to identity the original build process uniquely. Thus other tools may +dependent symbolically on a particular build instance. + +* Command-line tool "isabelle build_docker" supports Docker within Snap +more robustly; see also option -W. + +* Command-line tool "isabelle scala_project" supports Gradle as +alternative to Maven: either option -G or -M needs to be specified +explicitly. This increases the chances that the Java/Scala IDE project +works properly. + +* Command-line tool "isabelle hg_sync" synchronizes the working +directory of a local Mercurial repository with a target directory, using +rsync notation for destinations. + +* Command-line tool "isabelle sync" synchronizes Isabelle + AFP +repositories with a target directory, based on "isabelle hg_sync". Local +jars and sessions images may be uploaded as well, to avoid redundant +builds on the remote side. This tool requires a Mercurial clone of the +Isabelle repository: a regular download of the distribution will not +work! + +* Command-line tool "isabelle log" has been refined to support multiple +sessions, and to match messages against regular expressions (using Java +Pattern syntax). + +* System option "show_states" controls output of toplevel command states +(notably proof states) in batch-builds; in interactive mode this is +always done on demand. The option is relevant for tools that operate on +exported PIDE markup, e.g. document presentation or diagnostics. For +example: + + isabelle build -o show_states FOL-ex + isabelle log -v -U FOL-ex + +Option "show_states" is also the default for the configuration option +"show_results" within the formal context. + +Note that printing intermediate states may cause considerable slowdown +in building a session. + +* Session ROOT entries support 'export_classpath' to augment the +Java/Scala name space for tools that allow dynamic loading of service +classes within a session context. A notable example is document +preparation, which works via the class isabelle.Document_Build.Engine +and is configured by the corresponding system option "document_build". +The Isabelle/Isar command 'scala_build_generated_files' helps to produce +a suitable .jar module for inclusion via 'export_classpath'. + +* Isabelle/Scala SSH connections now use regular OpenSSH executables +from the local system: ssh, scp, sftp; the old ssh-java component has +been discontinued. This has various practical consequences: + + - Authentication and configuration works accurately via the official + .ssh/known_hosts and .ssh/config files. + + - Host connections are usually shared (via multiplexed channels), to + reduce the overhead for multiple commands. This also works for SSH + connections for rsync (e.g. "isabelle sync"). Windows/Cygwin does + not support multiplexing: the functionality should be the same, but + slower, with a new connection for each command. + + - Multiple hops via "bastion hosts" can be easily configured in + .ssh/config via ProxyJump declarations. The former Isabelle/Scala + parameters for proxy_host etc. have been discontinued: minor + INCOMPATIBILITY. + +* The MLton compiler for x86_64-linux has been bundled as Isabelle +component, since Ubuntu 22.04 no longer provides a suitable package. +Note that on macOS, MLton is readily available via Homebrew: +https://formulae.brew.sh/formula/mlton + +The Isabelle settings refer to an executable "$ISABELLE_MLTON" and +command-line options $ISABELLE_MLTON_OPTIONS, which need to fit +together. Potential INCOMPATIBILITY for existing +$ISABELLE_HOME_USER/etc/settings. + + + +New in Isabelle2021-1 (December 2021) +------------------------------------- + +*** General *** + +* The Isabelle/Haskell library ($ISABELLE_HOME/src/Tools/Haskell) has +been significantly improved. In particular, module Isabelle.Bytes +provides type Bytes for light-weight byte strings (with optional UTF8 +interpretation), similar to type string in Isabelle/ML. Isabelle symbols +now work uniformly in Isabelle/Haskell vs. Isabelle/ML vs. +Isabelle/Scala/PIDE. + +* Configuration option "show_results" controls output of final results +in commands like 'definition' or 'theorem'. Output is normally enabled +in interactive mode, but it could occasionally cause unnecessary +slowdown. It can be disabled like this: + + context notes [[show_results = false]] + begin + definition "test = True" + theorem test by (simp add: test_def) + end + +* Theory_Data / Generic_Data: "val extend = I" has been removed; +obsolete since Isabelle2021. + +* More symbol definitions for the Z Notation (Isabelle fonts and LaTeX). +See also the group "Z Notation" in the Symbols dockable of +Isabelle/jEdit. + + +*** Isar *** + +* Commands 'syntax' and 'no_syntax' now work in a local theory context, +but in contrast to 'notation' and 'no_notation' there is no proper way +to refer to local entities. Note that local syntax works well with +'bundle', e.g. see "lattice_syntax" vs. "no_lattice_syntax" in theory +Main of Isabelle/HOL. + +* The improper proof command 'guess' is no longer part of by Pure, but +provided by the separate theory "Pure-ex.Guess". INCOMPATIBILITY, +existing applications need to import session "Pure-ex" and theory +"Pure-ex.Guess". Afterwards it is usually better eliminate the 'guess' +command, using explicit 'obtain' instead. + +* More robust 'proof' outline for method "induct": support nested cases. + + +*** Isabelle/jEdit Prover IDE *** + +* The main plugin for Isabelle/jEdit can be deactivated and reactivated +as documented --- was broken at least since Isabelle2018. + +* Isabelle/jEdit is now composed more conventionally from the original +jEdit text editor in $JEDIT_HOME (with minor patches), plus two Isabelle +plugins that are produced in $JEDIT_SETTINGS/jars on demand. The main +isabelle.jedit module is now part of Isabelle/Scala (as one big +$ISABELLE_SCALA_JAR). + +* Add-on components may provide their own jEdit plugins, via the new +Scala/Java module concept: instances of class +isabelle.Scala_Project.Plugin that are declared as "services" within +etc/build.props are activated on Isabelle/jEdit startup. E.g. see +existing isabelle.jedit.JEdit_Plugin0 (for isabelle_jedit_base.jar) and +isabelle.jedit.JEdit_Plugin1 (for isabelle_jedit_main.jar). + +* Support for built-in font substitution of jEdit text area. + + +*** Document preparation *** + +* HTML presentation now includes links to formal entities. + +* High-quality blackboard-bold symbols from font "txmia" (LaTeX package +"pxfonts"): \\\\\\\\\\\\\\\\\\\\\\\\\\. + +* More predefined symbols: \ \ \ (package "stmaryrd"), \ \ (LaTeX +package "pifont"). + +* Document antiquotations for ML text have been refined: "def" and "ref" +variants support index entries, e.g. @{ML} (no entry) vs. @{ML_def} +(bold entry) vs. @{ML_ref} (regular entry); @{ML_type} supports explicit +type arguments for constructors (only relevant for index), e.g. +@{ML_type \'a list\} vs. @{ML_type 'a \list\}; @{ML_op} has been renamed +to @{ML_infix}. Minor INCOMPATIBILITY concerning name and syntax. + +* Option "document_logo" determines if an instance of the Isabelle logo +should be created in the document output directory. The given string +specifies the name of the logo variant, while "_" (underscore) refers to +the unnamed variant. The output file name is always "isabelle_logo.pdf". + +* Option "document_build" determines the document build engine, as +defined in Isabelle/Scala (as system service). The subsequent engines +are provided by the Isabelle distribution: + + - "lualatex" (default): use ISABELLE_LUALATEX for a standard LaTeX + build with optional ISABELLE_BIBTEX and ISABELLE_MAKEINDEX + + - "pdflatex": as above, but use ISABELLE_PDFLATEX (legacy mode for + special LaTeX styles) + + - "build": delegate to the executable "./build pdf" + +The presence of a "build" command within the document output directory +explicitly requires document_build=build. Minor INCOMPATIBILITY, need to +adjust session ROOT options. + +* Option "document_comment_latex" enables regular LaTeX comment.sty, +instead of the historic version for plain TeX (default). The latter is +much faster, but in conflict with LaTeX classes like Dagstuhl LIPIcs. + +* Option "document_echo" informs about document file names during +session presentation. + +* Option "document_heading_prefix" specifies a prefix for the LaTeX +macro names generated from document heading commands like 'chapter', +'section' etc. The default is "isamarkup", so 'section' becomes +"\isamarkupsection" for example. + +* The command-line tool "isabelle latex" has been discontinued, +INCOMPATIBILITY for old document build scripts. + + - Former "isabelle latex -o sty" has become obsolete: Isabelle .sty + files are automatically generated within the document output + directory. + + - Former "isabelle latex -o pdf" should be replaced by + "$ISABELLE_LUALATEX root" or "$ISABELLE_PDFLATEX root" (without + quotes), according to the intended LaTeX engine. + + - Former "isabelle latex -o bbl" should be replaced by + "$ISABELLE_BIBTEX root" (without quotes). + + - Former "isabelle latex -o idx" should be replaced by + "$ISABELLE_MAKEINDEX root" (without quotes). + +* Option "document_bibliography" explicitly enables the use of bibtex; +the default is to check the presence of root.bib, but it could have a +different name. + +* Improved LaTeX typesetting of \...\ using \guilsinglleft ... +\guilsinglright. INCOMPATIBILITY, need to use \usepackage[T1]{fontenc} +(which is now also the default in "isabelle mkroot"). + +* Simplified typesetting of \...\ using \guillemotleft ... +\guillemotright from \usepackage[T1]{fontenc} --- \usepackage{babel} is +no longer required. + + +*** Pure *** + +* "global_interpretation" is applicable in instantiation and overloading +targets and in any nested target built on top of a target supporting +"global_interpretation". + + +*** HOL *** + +* A new, verified order prover for partial and linear orders. The order +prover rearranges the goal to prove False, then retrieves order literals +(i.e. x = y, x <= y, x < y, and their negated versions) from the premises +and finally tries to derive a contradiction. Its main use case is as a +solver to the simplifier, where it e.g. solves premises of conditional +rewrite rules. + +The new prover (src/Provers/order_tac.ML) replaces the old prover +(src/Provers/order.ML) and improves upon the old one in several ways: + + - The completeness of the prover is verified in Isabelle (see the + ATVA 2021 paper "A Verified Decision Procedure for Orders in Isabelle/HOL"). + + - The new prover is complete for partial orders. + + - The interface to register new orders was reworked to reduce boilerplate. + +The prover has two configuration attributes that control its behaviour: + + - order_trace (default: false): Enables tracing for the solver. + + - order_split_limit (default: 8): Limits the number of order + literals of the form ~ x < y that are passed to the solver since + those lead to case splitting and thus exponential runtime. This + only applies to partial orders. + +The prover is agnostic to the object logic. For HOL, the setup for the +prover is performed in src/HOL/Orderings.thy where the structure +HOL_Order_Tac is obtained. The structure allows us to register new +orders with the functions HOL_Order_Tac.declare_order and +HOL_Order_Tac.declare_linorder. Using these functions, we register the +orders of the type classes order and linorder with the solver. If +possible, one should instantiate these type classes instead of adding +new orders to the solver. One can also interpret the type class locale +as in src/HOL/Library/Sublist.thy, which contains e.g. the prefix +order on lists. + +The method order calls the prover in a standalone fashion. + +The diagnostic command print_orders shows all orders known to the prover +in the current context. + +* Theorems "antisym" and "eq_iff" in class "order" have been renamed to +"order.antisym" and "order.eq_iff", to coexist locally with "antisym" +and "eq_iff" from locale "ordering". INCOMPATIBILITY: significant +potential for change can be avoided if interpretations of type class +"order" are replaced or augmented by interpretations of locale +"ordering". + +* Theorem "swap_def" now is always qualified as "Fun.swap_def". Minor +INCOMPATIBILITY; note that for most applications less elementary lemmas +exists. + +* Lemma "permutes_induct" has been given stronger hypotheses and named +premises. INCOMPATIBILITY. + +* Combinator "Fun.swap" resolved into a mere input abbreviation in +separate theory "Transposition" in HOL-Combinatorics. INCOMPATIBILITY. + +* Theory Bit_Operations is now part of HOL-Main. Minor INCOMPATIBILITY. + +* Infix syntax for bit operations AND, OR, XOR, NOT is now organized in +bundle bit_operations_syntax. INCOMPATIBILITY. + +* Bit operations set_bit, unset_bit and flip_bit are now class +operations. INCOMPATIBILITY. + +* Simplified class hierarchy for bit operations: bit operations reside +in classes (semi)ring_bit_operations, class semiring_bit_shifts is gone. + +* Consecutive conversions to and from words are not collapsed in any +case: rules unsigned_of_nat, unsigned_of_int, signed_of_int, +signed_of_nat, word_of_nat_eq_0_iff, word_of_int_eq_0_iff are not simp +by default any longer. INCOMPATIBILITY. + +* Abbreviation "max_word" has been moved to session Word_Lib in the AFP, +as also have constants "shiftl1", "shiftr1", "sshiftr1", "bshiftr1", +"setBit", "clearBit". See there further the changelog in theory Guide. +INCOMPATIBILITY. + +* Reorganized classes and locales for boolean algebras. INCOMPATIBILITY. + +* New simp rules: less_exp, min.absorb1, min.absorb2, min.absorb3, +min.absorb4, max.absorb1, max.absorb2, max.absorb3, max.absorb4. Minor +INCOMPATIBILITY. + +* The Mirabelle testing tool is now part of Main HOL, and accessible via +the command-line tool "isabelle mirabelle" (implemented in +Isabelle/Scala). It has become more robust and supports parallelism +within Isabelle/ML. + +* Nitpick: External solver "MiniSat" is available for all supported +Isabelle platforms (including 64bit Windows and ARM); while +"MiniSat_JNI" only works for Intel Linux and macOS. + +* Nitpick/Kodkod: default is back to external Java process (option +kodkod_scala = false), both for PIDE and batch builds. This reduces +confusion and increases robustness of timeouts, despite substantial +overhead to run an external JVM. For more fine-grained control, the +kodkod_scala option can be modified within the formal theory context +like this: + + declare [[kodkod_scala = false]] + +* Sledgehammer: + - Update of bundled provers: + . E 2.6 + . Vampire 4.6 (with Open Source license) + . veriT 2021.06.1-rmx + . Zipperposition 2.1 + . Z3 4.4.1 for arm64-linux, which approximates Z3 4.4.0pre, + but sometimes fails or crashes + - Adjusted default provers: + cvc4 vampire verit e spass z3 zipperposition + - Adjusted Zipperposition's slicing. + - Removed legacy "lam_lifting" (synonym for "lifting") from option + "lam_trans". Minor INCOMPATIBILITY. + - Renamed "hide_lams" to "opaque_lifting" in option "lam_trans". Minor + INCOMPATIBILITY. + - Added "opaque_combs" to option "lam_trans": lambda expressions are + rewritten using combinators, but the combinators are kept opaque, + i.e. without definitions. + +* Metis: + - Renamed option "hide_lams" to "opaque_lifting". Minor INCOMPATIBILITY. + - Updated the Metis prover underlying the "metis" proof method to + version 2.4 (release 20200713). The new version fixes one + implementation defect. Very slight INCOMPATIBILITY. + +* Theory HOL-Library.Lattice_Syntax has been superseded by bundle +"lattice_syntax": it can be used in a local context via 'include' or in +a global theory via 'unbundle'. The opposite declarations are bundled as +"no_lattice_syntax". Minor INCOMPATIBILITY. + +* Theory "HOL-Library.Multiset": dedicated predicate "multiset" is gone, +use explict expression instead. Minor INCOMPATIBILITY. + +* Theory "HOL-Library.Multiset": consolidated abbreviations Mempty, +Melem, not_Melem to empty_mset, member_mset, not_member_mset +respectively. Minor INCOMPATIBILITY. + +* Theory "HOL-Library.Multiset": consolidated operation and fact names: + inf_subset_mset ~> inter_mset + sup_subset_mset ~> union_mset + multiset_inter_def ~> inter_mset_def + sup_subset_mset_def ~> union_mset_def + multiset_inter_count ~> count_inter_mset + sup_subset_mset_count ~> count_union_mset + +* Theory "HOL-Library.Complex_Order": Defines less, less_eq on complex +numbers. Not imported by default. + +* Theory "HOL-Library.Multiset": syntax precendence for membership +operations has been adjusted to match the corresponding precendences on +sets. Rare INCOMPATIBILITY. + +* Theory "HOL-Library.Cardinality": code generator setup based on the +type classes finite_UNIV and card_UNIV has been moved to +"HOL-Library.Code_Cardinality", to avoid incompatibilities with +other code setups for sets in AFP/Containers. Applications relying on +this code setup should import "HOL-Library.Code_Cardinality". Minor +INCOMPATIBILITY. + +* Theory "HOL-Library.Permutation" has been renamed to the more specific +"HOL-Library.List_Permutation". Note that most notions from that theory +are already present in theory "HOL-Combinatorics.Permutations". +INCOMPATIBILITY. + +* Dedicated session "HOL-Combinatorics". INCOMPATIBILITY: theories +"Permutations", "List_Permutation" (formerly "Permutation"), "Stirling", +"Multiset_Permutations", "Perm" have been moved there from session +HOL-Library. + +* Theory "HOL-Combinatorics.Transposition" provides elementary swap +operation "transpose". + +* Theory "HOL-Analysis.Infinite_Sum": new theory for infinite sums with +a more general definition than the existing theory Infinite_Set_Sum. +(Infinite_Set_Sum contains theorems relating the two definitions.) + +* Theory "HOL-Analysis.Product_Vector": Instantiation of the product of +uniform spaces as a uniform space. Minor INCOMPATIBILITY: the old +definition "uniformity_prod_def" is available as a derived fact +"uniformity_dist". + +* Session "HOL-Analysis" and "HOL-Probability": indexed products of +discrete distributions, negative binomial distribution, Hoeffding's +inequality, Chernoff bounds, Cauchy–Schwarz inequality for nn_integral, +and some more small lemmas. Some theorems that were stated awkwardly +before were corrected. Minor INCOMPATIBILITY. + +* Session "HOL-Analysis": the complex Arg function has been identified +with the function "arg" of Complex_Main, renaming arg ~> Arg also in the +names of arg_bounded. Minor INCOMPATIBILITY. + +* Session "HOL-Statespace": various improvements and cleanup. + + +*** ML *** + +* External bash processes are always managed by Isabelle/Scala, in +contrast to Isabelle2021 where this was only done for macOS on Apple +Silicon. + +The main Isabelle/ML interface is Isabelle_System.bash_process with +result type Process_Result.T (resembling class Process_Result in Scala); +derived operations Isabelle_System.bash and Isabelle_System.bash_output +provide similar functionality as before. The underlying TCP/IP server +within Isabelle/Scala is available to other programming languages as +well, notably Isabelle/Haskell. + +Rare INCOMPATIBILITY due to subtle semantic differences: + + - Processes invoked from Isabelle/ML actually run in the context of + the Java VM of Isabelle/Scala. The settings environment and current + working directory are usually the same on both sides, but there can be + subtle corner cases (e.g. unexpected uses of "cd" or "putenv" in ML). + + - Output via stdout and stderr is line-oriented: Unix vs. Windows + line-endings are normalized towards Unix; presence or absence of a + final newline is irrelevant. The original lines are available as + Process_Result.out_lines/err_lines; the concatenated versions + Process_Result.out/err *omit* a trailing newline (using + Library.trim_line, which was occasional seen in applications before, + but is no longer necessary). + + - Output needs to be plain text encoded in UTF-8: Isabelle/Scala + recodes it temporarily as UTF-16. This works for well-formed Unicode + text, but not for arbitrary byte strings. In such cases, the bash + script should write tempory files, managed by Isabelle/ML operations + like Isabelle_System.with_tmp_file to create a file name and + File.read to retrieve its content. + + - The Isabelle/Scala "bash_process" server requires a PIDE session + context. This could be a regular batch session (e.g. "isabelle + build"), a PIDE editor session (e.g. "isabelle jedit"), or headless + PIDE (e.g. "isabelle dump" or "isabelle server"). Note that old + "isabelle console" or raw "isabelle process" don't have that. + +New Process_Result.timing works as in Isabelle/Scala, based on direct +measurements of the bash_process wrapper in C: elapsed time is always +available, CPU time is only available on Linux and macOS, GC time is +unavailable. + +* The following Isabelle/ML system operations are run in the context of +Isabelle/Scala, within a PIDE session context: + + - Isabelle_System.make_directory + - Isabelle_System.copy_dir + - Isabelle_System.copy_file + - Isabelle_System.copy_base_file + - Isabelle_System.rm_tree + - Isabelle_System.download + +* Term operations under abstractions are now more robust (and more +strict) by using the formal proof context in subsequent operations: + + Variable.dest_abs + Variable.dest_abs_cterm + Variable.dest_all + Variable.dest_all_cterm + +This works under the assumption that terms are always properly declared +to the proof context (e.g. via Variable.declare_term). Failure to do so, +or working with the wrong context, will cause an error (exception Fail, +based on Term.USED_FREE from Term.dest_abs_fresh). + +The Simplifier and equational conversions now use the above operations +routinely, and thus require user-space tools to be serious about the +proof context (notably in their use of Goal.prove, SUBPROOF etc.). +INCOMPATIBILITY in add-on tools is to be expected occasionally: a proper +context discipline needs to be followed. + +* Former operations Term.dest_abs and Logic.dest_all (without a proper +context) have been discontinued. INCOMPATIBILITY, either use +Variable.dest_abs etc. above, or the following operations that imitate +the old behavior to a great extent: + + Term.dest_abs_global + Logic.dest_all_global + +This works under the assumption that the given (sub-)term directly shows +all free variables that need to be avoided when generating a fresh name. +A violation of the assumption are variables stemming from the enclosing +context that get involved in a proof only later. + +* ML structures TFrees, TVars, Frees, Vars, Names provide scalable +operations to accumulate items from types and terms, using a fast +syntactic order. The original order of occurrences may be recovered as +well, e.g. via TFrees.list_set. + +* Thm.instantiate, Thm.generalize and related operations (e.g. +Variable.import) now use scalable data structures from structure TVars, +Vars, Names etc. INCOMPATIBILITY: e.g. use TVars.empty and TVars.make +for immediate adoption; better use TVars.add, TVars.add_tfrees etc. for +scalable accumulation of items. + +* Thm.instantiate_beta applies newly emerging abstractions to their +arguments in the term, but leaves other beta-redexes unchanged --- in +contrast to Drule.instantiate_normalize. + +* ML antiquotation "instantiate" allows to instantiate formal entities +(types, terms, theorems) with values given ML. This works uniformly for +"typ", "term", "prop", "ctyp", "cterm", "cprop", "lemma" --- given as a +keyword after the instantiation. + +A mode "(schematic)" behind the keyword means that some variables may +remain uninstantiated (fixed in the specification and schematic in the +result); by default, all variables need to be instantiated. + +Newly emerging abstractions are applied to their arguments in the term +(using Thm.instantiate_beta). + +Examples in HOL: + + fun make_assoc_type (A, B) = + \<^instantiate>\'a = A and 'b = B in typ \('a \ 'b) list\\; + + val make_assoc_list = + map (fn (x, y) => + \<^instantiate>\'a = \fastype_of x\ and 'b = \fastype_of y\ and + x and y in term \(x, y)\ for x :: 'a and y :: 'b\); + + fun symmetry x y = + \<^instantiate>\'a = \Thm.ctyp_of_cterm x\ and x and y in + lemma \x = y \ y = x\ for x y :: 'a by simp\ + + fun symmetry_schematic A = + \<^instantiate>\'a = A in + lemma (schematic) \x = y \ y = x\ for x y :: 'a by simp\ + +* ML antiquotation for embedded lemma supports local fixes, as usual in +many other Isar language elements. For example: + + @{lemma "x = x" for x :: nat by (rule refl)} + +* ML antiquotations for type constructors and term constants: + + \<^Type>\c\ + \<^Type>\c T \\ \ \same with type arguments\ + \<^Type_fn>\c T \\ \ \fn abstraction, failure via exception TYPE\ + \<^Const>\c\ + \<^Const>\c T \\ \ \same with type arguments\ + \<^Const>\c for t \\ \ \same with term arguments\ + \<^Const_>\c \\ \ \same for patterns: case, let, fn\ + \<^Const_fn>\c T \\ \ \fn abstraction, failure via exception TERM\ + +The type/term arguments refer to nested ML source, which may contain +antiquotations recursively. The following argument syntax is supported: + + - an underscore (dummy pattern) + - an atomic item of "embedded" syntax, e.g. identifier or cartouche + - an antiquotation in control-symbol/cartouche form, e.g. \<^Type>\c\ + as short form of \\<^Type>\c\\. + +Examples in HOL: + + val natT = \<^Type>\nat\; + fun mk_funT (A, B) = \<^Type>\fun A B\; + val dest_funT = \<^Type_fn>\fun A B => \(A, B)\\; + fun mk_conj (A, B) = \<^Const>\conj for A B\; + val dest_conj = \<^Const_fn>\conj for A B => \(A, B)\\; + fun mk_eq T (t, u) = \<^Const>\HOL.eq T for t u\; + val dest_eq = \<^Const_fn>\HOL.eq T for t u => \(T, (t, u))\\; + +* ML antiquotations \<^make_judgment> and \<^dest_judgment> refer to +corresponding functions for the object-logic of the ML compilation +context. This supersedes older mk_Trueprop / dest_Trueprop operations. + +* The "build" combinators of various data structures help to build +content from bottom-up, by applying an "add" function the "empty" value. +For example: + + - type 'a Symtab.table etc.: build + - type 'a Names.table etc.: build + - type 'a list: build and build_rev + - type Buffer.T: build and build_content + +For example, see src/Pure/PIDE/xml.ML: + + val content_of = Buffer.build_content o fold add_content; + +* ML antiquotations \<^try>\expr\ and \<^can>\expr\ operate directly on +the given ML expression, in contrast to functions "try" and "can" that +modify application of a function. + +* ML antiquotations for conditional ML text: + + \<^if_linux>\...\ + \<^if_macos>\...\ + \<^if_windows>\...\ + \<^if_unix>\...\ + +* ML profiling has been updated and reactivated, after some degration in +Isabelle2021: + + - "isabelle build -o threads=1 -o profiling=..." works properly + within the PIDE session context; + + - "isabelle profiling_report" now uses the session build database + (like "isabelle log"); + + - output uses non-intrusive tracing messages, instead of warnings. + + +*** System *** + +* Almost complete support for arm64-linux platform. The reference +platform is Raspberry Pi 4 with 8 GB RAM running Pi OS (64 bit). + +* Update to OpenJDK 17: the current long-term support version of Java. + +* Update to Poly/ML 5.9 with improved support for ARM on Linux. On +macOS, the Intel version works more smoothly with Rosetta 2, as already +used in Isabelle2021. Further changes to Poly/ML are documented here: +http://lists.inf.ed.ac.uk/pipermail/polyml/2021-May/002451.html + +* Perl is no longer required by Isabelle proper, and no longer provided +by specific Isabelle execution environments (Docker, Cygwin on Windows). +Minor INCOMPATIBILITY, add-on applications involving perl need to +provide it by different means. (Note that proper Isabelle systems +programming works via Scala/Java, without perl, python, ruby etc.). + +* Each Isabelle component may specify a Scala/Java jar module +declaratively via etc/build.props (file names are relative to the +component directory). E.g. see $ISABELLE_HOME/etc/build.props with +further explanations in the "system" manual. + +* Command-line tool "isabelle scala_build" allows to invoke the build +process of all Scala/Java modules explicitly. Normally this is done +implicitly on demand, e.g. for "isabelle scala" or "isabelle jedit". + +* Command-line tool "isabelle scala_project" is has been improved in +various ways: + - sources from all components with etc/build.props are included, + - sources of for the jEdit text editor and the Isabelle/jEdit + plugins (jedit_base and jedit_main) are included by default, + - more sources may be given on the command-line, + - options -f and -D make the tool more convenient, + - Gradle has been replaced by Maven (less ambitious and more robust). + +* Remote provers from SystemOnTPTP (notably for Sledgehammer) are now +managed via Isabelle/Scala instead of perl; the dependency on +libwww-perl has been eliminated (notably on Linux). Rare +INCOMPATIBILITY: HTTP proxy configuration now works via JVM properties +https://docs.oracle.com/en/java/javase/11/docs/api/java.base/java/net/doc-files/net-properties.html + +* System options may declare an implicit standard value, which is used +when the option is activated without providing an explicit value, e.g. +"isabelle build -o document -o document_output" instead of +"isabelle build -o document=true -o document_output=output". For options +of type "bool", the standard is always "true" and cannot be specified +differently. + +* System option "document=true" is an alias for "document=pdf", and +"document=false" is an alias for "document=" (empty string). + +* System option "system_log" specifies an optional log file for internal +messages produced by Output.system_message in Isabelle/ML; the standard +value "-" refers to console progress of the build job. This works for +"isabelle build" or any derivative of it. + +* Command-line tool "isabelle version" supports repository archives +(without full .hg directory). It also provides more options. + +* Obsolete settings variable ISABELLE_PLATFORM32 has been discontinued. +Note that only Windows supports old 32 bit executables, via settings +variable ISABELLE_WINDOWS_PLATFORM32. Everything else should be +ISABELLE_PLATFORM64 (generic Posix) or ISABELLE_WINDOWS_PLATFORM64 +(native Windows) or ISABELLE_APPLE_PLATFORM64 (Apple Silicon). + +* Timeouts for Isabelle/ML tools are subject to system option +"timeout_scale", to support adjustments to slow machines. Before, +timeout_scale was only used for the overall session build process, now +it affects the underlying Timeout.apply in Isabelle/ML as well. It +treats a timeout specification 0 as "no timeout", instead of "immediate +timeout". Rare INCOMPATIBILITY in boundary cases. + + + +New in Isabelle2021 (February 2021) +----------------------------------- + +*** General *** + +* On macOS, the IsabelleXYZ.app directory layout now follows the other +platforms, without indirection via Contents/Resources/. INCOMPATIBILITY, +use e.g. IsabelleXYZ.app/bin/isabelle instead of former +IsabelleXYZ.app/Isabelle/bin/isabelle or +IsabelleXYZ.app/Isabelle/Contents/Resources/IsabelleXYZ/bin/isabelle. + +* HTML presentation uses rich markup produced by Isabelle/PIDE, +resulting in more colors and links. + +* HTML presentation includes auxiliary files (e.g. ML) for each theory. + +* Proof method "subst" is confined to the original subgoal range: its +included distinct_subgoals_tac no longer affects unrelated subgoals. +Rare INCOMPATIBILITY. + +* Theory_Data extend operation is obsolete and needs to be the identity +function; merge should be conservative and not reset to the empty value. +Subtle INCOMPATIBILITY and change of semantics (due to +Theory.join_theory from Isabelle2020). Special extend/merge behaviour at +the begin of a new theory can be achieved via Theory.at_begin. + + +*** Isabelle/jEdit Prover IDE *** + +* Improved GUI look-and-feel: the portable and scalable "FlatLaf Light" +is used by default on all platforms (appearance similar to IntelliJ +IDEA). + +* Improved markup for theory header imports: hyperlinks for theory files +work without formal checking of content. + +* The prover process can download auxiliary files (e.g. 'ML_file') for +theories with remote URL. This requires the external "curl" program. + +* Action "isabelle.goto-entity" (shortcut CS+d) jumps to the definition +of the formal entity at the caret position. + +* The visual feedback on caret entity focus is normally restricted to +definitions within the visible text area. The keyboard modifier "CS" +overrides this: then all defining and referencing positions are shown. +See also option "jedit_focus_modifier". + +* The jEdit status line includes widgets both for JVM and ML heap usage. +Ongoing ML ongoing garbage collection is shown as "ML cleanup". + +* The Monitor dockable provides buttons to request a full garbage +collection and sharing of live data on the ML heap. It also includes +information about the Java Runtime system. + +* PIDE support for session ROOTS: markup for directories. + +* Update to jedit-5.6.0, the latest release. This version works properly +on macOS by default, without the special MacOSX plugin. + +* Action "full-screen-mode" (shortcut F11 or S+F11) has been modified +for better approximate window size on macOS and Linux/X11. + +* Improved GUI support for macOS 11.1 Big Sur: native fullscreen mode, +but non-native look-and-feel (FlatLaf). + +* Hyperlinks to various file-formats (.pdf, .png, etc.) open an external +viewer, instead of re-using the jEdit text editor. + +* IDE support for Naproche-SAD: Proof Checking of Natural Mathematical +Documents. See also $NAPROCHE_HOME/examples for files with .ftl or +.ftl.tex extension. The corresponding Naproche-SAD server process can be +disabled by setting the system option naproche_server=false and +restarting the Isabelle application. + + +*** Document preparation *** + +* Keyword 'document_theories' within ROOT specifies theories from other +sessions that should be included in the generated document source +directory. This does not affect the generated session.tex: \input{...} +needs to be used separately. + +* The standard LaTeX engine is now lualatex, according to settings +variable ISABELLE_PDFLATEX. This is mostly upwards compatible with old +pdflatex, but text encoding needs to conform strictly to utf8. Rare +INCOMPATIBILITY. + +* Discontinued obsolete DVI format and ISABELLE_LATEX settings variable: +document output is always PDF. + +* Antiquotation @{tool} refers to Isabelle command-line tools, with +completion and formal reference to the source (external script or +internal Scala function). + +* Antiquotation @{bash_function} refers to GNU bash functions that are +checked within the Isabelle settings environment. + +* Antiquotations @{scala}, @{scala_object}, @{scala_type}, +@{scala_method} refer to checked Isabelle/Scala entities. + + +*** Pure *** + +* Session Pure-Examples contains notable examples for Isabelle/Pure +(former entries of HOL-Isar_Examples). + +* Named contexts (locale and class specifications, locale and class +context blocks) allow bundle mixins for the surface context. This allows +syntax notations to be organized within bundles conveniently. See theory +"HOL-ex.Specifications_with_bundle_mixins" for examples and the isar-ref +manual for syntax descriptions. + +* Definitions in locales produce rule which can be added as congruence +rule to protect foundational terms during simplification. + +* Consolidated terminology and function signatures for nested targets: + + - Local_Theory.begin_nested replaces Local_Theory.open_target + + - Local_Theory.end_nested replaces Local_Theory.close_target + + - Combination of Local_Theory.begin_nested and + Local_Theory.end_nested(_result) replaces + Local_Theory.subtarget(_result) + +INCOMPATIBILITY. + +* Local_Theory.init replaces Generic_Target.init. Minor INCOMPATIBILITY. + + +*** HOL *** + +* Session HOL-Examples contains notable examples for Isabelle/HOL +(former entries of HOL-Isar_Examples, HOL-ex etc.). + +* An updated version of the veriT solver is now included as Isabelle +component. It can be used in the "smt" proof method via "smt (verit)" or +via "declare [[smt_solver = verit]]" in the context; see also session +HOL-Word-SMT_Examples. + +* Zipperposition 2.0 is now included as Isabelle component for +experimentation, e.g. in "sledgehammer [prover = zipperposition]". + +* Sledgehammer: + - support veriT in proof preplay + - take adventage of more cores in proof preplay + +* Updated the Metis prover underlying the "metis" proof method to +version 2.4 (release 20180810). The new version fixes one soundness +defect and two incompleteness defects. Very slight INCOMPATIBILITY. + +* Nitpick/Kodkod may be invoked directly within the running +Isabelle/Scala session (instead of an external Java process): this +improves reactivity and saves resources. This experimental feature is +guarded by system option "kodkod_scala" (default: true in PIDE +interaction, false in batch builds). + +* Simproc "defined_all" and rewrite rule "subst_all" perform more +aggressive substitution with variables from assumptions. +INCOMPATIBILITY, consider repairing proofs locally like this: + + supply subst_all [simp del] [[simproc del: defined_all]] + +* Simproc "datatype_no_proper_subterm" rewrites equalities "lhs = rhs" +on datatypes to "False" if either side is a proper subexpression of the +other (for any datatype with a reasonable size function). + +* Syntax for state monad combinators fcomp and scomp is organized in +bundle state_combinator_syntax. Minor INCOMPATIBILITY. + +* Syntax for reflected term syntax is organized in bundle term_syntax, +discontinuing previous locale term_syntax. Minor INCOMPATIBILITY. + +* New constant "power_int" for exponentiation with integer exponent, +written as "x powi n". + +* Added the "at most 1" quantifier, Uniq. + +* For the natural numbers, "Sup {} = 0". + +* New constant semiring_char gives the characteristic of any type of +class semiring_1, with the convenient notation CHAR('a). For example, +CHAR(nat) = CHAR(int) = CHAR(real) = 0, CHAR(17) = 17. + +* HOL-Computational_Algebra.Polynomial: Definition and basic properties +of algebraic integers. + +* Library theory "Bit_Operations" with generic bit operations. + +* Library theory "Signed_Division" provides operations for signed +division, instantiated for type int. + +* Theory "Multiset": removed misleading notation \# for sum_mset; +replaced with \\<^sub>#. Analogous notation for prod_mset also exists now. + +* New theory "HOL-Library.Word" takes over material from former session +"HOL-Word". INCOMPATIBILITY: need to adjust imports. + +* Theory "HOL-Library.Word": Type word is restricted to bit strings +consisting of at least one bit. INCOMPATIBILITY. + +* Theory "HOL-Library.Word": Bit operations NOT, AND, OR, XOR are based +on generic algebraic bit operations from theory +"HOL-Library.Bit_Operations". INCOMPATIBILITY. + +* Theory "HOL-Library.Word": Most operations on type word are set up for +transfer and lifting. INCOMPATIBILITY. + +* Theory "HOL-Library.Word": Generic type conversions. INCOMPATIBILITY, +sometimes additional rewrite rules must be added to applications to get +a confluent system again. + +* Theory "HOL-Library.Word": Uniform polymorphic "mask" operation for +both types int and word. INCOMPATIBILITY. + +* Theory "HOL-Library.Word": Syntax for signed compare operators has +been consolidated with syntax of regular compare operators. Minor +INCOMPATIBILITY. + +* Former session "HOL-Word": Various operations dealing with bit values +represented as reversed lists of bools are separated into theory +Reversed_Bit_Lists in session Word_Lib in the AFP. INCOMPATIBILITY. + +* Former session "HOL-Word": Theory "Word_Bitwise" has been moved to AFP +entry Word_Lib as theory "Bitwise". INCOMPATIBILITY. + +* Former session "HOL-Word": Compound operation "bin_split" simplifies +by default into its components "drop_bit" and "take_bit". +INCOMPATIBILITY. + +* Former session "HOL-Word": Operations lsb, msb and set_bit are +separated into theories Least_significant_bit, Most_significant_bit and +Generic_set_bit respectively in session Word_Lib in the AFP. +INCOMPATIBILITY. + +* Former session "HOL-Word": Ancient int numeral representation has been +factored out in separate theory "Ancient_Numeral" in session Word_Lib in +the AFP. INCOMPATIBILITY. + +* Former session "HOL-Word": Operations "bin_last", "bin_rest", +"bin_nth", "bintrunc", "sbintrunc", "norm_sint", "bin_cat" and +"max_word" are now mere input abbreviations. Minor INCOMPATIBILITY. + +* Former session "HOL-Word": Misc ancient material has been factored out +into separate theories and moved to session Word_Lib in the AFP. See +theory "Guide" there for further information. INCOMPATIBILITY. + +* Session HOL-TPTP: The "tptp_isabelle" and "tptp_sledgehammer" commands +are in working order again, as opposed to outputting "GaveUp" on nearly +all problems. + +* Session "HOL-Hoare": concrete syntax only for Hoare triples, not +abstract language constructors. + +* Session "HOL-Hoare": now provides a total correctness logic as well. + + +*** FOL *** + +* Added the "at most 1" quantifier, Uniq, as in HOL. + +* Simproc "defined_all" and rewrite rule "subst_all" have been changed +as in HOL. + + +*** ML *** + +* Antiquotations @{scala_function}, @{scala}, @{scala_thread} refer to +registered Isabelle/Scala functions (of type String => String): +invocation works via the PIDE protocol. + +* Path.append is available as overloaded "+" operator, similar to +corresponding Isabelle/Scala operation. + +* ML statistics via an external Poly/ML process: this allows monitoring +the runtime system while the ML program sleeps. + + +*** System *** + +* Isabelle server allows user-defined commands via +isabelle_scala_service. + +* Update/rebuild external provers on currently supported OS platforms, +notably CVC4 1.8, E prover 2.5, SPASS 3.8ds, CSDP 6.1.1. + +* The command-line tool "isabelle log" prints prover messages from the +build database of the given session, following the the order of theory +sources, instead of erratic parallel evaluation. Consequently, the +session log file is restricted to system messages of the overall build +process, and thus becomes more informative. + +* Discontinued obsolete isabelle display tool, and DVI_VIEWER settings +variable. + +* The command-line tool "isabelle logo" only outputs PDF; obsolete EPS +(for DVI documents) has been discontinued. Former option -n has been +turned into -o with explicit file name. Minor INCOMPATIBILITY. + +* The command-line tool "isabelle components" supports new options -u +and -x to manage $ISABELLE_HOME_USER/etc/components without manual +editing of Isabelle configuration files. + +* The shell function "isabelle_directory" (within etc/settings of +components) augments the list of special directories for persistent +symbolic path names. This improves portability of heap images and +session databases. It used to be hard-wired for Isabelle + AFP, but +other projects may now participate on equal terms. + +* The command-line tool "isabelle process" now prints output to +stdout/stderr separately and incrementally, instead of just one bulk to +stdout after termination. Potential INCOMPATIBILITY for external tools. + +* The command-line tool "isabelle console" now supports interrupts +properly (on Linux and macOS). + +* Batch-builds via "isabelle build" use a PIDE session with special +protocol: this allows to invoke Isabelle/Scala operations from +Isabelle/ML. Big build jobs (e.g. AFP) require extra heap space for the +java process, e.g. like this in $ISABELLE_HOME_USER/etc/settings: + + ISABELLE_TOOL_JAVA_OPTIONS="$ISABELLE_TOOL_JAVA_OPTIONS -Xmx8g" + +This includes full PIDE markup, if option "build_pide_reports" is +enabled. + +* The command-line tool "isabelle build" provides option -P DIR to +produce PDF/HTML presentation in the specified directory; -P: refers to +the standard directory according to ISABELLE_BROWSER_INFO / +ISABELLE_BROWSER_INFO_SYSTEM settings. Generated PDF documents are taken +from the build database -- from this or earlier builds with option +document=pdf. + +* The command-line tool "isabelle document" generates theory documents +on the spot, using the underlying session build database (exported +LaTeX sources or existing PDF files). INCOMPATIBILITY, the former +"isabelle document" tool was rather different and has been discontinued. + +* The command-line tool "isabelle sessions" explores the structure of +Isabelle sessions and prints result names in topological order (on +stdout). + +* The Isabelle/Scala "Progress" interface changed slightly and +"No_Progress" has been discontinued. INCOMPATIBILITY, use "new Progress" +instead. + +* General support for Isabelle/Scala system services, configured via the +shell function "isabelle_scala_service" in etc/settings (e.g. of an +Isabelle component); see implementations of class +Isabelle_System.Service in Isabelle/Scala. This supersedes former +"isabelle_scala_tools" and "isabelle_file_format": minor +INCOMPATIBILITY. + +* The syntax of theory load commands (for auxiliary files) is now +specified in Isabelle/Scala, as instance of class +isabelle.Command_Span.Load_Command registered via isabelle_scala_service +in etc/settings. This allows more flexible schemes than just a list of +file extensions. Minor INCOMPATIBILITY, e.g. see theory +HOL-SPARK.SPARK_Setup to emulate the old behaviour. + +* JVM system property "isabelle.laf" has been discontinued; the default +Swing look-and-feel is ""FlatLaf Light". + +* Isabelle/Phabricator supports Ubuntu 20.04 LTS. + +* Isabelle/Phabricator setup has been updated to follow ongoing +development: libphutil has been discontinued. Minor INCOMPATIBILITY: +existing server installations should remove libphutil from +/usr/local/bin/isabelle-phabricator-upgrade and each installation root +directory (e.g. /var/www/phabricator-vcs/libphutil). + +* Experimental support for arm64-linux platform. The reference platform +is Raspberry Pi 4 with 8 GB RAM running Pi OS (64 bit). + +* Support for Apple Silicon, using mostly x86_64-darwin runtime +translation via Rosetta 2 (e.g. Poly/ML and external provers), but also +some native arm64-darwin executables (e.g. Java). + + + +New in Isabelle2020 (April 2020) +-------------------------------- + +*** General *** + +* Session ROOT files need to specify explicit 'directories' for import +of theory files. Directories cannot be shared by different sessions. +(Recall that import of theories from other sessions works via +session-qualified theory names, together with suitable 'sessions' +declarations in the ROOT.) + +* Internal derivations record dependencies on oracles and other theorems +accurately, including the implicit type-class reasoning wrt. proven +class relations and type arities. In particular, the formal tagging with +"Pure.skip_proofs" of results stemming from "instance ... sorry" is now +propagated properly to theorems depending on such type instances. + +* Command 'sorry' (oracle "Pure.skip_proofs") is more precise about the +actual proposition that is assumed in the goal and proof context. This +requires at least Proofterm.proofs = 1 to show up in theorem +dependencies. + +* Command 'thm_oracles' prints all oracles used in given theorems, +covering the full graph of transitive dependencies. + +* Command 'thm_deps' prints immediate theorem dependencies of the given +facts. The former graph visualization has been discontinued, because it +was hardly usable. + +* Refined treatment of proof terms, including type-class proofs for +minor object-logics (FOL, FOLP, Sequents). + +* The inference kernel is now confined to one main module: structure +Thm, without the former circular dependency on structure Axclass. + +* Mixfix annotations may use "' " (single quote followed by space) to +separate delimiters (as documented in the isar-ref manual), without +requiring an auxiliary empty block. A literal single quote needs to be +escaped properly. Minor INCOMPATIBILITY. + + +*** Isar *** + +* The proof method combinator (subproofs m) applies the method +expression m consecutively to each subgoal, constructing individual +subproofs internally. This impacts the internal construction of proof +terms: it makes a cascade of let-expressions within the derivation tree +and may thus improve scalability. + +* Attribute "trace_locales" activates tracing of locale instances during +roundup. It replaces the diagnostic command 'print_dependencies', which +has been discontinued. + + +*** Isabelle/jEdit Prover IDE *** + +* Prover IDE startup is now much faster, because theory dependencies are +no longer explored in advance. The overall session structure with its +declarations of 'directories' is sufficient to locate theory files. Thus +the "session focus" of option "isabelle jedit -S" has become obsolete +(likewise for "isabelle vscode_server -S"). Existing option "-R" is both +sufficient and more convenient to start editing a particular session. + +* Actions isabelle.tooltip (CS+b) and isabelle.message (CS+m) display +tooltip message popups, corresponding to mouse hovering with/without the +CONTROL/COMMAND key pressed. + +* The following actions allow to navigate errors within the current +document snapshot: + + isabelle.first-error (CS+a) + isabelle.last-error (CS+z) + isabelle.next-error (CS+n) + isabelle.prev-error (CS+p) + +* Support more brackets: \ \ (intended for implicit argument syntax). + +* Action isabelle.jconsole (menu item Plugins / Isabelle / Java/VM +Monitor) applies the jconsole tool on the running Isabelle/jEdit +process. This allows to monitor resource usage etc. + +* More adequate default font sizes for Linux on HD / UHD displays: +automatic font scaling is usually absent on Linux, in contrast to +Windows and macOS. + +* The default value for the jEdit property "view.antiAlias" (menu item +Utilities / Global Options / Text Area / Anti Aliased smooth text) is +now "subpixel HRGB", instead of former "standard". Especially on Linux +this often leads to faster text rendering, but can also cause problems +with odd color shades. An alternative is to switch back to "standard" +here, and set the following Java system property: + + isabelle jedit -Dsun.java2d.opengl=true + +This can be made persistent via JEDIT_JAVA_OPTIONS in +$ISABELLE_HOME_USER/etc/settings. For the "Isabelle2020" desktop +application there is a corresponding options file in the same directory. + + +*** Isabelle/VSCode Prover IDE *** + +* Update of State and Preview panels to use new WebviewPanel API of +VSCode. + + +*** HOL *** + +* Improvements of the 'lift_bnf' command: + - Add support for quotient types. + - Generate transfer rules for the lifted map/set/rel/pred constants + (theorems "._transfer_raw"). + +* Term_XML.Encode/Decode.term uses compact representation of Const +"typargs" from the given declaration environment. This also makes more +sense for translations to lambda-calculi with explicit polymorphism. +INCOMPATIBILITY, use Term_XML.Encode/Decode.term_raw in special +applications. + +* ASCII membership syntax concerning big operators for infimum and +supremum has been discontinued. INCOMPATIBILITY. + +* Removed multiplicativity assumption from class +"normalization_semidom". Introduced various new intermediate classes +with the multiplicativity assumption; many theorem statements +(especially involving GCD/LCM) had to be adapted. This allows for a more +natural instantiation of the algebraic typeclasses for e.g. Gaussian +integers. INCOMPATIBILITY. + +* Clear distinction between types for bits (False / True) and Z2 (0 / +1): theory HOL-Library.Bit has been renamed accordingly. +INCOMPATIBILITY. + +* Dynamic facts "algebra_split_simps" and "field_split_simps" correspond +to algebra_simps and field_simps but contain more aggressive rules +potentially splitting goals; algebra_split_simps roughly replaces +sign_simps and field_split_simps can be used instead of divide_simps. +INCOMPATIBILITY. + +* Theory HOL.Complete_Lattices: +renamed Inf_Sup -> Inf_eq_Sup and Sup_Inf -> Sup_eq_Inf + +* Theory HOL-Library.Monad_Syntax: infix operation "bind" (\) +associates to the left now as is customary. + +* Theory HOL-Library.Ramsey: full finite Ramsey's theorem with +multiple colours and arbitrary exponents. + +* Session HOL-Proofs: build faster thanks to better treatment of proof +terms in Isabelle/Pure. + +* Session HOL-Word: bitwise NOT-operator has proper prefix syntax. Minor +INCOMPATIBILITY. + +* Session HOL-Analysis: proof method "metric" implements a decision +procedure for simple linear statements in metric spaces. + +* Session HOL-Complex_Analysis has been split off from HOL-Analysis. + + +*** ML *** + +* Theory construction may be forked internally, the operation +Theory.join_theory recovers a single result theory. See also the example +in theory "HOL-ex.Join_Theory". + +* Antiquotation @{oracle_name} inlines a formally checked oracle name. + +* Minimal support for a soft-type system within the Isabelle logical +framework (module Soft_Type_System). + +* Former Variable.auto_fixes has been replaced by slightly more general +Proof_Context.augment: it is subject to an optional soft-type system of +the underlying object-logic. Minor INCOMPATIBILITY. + +* More scalable Export.export using XML.tree to avoid premature string +allocations, with convenient shortcut XML.blob. Minor INCOMPATIBILITY. + +* Prover IDE support for the underlying Poly/ML compiler (not the basis +library). Open $ML_SOURCES/ROOT.ML in Isabelle/jEdit to browse the +implementation with full markup. + + +*** System *** + +* Standard rendering for more Isabelle symbols: \ \ \ \ + +* The command-line tool "isabelle scala_project" creates a Gradle +project configuration for Isabelle/Scala/jEdit, to support Scala IDEs +such as IntelliJ IDEA. + +* The command-line tool "isabelle phabricator_setup" facilitates +self-hosting of the Phabricator software-development platform, with +support for Git, Mercurial, Subversion repositories. This helps to avoid +monoculture and to escape the gravity of centralized version control by +Github and/or Bitbucket. For further documentation, see chapter +"Phabricator server administration" in the "system" manual. A notable +example installation is https://isabelle-dev.sketis.net/. + +* The command-line tool "isabelle hg_setup" simplifies the setup of +Mercurial repositories, with hosting via Phabricator or SSH file server +access. + +* The command-line tool "isabelle imports" has been discontinued: strict +checking of session directories enforces session-qualified theory names +in applications -- users are responsible to specify session ROOT entries +properly. + +* The command-line tool "isabelle dump" and its underlying +Isabelle/Scala module isabelle.Dump has become more scalable, by +splitting sessions and supporting a base logic image. Minor +INCOMPATIBILITY in options and parameters. + +* The command-line tool "isabelle build_docker" has been slightly +improved: it is now properly documented in the "system" manual. + +* Isabelle/Scala support for the Linux platform (Ubuntu): packages, +users, system services. + +* Isabelle/Scala support for proof terms (with full type/term +information) in module isabelle.Term. + +* Isabelle/Scala: more scalable output of YXML files, e.g. relevant for +"isabelle dump". + +* Theory export via Isabelle/Scala has been reworked. The former "fact" +name space is now split into individual "thm" items: names are +potentially indexed, such as "foo" for singleton facts, or "bar(1)", +"bar(2)", "bar(3)" for multi-facts. Theorem dependencies are now +exported as well: this spans an overall dependency graph of internal +inferences; it might help to reconstruct the formal structure of theory +libraries. See also the module isabelle.Export_Theory in Isabelle/Scala. + +* Theory export of structured specifications, based on internal +declarations of Spec_Rules by packages like 'definition', 'inductive', +'primrec', 'function'. + +* Old settings variables ISABELLE_PLATFORM and ISABELLE_WINDOWS_PLATFORM +have been discontinued -- deprecated since Isabelle2018. + +* More complete x86_64 platform support on macOS, notably Catalina where +old x86 has been discontinued. + +* Update to GHC stack 2.1.3 with stackage lts-13.19/ghc-8.6.4. + +* Update to OCaml Opam 2.0.6 (using ocaml 4.05.0 as before). + + + +New in Isabelle2019 (June 2019) +------------------------------- + +*** General *** + +* The font collection "Isabelle DejaVu" is systematically derived from +the existing "DejaVu" fonts, with variants "Sans Mono", "Sans", "Serif" +and styles "Normal", "Bold", "Italic/Oblique", "Bold-Italic/Oblique". +The DejaVu base fonts are retricted to well-defined Unicode ranges and +augmented by special Isabelle symbols, taken from the former +"IsabelleText" font (which is no longer provided separately). The line +metrics and overall rendering quality is closer to original DejaVu. +INCOMPATIBILITY with display configuration expecting the old +"IsabelleText" font: use e.g. "Isabelle DejaVu Sans Mono" instead. + +* The Isabelle fonts render "\" properly as superscript "-1". + +* Old-style inner comments (* ... *) within the term language are no +longer supported (legacy feature in Isabelle2018). + +* Old-style {* verbatim *} tokens are explicitly marked as legacy +feature and will be removed soon. Use \cartouche\ syntax instead, e.g. +via "isabelle update_cartouches -t" (available since Isabelle2015). + +* Infix operators that begin or end with a "*" are now parenthesized +without additional spaces, e.g. "(*)" instead of "( * )". Minor +INCOMPATIBILITY. + +* Mixfix annotations may use cartouches instead of old-style double +quotes, e.g. (infixl \+\ 60). The command-line tool "isabelle update -u +mixfix_cartouches" allows to update existing theory sources +automatically. + +* ML setup commands (e.g. 'setup', 'method_setup', 'parse_translation') +need to provide a closed expression -- without trailing semicolon. Minor +INCOMPATIBILITY. + +* Commands 'generate_file', 'export_generated_files', and +'compile_generated_files' support a stateless (PIDE-conformant) model +for generated sources and compiled binaries of other languages. The +compilation process is managed in Isabelle/ML, and results exported to +the session database for further use (e.g. with "isabelle export" or +"isabelle build -e"). + + +*** Isabelle/jEdit Prover IDE *** + +* Fonts for the text area, gutter, GUI elements etc. use the "Isabelle +DejaVu" collection by default, which provides uniform rendering quality +with the usual Isabelle symbols. Line spacing no longer needs to be +adjusted: properties for the old IsabelleText font had "Global Options / +Text Area / Extra vertical line spacing (in pixels): -2", it now +defaults to 1, but 0 works as well. + +* The jEdit File Browser is more prominent in the default GUI layout of +Isabelle/jEdit: various virtual file-systems provide access to Isabelle +resources, notably via "favorites:" (or "Edit Favorites"). + +* Further markup and rendering for "plain text" (e.g. informal prose) +and "raw text" (e.g. verbatim sources). This improves the visual +appearance of formal comments inside the term language, or in general +for repeated alternation of formal and informal text. + +* Action "isabelle-export-browser" points the File Browser to the theory +exports of the current buffer, based on the "isabelle-export:" virtual +file-system. The directory view needs to be reloaded manually to follow +ongoing document processing. + +* Action "isabelle-session-browser" points the File Browser to session +information, based on the "isabelle-session:" virtual file-system. Its +entries are structured according to chapter / session names, the open +operation is redirected to the session ROOT file. + +* Support for user-defined file-formats via class isabelle.File_Format +in Isabelle/Scala (e.g. see isabelle.Bibtex.File_Format), configured via +the shell function "isabelle_file_format" in etc/settings (e.g. of an +Isabelle component). + +* System option "jedit_text_overview" allows to disable the text +overview column. + +* Command-line options "-s" and "-u" of "isabelle jedit" override the +default for system option "system_heaps" that determines the heap +storage directory for "isabelle build". Option "-n" is now clearly +separated from option "-s". + +* The Isabelle/jEdit desktop application uses the same options as +"isabelle jedit" for its internal "isabelle build" process: the implicit +option "-o system_heaps" (or "-s") has been discontinued. This reduces +the potential for surprise wrt. command-line tools. + +* The official download of the Isabelle/jEdit application already +contains heap images for Isabelle/HOL within its main directory: thus +the first encounter becomes faster and more robust (e.g. when run from a +read-only directory). + +* Isabelle DejaVu fonts are available with hinting by default, which is +relevant for low-resolution displays. This may be disabled via system +option "isabelle_fonts_hinted = false" in +$ISABELLE_HOME_USER/etc/preferences -- it occasionally yields better +results. + +* OpenJDK 11 has quite different font rendering, with better glyph +shapes and improved sub-pixel anti-aliasing. In some situations results +might be *worse* than Oracle Java 8, though -- a proper HiDPI / UHD +display is recommended. + +* OpenJDK 11 supports GTK version 2.2 and 3 (according to system +property jdk.gtk.version). The factory default is version 3, but +ISABELLE_JAVA_SYSTEM_OPTIONS includes "-Djdk.gtk.version=2.2" to make +this more conservative (as in Java 8). Depending on the GTK theme +configuration, "-Djdk.gtk.version=3" might work better or worse. + + +*** Document preparation *** + +* Document markers are formal comments of the form \<^marker>\marker_body\ that +are stripped from document output: the effect is to modify the semantic +presentation context or to emit markup to the PIDE document. Some +predefined markers are taken from the Dublin Core Metadata Initiative, +e.g. \<^marker>\contributor arg\ or \<^marker>\license arg\ and produce PIDE markup that +can be retrieved from the document database. + +* Old-style command tags %name are re-interpreted as markers with +proof-scope \<^marker>\tag (proof) name\ and produce LaTeX environments as +before. Potential INCOMPATIBILITY: multiple markers are composed in +canonical order, resulting in a reversed list of tags in the +presentation context. + +* Marker \<^marker>\tag name\ does not apply to the proof of a top-level goal +statement by default (e.g. 'theorem', 'lemma'). This is a subtle change +of semantics wrt. old-style %name. + +* In Isabelle/jEdit, the string "\tag" may be completed to a "\<^marker>\tag \" +template. + +* Document antiquotation option "cartouche" indicates if the output +should be delimited as cartouche; this takes precedence over the +analogous option "quotes". + +* Many document antiquotations are internally categorized as "embedded" +and expect one cartouche argument, which is typically used with the +\<^control>\cartouche\ notation (e.g. \<^term>\\x y. x\). The cartouche +delimiters are stripped in output of the source (antiquotation option +"source"), but it is possible to enforce delimiters via option +"source_cartouche", e.g. @{term [source_cartouche] \\x y. x\}. + + +*** Isar *** + +* Implicit cases goal1, goal2, goal3, etc. have been discontinued +(legacy feature since Isabelle2016). + +* More robust treatment of structural errors: begin/end blocks take +precedence over goal/proof. This is particularly relevant for the +headless PIDE session and server. + +* Command keywords of kind thy_decl / thy_goal may be more specifically +fit into the traditional document model of "definition-statement-proof" +via thy_defn / thy_stmt / thy_goal_defn / thy_goal_stmt. + + +*** HOL *** + +* Command 'export_code' produces output as logical files within the +theory context, as well as formal session exports that can be +materialized via command-line tools "isabelle export" or "isabelle build +-e" (with 'export_files' in the session ROOT). Isabelle/jEdit also +provides a virtual file-system "isabelle-export:" that can be explored +in the regular file-browser. A 'file_prefix' argument allows to specify +an explicit name prefix for the target file (SML, OCaml, Scala) or +directory (Haskell); the default is "export" with a consecutive number +within each theory. + +* Command 'export_code': the 'file' argument is now legacy and will be +removed soon: writing to the physical file-system is not well-defined in +a reactive/parallel application like Isabelle. The empty 'file' argument +has been discontinued already: it is superseded by the file-browser in +Isabelle/jEdit on "isabelle-export:". Minor INCOMPATIBILITY. + +* Command 'code_reflect' no longer supports the 'file' argument: it has +been superseded by 'file_prefix' for stateless file management as in +'export_code'. Minor INCOMPATIBILITY. + +* Code generation for OCaml: proper strings are used for literals. +Minor INCOMPATIBILITY. + +* Code generation for OCaml: Zarith supersedes Nums as library for +proper integer arithmetic. The library is located via standard +invocations of "ocamlfind" (via ISABELLE_OCAMLFIND settings variable). +The environment provided by "isabelle ocaml_setup" already contains this +tool and the required packages. Minor INCOMPATIBILITY. + +* Code generation for Haskell: code includes for Haskell must contain +proper module frame, nothing is added magically any longer. +INCOMPATIBILITY. + +* Code generation: slightly more conventional syntax for 'code_stmts' +antiquotation. Minor INCOMPATIBILITY. + +* Theory List: the precedence of the list_update operator has changed: +"f a [n := x]" now needs to be written "(f a)[n := x]". + +* The functions \, \, \, \ (not the corresponding binding operators) +now have the same precedence as any other prefix function symbol. Minor +INCOMPATIBILITY. + +* Simplified syntax setup for big operators under image. In rare +situations, type conversions are not inserted implicitly any longer +and need to be given explicitly. Auxiliary abbreviations INFIMUM, +SUPREMUM, UNION, INTER should now rarely occur in output and are just +retained as migration auxiliary. Abbreviations MINIMUM and MAXIMUM +are gone INCOMPATIBILITY. + +* The simplifier uses image_cong_simp as a congruence rule. The historic +and not really well-formed congruence rules INF_cong*, SUP_cong*, are +not used by default any longer. INCOMPATIBILITY; consider using declare +image_cong_simp [cong del] in extreme situations. + +* INF_image and SUP_image are no default simp rules any longer. +INCOMPATIBILITY, prefer image_comp as simp rule if needed. + +* Strong congruence rules (with =simp=> in the premises) for constant f +are now uniformly called f_cong_simp, in accordance with congruence +rules produced for mappers by the datatype package. INCOMPATIBILITY. + +* Retired lemma card_Union_image; use the simpler card_UN_disjoint +instead. INCOMPATIBILITY. + +* Facts sum_mset.commute and prod_mset.commute have been renamed to +sum_mset.swap and prod_mset.swap, similarly to sum.swap and prod.swap. +INCOMPATIBILITY. + +* ML structure Inductive: slightly more conventional naming schema. +Minor INCOMPATIBILITY. + +* ML: Various _global variants of specification tools have been removed. +Minor INCOMPATIBILITY, prefer combinators +Named_Target.theory_map[_result] to lift specifications to the global +theory level. + +* Theory HOL-Library.Simps_Case_Conv: 'case_of_simps' now supports +overlapping and non-exhaustive patterns and handles arbitrarily nested +patterns. It uses on the same algorithm as HOL-Library.Code_Lazy, which +assumes sequential left-to-right pattern matching. The generated +equation no longer tuples the arguments on the right-hand side. +INCOMPATIBILITY. + +* Theory HOL-Library.Multiset: the \# operator now has the same +precedence as any other prefix function symbol. + +* Theory HOL-Library.Cardinal_Notations has been discontinued in favor +of the bundle cardinal_syntax (available in theory Main). Minor +INCOMPATIBILITY. + +* Session HOL-Library and HOL-Number_Theory: Exponentiation by squaring, +used for computing powers in class "monoid_mult" and modular +exponentiation. + +* Session HOL-Computational_Algebra: Formal Laurent series and overhaul +of Formal power series. + +* Session HOL-Number_Theory: More material on residue rings in +Carmichael's function, primitive roots, more properties for "ord". + +* Session HOL-Analysis: Better organization and much more material +at the level of abstract topological spaces. + +* Session HOL-Algebra: Free abelian groups, etc., ported from HOL Light; + algebraic closure of a field by de Vilhena and Baillon. + +* Session HOL-Homology has been added. It is a port of HOL Light's +homology library, with new proofs of "invariance of domain" and related +results. + +* Session HOL-SPARK: .prv files are no longer written to the +file-system, but exported to the session database. Results may be +retrieved via "isabelle build -e HOL-SPARK-Examples" on the +command-line. + +* Sledgehammer: + - The URL for SystemOnTPTP, which is used by remote provers, has been + updated. + - The machine-learning-based filter MaSh has been optimized to take + less time (in most cases). + +* SMT: reconstruction is now possible using the SMT solver veriT. + +* Session HOL-Word: + * New theory More_Word as comprehensive entrance point. + * Merged type class bitss into type class bits. + INCOMPATIBILITY. + + +*** ML *** + +* Command 'generate_file' allows to produce sources for other languages, +with antiquotations in the Isabelle context (only the control-cartouche +form). The default "cartouche" antiquotation evaluates an ML expression +of type string and inlines the result as a string literal of the target +language. For example, this works for Haskell as follows: + + generate_file "Pure.hs" = \ + module Isabelle.Pure where + allConst, impConst, eqConst :: String + allConst = \\<^const_name>\Pure.all\\ + impConst = \\<^const_name>\Pure.imp\\ + eqConst = \\<^const_name>\Pure.eq\\ + \ + +See also commands 'export_generated_files' and 'compile_generated_files' +to use the results. + +* ML evaluation (notably via command 'ML' or 'ML_file') is subject to +option ML_environment to select a named environment, such as "Isabelle" +for Isabelle/ML, or "SML" for official Standard ML. + +* ML antiquotation @{master_dir} refers to the master directory of the +underlying theory, i.e. the directory of the theory file. + +* ML antiquotation @{verbatim} inlines its argument as string literal, +preserving newlines literally. The short form \<^verbatim>\abc\ is particularly +useful. + +* Local_Theory.reset is no longer available in user space. Regular +definitional packages should use balanced blocks of +Local_Theory.open_target versus Local_Theory.close_target instead, or +the Local_Theory.subtarget(_result) combinator. Rare INCOMPATIBILITY. + +* Original PolyML.pointerEq is retained as a convenience for tools that +don't use Isabelle/ML (where this is called "pointer_eq"). + + +*** System *** + +* Update to OpenJDK 11: the current long-term support version of Java. + +* Update to Poly/ML 5.8 allows to use the native x86_64 platform without +the full overhead of 64-bit values everywhere. This special x86_64_32 +mode provides up to 16GB ML heap, while program code and stacks are +allocated elsewhere. Thus approx. 5 times more memory is available for +applications compared to old x86 mode (which is no longer used by +Isabelle). The switch to the x86_64 CPU architecture also avoids +compatibility problems with Linux and macOS, where 32-bit applications +are gradually phased out. + +* System option "checkpoint" has been discontinued: obsolete thanks to +improved memory management in Poly/ML. + +* System option "system_heaps" determines where to store the session +image of "isabelle build" (and other tools using that internally). +Former option "-s" is superseded by option "-o system_heaps". +INCOMPATIBILITY in command-line syntax. + +* Session directory $ISABELLE_HOME/src/Tools/Haskell provides some +source modules for Isabelle tools implemented in Haskell, notably for +Isabelle/PIDE. + +* The command-line tool "isabelle build -e" retrieves theory exports +from the session build database, using 'export_files' in session ROOT +entries. + +* The command-line tool "isabelle update" uses Isabelle/PIDE in +batch-mode to update theory sources based on semantic markup produced in +Isabelle/ML. Actual updates depend on system options that may be enabled +via "-u OPT" (for "update_OPT"), see also $ISABELLE_HOME/etc/options +section "Theory update". Theory sessions are specified as in "isabelle +dump". + +* The command-line tool "isabelle update -u control_cartouches" changes +antiquotations into control-symbol format (where possible): @{NAME} +becomes \<^NAME> and @{NAME ARG} becomes \<^NAME>\ARG\. + +* Support for Isabelle command-line tools defined in Isabelle/Scala. +Instances of class Isabelle_Scala_Tools may be configured via the shell +function "isabelle_scala_tools" in etc/settings (e.g. of an Isabelle +component). + +* Isabelle Server command "use_theories" supports "nodes_status_delay" +for continuous output of node status information. The time interval is +specified in seconds; a negative value means it is disabled (default). + +* Isabelle Server command "use_theories" terminates more robustly in the +presence of structurally broken sources: full consolidation of theories +is no longer required. + +* OCaml tools and libraries are now accesed via ISABELLE_OCAMLFIND, +which needs to point to a suitable version of "ocamlfind" (e.g. via +OPAM, see below). INCOMPATIBILITY: settings variables ISABELLE_OCAML and +ISABELLE_OCAMLC are no longer supported. + +* Support for managed installations of Glasgow Haskell Compiler and +OCaml via the following command-line tools: + + isabelle ghc_setup + isabelle ghc_stack + + isabelle ocaml_setup + isabelle ocaml_opam + +The global installation state is determined by the following settings +(and corresponding directory contents): + + ISABELLE_STACK_ROOT + ISABELLE_STACK_RESOLVER + ISABELLE_GHC_VERSION + + ISABELLE_OPAM_ROOT + ISABELLE_OCAML_VERSION + +After setup, the following Isabelle settings are automatically +redirected (overriding existing user settings): + + ISABELLE_GHC + + ISABELLE_OCAMLFIND + +The old meaning of these settings as locally installed executables may +be recovered by purging the directories ISABELLE_STACK_ROOT / +ISABELLE_OPAM_ROOT, or by resetting these variables in +$ISABELLE_HOME_USER/etc/settings. + + + +New in Isabelle2018 (August 2018) +--------------------------------- + +*** General *** + +* Session-qualified theory names are mandatory: it is no longer possible +to refer to unqualified theories from the parent session. +INCOMPATIBILITY for old developments that have not been updated to +Isabelle2017 yet (using the "isabelle imports" tool). + +* Only the most fundamental theory names are global, usually the entry +points to major logic sessions: Pure, Main, Complex_Main, HOLCF, IFOL, +FOL, ZF, ZFC etc. INCOMPATIBILITY, need to use qualified names for +formerly global "HOL-Probability.Probability" and "HOL-SPARK.SPARK". + +* Global facts need to be closed: no free variables and no hypotheses. +Rare INCOMPATIBILITY. + +* Facts stemming from locale interpretation are subject to lazy +evaluation for improved performance. Rare INCOMPATIBILITY: errors +stemming from interpretation morphisms might be deferred and thus +difficult to locate; enable system option "strict_facts" temporarily to +avoid this. + +* Marginal comments need to be written exclusively in the new-style form +"\ \text\", old ASCII variants like "-- {* ... *}" are no longer +supported. INCOMPATIBILITY, use the command-line tool "isabelle +update_comments" to update existing theory files. + +* Old-style inner comments (* ... *) within the term language are legacy +and will be discontinued soon: use formal comments "\ \...\" or "\<^cancel>\...\" +instead. + +* The "op " syntax for infix operators has been replaced by +"()". If begins or ends with a "*", there needs to +be a space between the "*" and the corresponding parenthesis. +INCOMPATIBILITY, use the command-line tool "isabelle update_op" to +convert theory and ML files to the new syntax. Because it is based on +regular expression matching, the result may need a bit of manual +postprocessing. Invoking "isabelle update_op" converts all files in the +current directory (recursively). In case you want to exclude conversion +of ML files (because the tool frequently also converts ML's "op" +syntax), use option "-m". + +* Theory header 'abbrevs' specifications need to be separated by 'and'. +INCOMPATIBILITY. + +* Command 'external_file' declares the formal dependency on the given +file name, such that the Isabelle build process knows about it, but +without specific Prover IDE management. + +* Session ROOT entries no longer allow specification of 'files'. Rare +INCOMPATIBILITY, use command 'external_file' within a proper theory +context. + +* Session root directories may be specified multiple times: each +accessible ROOT file is processed only once. This facilitates +specification of $ISABELLE_HOME_USER/ROOTS or command-line options like +-d or -D for "isabelle build" and "isabelle jedit". Example: + + isabelle build -D '~~/src/ZF' + +* The command 'display_drafts' has been discontinued. INCOMPATIBILITY, +use action "isabelle.draft" (or "print") in Isabelle/jEdit instead. + +* In HTML output, the Isabelle symbol "\" is rendered as explicit +Unicode hyphen U+2010, to avoid unclear meaning of the old "soft hyphen" +U+00AD. Rare INCOMPATIBILITY, e.g. copy-paste of historic Isabelle HTML +output. + + +*** Isabelle/jEdit Prover IDE *** + +* The command-line tool "isabelle jedit" provides more flexible options +for session management: + + - option -R builds an auxiliary logic image with all theories from + other sessions that are not already present in its parent + + - option -S is like -R, with a focus on the selected session and its + descendants (this reduces startup time for big projects like AFP) + + - option -A specifies an alternative ancestor session for options -R + and -S + + - option -i includes additional sessions into the name-space of + theories + + Examples: + isabelle jedit -R HOL-Number_Theory + isabelle jedit -R HOL-Number_Theory -A HOL + isabelle jedit -d '$AFP' -S Formal_SSA -A HOL + isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis + isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis -i CryptHOL + +* PIDE markup for session ROOT files: allows to complete session names, +follow links to theories and document files etc. + +* Completion supports theory header imports, using theory base name. +E.g. "Prob" may be completed to "HOL-Probability.Probability". + +* Named control symbols (without special Unicode rendering) are shown as +bold-italic keyword. This is particularly useful for the short form of +antiquotations with control symbol: \<^name>\argument\. The action +"isabelle.antiquoted_cartouche" turns an antiquotation with 0 or 1 +arguments into this format. + +* Completion provides templates for named symbols with arguments, +e.g. "\ \ARGUMENT\" or "\<^emph>\ARGUMENT\". + +* Slightly more parallel checking, notably for high priority print +functions (e.g. State output). + +* The view title is set dynamically, according to the Isabelle +distribution and the logic session name. The user can override this via +set-view-title (stored persistently in $JEDIT_SETTINGS/perspective.xml). + +* System options "spell_checker_include" and "spell_checker_exclude" +supersede former "spell_checker_elements" to determine regions of text +that are subject to spell-checking. Minor INCOMPATIBILITY. + +* Action "isabelle.preview" is able to present more file formats, +notably bibtex database files and ML files. + +* Action "isabelle.draft" is similar to "isabelle.preview", but shows a +plain-text document draft. Both are available via the menu "Plugins / +Isabelle". + +* When loading text files, the Isabelle symbols encoding UTF-8-Isabelle +is only used if there is no conflict with existing Unicode sequences in +the file. Otherwise, the fallback encoding is plain UTF-8 and Isabelle +symbols remain in literal \ form. This avoids accidental loss of +Unicode content when saving the file. + +* Bibtex database files (.bib) are semantically checked. + +* Update to jedit-5.5.0, the latest release. + + +*** Isabelle/VSCode Prover IDE *** + +* HTML preview of theories and other file-formats similar to +Isabelle/jEdit. + +* Command-line tool "isabelle vscode_server" accepts the same options +-A, -R, -S, -i for session selection as "isabelle jedit". This is +relevant for isabelle.args configuration settings in VSCode. The former +option -A (explore all known session files) has been discontinued: it is +enabled by default, unless option -S is used to focus on a particular +spot in the session structure. INCOMPATIBILITY. + + +*** Document preparation *** + +* Formal comments work uniformly in outer syntax, inner syntax (term +language), Isabelle/ML and some other embedded languages of Isabelle. +See also "Document comments" in the isar-ref manual. The following forms +are supported: + + - marginal text comment: \ \\\ + - canceled source: \<^cancel>\\\ + - raw LaTeX: \<^latex>\\\ + +* Outside of the inner theory body, the default presentation context is +theory Pure. Thus elementary antiquotations may be used in markup +commands (e.g. 'chapter', 'section', 'text') and formal comments. + +* System option "document_tags" specifies alternative command tags. This +is occasionally useful to control the global visibility of commands via +session options (e.g. in ROOT). + +* Document markup commands ('section', 'text' etc.) are implicitly +tagged as "document" and visible by default. This avoids the application +of option "document_tags" to these commands. + +* Isabelle names are mangled into LaTeX macro names to allow the full +identifier syntax with underscore, prime, digits. This is relevant for +antiquotations in control symbol notation, e.g. \<^const_name> becomes +\isactrlconstUNDERSCOREname. + +* Document preparation with skip_proofs option now preserves the content +more accurately: only terminal proof steps ('by' etc.) are skipped. + +* Document antiquotation @{theory name} requires the long +session-qualified theory name: this is what users reading the text +normally need to import. + +* Document antiquotation @{session name} checks and prints the given +session name verbatim. + +* Document antiquotation @{cite} now checks the given Bibtex entries +against the Bibtex database files -- only in batch-mode session builds. + +* Command-line tool "isabelle document" has been re-implemented in +Isabelle/Scala, with simplified arguments and explicit errors from the +latex and bibtex process. Minor INCOMPATIBILITY. + +* Session ROOT entry: empty 'document_files' means there is no document +for this session. There is no need to specify options [document = false] +anymore. + + +*** Isar *** + +* Command 'interpret' no longer exposes resulting theorems as literal +facts, notably for the \prop\ notation or the "fact" proof method. This +improves modularity of proofs and scalability of locale interpretation. +Rare INCOMPATIBILITY, need to refer to explicitly named facts instead +(e.g. use 'find_theorems' or 'try' to figure this out). + +* The old 'def' command has been discontinued (legacy since +Isbelle2016-1). INCOMPATIBILITY, use 'define' instead -- usually with +object-logic equality or equivalence. + + +*** Pure *** + +* The inner syntax category "sort" now includes notation "_" for the +dummy sort: it is effectively ignored in type-inference. + +* Rewrites clauses (keyword 'rewrites') were moved into the locale +expression syntax, where they are part of locale instances. In +interpretation commands rewrites clauses now need to occur before 'for' +and 'defines'. Rare INCOMPATIBILITY; definitions immediately subject to +rewriting may need to be pulled up into the surrounding theory. + +* For 'rewrites' clauses, if activating a locale instance fails, fall +back to reading the clause first. This helps avoid qualification of +locale instances where the qualifier's sole purpose is avoiding +duplicate constant declarations. + +* Proof method "simp" now supports a new modifier "flip:" followed by a +list of theorems. Each of these theorems is removed from the simpset +(without warning if it is not there) and the symmetric version of the +theorem (i.e. lhs and rhs exchanged) is added to the simpset. For "auto" +and friends the modifier is "simp flip:". + + +*** HOL *** + +* Sledgehammer: bundled version of "vampire" (for non-commercial users) +helps to avoid fragility of "remote_vampire" service. + +* Clarified relationship of characters, strings and code generation: + + - Type "char" is now a proper datatype of 8-bit values. + + - Conversions "nat_of_char" and "char_of_nat" are gone; use more + general conversions "of_char" and "char_of" with suitable type + constraints instead. + + - The zero character is just written "CHR 0x00", not "0" any longer. + + - Type "String.literal" (for code generation) is now isomorphic to + lists of 7-bit (ASCII) values; concrete values can be written as + "STR ''...''" for sequences of printable characters and "STR 0x..." + for one single ASCII code point given as hexadecimal numeral. + + - Type "String.literal" supports concatenation "... + ..." for all + standard target languages. + + - Theory HOL-Library.Code_Char is gone; study the explanations + concerning "String.literal" in the tutorial on code generation to + get an idea how target-language string literals can be converted to + HOL string values and vice versa. + + - Session Imperative-HOL: operation "raise" directly takes a value of + type "String.literal" as argument, not type "string". + +INCOMPATIBILITY. + +* Code generation: Code generation takes an explicit option +"case_insensitive" to accomodate case-insensitive file systems. + +* Abstract bit operations as part of Main: push_bit, take_bit, drop_bit. + +* New, more general, axiomatization of complete_distrib_lattice. The +former axioms: + + "sup x (Inf X) = Inf (sup x ` X)" and "inf x (Sup X) = Sup (inf x ` X)" + +are replaced by: + + "Inf (Sup ` A) <= Sup (Inf ` {f ` A | f . (! Y \ A . f Y \ Y)})" + +The instantiations of sets and functions as complete_distrib_lattice are +moved to Hilbert_Choice.thy because their proofs need the Hilbert choice +operator. The dual of this property is also proved in theory +HOL.Hilbert_Choice. + +* New syntax for the minimum/maximum of a function over a finite set: +MIN x\A. B and even MIN x. B (only useful for finite types), also MAX. + +* Clarifed theorem names: + + Min.antimono ~> Min.subset_imp + Max.antimono ~> Max.subset_imp + +Minor INCOMPATIBILITY. + +* SMT module: + + - The 'smt_oracle' option is now necessary when using the 'smt' method + with a solver other than Z3. INCOMPATIBILITY. + + - The encoding to first-order logic is now more complete in the + presence of higher-order quantifiers. An 'smt_explicit_application' + option has been added to control this. INCOMPATIBILITY. + +* Facts sum.commute(_restrict) and prod.commute(_restrict) renamed to +sum.swap(_restrict) and prod.swap(_restrict), to avoid name clashes on +interpretation of abstract locales. INCOMPATIBILITY. + +* Predicate coprime is now a real definition, not a mere abbreviation. +INCOMPATIBILITY. + +* Predicate pairwise_coprime abolished, use "pairwise coprime" instead. +INCOMPATIBILITY. + +* The relator rel_filter on filters has been strengthened to its +canonical categorical definition with better properties. +INCOMPATIBILITY. + +* Generalized linear algebra involving linear, span, dependent, dim +from type class real_vector to locales module and vector_space. +Renamed: + + span_inc ~> span_superset + span_superset ~> span_base + span_eq ~> span_eq_iff + +INCOMPATIBILITY. + +* Class linordered_semiring_1 covers zero_less_one also, ruling out +pathologic instances. Minor INCOMPATIBILITY. + +* Theory HOL.List: functions "sorted_wrt" and "sorted" now compare every +element in a list to all following elements, not just the next one. + +* Theory HOL.List syntax: + + - filter-syntax "[x <- xs. P]" is no longer output syntax, but only + input syntax + + - list comprehension syntax now supports tuple patterns in "pat <- xs" + +* Theory Map: "empty" must now be qualified as "Map.empty". + +* Removed nat-int transfer machinery. Rare INCOMPATIBILITY. + +* Fact mod_mult_self4 (on nat) renamed to Suc_mod_mult_self3, to avoid +clash with fact mod_mult_self4 (on more generic semirings). +INCOMPATIBILITY. + +* Eliminated some theorem aliasses: + even_times_iff ~> even_mult_iff + mod_2_not_eq_zero_eq_one_nat ~> not_mod_2_eq_0_eq_1 + even_of_nat ~> even_int_iff + +INCOMPATIBILITY. + +* Eliminated some theorem duplicate variations: + + - dvd_eq_mod_eq_0_numeral can be replaced by dvd_eq_mod_eq_0 + - mod_Suc_eq_Suc_mod can be replaced by mod_Suc + - mod_Suc_eq_Suc_mod [symmetrict] can be replaced by mod_simps + - mod_eq_0_iff can be replaced by mod_eq_0_iff_dvd and dvd_def + - the witness of mod_eqD can be given directly as "_ div _" + +INCOMPATIBILITY. + +* Classical setup: Assumption "m mod d = 0" (for m d :: nat) is no +longer aggresively destroyed to "\q. m = d * q". INCOMPATIBILITY, adding +"elim!: dvd" to classical proof methods in most situations restores +broken proofs. + +* Theory HOL-Library.Conditional_Parametricity provides command +'parametric_constant' for proving parametricity of non-recursive +definitions. For constants that are not fully parametric the command +will infer conditions on relations (e.g., bi_unique, bi_total, or type +class conditions such as "respects 0") sufficient for parametricity. See +theory HOL-ex.Conditional_Parametricity_Examples for some examples. + +* Theory HOL-Library.Code_Lazy provides a new preprocessor for the code +generator to generate code for algebraic types with lazy evaluation +semantics even in call-by-value target languages. See the theories +HOL-ex.Code_Lazy_Demo and HOL-Codegenerator_Test.Code_Lazy_Test for some +examples. + +* Theory HOL-Library.Landau_Symbols has been moved here from AFP. + +* Theory HOL-Library.Old_Datatype no longer provides the legacy command +'old_datatype'. INCOMPATIBILITY. + +* Theory HOL-Computational_Algebra.Polynomial_Factorial does not provide +instances of rat, real, complex as factorial rings etc. Import +HOL-Computational_Algebra.Field_as_Ring explicitly in case of need. +INCOMPATIBILITY. + +* Session HOL-Algebra: renamed (^) to [^] to avoid conflict with new +infix/prefix notation. + +* Session HOL-Algebra: revamped with much new material. The set of +isomorphisms between two groups is now denoted iso rather than iso_set. +INCOMPATIBILITY. + +* Session HOL-Analysis: the Arg function now respects the same interval +as Ln, namely (-pi,pi]; the old Arg function has been renamed Arg2pi. +INCOMPATIBILITY. + +* Session HOL-Analysis: the functions zorder, zer_poly, porder and +pol_poly have been redefined. All related lemmas have been reworked. +INCOMPATIBILITY. + +* Session HOL-Analysis: infinite products, Moebius functions, the +Riemann mapping theorem, the Vitali covering theorem, +change-of-variables results for integration and measures. + +* Session HOL-Real_Asymp: proof method "real_asymp" proves asymptotics +or real-valued functions (limits, "Big-O", etc.) automatically. +See also ~~/src/HOL/Real_Asymp/Manual for some documentation. + +* Session HOL-Types_To_Sets: more tool support (unoverload_type combines +internalize_sorts and unoverload) and larger experimental application +(type based linear algebra transferred to linear algebra on subspaces). + + +*** ML *** + +* Operation Export.export emits theory exports (arbitrary blobs), which +are stored persistently in the session build database. + +* Command 'ML_export' exports ML toplevel bindings to the global +bootstrap environment of the ML process. This allows ML evaluation +without a formal theory context, e.g. in command-line tools like +"isabelle process". + + +*** System *** + +* Mac OS X 10.10 Yosemite is now the baseline version; Mavericks is no +longer supported. + +* Linux and Windows/Cygwin is for x86_64 only, old 32bit platform +support has been discontinued. + +* Java runtime is for x86_64 only. Corresponding Isabelle settings have +been renamed to ISABELLE_TOOL_JAVA_OPTIONS and JEDIT_JAVA_OPTIONS, +instead of former 32/64 variants. INCOMPATIBILITY. + +* Old settings ISABELLE_PLATFORM and ISABELLE_WINDOWS_PLATFORM should be +phased out due to unclear preference of 32bit vs. 64bit architecture. +Explicit GNU bash expressions are now preferred, for example (with +quotes): + + #Posix executables (Unix or Cygwin), with preference for 64bit + "${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}" + + #native Windows or Unix executables, with preference for 64bit + "${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}" + + #native Windows (32bit) or Unix executables (preference for 64bit) + "${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}" + +* Command-line tool "isabelle build" supports new options: + - option -B NAME: include session NAME and all descendants + - option -S: only observe changes of sources, not heap images + - option -f: forces a fresh build + +* Command-line tool "isabelle build" options -c -x -B refer to +descendants wrt. the session parent or import graph. Subtle +INCOMPATIBILITY: options -c -x used to refer to the session parent graph +only. + +* Command-line tool "isabelle build" takes "condition" options with the +corresponding environment values into account, when determining the +up-to-date status of a session. + +* The command-line tool "dump" dumps information from the cumulative +PIDE session database: many sessions may be loaded into a given logic +image, results from all loaded theories are written to the output +directory. + +* Command-line tool "isabelle imports -I" also reports actual session +imports. This helps to minimize the session dependency graph. + +* The command-line tool "export" and 'export_files' in session ROOT +entries retrieve theory exports from the session build database. + +* The command-line tools "isabelle server" and "isabelle client" provide +access to the Isabelle Server: it supports responsive session management +and concurrent use of theories, based on Isabelle/PIDE infrastructure. +See also the "system" manual. + +* The command-line tool "isabelle update_comments" normalizes formal +comments in outer syntax as follows: \ \text\ (whith a single space to +approximate the appearance in document output). This is more specific +than former "isabelle update_cartouches -c": the latter tool option has +been discontinued. + +* The command-line tool "isabelle mkroot" now always produces a document +outline: its options have been adapted accordingly. INCOMPATIBILITY. + +* The command-line tool "isabelle mkroot -I" initializes a Mercurial +repository for the generated session files. + +* Settings ISABELLE_HEAPS + ISABELLE_BROWSER_INFO (or +ISABELLE_HEAPS_SYSTEM + ISABELLE_BROWSER_INFO_SYSTEM in "system build +mode") determine the directory locations of the main build artefacts -- +instead of hard-wired directories in ISABELLE_HOME_USER (or +ISABELLE_HOME). + +* Settings ISABELLE_PATH and ISABELLE_OUTPUT have been discontinued: +heap images and session databases are always stored in +$ISABELLE_HEAPS/$ML_IDENTIFIER (command-line default) or +$ISABELLE_HEAPS_SYSTEM/$ML_IDENTIFIER (main Isabelle application or +"isabelle jedit -s" or "isabelle build -s"). + +* ISABELLE_LATEX and ISABELLE_PDFLATEX now include platform-specific +options for improved error reporting. Potential INCOMPATIBILITY with +unusual LaTeX installations, may have to adapt these settings. + +* Update to Poly/ML 5.7.1 with slightly improved performance and PIDE +markup for identifier bindings. It now uses The GNU Multiple Precision +Arithmetic Library (libgmp) on all platforms, notably Mac OS X with +32/64 bit. + + + +New in Isabelle2017 (October 2017) +---------------------------------- + +*** General *** + +* Experimental support for Visual Studio Code (VSCode) as alternative +Isabelle/PIDE front-end, see also +https://marketplace.visualstudio.com/items?itemName=makarius.Isabelle2017 + +VSCode is a new type of application that continues the concepts of +"programmer's editor" and "integrated development environment" towards +fully semantic editing and debugging -- in a relatively light-weight +manner. Thus it fits nicely on top of the Isabelle/PIDE infrastructure. +Technically, VSCode is based on the Electron application framework +(Node.js + Chromium browser + V8), which is implemented in JavaScript +and TypeScript, while Isabelle/VSCode mainly consists of Isabelle/Scala +modules around a Language Server implementation. + +* Theory names are qualified by the session name that they belong to. +This affects imports, but not the theory name space prefix (which is +just the theory base name as before). + +In order to import theories from other sessions, the ROOT file format +provides a new 'sessions' keyword. In contrast, a theory that is +imported in the old-fashioned manner via an explicit file-system path +belongs to the current session, and might cause theory name conflicts +later on. Theories that are imported from other sessions are excluded +from the current session document. The command-line tool "isabelle +imports" helps to update theory imports. + +* The main theory entry points for some non-HOL sessions have changed, +to avoid confusion with the global name "Main" of the session HOL. This +leads to the follow renamings: + + CTT/Main.thy ~> CTT/CTT.thy + ZF/Main.thy ~> ZF/ZF.thy + ZF/Main_ZF.thy ~> ZF/ZF.thy + ZF/Main_ZFC.thy ~> ZF/ZFC.thy + ZF/ZF.thy ~> ZF/ZF_Base.thy + +INCOMPATIBILITY. + +* Commands 'alias' and 'type_alias' introduce aliases for constants and +type constructors, respectively. This allows adhoc changes to name-space +accesses within global or local theory contexts, e.g. within a 'bundle'. + +* Document antiquotations @{prf} and @{full_prf} output proof terms +(again) in the same way as commands 'prf' and 'full_prf'. + +* Computations generated by the code generator can be embedded directly +into ML, alongside with @{code} antiquotations, using the following +antiquotations: + + @{computation ... terms: ... datatypes: ...} : + ((term -> term) -> 'ml option -> 'a) -> Proof.context -> term -> 'a + @{computation_conv ... terms: ... datatypes: ...} : + (Proof.context -> 'ml -> conv) -> Proof.context -> conv + @{computation_check terms: ... datatypes: ...} : Proof.context -> conv + +See src/HOL/ex/Computations.thy, +src/HOL/Decision_Procs/Commutative_Ring.thy and +src/HOL/Decision_Procs/Reflective_Field.thy for examples and the +tutorial on code generation. + + +*** Prover IDE -- Isabelle/Scala/jEdit *** + +* Session-qualified theory imports allow the Prover IDE to process +arbitrary theory hierarchies independently of the underlying logic +session image (e.g. option "isabelle jedit -l"), but the directory +structure needs to be known in advance (e.g. option "isabelle jedit -d" +or a line in the file $ISABELLE_HOME_USER/ROOTS). + +* The PIDE document model maintains file content independently of the +status of jEdit editor buffers. Reloading jEdit buffers no longer causes +changes of formal document content. Theory dependencies are always +resolved internally, without the need for corresponding editor buffers. +The system option "jedit_auto_load" has been discontinued: it is +effectively always enabled. + +* The Theories dockable provides a "Purge" button, in order to restrict +the document model to theories that are required for open editor +buffers. + +* The Theories dockable indicates the overall status of checking of each +entry. When all forked tasks of a theory are finished, the border is +painted with thick lines; remaining errors in this situation are +represented by a different border color. + +* Automatic indentation is more careful to avoid redundant spaces in +intermediate situations. Keywords are indented after input (via typed +characters or completion); see also option "jedit_indent_input". + +* Action "isabelle.preview" opens an HTML preview of the current theory +document in the default web browser. + +* Command-line invocation "isabelle jedit -R -l LOGIC" opens the ROOT +entry of the specified logic session in the editor, while its parent is +used for formal checking. + +* The main Isabelle/jEdit plugin may be restarted manually (using the +jEdit Plugin Manager), as long as the "Isabelle Base" plugin remains +enabled at all times. + +* Update to current jedit-5.4.0. + + +*** Pure *** + +* Deleting the last code equations for a particular function using +[code del] results in function with no equations (runtime abort) rather +than an unimplemented function (generation time abort). Use explicit +[[code drop:]] to enforce the latter. Minor INCOMPATIBILITY. + +* Proper concept of code declarations in code.ML: + - Regular code declarations act only on the global theory level, being + ignored with warnings if syntactically malformed. + - Explicitly global code declarations yield errors if syntactically + malformed. + - Default code declarations are silently ignored if syntactically + malformed. +Minor INCOMPATIBILITY. + +* Clarified and standardized internal data bookkeeping of code +declarations: history of serials allows to track potentially +non-monotonous declarations appropriately. Minor INCOMPATIBILITY. + + +*** HOL *** + +* The Nunchaku model finder is now part of "Main". + +* SMT module: + - A new option, 'smt_nat_as_int', has been added to translate 'nat' to + 'int' and benefit from the SMT solver's theory reasoning. It is + disabled by default. + - The legacy module "src/HOL/Library/Old_SMT.thy" has been removed. + - Several small issues have been rectified in the 'smt' command. + +* (Co)datatype package: The 'size_gen_o_map' lemma is no longer +generated for datatypes with type class annotations. As a result, the +tactic that derives it no longer fails on nested datatypes. Slight +INCOMPATIBILITY. + +* Command and antiquotation "value" with modified default strategy: +terms without free variables are always evaluated using plain evaluation +only, with no fallback on normalization by evaluation. Minor +INCOMPATIBILITY. + +* Theories "GCD" and "Binomial" are already included in "Main" (instead +of "Complex_Main"). + +* Constant "surj" is a full input/output abbreviation (again). +Minor INCOMPATIBILITY. + +* Dropped aliasses RangeP, DomainP for Rangep, Domainp respectively. +INCOMPATIBILITY. + +* Renamed ii to imaginary_unit in order to free up ii as a variable +name. The syntax \ remains available. INCOMPATIBILITY. + +* Dropped abbreviations transP, antisymP, single_valuedP; use constants +transp, antisymp, single_valuedp instead. INCOMPATIBILITY. + +* Constant "subseq" in Topological_Spaces has been removed -- it is +subsumed by "strict_mono". Some basic lemmas specific to "subseq" have +been renamed accordingly, e.g. "subseq_o" -> "strict_mono_o" etc. + +* Theory List: "sublist" renamed to "nths" in analogy with "nth", and +"sublisteq" renamed to "subseq". Minor INCOMPATIBILITY. + +* Theory List: new generic function "sorted_wrt". + +* Named theorems mod_simps covers various congruence rules concerning +mod, replacing former zmod_simps. INCOMPATIBILITY. + +* Swapped orientation of congruence rules mod_add_left_eq, +mod_add_right_eq, mod_add_eq, mod_mult_left_eq, mod_mult_right_eq, +mod_mult_eq, mod_minus_eq, mod_diff_left_eq, mod_diff_right_eq, +mod_diff_eq. INCOMPATIBILITY. + +* Generalized some facts: + measure_induct_rule + measure_induct + zminus_zmod ~> mod_minus_eq + zdiff_zmod_left ~> mod_diff_left_eq + zdiff_zmod_right ~> mod_diff_right_eq + zmod_eq_dvd_iff ~> mod_eq_dvd_iff +INCOMPATIBILITY. + +* Algebraic type class hierarchy of euclidean (semi)rings in HOL: +euclidean_(semi)ring, euclidean_(semi)ring_cancel, +unique_euclidean_(semi)ring; instantiation requires provision of a +euclidean size. + +* Theory "HOL-Number_Theory.Euclidean_Algorithm" has been reworked: + - Euclidean induction is available as rule eucl_induct. + - Constants Euclidean_Algorithm.gcd, Euclidean_Algorithm.lcm, + Euclidean_Algorithm.Gcd and Euclidean_Algorithm.Lcm allow + easy instantiation of euclidean (semi)rings as GCD (semi)rings. + - Coefficients obtained by extended euclidean algorithm are + available as "bezout_coefficients". +INCOMPATIBILITY. + +* Theory "Number_Theory.Totient" introduces basic notions about Euler's +totient function previously hidden as solitary example in theory +Residues. Definition changed so that "totient 1 = 1" in agreement with +the literature. Minor INCOMPATIBILITY. + +* New styles in theory "HOL-Library.LaTeXsugar": + - "dummy_pats" for printing equations with "_" on the lhs; + - "eta_expand" for printing eta-expanded terms. + +* Theory "HOL-Library.Permutations": theorem bij_swap_ompose_bij has +been renamed to bij_swap_compose_bij. INCOMPATIBILITY. + +* New theory "HOL-Library.Going_To_Filter" providing the "f going_to F" +filter for describing points x such that f(x) is in the filter F. + +* Theory "HOL-Library.Formal_Power_Series": constants X/E/L/F have been +renamed to fps_X/fps_exp/fps_ln/fps_hypergeo to avoid polluting the name +space. INCOMPATIBILITY. + +* Theory "HOL-Library.FinFun" has been moved to AFP (again). +INCOMPATIBILITY. + +* Theory "HOL-Library.FuncSet": some old and rarely used ASCII +replacement syntax has been removed. INCOMPATIBILITY, standard syntax +with symbols should be used instead. The subsequent commands help to +reproduce the old forms, e.g. to simplify porting old theories: + +syntax (ASCII) + "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3PIE _:_./ _)" 10) + "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3PI _:_./ _)" 10) + "_lam" :: "pttrn \ 'a set \ 'a \ 'b \ ('a \ 'b)" ("(3%_:_./ _)" [0,0,3] 3) + +* Theory "HOL-Library.Multiset": the simprocs on subsets operators of +multisets have been renamed: + + msetless_cancel_numerals ~> msetsubset_cancel + msetle_cancel_numerals ~> msetsubset_eq_cancel + +INCOMPATIBILITY. + +* Theory "HOL-Library.Pattern_Aliases" provides input and output syntax +for pattern aliases as known from Haskell, Scala and ML. + +* Theory "HOL-Library.Uprod" formalizes the type of unordered pairs. + +* Session HOL-Analysis: more material involving arcs, paths, covering +spaces, innessential maps, retracts, infinite products, simplicial +complexes. Baire Category theorem. Major results include the Jordan +Curve Theorem and the Great Picard Theorem. + +* Session HOL-Algebra has been extended by additional lattice theory: +the Knaster-Tarski fixed point theorem and Galois Connections. + +* Sessions HOL-Computational_Algebra and HOL-Number_Theory: new notions +of squarefreeness, n-th powers, and prime powers. + +* Session "HOL-Computional_Algebra" covers many previously scattered +theories, notably Euclidean_Algorithm, Factorial_Ring, +Formal_Power_Series, Fraction_Field, Fundamental_Theorem_Algebra, +Normalized_Fraction, Polynomial_FPS, Polynomial, Primes. Minor +INCOMPATIBILITY. + + +*** System *** + +* Isabelle/Scala: the SQL module supports access to relational +databases, either as plain file (SQLite) or full-scale server +(PostgreSQL via local port or remote ssh connection). + +* Results of "isabelle build" are recorded as SQLite database (i.e. +"Application File Format" in the sense of +https://www.sqlite.org/appfileformat.html). This allows systematic +access via operations from module Sessions.Store in Isabelle/Scala. + +* System option "parallel_proofs" is 1 by default (instead of more +aggressive 2). This requires less heap space and avoids burning parallel +CPU cycles, while full subproof parallelization is enabled for repeated +builds (according to parallel_subproofs_threshold). + +* System option "record_proofs" allows to change the global +Proofterm.proofs variable for a session. Regular values are are 0, 1, 2; +a negative value means the current state in the ML heap image remains +unchanged. + +* Isabelle settings variable ISABELLE_SCALA_BUILD_OPTIONS has been +renamed to ISABELLE_SCALAC_OPTIONS. Rare INCOMPATIBILITY. + +* Isabelle settings variables ISABELLE_WINDOWS_PLATFORM, +ISABELLE_WINDOWS_PLATFORM32, ISABELLE_WINDOWS_PLATFORM64 indicate the +native Windows platform (independently of the Cygwin installation). This +is analogous to ISABELLE_PLATFORM, ISABELLE_PLATFORM32, +ISABELLE_PLATFORM64. + +* Command-line tool "isabelle build_docker" builds a Docker image from +the Isabelle application bundle for Linux. See also +https://hub.docker.com/r/makarius/isabelle + +* Command-line tool "isabelle vscode_server" provides a Language Server +Protocol implementation, e.g. for the Visual Studio Code editor. It +serves as example for alternative PIDE front-ends. + +* Command-line tool "isabelle imports" helps to maintain theory imports +wrt. session structure. Examples for the main Isabelle distribution: + + isabelle imports -I -a + isabelle imports -U -a + isabelle imports -U -i -a + isabelle imports -M -a -d '~~/src/Benchmarks' + + + +New in Isabelle2016-1 (December 2016) +------------------------------------- + +*** General *** + +* Splitter in proof methods "simp", "auto" and friends: + - The syntax "split add" has been discontinued, use plain "split", + INCOMPATIBILITY. + - For situations with many conditional or case expressions, there is + an alternative splitting strategy that can be much faster. It is + selected by writing "split!" instead of "split". It applies safe + introduction and elimination rules after each split rule. As a + result the subgoal may be split into several subgoals. + +* Command 'bundle' provides a local theory target to define a bundle +from the body of specification commands (such as 'declare', +'declaration', 'notation', 'lemmas', 'lemma'). For example: + +bundle foo +begin + declare a [simp] + declare b [intro] +end + +* Command 'unbundle' is like 'include', but works within a local theory +context. Unlike "context includes ... begin", the effect of 'unbundle' +on the target context persists, until different declarations are given. + +* Simplified outer syntax: uniform category "name" includes long +identifiers. Former "xname" / "nameref" / "name reference" has been +discontinued. + +* Embedded content (e.g. the inner syntax of types, terms, props) may be +delimited uniformly via cartouches. This works better than old-fashioned +quotes when sub-languages are nested. + +* Mixfix annotations support general block properties, with syntax +"(\x=a y=b z \\". Notable property names are "indent", "consistent", +"unbreakable", "markup". The existing notation "(DIGITS" is equivalent +to "(\indent=DIGITS\". The former notation "(00" for unbreakable blocks +is superseded by "(\unbreabable\" --- rare INCOMPATIBILITY. + +* Proof method "blast" is more robust wrt. corner cases of Pure +statements without object-logic judgment. + +* Commands 'prf' and 'full_prf' are somewhat more informative (again): +proof terms are reconstructed and cleaned from administrative thm nodes. + +* Code generator: config option "code_timing" triggers measurements of +different phases of code generation. See src/HOL/ex/Code_Timing.thy for +examples. + +* Code generator: implicits in Scala (stemming from type class +instances) are generated into companion object of corresponding type +class, to resolve some situations where ambiguities may occur. + +* Solve direct: option "solve_direct_strict_warnings" gives explicit +warnings for lemma statements with trivial proofs. + + +*** Prover IDE -- Isabelle/Scala/jEdit *** + +* More aggressive flushing of machine-generated input, according to +system option editor_generated_input_delay (in addition to existing +editor_input_delay for regular user edits). This may affect overall PIDE +reactivity and CPU usage. + +* Syntactic indentation according to Isabelle outer syntax. Action +"indent-lines" (shortcut C+i) indents the current line according to +command keywords and some command substructure. Action +"isabelle.newline" (shortcut ENTER) indents the old and the new line +according to command keywords only; see also option +"jedit_indent_newline". + +* Semantic indentation for unstructured proof scripts ('apply' etc.) via +number of subgoals. This requires information of ongoing document +processing and may thus lag behind, when the user is editing too +quickly; see also option "jedit_script_indent" and +"jedit_script_indent_limit". + +* Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed' +are treated as delimiters for fold structure; 'begin' and 'end' +structure of theory specifications is treated as well. + +* Command 'proof' provides information about proof outline with cases, +e.g. for proof methods "cases", "induct", "goal_cases". + +* Completion templates for commands involving "begin ... end" blocks, +e.g. 'context', 'notepad'. + +* Sidekick parser "isabelle-context" shows nesting of context blocks +according to 'begin' and 'end' structure. + +* Highlighting of entity def/ref positions wrt. cursor. + +* Action "isabelle.select-entity" (shortcut CS+ENTER) selects all +occurrences of the formal entity at the caret position. This facilitates +systematic renaming. + +* PIDE document markup works across multiple Isar commands, e.g. the +results established at the end of a proof are properly identified in the +theorem statement. + +* Cartouche abbreviations work both for " and ` to accomodate typical +situations where old ASCII notation may be updated. + +* Dockable window "Symbols" also provides access to 'abbrevs' from the +outer syntax of the current theory buffer. This provides clickable +syntax templates, including entries with empty abbrevs name (which are +inaccessible via keyboard completion). + +* IDE support for the Isabelle/Pure bootstrap process, with the +following independent stages: + + src/Pure/ROOT0.ML + src/Pure/ROOT.ML + src/Pure/Pure.thy + src/Pure/ML_Bootstrap.thy + +The ML ROOT files act like quasi-theories in the context of theory +ML_Bootstrap: this allows continuous checking of all loaded ML files. +The theory files are presented with a modified header to import Pure +from the running Isabelle instance. Results from changed versions of +each stage are *not* propagated to the next stage, and isolated from the +actual Isabelle/Pure that runs the IDE itself. The sequential +dependencies of the above files are only observed for batch build. + +* Isabelle/ML and Standard ML files are presented in Sidekick with the +tree structure of section headings: this special comment format is +described in "implementation" chapter 0, e.g. (*** section ***). + +* Additional abbreviations for syntactic completion may be specified +within the theory header as 'abbrevs'. The theory syntax for 'keywords' +has been simplified accordingly: optional abbrevs need to go into the +new 'abbrevs' section. + +* Global abbreviations via $ISABELLE_HOME/etc/abbrevs and +$ISABELLE_HOME_USER/etc/abbrevs are no longer supported. Minor +INCOMPATIBILITY, use 'abbrevs' within theory header instead. + +* Action "isabelle.keymap-merge" asks the user to resolve pending +Isabelle keymap changes that are in conflict with the current jEdit +keymap; non-conflicting changes are always applied implicitly. This +action is automatically invoked on Isabelle/jEdit startup and thus +increases chances that users see new keyboard shortcuts when re-using +old keymaps. + +* ML and document antiquotations for file-systems paths are more uniform +and diverse: + + @{path NAME} -- no file-system check + @{file NAME} -- check for plain file + @{dir NAME} -- check for directory + +Minor INCOMPATIBILITY, former uses of @{file} and @{file_unchecked} may +have to be changed. + + +*** Document preparation *** + +* New symbol \, e.g. for temporal operator. + +* New document and ML antiquotation @{locale} for locales, similar to +existing antiquotation @{class}. + +* Mixfix annotations support delimiters like \<^control>\cartouche\ -- +this allows special forms of document output. + +* Raw LaTeX output now works via \<^latex>\...\ instead of raw control +symbol \<^raw:...>. INCOMPATIBILITY, notably for LaTeXsugar.thy and its +derivatives. + +* \<^raw:...> symbols are no longer supported. + +* Old 'header' command is no longer supported (legacy since +Isabelle2015). + + +*** Isar *** + +* Many specification elements support structured statements with 'if' / +'for' eigen-context, e.g. 'axiomatization', 'abbreviation', +'definition', 'inductive', 'function'. + +* Toplevel theorem statements support eigen-context notation with 'if' / +'for' (in postfix), which corresponds to 'assumes' / 'fixes' in the +traditional long statement form (in prefix). Local premises are called +"that" or "assms", respectively. Empty premises are *not* bound in the +context: INCOMPATIBILITY. + +* Command 'define' introduces a local (non-polymorphic) definition, with +optional abstraction over local parameters. The syntax resembles +'definition' and 'obtain'. It fits better into the Isar language than +old 'def', which is now a legacy feature. + +* Command 'obtain' supports structured statements with 'if' / 'for' +context. + +* Command '\' is an alias for 'sorry', with different +typesetting. E.g. to produce proof holes in examples and documentation. + +* The defining position of a literal fact \prop\ is maintained more +carefully, and made accessible as hyperlink in the Prover IDE. + +* Commands 'finally' and 'ultimately' used to expose the result as +literal fact: this accidental behaviour has been discontinued. Rare +INCOMPATIBILITY, use more explicit means to refer to facts in Isar. + +* Command 'axiomatization' has become more restrictive to correspond +better to internal axioms as singleton facts with mandatory name. Minor +INCOMPATIBILITY. + +* Proof methods may refer to the main facts via the dynamic fact +"method_facts". This is particularly useful for Eisbach method +definitions. + +* Proof method "use" allows to modify the main facts of a given method +expression, e.g. + + (use facts in simp) + (use facts in \simp add: ...\) + +* The old proof method "default" has been removed (legacy since +Isabelle2016). INCOMPATIBILITY, use "standard" instead. + + +*** Pure *** + +* Pure provides basic versions of proof methods "simp" and "simp_all" +that only know about meta-equality (==). Potential INCOMPATIBILITY in +theory imports that merge Pure with e.g. Main of Isabelle/HOL: the order +is relevant to avoid confusion of Pure.simp vs. HOL.simp. + +* The command 'unfolding' and proof method "unfold" include a second +stage where given equations are passed through the attribute "abs_def" +before rewriting. This ensures that definitions are fully expanded, +regardless of the actual parameters that are provided. Rare +INCOMPATIBILITY in some corner cases: use proof method (simp only:) +instead, or declare [[unfold_abs_def = false]] in the proof context. + +* Type-inference improves sorts of newly introduced type variables for +the object-logic, using its base sort (i.e. HOL.type for Isabelle/HOL). +Thus terms like "f x" or "\x. P x" without any further syntactic context +produce x::'a::type in HOL instead of x::'a::{} in Pure. Rare +INCOMPATIBILITY, need to provide explicit type constraints for Pure +types where this is really intended. + + +*** HOL *** + +* New proof method "argo" using the built-in Argo solver based on SMT +technology. The method can be used to prove goals of quantifier-free +propositional logic, goals based on a combination of quantifier-free +propositional logic with equality, and goals based on a combination of +quantifier-free propositional logic with linear real arithmetic +including min/max/abs. See HOL/ex/Argo_Examples.thy for examples. + +* The new "nunchaku" command integrates the Nunchaku model finder. The +tool is experimental. See ~~/src/HOL/Nunchaku/Nunchaku.thy for details. + +* Metis: The problem encoding has changed very slightly. This might +break existing proofs. INCOMPATIBILITY. + +* Sledgehammer: + - The MaSh relevance filter is now faster than before. + - Produce syntactically correct Vampire 4.0 problem files. + +* (Co)datatype package: + - New commands for defining corecursive functions and reasoning about + them in "~~/src/HOL/Library/BNF_Corec.thy": 'corec', 'corecursive', + 'friend_of_corec', and 'corecursion_upto'; and 'corec_unique' proof + method. See 'isabelle doc corec'. + - The predicator :: ('a \ bool) \ 'a F \ bool is now a first-class + citizen in bounded natural functors. + - 'primrec' now allows nested calls through the predicator in addition + to the map function. + - 'bnf' automatically discharges reflexive proof obligations. + - 'bnf' outputs a slightly modified proof obligation expressing rel in + terms of map and set + (not giving a specification for rel makes this one reflexive). + - 'bnf' outputs a new proof obligation expressing pred in terms of set + (not giving a specification for pred makes this one reflexive). + INCOMPATIBILITY: manual 'bnf' declarations may need adjustment. + - Renamed lemmas: + rel_prod_apply ~> rel_prod_inject + pred_prod_apply ~> pred_prod_inject + INCOMPATIBILITY. + - The "size" plugin has been made compatible again with locales. + - The theorems about "rel" and "set" may have a slightly different (but + equivalent) form. + INCOMPATIBILITY. + +* The 'coinductive' command produces a proper coinduction rule for +mutual coinductive predicates. This new rule replaces the old rule, +which exposed details of the internal fixpoint construction and was +hard to use. INCOMPATIBILITY. + +* New abbreviations for negated existence (but not bounded existence): + + \x. P x \ \ (\x. P x) + \!x. P x \ \ (\!x. P x) + +* The print mode "HOL" for ASCII syntax of binders "!", "?", "?!", "@" +has been removed for output. It is retained for input only, until it is +eliminated altogether. + +* The unique existence quantifier no longer provides 'binder' syntax, +but uses syntax translations (as for bounded unique existence). Thus +iterated quantification \!x y. P x y with its slightly confusing +sequential meaning \!x. \!y. P x y is no longer possible. Instead, +pattern abstraction admits simultaneous unique existence \!(x, y). P x y +(analogous to existing notation \!(x, y)\A. P x y). Potential +INCOMPATIBILITY in rare situations. + +* Conventional syntax "%(). t" for unit abstractions. Slight syntactic +INCOMPATIBILITY. + +* Renamed constants and corresponding theorems: + + setsum ~> sum + setprod ~> prod + listsum ~> sum_list + listprod ~> prod_list + +INCOMPATIBILITY. + +* Sligthly more standardized theorem names: + sgn_times ~> sgn_mult + sgn_mult' ~> Real_Vector_Spaces.sgn_mult + divide_zero_left ~> div_0 + zero_mod_left ~> mod_0 + divide_zero ~> div_by_0 + divide_1 ~> div_by_1 + nonzero_mult_divide_cancel_left ~> nonzero_mult_div_cancel_left + div_mult_self1_is_id ~> nonzero_mult_div_cancel_left + nonzero_mult_divide_cancel_right ~> nonzero_mult_div_cancel_right + div_mult_self2_is_id ~> nonzero_mult_div_cancel_right + is_unit_divide_mult_cancel_left ~> is_unit_div_mult_cancel_left + is_unit_divide_mult_cancel_right ~> is_unit_div_mult_cancel_right + mod_div_equality ~> div_mult_mod_eq + mod_div_equality2 ~> mult_div_mod_eq + mod_div_equality3 ~> mod_div_mult_eq + mod_div_equality4 ~> mod_mult_div_eq + minus_div_eq_mod ~> minus_div_mult_eq_mod + minus_div_eq_mod2 ~> minus_mult_div_eq_mod + minus_mod_eq_div ~> minus_mod_eq_div_mult + minus_mod_eq_div2 ~> minus_mod_eq_mult_div + div_mod_equality' ~> minus_mod_eq_div_mult [symmetric] + mod_div_equality' ~> minus_div_mult_eq_mod [symmetric] + zmod_zdiv_equality ~> mult_div_mod_eq [symmetric] + zmod_zdiv_equality' ~> minus_div_mult_eq_mod [symmetric] + Divides.mult_div_cancel ~> minus_mod_eq_mult_div [symmetric] + mult_div_cancel ~> minus_mod_eq_mult_div [symmetric] + zmult_div_cancel ~> minus_mod_eq_mult_div [symmetric] + div_1 ~> div_by_Suc_0 + mod_1 ~> mod_by_Suc_0 +INCOMPATIBILITY. + +* New type class "idom_abs_sgn" specifies algebraic properties +of sign and absolute value functions. Type class "sgn_if" has +disappeared. Slight INCOMPATIBILITY. + +* Dedicated syntax LENGTH('a) for length of types. + +* Characters (type char) are modelled as finite algebraic type +corresponding to {0..255}. + + - Logical representation: + * 0 is instantiated to the ASCII zero character. + * All other characters are represented as "Char n" + with n being a raw numeral expression less than 256. + * Expressions of the form "Char n" with n greater than 255 + are non-canonical. + - Printing and parsing: + * Printable characters are printed and parsed as "CHR ''\''" + (as before). + * The ASCII zero character is printed and parsed as "0". + * All other canonical characters are printed as "CHR 0xXX" + with XX being the hexadecimal character code. "CHR n" + is parsable for every numeral expression n. + * Non-canonical characters have no special syntax and are + printed as their logical representation. + - Explicit conversions from and to the natural numbers are + provided as char_of_nat, nat_of_char (as before). + - The auxiliary nibble type has been discontinued. + +INCOMPATIBILITY. + +* Type class "div" with operation "mod" renamed to type class "modulo" +with operation "modulo", analogously to type class "divide". This +eliminates the need to qualify any of those names in the presence of +infix "mod" syntax. INCOMPATIBILITY. + +* Statements and proofs of Knaster-Tarski fixpoint combinators lfp/gfp +have been clarified. The fixpoint properties are lfp_fixpoint, its +symmetric lfp_unfold (as before), and the duals for gfp. Auxiliary items +for the proof (lfp_lemma2 etc.) are no longer exported, but can be +easily recovered by composition with eq_refl. Minor INCOMPATIBILITY. + +* Constant "surj" is a mere input abbreviation, to avoid hiding an +equation in term output. Minor INCOMPATIBILITY. + +* Command 'code_reflect' accepts empty constructor lists for datatypes, +which renders those abstract effectively. + +* Command 'export_code' checks given constants for abstraction +violations: a small guarantee that given constants specify a safe +interface for the generated code. + +* Code generation for Scala: ambiguous implicts in class diagrams are +spelt out explicitly. + +* Static evaluators (Code_Evaluation.static_* in Isabelle/ML) rely on +explicitly provided auxiliary definitions for required type class +dictionaries rather than half-working magic. INCOMPATIBILITY, see the +tutorial on code generation for details. + +* Theory Set_Interval: substantial new theorems on indexed sums and +products. + +* Locale bijection establishes convenient default simp rules such as +"inv f (f a) = a" for total bijections. + +* Abstract locales semigroup, abel_semigroup, semilattice, +semilattice_neutr, ordering, ordering_top, semilattice_order, +semilattice_neutr_order, comm_monoid_set, semilattice_set, +semilattice_neutr_set, semilattice_order_set, +semilattice_order_neutr_set monoid_list, comm_monoid_list, +comm_monoid_list_set, comm_monoid_mset, comm_monoid_fun use boldified +syntax uniformly that does not clash with corresponding global syntax. +INCOMPATIBILITY. + +* Former locale lifting_syntax is now a bundle, which is easier to +include in a local context or theorem statement, e.g. "context includes +lifting_syntax begin ... end". Minor INCOMPATIBILITY. + +* Some old / obsolete theorems have been renamed / removed, potential +INCOMPATIBILITY. + + nat_less_cases -- removed, use linorder_cases instead + inv_image_comp -- removed, use image_inv_f_f instead + image_surj_f_inv_f ~> image_f_inv_f + +* Some theorems about groups and orders have been generalised from + groups to semi-groups that are also monoids: + le_add_same_cancel1 + le_add_same_cancel2 + less_add_same_cancel1 + less_add_same_cancel2 + add_le_same_cancel1 + add_le_same_cancel2 + add_less_same_cancel1 + add_less_same_cancel2 + +* Some simplifications theorems about rings have been removed, since + superseeded by a more general version: + less_add_cancel_left_greater_zero ~> less_add_same_cancel1 + less_add_cancel_right_greater_zero ~> less_add_same_cancel2 + less_eq_add_cancel_left_greater_eq_zero ~> le_add_same_cancel1 + less_eq_add_cancel_right_greater_eq_zero ~> le_add_same_cancel2 + less_eq_add_cancel_left_less_eq_zero ~> add_le_same_cancel1 + less_eq_add_cancel_right_less_eq_zero ~> add_le_same_cancel2 + less_add_cancel_left_less_zero ~> add_less_same_cancel1 + less_add_cancel_right_less_zero ~> add_less_same_cancel2 +INCOMPATIBILITY. + +* Renamed split_if -> if_split and split_if_asm -> if_split_asm to +resemble the f.split naming convention, INCOMPATIBILITY. + +* Added class topological_monoid. + +* The following theorems have been renamed: + + setsum_left_distrib ~> sum_distrib_right + setsum_right_distrib ~> sum_distrib_left + +INCOMPATIBILITY. + +* Compound constants INFIMUM and SUPREMUM are mere abbreviations now. +INCOMPATIBILITY. + +* "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional +comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f ` +A)". + +* Class semiring_Lcd merged into semiring_Gcd. INCOMPATIBILITY. + +* The type class ordered_comm_monoid_add is now called +ordered_cancel_comm_monoid_add. A new type class ordered_comm_monoid_add +is introduced as the combination of ordered_ab_semigroup_add + +comm_monoid_add. INCOMPATIBILITY. + +* Introduced the type classes canonically_ordered_comm_monoid_add and +dioid. + +* Introduced the type class ordered_ab_semigroup_monoid_add_imp_le. When +instantiating linordered_semiring_strict and ordered_ab_group_add, an +explicit instantiation of ordered_ab_semigroup_monoid_add_imp_le might +be required. INCOMPATIBILITY. + +* Dropped various legacy fact bindings, whose replacements are often +of a more general type also: + lcm_left_commute_nat ~> lcm.left_commute + lcm_left_commute_int ~> lcm.left_commute + gcd_left_commute_nat ~> gcd.left_commute + gcd_left_commute_int ~> gcd.left_commute + gcd_greatest_iff_nat ~> gcd_greatest_iff + gcd_greatest_iff_int ~> gcd_greatest_iff + coprime_dvd_mult_nat ~> coprime_dvd_mult + coprime_dvd_mult_int ~> coprime_dvd_mult + zpower_numeral_even ~> power_numeral_even + gcd_mult_cancel_nat ~> gcd_mult_cancel + gcd_mult_cancel_int ~> gcd_mult_cancel + div_gcd_coprime_nat ~> div_gcd_coprime + div_gcd_coprime_int ~> div_gcd_coprime + zpower_numeral_odd ~> power_numeral_odd + zero_less_int_conv ~> of_nat_0_less_iff + gcd_greatest_nat ~> gcd_greatest + gcd_greatest_int ~> gcd_greatest + coprime_mult_nat ~> coprime_mult + coprime_mult_int ~> coprime_mult + lcm_commute_nat ~> lcm.commute + lcm_commute_int ~> lcm.commute + int_less_0_conv ~> of_nat_less_0_iff + gcd_commute_nat ~> gcd.commute + gcd_commute_int ~> gcd.commute + Gcd_insert_nat ~> Gcd_insert + Gcd_insert_int ~> Gcd_insert + of_int_int_eq ~> of_int_of_nat_eq + lcm_least_nat ~> lcm_least + lcm_least_int ~> lcm_least + lcm_assoc_nat ~> lcm.assoc + lcm_assoc_int ~> lcm.assoc + int_le_0_conv ~> of_nat_le_0_iff + int_eq_0_conv ~> of_nat_eq_0_iff + Gcd_empty_nat ~> Gcd_empty + Gcd_empty_int ~> Gcd_empty + gcd_assoc_nat ~> gcd.assoc + gcd_assoc_int ~> gcd.assoc + zero_zle_int ~> of_nat_0_le_iff + lcm_dvd2_nat ~> dvd_lcm2 + lcm_dvd2_int ~> dvd_lcm2 + lcm_dvd1_nat ~> dvd_lcm1 + lcm_dvd1_int ~> dvd_lcm1 + gcd_zero_nat ~> gcd_eq_0_iff + gcd_zero_int ~> gcd_eq_0_iff + gcd_dvd2_nat ~> gcd_dvd2 + gcd_dvd2_int ~> gcd_dvd2 + gcd_dvd1_nat ~> gcd_dvd1 + gcd_dvd1_int ~> gcd_dvd1 + int_numeral ~> of_nat_numeral + lcm_ac_nat ~> ac_simps + lcm_ac_int ~> ac_simps + gcd_ac_nat ~> ac_simps + gcd_ac_int ~> ac_simps + abs_int_eq ~> abs_of_nat + zless_int ~> of_nat_less_iff + zdiff_int ~> of_nat_diff + zadd_int ~> of_nat_add + int_mult ~> of_nat_mult + int_Suc ~> of_nat_Suc + inj_int ~> inj_of_nat + int_1 ~> of_nat_1 + int_0 ~> of_nat_0 + Lcm_empty_nat ~> Lcm_empty + Lcm_empty_int ~> Lcm_empty + Lcm_insert_nat ~> Lcm_insert + Lcm_insert_int ~> Lcm_insert + comp_fun_idem_gcd_nat ~> comp_fun_idem_gcd + comp_fun_idem_gcd_int ~> comp_fun_idem_gcd + comp_fun_idem_lcm_nat ~> comp_fun_idem_lcm + comp_fun_idem_lcm_int ~> comp_fun_idem_lcm + Lcm_eq_0 ~> Lcm_eq_0_I + Lcm0_iff ~> Lcm_0_iff + Lcm_dvd_int ~> Lcm_least + divides_mult_nat ~> divides_mult + divides_mult_int ~> divides_mult + lcm_0_nat ~> lcm_0_right + lcm_0_int ~> lcm_0_right + lcm_0_left_nat ~> lcm_0_left + lcm_0_left_int ~> lcm_0_left + dvd_gcd_D1_nat ~> dvd_gcdD1 + dvd_gcd_D1_int ~> dvd_gcdD1 + dvd_gcd_D2_nat ~> dvd_gcdD2 + dvd_gcd_D2_int ~> dvd_gcdD2 + coprime_dvd_mult_iff_nat ~> coprime_dvd_mult_iff + coprime_dvd_mult_iff_int ~> coprime_dvd_mult_iff + realpow_minus_mult ~> power_minus_mult + realpow_Suc_le_self ~> power_Suc_le_self + dvd_Gcd, dvd_Gcd_nat, dvd_Gcd_int removed in favour of Gcd_greatest +INCOMPATIBILITY. + +* Renamed HOL/Quotient_Examples/FSet.thy to +HOL/Quotient_Examples/Quotient_FSet.thy INCOMPATIBILITY. + +* Session HOL-Library: theory FinFun bundles "finfun_syntax" and +"no_finfun_syntax" allow to control optional syntax in local contexts; +this supersedes former theory FinFun_Syntax. INCOMPATIBILITY, e.g. use +"unbundle finfun_syntax" to imitate import of +"~~/src/HOL/Library/FinFun_Syntax". + +* Session HOL-Library: theory Multiset_Permutations (executably) defines +the set of permutations of a given set or multiset, i.e. the set of all +lists that contain every element of the carrier (multi-)set exactly +once. + +* Session HOL-Library: multiset membership is now expressed using +set_mset rather than count. + + - Expressions "count M a > 0" and similar simplify to membership + by default. + + - Converting between "count M a = 0" and non-membership happens using + equations count_eq_zero_iff and not_in_iff. + + - Rules count_inI and in_countE obtain facts of the form + "count M a = n" from membership. + + - Rules count_in_diffI and in_diff_countE obtain facts of the form + "count M a = n + count N a" from membership on difference sets. + +INCOMPATIBILITY. + +* Session HOL-Library: theory LaTeXsugar uses new-style "dummy_pats" for +displaying equations in functional programming style --- variables +present on the left-hand but not on the righ-hand side are replaced by +underscores. + +* Session HOL-Library: theory Combinator_PER provides combinator to +build partial equivalence relations from a predicate and an equivalence +relation. + +* Session HOL-Library: theory Perm provides basic facts about almost +everywhere fix bijections. + +* Session HOL-Library: theory Normalized_Fraction allows viewing an +element of a field of fractions as a normalized fraction (i.e. a pair of +numerator and denominator such that the two are coprime and the +denominator is normalized wrt. unit factors). + +* Session HOL-NSA has been renamed to HOL-Nonstandard_Analysis. + +* Session HOL-Multivariate_Analysis has been renamed to HOL-Analysis. + +* Session HOL-Analysis: measure theory has been moved here from +HOL-Probability. When importing HOL-Analysis some theorems need +additional name spaces prefixes due to name clashes. INCOMPATIBILITY. + +* Session HOL-Analysis: more complex analysis including Cauchy's +inequality, Liouville theorem, open mapping theorem, maximum modulus +principle, Residue theorem, Schwarz Lemma. + +* Session HOL-Analysis: Theory of polyhedra: faces, extreme points, +polytopes, and the Krein–Milman Minkowski theorem. + +* Session HOL-Analysis: Numerous results ported from the HOL Light +libraries: homeomorphisms, continuous function extensions, invariance of +domain. + +* Session HOL-Probability: the type of emeasure and nn_integral was +changed from ereal to ennreal, INCOMPATIBILITY. + + emeasure :: 'a measure \ 'a set \ ennreal + nn_integral :: 'a measure \ ('a \ ennreal) \ ennreal + +* Session HOL-Probability: Code generation and QuickCheck for +Probability Mass Functions. + +* Session HOL-Probability: theory Random_Permutations contains some +theory about choosing a permutation of a set uniformly at random and +folding over a list in random order. + +* Session HOL-Probability: theory SPMF formalises discrete +subprobability distributions. + +* Session HOL-Library: the names of multiset theorems have been +normalised to distinguish which ordering the theorems are about + + mset_less_eqI ~> mset_subset_eqI + mset_less_insertD ~> mset_subset_insertD + mset_less_eq_count ~> mset_subset_eq_count + mset_less_diff_self ~> mset_subset_diff_self + mset_le_exists_conv ~> mset_subset_eq_exists_conv + mset_le_mono_add_right_cancel ~> mset_subset_eq_mono_add_right_cancel + mset_le_mono_add_left_cancel ~> mset_subset_eq_mono_add_left_cancel + mset_le_mono_add ~> mset_subset_eq_mono_add + mset_le_add_left ~> mset_subset_eq_add_left + mset_le_add_right ~> mset_subset_eq_add_right + mset_le_single ~> mset_subset_eq_single + mset_le_multiset_union_diff_commute ~> mset_subset_eq_multiset_union_diff_commute + diff_le_self ~> diff_subset_eq_self + mset_leD ~> mset_subset_eqD + mset_lessD ~> mset_subsetD + mset_le_insertD ~> mset_subset_eq_insertD + mset_less_of_empty ~> mset_subset_of_empty + mset_less_size ~> mset_subset_size + wf_less_mset_rel ~> wf_subset_mset_rel + count_le_replicate_mset_le ~> count_le_replicate_mset_subset_eq + mset_remdups_le ~> mset_remdups_subset_eq + ms_lesseq_impl ~> subset_eq_mset_impl + +Some functions have been renamed: + ms_lesseq_impl -> subset_eq_mset_impl + +* HOL-Library: multisets are now ordered with the multiset ordering + #\# ~> \ + #\# ~> < + le_multiset ~> less_eq_multiset + less_multiset ~> le_multiset +INCOMPATIBILITY. + +* Session HOL-Library: the prefix multiset_order has been discontinued: +the theorems can be directly accessed. As a consequence, the lemmas +"order_multiset" and "linorder_multiset" have been discontinued, and the +interpretations "multiset_linorder" and "multiset_wellorder" have been +replaced by instantiations. INCOMPATIBILITY. + +* Session HOL-Library: some theorems about the multiset ordering have +been renamed: + + le_multiset_def ~> less_eq_multiset_def + less_multiset_def ~> le_multiset_def + less_eq_imp_le_multiset ~> subset_eq_imp_le_multiset + mult_less_not_refl ~> mset_le_not_refl + mult_less_trans ~> mset_le_trans + mult_less_not_sym ~> mset_le_not_sym + mult_less_asym ~> mset_le_asym + mult_less_irrefl ~> mset_le_irrefl + union_less_mono2{,1,2} ~> union_le_mono2{,1,2} + + le_multiset\<^sub>H\<^sub>O ~> less_eq_multiset\<^sub>H\<^sub>O + le_multiset_total ~> less_eq_multiset_total + less_multiset_right_total ~> subset_eq_imp_le_multiset + le_multiset_empty_left ~> less_eq_multiset_empty_left + le_multiset_empty_right ~> less_eq_multiset_empty_right + less_multiset_empty_right ~> le_multiset_empty_left + less_multiset_empty_left ~> le_multiset_empty_right + union_less_diff_plus ~> union_le_diff_plus + ex_gt_count_imp_less_multiset ~> ex_gt_count_imp_le_multiset + less_multiset_plus_left_nonempty ~> le_multiset_plus_left_nonempty + le_multiset_plus_right_nonempty ~> le_multiset_plus_right_nonempty +INCOMPATIBILITY. + +* Session HOL-Library: the lemma mset_map has now the attribute [simp]. +INCOMPATIBILITY. + +* Session HOL-Library: some theorems about multisets have been removed. +INCOMPATIBILITY, use the following replacements: + + le_multiset_plus_plus_left_iff ~> add_less_cancel_right + less_multiset_plus_plus_left_iff ~> add_less_cancel_right + le_multiset_plus_plus_right_iff ~> add_less_cancel_left + less_multiset_plus_plus_right_iff ~> add_less_cancel_left + add_eq_self_empty_iff ~> add_cancel_left_right + mset_subset_add_bothsides ~> subset_mset.add_less_cancel_right + mset_less_add_bothsides ~> subset_mset.add_less_cancel_right + mset_le_add_bothsides ~> subset_mset.add_less_cancel_right + empty_inter ~> subset_mset.inf_bot_left + inter_empty ~> subset_mset.inf_bot_right + empty_sup ~> subset_mset.sup_bot_left + sup_empty ~> subset_mset.sup_bot_right + bdd_below_multiset ~> subset_mset.bdd_above_bot + subset_eq_empty ~> subset_mset.le_zero_eq + le_empty ~> subset_mset.le_zero_eq + mset_subset_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero + mset_less_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero + +* Session HOL-Library: some typeclass constraints about multisets have +been reduced from ordered or linordered to preorder. Multisets have the +additional typeclasses order_bot, no_top, +ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add, +linordered_cancel_ab_semigroup_add, and +ordered_ab_semigroup_monoid_add_imp_le. INCOMPATIBILITY. + +* Session HOL-Library: there are some new simplification rules about +multisets, the multiset ordering, and the subset ordering on multisets. +INCOMPATIBILITY. + +* Session HOL-Library: the subset ordering on multisets has now the +interpretations ordered_ab_semigroup_monoid_add_imp_le and +bounded_lattice_bot. INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: single has been removed in favor +of add_mset that roughly corresponds to Set.insert. Some theorems have +removed or changed: + + single_not_empty ~> add_mset_not_empty or empty_not_add_mset + fold_mset_insert ~> fold_mset_add_mset + image_mset_insert ~> image_mset_add_mset + union_single_eq_diff + multi_self_add_other_not_self + diff_single_eq_union +INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: some theorems have been changed +to use add_mset instead of single: + + mset_add + multi_self_add_other_not_self + diff_single_eq_union + union_single_eq_diff + union_single_eq_member + add_eq_conv_diff + insert_noteq_member + add_eq_conv_ex + multi_member_split + multiset_add_sub_el_shuffle + mset_subset_eq_insertD + mset_subset_insertD + insert_subset_eq_iff + insert_union_subset_iff + multi_psub_of_add_self + inter_add_left1 + inter_add_left2 + inter_add_right1 + inter_add_right2 + sup_union_left1 + sup_union_left2 + sup_union_right1 + sup_union_right2 + size_eq_Suc_imp_eq_union + multi_nonempty_split + mset_insort + mset_update + mult1I + less_add + mset_zip_take_Cons_drop_twice + rel_mset_Zero + msed_map_invL + msed_map_invR + msed_rel_invL + msed_rel_invR + le_multiset_right_total + multiset_induct + multiset_induct2_size + multiset_induct2 +INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: the definitions of some +constants have changed to use add_mset instead of adding a single +element: + + image_mset + mset + replicate_mset + mult1 + pred_mset + rel_mset' + mset_insort + +INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: due to the above changes, the +attributes of some multiset theorems have been changed: + + insert_DiffM [] ~> [simp] + insert_DiffM2 [simp] ~> [] + diff_add_mset_swap [simp] + fold_mset_add_mset [simp] + diff_diff_add [simp] (for multisets only) + diff_cancel [simp] ~> [] + count_single [simp] ~> [] + set_mset_single [simp] ~> [] + size_multiset_single [simp] ~> [] + size_single [simp] ~> [] + image_mset_single [simp] ~> [] + mset_subset_eq_mono_add_right_cancel [simp] ~> [] + mset_subset_eq_mono_add_left_cancel [simp] ~> [] + fold_mset_single [simp] ~> [] + subset_eq_empty [simp] ~> [] + empty_sup [simp] ~> [] + sup_empty [simp] ~> [] + inter_empty [simp] ~> [] + empty_inter [simp] ~> [] +INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: the order of the variables in +the second cases of multiset_induct, multiset_induct2_size, +multiset_induct2 has been changed (e.g. Add A a ~> Add a A). +INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: there is now a simplification +procedure on multisets. It mimics the behavior of the procedure on +natural numbers. INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: renamed sums and products of +multisets: + + msetsum ~> sum_mset + msetprod ~> prod_mset + +* Session HOL-Library, theory Multiset: the notation for intersection +and union of multisets have been changed: + + #\ ~> \# + #\ ~> \# + +INCOMPATIBILITY. + +* Session HOL-Library, theory Multiset: the lemma +one_step_implies_mult_aux on multisets has been removed, use +one_step_implies_mult instead. INCOMPATIBILITY. + +* Session HOL-Library: theory Complete_Partial_Order2 provides reasoning +support for monotonicity and continuity in chain-complete partial orders +and about admissibility conditions for fixpoint inductions. + +* Session HOL-Library: theory Library/Polynomial contains also +derivation of polynomials (formerly in Library/Poly_Deriv) but not +gcd/lcm on polynomials over fields. This has been moved to a separate +theory Library/Polynomial_GCD_euclidean.thy, to pave way for a possible +future different type class instantiation for polynomials over factorial +rings. INCOMPATIBILITY. + +* Session HOL-Library: theory Sublist provides function "prefixes" with +the following renaming + + prefixeq -> prefix + prefix -> strict_prefix + suffixeq -> suffix + suffix -> strict_suffix + +Added theory of longest common prefixes. + +* Session HOL-Number_Theory: algebraic foundation for primes: +Generalisation of predicate "prime" and introduction of predicates +"prime_elem", "irreducible", a "prime_factorization" function, and the +"factorial_ring" typeclass with instance proofs for nat, int, poly. Some +theorems now have different names, most notably "prime_def" is now +"prime_nat_iff". INCOMPATIBILITY. + +* Session Old_Number_Theory has been removed, after porting remaining +theories. + +* Session HOL-Types_To_Sets provides an experimental extension of +Higher-Order Logic to allow translation of types to sets. + + +*** ML *** + +* Integer.gcd and Integer.lcm use efficient operations from the Poly/ML +library (notably for big integers). Subtle change of semantics: +Integer.gcd and Integer.lcm both normalize the sign, results are never +negative. This coincides with the definitions in HOL/GCD.thy. +INCOMPATIBILITY. + +* Structure Rat for rational numbers is now an integral part of +Isabelle/ML, with special notation @int/nat or @int for numerals (an +abbreviation for antiquotation @{Pure.rat argument}) and ML pretty +printing. Standard operations on type Rat.rat are provided via ad-hoc +overloading of + - * / < <= > >= ~ abs. INCOMPATIBILITY, need to +use + instead of +/ etc. Moreover, exception Rat.DIVZERO has been +superseded by General.Div. + +* ML antiquotation @{path} is superseded by @{file}, which ensures that +the argument is a plain file. Minor INCOMPATIBILITY. + +* Antiquotation @{make_string} is available during Pure bootstrap -- +with approximative output quality. + +* Low-level ML system structures (like PolyML and RunCall) are no longer +exposed to Isabelle/ML user-space. Potential INCOMPATIBILITY. + +* The ML function "ML" provides easy access to run-time compilation. +This is particularly useful for conditional compilation, without +requiring separate files. + +* Option ML_exception_debugger controls detailed exception trace via the +Poly/ML debugger. Relevant ML modules need to be compiled beforehand +with ML_file_debug, or with ML_file and option ML_debugger enabled. Note +debugger information requires consirable time and space: main +Isabelle/HOL with full debugger support may need ML_system_64. + +* Local_Theory.restore has been renamed to Local_Theory.reset to +emphasize its disruptive impact on the cumulative context, notably the +scope of 'private' or 'qualified' names. Note that Local_Theory.reset is +only appropriate when targets are managed, e.g. starting from a global +theory and returning to it. Regular definitional packages should use +balanced blocks of Local_Theory.open_target versus +Local_Theory.close_target instead. Rare INCOMPATIBILITY. + +* Structure TimeLimit (originally from the SML/NJ library) has been +replaced by structure Timeout, with slightly different signature. +INCOMPATIBILITY. + +* Discontinued cd and pwd operations, which are not well-defined in a +multi-threaded environment. Note that files are usually located +relatively to the master directory of a theory (see also +File.full_path). Potential INCOMPATIBILITY. + +* Binding.empty_atts supersedes Thm.empty_binding and +Attrib.empty_binding. Minor INCOMPATIBILITY. + + +*** System *** + +* SML/NJ and old versions of Poly/ML are no longer supported. + +* Poly/ML heaps now follow the hierarchy of sessions, and thus require +much less disk space. + +* The Isabelle ML process is now managed directly by Isabelle/Scala, and +shell scripts merely provide optional command-line access. In +particular: + + . Scala module ML_Process to connect to the raw ML process, + with interaction via stdin/stdout/stderr or in batch mode; + . command-line tool "isabelle console" as interactive wrapper; + . command-line tool "isabelle process" as batch mode wrapper. + +* The executable "isabelle_process" has been discontinued. Tools and +prover front-ends should use ML_Process or Isabelle_Process in +Isabelle/Scala. INCOMPATIBILITY. + +* New command-line tool "isabelle process" supports ML evaluation of +literal expressions (option -e) or files (option -f) in the context of a +given heap image. Errors lead to premature exit of the ML process with +return code 1. + +* The command-line tool "isabelle build" supports option -N for cyclic +shuffling of NUMA CPU nodes. This may help performance tuning on Linux +servers with separate CPU/memory modules. + +* System option "threads" (for the size of the Isabelle/ML thread farm) +is also passed to the underlying ML runtime system as --gcthreads, +unless there is already a default provided via ML_OPTIONS settings. + +* System option "checkpoint" helps to fine-tune the global heap space +management of isabelle build. This is relevant for big sessions that may +exhaust the small 32-bit address space of the ML process (which is used +by default). + +* System option "profiling" specifies the mode for global ML profiling +in "isabelle build". Possible values are "time", "allocations". The +command-line tool "isabelle profiling_report" helps to digest the +resulting log files. + +* System option "ML_process_policy" specifies an optional command prefix +for the underlying ML process, e.g. to control CPU affinity on +multiprocessor systems. The "isabelle jedit" tool allows to override the +implicit default via option -p. + +* Command-line tool "isabelle console" provides option -r to help to +bootstrapping Isabelle/Pure interactively. + +* Command-line tool "isabelle yxml" has been discontinued. +INCOMPATIBILITY, use operations from the modules "XML" and "YXML" in +Isabelle/ML or Isabelle/Scala. + +* Many Isabelle tools that require a Java runtime system refer to the +settings ISABELLE_TOOL_JAVA_OPTIONS32 / ISABELLE_TOOL_JAVA_OPTIONS64, +depending on the underlying platform. The settings for "isabelle build" +ISABELLE_BUILD_JAVA_OPTIONS32 / ISABELLE_BUILD_JAVA_OPTIONS64 have been +discontinued. Potential INCOMPATIBILITY. + +* The Isabelle system environment always ensures that the main +executables are found within the shell search $PATH: "isabelle" and +"isabelle_scala_script". + +* Isabelle tools may consist of .scala files: the Scala compiler is +invoked on the spot. The source needs to define some object that extends +Isabelle_Tool.Body. + +* File.bash_string, File.bash_path etc. represent Isabelle/ML and +Isabelle/Scala strings authentically within GNU bash. This is useful to +produce robust shell scripts under program control, without worrying +about spaces or special characters. Note that user output works via +Path.print (ML) or Path.toString (Scala). INCOMPATIBILITY, the old (and +less versatile) operations File.shell_quote, File.shell_path etc. have +been discontinued. + +* The isabelle_java executable allows to run a Java process within the +name space of Java and Scala components that are bundled with Isabelle, +but without the Isabelle settings environment. + +* Isabelle/Scala: the SSH module supports ssh and sftp connections, for +remote command-execution and file-system access. This resembles +operations from module File and Isabelle_System to some extent. Note +that Path specifications need to be resolved remotely via +ssh.remote_path instead of File.standard_path: the implicit process +environment is different, Isabelle settings are not available remotely. + +* Isabelle/Scala: the Mercurial module supports repositories via the +regular hg command-line interface. The repositroy clone and working +directory may reside on a local or remote file-system (via ssh +connection). + + + +New in Isabelle2016 (February 2016) +----------------------------------- + +*** General *** + +* Eisbach is now based on Pure instead of HOL. Objects-logics may import +either the theory ~~/src/HOL/Eisbach/Eisbach (for HOL etc.) or +~~/src/HOL/Eisbach/Eisbach_Old_Appl_Syntax (for FOL, ZF etc.). Note that +the HOL-Eisbach session located in ~~/src/HOL/Eisbach/ contains further +examples that do require HOL. + +* Better resource usage on all platforms (Linux, Windows, Mac OS X) for +both Isabelle/ML and Isabelle/Scala. Slightly reduced heap space usage. + +* Former "xsymbols" syntax with Isabelle symbols is used by default, +without any special print mode. Important ASCII replacement syntax +remains available under print mode "ASCII", but less important syntax +has been removed (see below). + +* Support for more arrow symbols, with rendering in LaTeX and Isabelle +fonts: \ \ \ \ \ \. + +* Special notation \ 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. \\<^bsub>A\<^esub>). + +* The glyph for \ in the IsabelleText font now corresponds better to its +counterpart \ as quantifier-like symbol. A small diamond is available as +\; the old symbol \ loses this rendering and any special +meaning. + +* Syntax for formal comments "-- text" now also supports the symbolic +form "\ text". Command-line tool "isabelle update_cartouches -c" helps +to update old sources. + +* Toplevel theorem statements have been simplified as follows: + + theorems ~> lemmas + schematic_lemma ~> schematic_goal + schematic_theorem ~> schematic_goal + schematic_corollary ~> schematic_goal + +Command-line tool "isabelle update_theorems" updates theory sources +accordingly. + +* Toplevel theorem statement 'proposition' is another alias for +'theorem'. + +* The old 'defs' command has been removed (legacy since Isabelle2014). +INCOMPATIBILITY, use regular 'definition' instead. Overloaded and/or +deferred definitions require a surrounding 'overloading' block. + + +*** Prover IDE -- Isabelle/Scala/jEdit *** + +* IDE support for the source-level debugger of Poly/ML, to work with +Isabelle/ML and official Standard ML. Option "ML_debugger" and commands +'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug', +'SML_file_no_debug' control compilation of sources with or without +debugging information. The Debugger panel allows to set breakpoints (via +context menu), step through stopped threads, evaluate local ML +expressions etc. At least one Debugger view needs to be active to have +any effect on the running ML program. + +* The State panel manages explicit proof state output, with dynamic +auto-update according to cursor movement. Alternatively, the jEdit +action "isabelle.update-state" (shortcut S+ENTER) triggers manual +update. + +* The Output panel no longer shows proof state output by default, to +avoid GUI overcrowding. INCOMPATIBILITY, use the State panel instead or +enable option "editor_output_state". + +* The text overview column (status of errors, warnings etc.) is updated +asynchronously, leading to much better editor reactivity. Moreover, the +full document node content is taken into account. The width of the +column is scaled according to the main text area font, for improved +visibility. + +* The main text area no longer changes its color hue in outdated +situations. The text overview column takes over the role to indicate +unfinished edits in the PIDE pipeline. This avoids flashing text display +due to ad-hoc updates by auxiliary GUI components, such as the State +panel. + +* Slightly improved scheduling for urgent print tasks (e.g. command +state output, interactive queries) wrt. long-running background tasks. + +* Completion of symbols via prefix of \ or \<^name> or \name is +always possible, independently of the language context. It is never +implicit: a popup will show up unconditionally. + +* Additional abbreviations for syntactic completion may be specified in +$ISABELLE_HOME/etc/abbrevs and $ISABELLE_HOME_USER/etc/abbrevs, with +support for simple templates using ASCII 007 (bell) as placeholder. + +* Symbols \, \, \, \, \, \, \, \ no longer provide abbreviations for +completion like "+o", "*o", ".o" etc. -- due to conflicts with other +ASCII syntax. INCOMPATIBILITY, use plain backslash-completion or define +suitable abbreviations in $ISABELLE_HOME_USER/etc/abbrevs. + +* Action "isabelle-emph" (with keyboard shortcut C+e LEFT) controls +emphasized text style; the effect is visible in document output, not in +the editor. + +* Action "isabelle-reset" now uses keyboard shortcut C+e BACK_SPACE, +instead of former C+e LEFT. + +* The command-line tool "isabelle jedit" and the isabelle.Main +application wrapper treat the default $USER_HOME/Scratch.thy more +uniformly, and allow the dummy file argument ":" to open an empty buffer +instead. + +* New command-line tool "isabelle jedit_client" allows to connect to an +already running Isabelle/jEdit process. This achieves the effect of +single-instance applications seen on common GUI desktops. + +* The default look-and-feel for Linux is the traditional "Metal", which +works better with GUI scaling for very high-resolution displays (e.g. +4K). Moreover, it is generally more robust than "Nimbus". + +* Update to jedit-5.3.0, with improved GUI scaling and support of +high-resolution displays (e.g. 4K). + +* The main Isabelle executable is managed as single-instance Desktop +application uniformly on all platforms: Linux, Windows, Mac OS X. + + +*** Document preparation *** + +* Commands 'paragraph' and 'subparagraph' provide additional section +headings. Thus there are 6 levels of standard headings, as in HTML. + +* Command 'text_raw' has been clarified: input text is processed as in +'text' (with antiquotations and control symbols). The key difference is +the lack of the surrounding isabelle markup environment in output. + +* Text is structured in paragraphs and nested lists, using notation that +is similar to Markdown. The control symbols for list items are as +follows: + + \<^item> itemize + \<^enum> enumerate + \<^descr> description + +* There is a new short form for antiquotations with a single argument +that is a cartouche: \<^name>\...\ is equivalent to @{name \...\} and +\...\ without control symbol is equivalent to @{cartouche \...\}. +\<^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, \...\ 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 \"foo"\}. 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>\...\ and \<^bold>\...\. + +* Document antiquotation @{footnote} outputs LaTeX source recursively, +marked as \footnote{}. This may be used in the short form \<^footnote>\...\. + +* Antiquotation @{verbatim [display]} supports option "indent". + +* Antiquotation @{theory_text} prints uninterpreted theory source text +(Isar outer syntax with command keywords etc.). This may be used in the +short form \<^theory_text>\...\. @{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 + + +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" + } + 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 "\x y::'a. A x \ B y \ C x y" and produces a general +result as usual: "A ?x \ B ?y \ 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 "\x. A x" 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 "\x. A x \ B x \ C x" can be solved completely as +follows: + + show "\x. A x \ B x \ 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 + then have something + proof cases + case a + then show ?thesis + next + case b + then show ?thesis + next + case c + then show ?thesis + 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 \ 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 "\x. A x \ B x \ C x" + and "\y z. U y \ V z \ W y z" +proof goal_cases + case (1 x) + then show ?case using \A x\ \B x\ sorry +next + case (2 y z) + then show ?case using \U y\ \V z\ sorry +qed + +lemma "\x. A x \ B x \ C x" + and "\y z. U y \ V z \ W y z" +proof goal_cases + case prems: 1 + then show ?case using prems sorry +next + case prems: 2 + then show ?case using prems sorry +qed + +* The undocumented feature of implicit cases goal1, goal2, goal3, etc. +is marked as legacy, and will be removed eventually. The proof method +"goals" achieves a similar effect within regular Isar; often it can be +done more adequately by other means (e.g. 'consider'). + +* The vacuous fact "TERM x" may be established "by fact" or as `TERM x` +as well, not just "by this" or "." as before. + +* Method "sleep" succeeds after a real-time delay (in seconds). This is +occasionally useful for demonstration and testing purposes. + + +*** Pure *** + +* Qualifiers in locale expressions default to mandatory ('!') regardless +of the command. Previously, for 'locale' and 'sublocale' the default was +optional ('?'). The old synatx '!' has been discontinued. +INCOMPATIBILITY, remove '!' and add '?' as required. + +* Keyword 'rewrites' identifies rewrite morphisms in interpretation +commands. Previously, the keyword was 'where'. INCOMPATIBILITY. + +* More gentle suppression of syntax along locale morphisms while +printing terms. Previously 'abbreviation' and 'notation' declarations +would be suppressed for morphisms except term identity. Now +'abbreviation' is also kept for morphims that only change the involved +parameters, and only 'notation' is suppressed. This can be of great help +when working with complex locale hierarchies, because proof states are +displayed much more succinctly. It also means that only notation needs +to be redeclared if desired, as illustrated by this example: + + locale struct = fixes composition :: "'a => 'a => 'a" (infixl "\" 65) + begin + definition derived (infixl "\" 65) where ... + end + + locale morphism = + left: struct composition + right: struct composition' + for composition (infix "\" 65) and composition' (infix "\''" 65) + begin + notation right.derived ("\''") + 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 "\" for type and sort constraints has been +removed: in LaTeX document output it looks the same as "::". +INCOMPATIBILITY, use plain "::" instead. + +* Commands 'inductive' and 'inductive_set' work better when names for +intro rules are omitted: the "cases" and "induct" rules no longer +declare empty case_names, but no case_names at all. This allows to use +numbered cases in proofs, without requiring method "goal_cases". + +* Inductive definitions ('inductive', 'coinductive', etc.) expose +low-level facts of the internal construction only if the option +"inductive_internals" is enabled. This refers to the internal predicate +definition and its monotonicity result. Rare INCOMPATIBILITY. + +* Recursive function definitions ('fun', 'function', 'partial_function') +expose low-level facts of the internal construction only if the option +"function_internals" is enabled. Its internal inductive definition is +also subject to "inductive_internals". Rare INCOMPATIBILITY. + +* BNF datatypes ('datatype', 'codatatype', etc.) expose low-level facts +of the internal construction only if the option "bnf_internals" is +enabled. This supersedes the former option "bnf_note_all". Rare +INCOMPATIBILITY. + +* Combinator to represent case distinction on products is named +"case_prod", uniformly, discontinuing any input aliasses. Very popular +theorem aliasses have been retained. + +Consolidated facts: + PairE ~> prod.exhaust + Pair_eq ~> prod.inject + pair_collapse ~> prod.collapse + Pair_fst_snd_eq ~> prod_eq_iff + split_twice ~> prod.case_distrib + split_weak_cong ~> prod.case_cong_weak + split_split ~> prod.split + split_split_asm ~> prod.split_asm + splitI ~> case_prodI + splitD ~> case_prodD + splitI2 ~> case_prodI2 + splitI2' ~> case_prodI2' + splitE ~> case_prodE + splitE' ~> case_prodE' + split_pair ~> case_prod_Pair + split_eta ~> case_prod_eta + split_comp ~> case_prod_comp + mem_splitI ~> mem_case_prodI + mem_splitI2 ~> mem_case_prodI2 + mem_splitE ~> mem_case_prodE + The_split ~> The_case_prod + cond_split_eta ~> cond_case_prod_eta + Collect_split_in_rel_leE ~> Collect_case_prod_in_rel_leE + Collect_split_in_rel_leI ~> Collect_case_prod_in_rel_leI + in_rel_Collect_split_eq ~> in_rel_Collect_case_prod_eq + Collect_split_Grp_eqD ~> Collect_case_prod_Grp_eqD + Collect_split_Grp_inD ~> Collect_case_prod_Grp_in + Domain_Collect_split ~> Domain_Collect_case_prod + Image_Collect_split ~> Image_Collect_case_prod + Range_Collect_split ~> Range_Collect_case_prod + Eps_split ~> Eps_case_prod + Eps_split_eq ~> Eps_case_prod_eq + split_rsp ~> case_prod_rsp + curry_split ~> curry_case_prod + split_curry ~> case_prod_curry + +Changes in structure HOLogic: + split_const ~> case_prod_const + mk_split ~> mk_case_prod + mk_psplits ~> mk_ptupleabs + strip_psplits ~> strip_ptupleabs + +INCOMPATIBILITY. + +* The coercions to type 'real' have been reorganised. The function +'real' is no longer overloaded, but has type 'nat => real' and +abbreviates of_nat for that type. Also 'real_of_int :: int => real' +abbreviates of_int for that type. Other overloaded instances of 'real' +have been replaced by 'real_of_ereal' and 'real_of_float'. + +Consolidated facts (among others): + real_of_nat_le_iff -> of_nat_le_iff + real_of_nat_numeral of_nat_numeral + real_of_int_zero of_int_0 + real_of_nat_zero of_nat_0 + real_of_one of_int_1 + real_of_int_add of_int_add + real_of_nat_add of_nat_add + real_of_int_diff of_int_diff + real_of_nat_diff of_nat_diff + floor_subtract floor_diff_of_int + real_of_int_inject of_int_eq_iff + real_of_int_gt_zero_cancel_iff of_int_0_less_iff + real_of_int_ge_zero_cancel_iff of_int_0_le_iff + real_of_nat_ge_zero of_nat_0_le_iff + real_of_int_ceiling_ge le_of_int_ceiling + ceiling_less_eq ceiling_less_iff + ceiling_le_eq ceiling_le_iff + less_floor_eq less_floor_iff + floor_less_eq floor_less_iff + floor_divide_eq_div floor_divide_of_int_eq + real_of_int_zero_cancel of_nat_eq_0_iff + ceiling_real_of_int ceiling_of_int + +INCOMPATIBILITY. + +* Theory Map: lemma map_of_is_SomeD was a clone of map_of_SomeD and has +been removed. INCOMPATIBILITY. + +* Quickcheck setup for finite sets. + +* Discontinued simp_legacy_precond. Potential INCOMPATIBILITY. + +* Sledgehammer: + - The MaSh relevance filter has been sped up. + - Proof reconstruction has been improved, to minimize the incidence of + cases where Sledgehammer gives a proof that does not work. + - Auto Sledgehammer now minimizes and preplays the results. + - Handle Vampire 4.0 proof output without raising exception. + - Eliminated "MASH" environment variable. Use the "MaSh" option in + Isabelle/jEdit instead. INCOMPATIBILITY. + - Eliminated obsolete "blocking" option and related subcommands. + +* Nitpick: + - Fixed soundness bug in translation of "finite" predicate. + - Fixed soundness bug in "destroy_constrs" optimization. + - Fixed soundness bug in translation of "rat" type. + - Removed "check_potential" and "check_genuine" options. + - Eliminated obsolete "blocking" option. + +* (Co)datatype package: + - New commands "lift_bnf" and "copy_bnf" for lifting (copying) a BNF + structure on the raw type to an abstract type defined using typedef. + - Always generate "case_transfer" theorem. + - For mutual types, generate slightly stronger "rel_induct", + "rel_coinduct", and "coinduct" theorems. INCOMPATIBILITY. + - Allow discriminators and selectors with the same name as the type + being defined. + - Avoid various internal name clashes (e.g., 'datatype f = f'). + +* Transfer: new methods for interactive debugging of 'transfer' and +'transfer_prover': 'transfer_start', 'transfer_step', 'transfer_end', +'transfer_prover_start' and 'transfer_prover_end'. + +* New diagnostic command print_record for displaying record definitions. + +* Division on integers is bootstrapped directly from division on +naturals and uses generic numeral algorithm for computations. Slight +INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes former +simprocs binary_int_div and binary_int_mod + +* Tightened specification of class semiring_no_zero_divisors. Minor +INCOMPATIBILITY. + +* Class algebraic_semidom introduces common algebraic notions of +integral (semi)domains, particularly units. Although logically subsumed +by fields, is is not a super class of these in order not to burden +fields with notions that are trivial there. + +* Class normalization_semidom specifies canonical representants for +equivalence classes of associated elements in an integral (semi)domain. +This formalizes associated elements as well. + +* Abstract specification of gcd/lcm operations in classes semiring_gcd, +semiring_Gcd, semiring_Lcd. Minor INCOMPATIBILITY: facts gcd_nat.commute +and gcd_int.commute are subsumed by gcd.commute, as well as +gcd_nat.assoc and gcd_int.assoc by gcd.assoc. + +* Former constants Fields.divide (_ / _) and Divides.div (_ div _) are +logically unified to Rings.divide in syntactic type class Rings.divide, +with infix syntax (_ div _). Infix syntax (_ / _) for field division is +added later as abbreviation in class Fields.inverse. INCOMPATIBILITY, +instantiations must refer to Rings.divide rather than the former +separate constants, hence infix syntax (_ / _) is usually not available +during instantiation. + +* New cancellation simprocs for boolean algebras to cancel complementary +terms for sup and inf. For example, "sup x (sup y (- x))" simplifies to +"top". INCOMPATIBILITY. + +* Class uniform_space introduces uniform spaces btw topological spaces +and metric spaces. Minor INCOMPATIBILITY: open__def needs to be +introduced in the form of an uniformity. Some constants are more general +now, it may be necessary to add type class constraints. + + open_real_def \ open_dist + open_complex_def \ open_dist + +* Library/Monad_Syntax: notation uses symbols \ and \. INCOMPATIBILITY. + +* Library/Multiset: + - Renamed multiset inclusion operators: + < ~> <# + > ~> ># + <= ~> <=# + >= ~> >=# + \ ~> \# + \ ~> \# + INCOMPATIBILITY. + - Added multiset inclusion operator syntax: + \# + \# + \# + \# + - "'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 \ 'a. + +* Library/Bourbaki_Witt_Fixpoint: Added formalisation of the +Bourbaki-Witt fixpoint theorem for increasing functions in +chain-complete partial orders. + +* Library/Old_Recdef: discontinued obsolete 'defer_recdef' command. +Minor INCOMPATIBILITY, use 'function' instead. + +* Library/Periodic_Fun: a locale that provides convenient lemmas for +periodic functions. + +* Library/Formal_Power_Series: proper definition of division (with +remainder) for formal power series; instances for Euclidean Ring and +GCD. + +* HOL-Imperative_HOL: obsolete theory Legacy_Mrec has been removed. + +* HOL-Statespace: command 'statespace' uses mandatory qualifier for +import of parent, as for general 'locale' expressions. INCOMPATIBILITY, +remove '!' and add '?' as required. + +* HOL-Decision_Procs: The "approximation" method works with "powr" +(exponentiation on real numbers) again. + +* HOL-Multivariate_Analysis: theory Cauchy_Integral_Thm with Contour +integrals (= complex path integrals), Cauchy's integral theorem, winding +numbers and Cauchy's integral formula, Liouville theorem, Fundamental +Theorem of Algebra. Ported from HOL Light. + +* HOL-Multivariate_Analysis: topological concepts such as connected +components, homotopic paths and the inside or outside of a set. + +* HOL-Multivariate_Analysis: radius of convergence of power series and +various summability tests; Harmonic numbers and the Euler–Mascheroni +constant; the Generalised Binomial Theorem; the complex and real +Gamma/log-Gamma/Digamma/ Polygamma functions and their most important +properties. + +* HOL-Probability: The central limit theorem based on Levy's uniqueness +and continuity theorems, weak convergence, and characterisitc functions. + +* HOL-Data_Structures: new and growing session of standard data +structures. + + +*** ML *** + +* The following combinators for low-level profiling of the ML runtime +system are available: + + profile_time (*CPU time*) + profile_time_thread (*CPU time on this thread*) + profile_allocations (*overall heap allocations*) + +* Antiquotation @{undefined} or \<^undefined> inlines (raise Match). + +* Antiquotation @{method NAME} inlines the (checked) name of the given +Isar proof method. + +* Pretty printing of Poly/ML compiler output in Isabelle has been +improved: proper treatment of break offsets and blocks with consistent +breaks. + +* The auxiliary module Pure/display.ML has been eliminated. Its +elementary thm print operations are now in Pure/more_thm.ML and thus +called Thm.pretty_thm, Thm.string_of_thm etc. INCOMPATIBILITY. + +* Simproc programming interfaces have been simplified: +Simplifier.make_simproc and Simplifier.define_simproc supersede various +forms of Simplifier.mk_simproc, Simplifier.simproc_global etc. Note that +term patterns for the left-hand sides are specified with implicitly +fixed variables, like top-level theorem statements. INCOMPATIBILITY. + +* Instantiation rules have been re-organized as follows: + + Thm.instantiate (*low-level instantiation with named arguments*) + Thm.instantiate' (*version with positional arguments*) + + Drule.infer_instantiate (*instantiation with type inference*) + Drule.infer_instantiate' (*version with positional arguments*) + +The LHS only requires variable specifications, instead of full terms. +Old cterm_instantiate is superseded by infer_instantiate. +INCOMPATIBILITY, need to re-adjust some ML names and types accordingly. + +* Old tactic shorthands atac, rtac, etac, dtac, ftac have been +discontinued. INCOMPATIBILITY, use regular assume_tac, resolve_tac etc. +instead (with proper context). + +* Thm.instantiate (and derivatives) no longer require the LHS of the +instantiation to be certified: plain variables are given directly. + +* Subgoal.SUBPROOF and Subgoal.FOCUS combinators use anonymous +quasi-bound variables (like the Simplifier), instead of accidentally +named local fixes. This has the potential to improve stability of proof +tools, but can also cause INCOMPATIBILITY for tools that don't observe +the proof context discipline. + +* Isar proof methods are based on a slightly more general type +context_tactic, which allows to change the proof context dynamically +(e.g. to update cases) and indicate explicit Seq.Error results. Former +METHOD_CASES is superseded by CONTEXT_METHOD; further combinators are +provided in src/Pure/Isar/method.ML for convenience. INCOMPATIBILITY. + + +*** System *** + +* Command-line tool "isabelle console" enables print mode "ASCII". + +* Command-line tool "isabelle update_then" expands old Isar command +conflations: + + hence ~> then have + thus ~> then show + +This syntax is more orthogonal and improves readability and +maintainability of proofs. + +* Global session timeout is multiplied by timeout_scale factor. This +allows to adjust large-scale tests (e.g. AFP) to overall hardware +performance. + +* Property values in etc/symbols may contain spaces, if written with the +replacement character "␣" (Unicode point 0x2324). For example: + + \ code: 0x0022c6 group: operator font: Deja␣Vu␣Sans␣Mono + +* Java runtime environment for x86_64-windows allows to use larger heap +space. + +* Java runtime options are determined separately for 32bit vs. 64bit +platforms as follows. + + - Isabelle desktop application: platform-specific files that are + associated with the main app bundle + + - isabelle jedit: settings + JEDIT_JAVA_SYSTEM_OPTIONS + JEDIT_JAVA_OPTIONS32 vs. JEDIT_JAVA_OPTIONS64 + + - isabelle build: settings + ISABELLE_BUILD_JAVA_OPTIONS32 vs. ISABELLE_BUILD_JAVA_OPTIONS64 + +* Bash shell function "jvmpath" has been renamed to "platform_path": it +is relevant both for Poly/ML and JVM processes. + +* Poly/ML default platform architecture may be changed from 32bit to +64bit via system option ML_system_64. A system restart (and rebuild) is +required after change. + +* Poly/ML 5.6 runs natively on x86-windows and x86_64-windows, which +both allow larger heap space than former x86-cygwin. + +* Heap images are 10-15% smaller due to less wasteful persistent theory +content (using ML type theory_id instead of theory); + + + +New in Isabelle2015 (May 2015) +------------------------------ + +*** General *** + +* Local theory specification commands may have a 'private' or +'qualified' modifier to restrict name space accesses to the local scope, +as provided by some "context begin ... end" block. For example: + + context + begin + + private definition ... + private lemma ... + + qualified definition ... + qualified lemma ... + + lemma ... + theorem ... + + end + +* Command 'experiment' opens an anonymous locale context with private +naming policy. + +* Command 'notepad' requires proper nesting of begin/end and its proof +structure in the body: 'oops' is no longer supported here. Minor +INCOMPATIBILITY, use 'sorry' instead. + +* Command 'named_theorems' declares a dynamic fact within the context, +together with an attribute to maintain the content incrementally. This +supersedes functor Named_Thms in Isabelle/ML, but with a subtle change +of semantics due to external visual order vs. internal reverse order. + +* 'find_theorems': search patterns which are abstractions are +schematically expanded before search. Search results match the naive +expectation more closely, particularly wrt. abbreviations. +INCOMPATIBILITY. + +* Commands 'method_setup' and 'attribute_setup' now work within a local +theory context. + +* Outer syntax commands are managed authentically within the theory +context, without implicit global state. Potential for accidental +INCOMPATIBILITY, make sure that required theories are really imported. + +* Historical command-line terminator ";" is no longer accepted (and +already used differently in Isar). Minor INCOMPATIBILITY, use "isabelle +update_semicolons" to remove obsolete semicolons from old theory +sources. + +* Structural composition of proof methods (meth1; meth2) in Isar +corresponds to (tac1 THEN_ALL_NEW tac2) in ML. + +* The Eisbach proof method language allows to define new proof methods +by combining existing ones with their usual syntax. The "match" proof +method provides basic fact/term matching in addition to +premise/conclusion matching through Subgoal.focus, and binds fact names +from matches as well as term patterns within matches. The Isabelle +documentation provides an entry "eisbach" for the Eisbach User Manual. +Sources and various examples are in ~~/src/HOL/Eisbach/. + + +*** Prover IDE -- Isabelle/Scala/jEdit *** + +* Improved folding mode "isabelle" based on Isar syntax. Alternatively, +the "sidekick" mode may be used for document structure. + +* Extended bracket matching based on Isar language structure. System +option jedit_structure_limit determines maximum number of lines to scan +in the buffer. + +* Support for BibTeX files: context menu, context-sensitive token +marker, SideKick parser. + +* Document antiquotation @{cite} provides formal markup, which is +interpreted semi-formally based on .bib files that happen to be open in +the editor (hyperlinks, completion etc.). + +* Less waste of vertical space via negative line spacing (see Global +Options / Text Area). + +* Improved graphview panel with optional output of PNG or PDF, for +display of 'thy_deps', 'class_deps' etc. + +* The commands 'thy_deps' and 'class_deps' allow optional bounds to +restrict the visualized hierarchy. + +* Improved scheduling for asynchronous print commands (e.g. provers +managed by the Sledgehammer panel) wrt. ongoing document processing. + + +*** Document preparation *** + +* Document markup commands 'chapter', 'section', 'subsection', +'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any +context, even before the initial 'theory' command. Obsolete proof +commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been +discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw' +instead. The old 'header' command is still retained for some time, but +should be replaced by 'chapter', 'section' etc. (using "isabelle +update_header"). Minor INCOMPATIBILITY. + +* Official support for "tt" style variants, via \isatt{...} or +\begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or +verbatim environment of LaTeX is no longer used. This allows @{ML} etc. +as argument to other macros (such as footnotes). + +* Document antiquotation @{verbatim} prints ASCII text literally in "tt" +style. + +* Discontinued obsolete option "document_graph": session_graph.pdf is +produced unconditionally for HTML browser_info and PDF-LaTeX document. + +* Diagnostic commands and document markup commands within a proof do not +affect the command tag for output. Thus commands like 'thm' are subject +to proof document structure, and no longer "stick out" accidentally. +Commands 'text' and 'txt' merely differ in the LaTeX style, not their +tags. Potential INCOMPATIBILITY in exotic situations. + +* System option "pretty_margin" is superseded by "thy_output_margin", +which is also accessible via document antiquotation option "margin". +Only the margin for document output may be changed, but not the global +pretty printing: that is 76 for plain console output, and adapted +dynamically in GUI front-ends. Implementations of document +antiquotations need to observe the margin explicitly according to +Thy_Output.string_of_margin. Minor INCOMPATIBILITY. + +* Specification of 'document_files' in the session ROOT file is +mandatory for document preparation. The legacy mode with implicit +copying of the document/ directory is no longer supported. Minor +INCOMPATIBILITY. + + +*** Pure *** + +* Proof methods with explicit instantiation ("rule_tac", "subgoal_tac" +etc.) allow an optional context of local variables ('for' declaration): +these variables become schematic in the instantiated theorem; this +behaviour is analogous to 'for' in attributes "where" and "of". +Configuration option rule_insts_schematic (default false) controls use +of schematic variables outside the context. Minor INCOMPATIBILITY, +declare rule_insts_schematic = true temporarily and update to use local +variable declarations or dummy patterns instead. + +* Explicit instantiation via attributes "where", "of", and proof methods +"rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns +("_") that stand for anonymous local variables. + +* Generated schematic variables in standard format of exported facts are +incremented to avoid material in the proof context. Rare +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: + <# ~> #<# + <=# ~> #<=# + \# ~> #\# + \# ~> #\# + INCOMPATIBILITY. + - Introduced abbreviations for ill-named multiset operations: + <#, \# abbreviate < (strict subset) + <=#, \#, \# 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 \ 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 "\" 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 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 + \ * 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 \ 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" + "\\ :: (nat => real) => 'a::euclidean_space" + +on euclidean spaces by using the inner product "_ \ _" with +vectors from the Basis set: "\\ 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 + \, \' 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 \ 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 \ (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 +\/\ 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 \ 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 \'. +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.\' -> 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 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 \ and \ +(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 +'definition' instead. + +* The "prems" fact, which refers to the accidental collection of +foundational premises in the context, is now explicitly marked as +legacy feature and will be discontinued soon. Consider using "assms" +of the head statement or reference facts by explicit names. + +* Document antiquotations @{class} and @{type} print classes and type +constructors. + +* Document antiquotation @{file} checks file/directory entries within +the local file system. + + +*** HOL *** + +* Coercive subtyping: functions can be declared as coercions and type +inference will add them as necessary upon input of a term. Theory +Complex_Main declares real :: nat => real and real :: int => real as +coercions. A coercion function f is declared like this: + + declare [[coercion f]] + +To lift coercions through type constructors (e.g. from nat => real to +nat list => real list), map functions can be declared, e.g. + + declare [[coercion_map map]] + +Currently coercion inference is activated only in theories including +real numbers, i.e. descendants of Complex_Main. This is controlled by +the configuration option "coercion_enabled", e.g. it can be enabled in +other theories like this: + + declare [[coercion_enabled]] + +* Command 'partial_function' provides basic support for recursive +function definitions over complete partial orders. Concrete instances +are provided for i) the option type, ii) tail recursion on arbitrary +types, and iii) the heap monad of Imperative_HOL. See +src/HOL/ex/Fundefs.thy and src/HOL/Imperative_HOL/ex/Linked_Lists.thy +for examples. + +* Function package: f.psimps rules are no longer implicitly declared +as [simp]. INCOMPATIBILITY. + +* Datatype package: theorems generated for executable equality (class +"eq") carry proper names and are treated as default code equations. + +* Inductive package: now offers command 'inductive_simps' to +automatically derive instantiated and simplified equations for +inductive predicates, similar to 'inductive_cases'. + +* Command 'enriched_type' allows to register properties of the +functorial structure of types. + +* Improved infrastructure for term evaluation using code generator +techniques, in particular static evaluation conversions. + +* Code generator: Scala (2.8 or higher) has been added to the target +languages. + +* Code generator: globbing constant expressions "*" and "Theory.*" +have been replaced by the more idiomatic "_" and "Theory._". +INCOMPATIBILITY. + +* Code generator: export_code without explicit file declaration prints +to standard output. INCOMPATIBILITY. + +* Code generator: do not print function definitions for case +combinators any longer. + +* Code generator: simplification with rules determined with +src/Tools/Code/code_simp.ML and method "code_simp". + +* Code generator for records: more idiomatic representation of record +types. Warning: records are not covered by ancient SML code +generation any longer. INCOMPATIBILITY. In cases of need, a suitable +rep_datatype declaration helps to succeed then: + + record 'a foo = ... + ... + rep_datatype foo_ext ... + +* Records: logical foundation type for records does not carry a +'_type' suffix any longer (obsolete due to authentic syntax). +INCOMPATIBILITY. + +* Quickcheck now by default uses exhaustive testing instead of random +testing. Random testing can be invoked by "quickcheck [random]", +exhaustive testing by "quickcheck [exhaustive]". + +* Quickcheck instantiates polymorphic types with small finite +datatypes by default. This enables a simple execution mechanism to +handle quantifiers and function equality over the finite datatypes. + +* Quickcheck random generator has been renamed from "code" to +"random". INCOMPATIBILITY. + +* Quickcheck now has a configurable time limit which is set to 30 +seconds by default. This can be changed by adding [timeout = n] to the +quickcheck command. The time limit for Auto Quickcheck is still set +independently. + +* Quickcheck in locales considers interpretations of that locale for +counter example search. + +* Sledgehammer: + - Added "smt" and "remote_smt" provers based on the "smt" proof + method. See the Sledgehammer manual for details ("isabelle doc + sledgehammer"). + - Renamed commands: + sledgehammer atp_info ~> sledgehammer running_provers + sledgehammer atp_kill ~> sledgehammer kill_provers + sledgehammer available_atps ~> sledgehammer available_provers + INCOMPATIBILITY. + - Renamed options: + sledgehammer [atps = ...] ~> sledgehammer [provers = ...] + sledgehammer [atp = ...] ~> sledgehammer [prover = ...] + sledgehammer [timeout = 77 s] ~> sledgehammer [timeout = 77] + (and "ms" and "min" are no longer supported) + INCOMPATIBILITY. + +* Nitpick: + - Renamed options: + nitpick [timeout = 77 s] ~> nitpick [timeout = 77] + nitpick [tac_timeout = 777 ms] ~> nitpick [tac_timeout = 0.777] + INCOMPATIBILITY. + - Added support for partial quotient types. + - Added local versions of the "Nitpick.register_xxx" functions. + - Added "whack" option. + - Allow registration of quotient types as codatatypes. + - Improved "merge_type_vars" option to merge more types. + - Removed unsound "fast_descrs" option. + - Added custom symmetry breaking for datatypes, making it possible to reach + higher cardinalities. + - Prevent the expansion of too large definitions. + +* Proof methods "metis" and "meson" now have configuration options +"meson_trace", "metis_trace", and "metis_verbose" that can be enabled +to diagnose these tools. E.g. + + using [[metis_trace = true]] + +* Auto Solve: Renamed "Auto Solve Direct". The tool is now available +manually as command 'solve_direct'. + +* The default SMT solver Z3 must be enabled explicitly (due to +licensing issues) by setting the environment variable +Z3_NON_COMMERCIAL in etc/settings of the component, for example. For +commercial applications, the SMT solver CVC3 is provided as fall-back; +changing the SMT solver is done via the configuration option +"smt_solver". + +* Remote SMT solvers need to be referred to by the "remote_" prefix, +i.e. "remote_cvc3" and "remote_z3". + +* Added basic SMT support for datatypes, records, and typedefs using +the oracle mode (no proofs). Direct support of pairs has been dropped +in exchange (pass theorems fst_conv snd_conv pair_collapse to the SMT +support for a similar behavior). Minor INCOMPATIBILITY. + +* Changed SMT configuration options: + - Renamed: + z3_proofs ~> smt_oracle (with inverted meaning) + z3_trace_assms ~> smt_trace_used_facts + INCOMPATIBILITY. + - Added: + smt_verbose + smt_random_seed + smt_datatypes + smt_infer_triggers + smt_monomorph_limit + cvc3_options + remote_cvc3_options + remote_z3_options + yices_options + +* Boogie output files (.b2i files) need to be declared in the theory +header. + +* Simplification procedure "list_to_set_comprehension" rewrites list +comprehensions applied to List.set to set comprehensions. Occasional +INCOMPATIBILITY, may be deactivated like this: + + declare [[simproc del: list_to_set_comprehension]] + +* Removed old version of primrec package. INCOMPATIBILITY. + +* Removed simplifier congruence rule of "prod_case", as has for long +been the case with "split". INCOMPATIBILITY. + +* String.literal is a type, but not a datatype. INCOMPATIBILITY. + +* Removed [split_format ... and ... and ...] version of +[split_format]. Potential INCOMPATIBILITY. + +* Predicate "sorted" now defined inductively, with nice induction +rules. INCOMPATIBILITY: former sorted.simps now named sorted_simps. + +* Constant "contents" renamed to "the_elem", to free the generic name +contents for other uses. INCOMPATIBILITY. + +* Renamed class eq and constant eq (for code generation) to class +equal and constant equal, plus renaming of related facts and various +tuning. INCOMPATIBILITY. + +* Dropped type classes mult_mono and mult_mono1. INCOMPATIBILITY. + +* Removed output syntax "'a ~=> 'b" for "'a => 'b option". +INCOMPATIBILITY. + +* Renamed theory Fset to Cset, type Fset.fset to Cset.set, in order to +avoid confusion with finite sets. INCOMPATIBILITY. + +* Abandoned locales equiv, congruent and congruent2 for equivalence +relations. INCOMPATIBILITY: use equivI rather than equiv_intro (same +for congruent(2)). + +* Some previously unqualified names have been qualified: + + types + bool ~> HOL.bool + nat ~> Nat.nat + + constants + Trueprop ~> HOL.Trueprop + True ~> HOL.True + False ~> HOL.False + op & ~> HOL.conj + op | ~> HOL.disj + op --> ~> HOL.implies + op = ~> HOL.eq + Not ~> HOL.Not + The ~> HOL.The + All ~> HOL.All + Ex ~> HOL.Ex + Ex1 ~> HOL.Ex1 + Let ~> HOL.Let + If ~> HOL.If + Ball ~> Set.Ball + Bex ~> Set.Bex + Suc ~> Nat.Suc + Pair ~> Product_Type.Pair + fst ~> Product_Type.fst + snd ~> Product_Type.snd + curry ~> Product_Type.curry + op : ~> Set.member + Collect ~> Set.Collect + +INCOMPATIBILITY. + +* More canonical naming convention for some fundamental definitions: + + bot_bool_eq ~> bot_bool_def + top_bool_eq ~> top_bool_def + inf_bool_eq ~> inf_bool_def + sup_bool_eq ~> sup_bool_def + bot_fun_eq ~> bot_fun_def + top_fun_eq ~> top_fun_def + inf_fun_eq ~> inf_fun_def + sup_fun_eq ~> sup_fun_def + +INCOMPATIBILITY. + +* More stylized fact names: + + expand_fun_eq ~> fun_eq_iff + expand_set_eq ~> set_eq_iff + set_ext ~> set_eqI + nat_number ~> eval_nat_numeral + +INCOMPATIBILITY. + +* Refactoring of code-generation specific operations in theory List: + + constants + null ~> List.null + + facts + mem_iff ~> member_def + null_empty ~> null_def + +INCOMPATIBILITY. Note that these were not supposed to be used +regularly unless for striking reasons; their main purpose was code +generation. + +Various operations from the Haskell prelude are used for generating +Haskell code. + +* Term "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". Term +"surj f" is now an abbreviation of "range f = UNIV". The theorems +bij_def and surj_def are unchanged. INCOMPATIBILITY. + +* Abolished some non-alphabetic type names: "prod" and "sum" replace +"*" and "+" respectively. INCOMPATIBILITY. + +* Name "Plus" of disjoint sum operator "<+>" is now hidden. Write +"Sum_Type.Plus" instead. + +* Constant "split" has been merged with constant "prod_case"; names of +ML functions, facts etc. involving split have been retained so far, +though. INCOMPATIBILITY. + +* Dropped old infix syntax "_ mem _" for List.member; use "_ : set _" +instead. INCOMPATIBILITY. + +* Removed lemma "Option.is_none_none" which duplicates "is_none_def". +INCOMPATIBILITY. + +* Former theory Library/Enum is now part of the HOL-Main image. +INCOMPATIBILITY: all constants of the Enum theory now have to be +referred to by its qualified name. + + enum ~> Enum.enum + nlists ~> Enum.nlists + product ~> Enum.product + +* Theory Library/Monad_Syntax provides do-syntax for monad types. +Syntax in Library/State_Monad has been changed to avoid ambiguities. +INCOMPATIBILITY. + +* Theory Library/SetsAndFunctions has been split into +Library/Function_Algebras and Library/Set_Algebras; canonical names +for instance definitions for functions; various improvements. +INCOMPATIBILITY. + +* Theory Library/Multiset provides stable quicksort implementation of +sort_key. + +* Theory Library/Multiset: renamed empty_idemp ~> empty_neutral. +INCOMPATIBILITY. + +* Session Multivariate_Analysis: introduced a type class for euclidean +space. Most theorems are now stated in terms of euclidean spaces +instead of finite cartesian products. + + types + real ^ 'n ~> 'a::real_vector + ~> 'a::euclidean_space + ~> 'a::ordered_euclidean_space + (depends on your needs) + + constants + _ $ _ ~> _ $$ _ + \ x. _ ~> \\ x. _ + CARD('n) ~> DIM('a) + +Also note that the indices are now natural numbers and not from some +finite type. Finite cartesian products of euclidean spaces, products +of euclidean spaces the real and complex numbers are instantiated to +be euclidean_spaces. INCOMPATIBILITY. + +* Session Probability: introduced pextreal as positive extended real +numbers. Use pextreal as value for measures. Introduce the +Radon-Nikodym derivative, product spaces and Fubini's theorem for +arbitrary sigma finite measures. Introduces Lebesgue measure based on +the integral in Multivariate Analysis. INCOMPATIBILITY. + +* Session Imperative_HOL: revamped, corrected dozens of inadequacies. +INCOMPATIBILITY. + +* Session SPARK (with image HOL-SPARK) provides commands to load and +prove verification conditions generated by the SPARK Ada program +verifier. See also src/HOL/SPARK and src/HOL/SPARK/Examples. + + +*** HOL-Algebra *** + +* Theorems for additive ring operations (locale abelian_monoid and +descendants) are generated by interpretation from their multiplicative +counterparts. Names (in particular theorem names) have the mandatory +qualifier 'add'. Previous theorem names are redeclared for +compatibility. + +* Structure "int_ring" is now an abbreviation (previously a +definition). This fits more natural with advanced interpretations. + + +*** HOLCF *** + +* The domain package now runs in definitional mode by default: The +former command 'new_domain' is now called 'domain'. To use the domain +package in its original axiomatic mode, use 'domain (unsafe)'. +INCOMPATIBILITY. + +* The new class "domain" is now the default sort. Class "predomain" +is an unpointed version of "domain". Theories can be updated by +replacing sort annotations as shown below. INCOMPATIBILITY. + + 'a::type ~> 'a::countable + 'a::cpo ~> 'a::predomain + 'a::pcpo ~> 'a::domain + +* The old type class "rep" has been superseded by class "domain". +Accordingly, users of the definitional package must remove any +"default_sort rep" declarations. INCOMPATIBILITY. + +* The domain package (definitional mode) now supports unpointed +predomain argument types, as long as they are marked 'lazy'. (Strict +arguments must be in class "domain".) For example, the following +domain definition now works: + + domain natlist = nil | cons (lazy "nat discr") (lazy "natlist") + +* Theory HOLCF/Library/HOL_Cpo provides cpo and predomain class +instances for types from main HOL: bool, nat, int, char, 'a + 'b, +'a option, and 'a list. Additionally, it configures fixrec and the +domain package to work with these types. For example: + + fixrec isInl :: "('a + 'b) u -> tr" + where "isInl$(up$(Inl x)) = TT" | "isInl$(up$(Inr y)) = FF" + + domain V = VFun (lazy "V -> V") | VCon (lazy "nat") (lazy "V list") + +* The "(permissive)" option of fixrec has been replaced with a +per-equation "(unchecked)" option. See +src/HOL/HOLCF/Tutorial/Fixrec_ex.thy for examples. INCOMPATIBILITY. + +* The "bifinite" class no longer fixes a constant "approx"; the class +now just asserts that such a function exists. INCOMPATIBILITY. + +* Former type "alg_defl" has been renamed to "defl". HOLCF no longer +defines an embedding of type 'a defl into udom by default; instances +of "bifinite" and "domain" classes are available in +src/HOL/HOLCF/Library/Defl_Bifinite.thy. + +* The syntax "REP('a)" has been replaced with "DEFL('a)". + +* The predicate "directed" has been removed. INCOMPATIBILITY. + +* The type class "finite_po" has been removed. INCOMPATIBILITY. + +* The function "cprod_map" has been renamed to "prod_map". +INCOMPATIBILITY. + +* The monadic bind operator on each powerdomain has new binder syntax +similar to sets, e.g. "\\x\xs. t" represents +"upper_bind\xs\(\ x. t)". + +* The infix syntax for binary union on each powerdomain has changed +from e.g. "+\" to "\\", for consistency with set +syntax. INCOMPATIBILITY. + +* The constant "UU" has been renamed to "bottom". The syntax "UU" is +still supported as an input translation. + +* Renamed some theorems (the original names are also still available). + + 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 + cont2cont_Rep_CFun ~> cont2cont_APP + +* The Abs and Rep functions for various types have changed names. +Related theorem names have also changed to match. INCOMPATIBILITY. + + Rep_CFun ~> Rep_cfun + Abs_CFun ~> Abs_cfun + Rep_Sprod ~> Rep_sprod + Abs_Sprod ~> Abs_sprod + Rep_Ssum ~> Rep_ssum + Abs_Ssum ~> Abs_ssum + +* Lemmas with names of the form *_defined_iff or *_strict_iff have +been renamed to *_bottom_iff. INCOMPATIBILITY. + +* Various changes to bisimulation/coinduction with domain package: + + - Definitions of "bisim" constants no longer mention definedness. + - With mutual recursion, "bisim" predicate is now curried. + - With mutual recursion, each type gets a separate coind theorem. + - Variable names in bisim_def and coinduct rules have changed. + +INCOMPATIBILITY. + +* Case combinators generated by the domain package for type "foo" are +now named "foo_case" instead of "foo_when". INCOMPATIBILITY. + +* Several theorems have been renamed to more accurately reflect the +names of constants and types involved. INCOMPATIBILITY. + + thelub_const ~> lub_const + lub_const ~> is_lub_const + thelubI ~> lub_eqI + is_lub_lub ~> is_lubD2 + lubI ~> is_lub_lub + unique_lub ~> is_lub_unique + is_ub_lub ~> is_lub_rangeD1 + lub_bin_chain ~> is_lub_bin_chain + lub_fun ~> is_lub_fun + thelub_fun ~> lub_fun + thelub_cfun ~> lub_cfun + thelub_Pair ~> lub_Pair + lub_cprod ~> is_lub_prod + thelub_cprod ~> lub_prod + minimal_cprod ~> minimal_prod + inst_cprod_pcpo ~> inst_prod_pcpo + UU_I ~> bottomI + compact_UU ~> compact_bottom + deflation_UU ~> deflation_bottom + finite_deflation_UU ~> finite_deflation_bottom + +* Many legacy theorem names have been discontinued. INCOMPATIBILITY. + + sq_ord_less_eq_trans ~> below_eq_trans + sq_ord_eq_less_trans ~> eq_below_trans + refl_less ~> below_refl + trans_less ~> below_trans + antisym_less ~> below_antisym + antisym_less_inverse ~> po_eq_conv [THEN iffD1] + box_less ~> box_below + rev_trans_less ~> rev_below_trans + not_less2not_eq ~> not_below2not_eq + less_UU_iff ~> below_UU_iff + flat_less_iff ~> flat_below_iff + adm_less ~> adm_below + adm_not_less ~> adm_not_below + adm_compact_not_less ~> adm_compact_not_below + less_fun_def ~> below_fun_def + expand_fun_less ~> fun_below_iff + less_fun_ext ~> fun_belowI + less_discr_def ~> below_discr_def + discr_less_eq ~> discr_below_eq + less_unit_def ~> below_unit_def + less_cprod_def ~> below_prod_def + prod_lessI ~> prod_belowI + Pair_less_iff ~> Pair_below_iff + fst_less_iff ~> fst_below_iff + snd_less_iff ~> snd_below_iff + expand_cfun_less ~> cfun_below_iff + less_cfun_ext ~> cfun_belowI + injection_less ~> injection_below + less_up_def ~> below_up_def + not_Iup_less ~> not_Iup_below + Iup_less ~> Iup_below + up_less ~> up_below + Def_inject_less_eq ~> Def_below_Def + Def_less_is_eq ~> Def_below_iff + spair_less_iff ~> spair_below_iff + less_sprod ~> below_sprod + spair_less ~> spair_below + sfst_less_iff ~> sfst_below_iff + ssnd_less_iff ~> ssnd_below_iff + fix_least_less ~> fix_least_below + dist_less_one ~> dist_below_one + less_ONE ~> below_ONE + ONE_less_iff ~> ONE_below_iff + less_sinlD ~> below_sinlD + less_sinrD ~> below_sinrD + + +*** FOL and ZF *** + +* All constant names are now qualified internally and use proper +identifiers, e.g. "IFOL.eq" instead of "op =". INCOMPATIBILITY. + + +*** ML *** + +* Antiquotation @{assert} inlines a function bool -> unit that raises +Fail if the argument is false. Due to inlining the source position of +failed assertions is included in the error output. + +* Discontinued antiquotation @{theory_ref}, which is obsolete since ML +text is in practice always evaluated with a stable theory checkpoint. +Minor INCOMPATIBILITY, use (Theory.check_thy @{theory}) instead. + +* Antiquotation @{theory A} refers to theory A from the ancestry of +the current context, not any accidental theory loader state as before. +Potential INCOMPATIBILITY, subtle change in semantics. + +* Syntax.pretty_priority (default 0) configures the required priority +of pretty-printed output and thus affects insertion of parentheses. + +* Syntax.default_root (default "any") configures the inner syntax +category (nonterminal symbol) for parsing of terms. + +* Former exception Library.UnequalLengths now coincides with +ListPair.UnequalLengths. + +* Renamed structure MetaSimplifier to Raw_Simplifier. Note that the +main functionality is provided by structure Simplifier. + +* Renamed raw "explode" function to "raw_explode" to emphasize its +meaning. Note that internally to Isabelle, Symbol.explode is used in +almost all situations. + +* Discontinued obsolete function sys_error and exception SYS_ERROR. +See implementation manual for further details on exceptions in +Isabelle/ML. + +* Renamed setmp_noncritical to Unsynchronized.setmp to emphasize its +meaning. + +* Renamed structure PureThy to Pure_Thy and moved most of its +operations to structure Global_Theory, to emphasize that this is +rarely-used global-only stuff. + +* Discontinued Output.debug. Minor INCOMPATIBILITY, use plain writeln +instead (or tracing for high-volume output). + +* Configuration option show_question_marks only affects regular pretty +printing of types and terms, not raw Term.string_of_vname. + +* ML_Context.thm and ML_Context.thms are no longer pervasive. Rare +INCOMPATIBILITY, superseded by static antiquotations @{thm} and +@{thms} for most purposes. + +* ML structure Unsynchronized is never opened, not even in Isar +interaction mode as before. Old Unsynchronized.set etc. have been +discontinued -- use plain := instead. This should be *rare* anyway, +since modern tools always work via official context data, notably +configuration options. + +* Parallel and asynchronous execution requires special care concerning +interrupts. Structure Exn provides some convenience functions that +avoid working directly with raw Interrupt. User code must not absorb +interrupts -- intermediate handling (for cleanup etc.) needs to be +followed by re-raising of the original exception. Another common +source of mistakes are "handle _" patterns, which make the meaning of +the program subject to physical effects of the environment. + + + +New in Isabelle2009-2 (June 2010) +--------------------------------- + +*** General *** + +* Authentic syntax for *all* logical entities (type classes, type +constructors, term constants): provides simple and robust +correspondence between formal entities and concrete syntax. Within +the parse tree / AST representations, "constants" are decorated by +their category (class, type, const) and spelled out explicitly with +their full internal name. + +Substantial INCOMPATIBILITY concerning low-level syntax declarations +and translations (translation rules and translation functions in ML). +Some hints on upgrading: + + - Many existing uses of 'syntax' and 'translations' can be replaced + by more modern 'type_notation', 'notation' and 'abbreviation', + which are independent of this issue. + + - 'translations' require markup within the AST; the term syntax + provides the following special forms: + + CONST c -- produces syntax version of constant c from context + XCONST c -- literally c, checked as constant from context + c -- literally c, if declared by 'syntax' + + Plain identifiers are treated as AST variables -- occasionally the + system indicates accidental variables via the error "rhs contains + extra variables". + + Type classes and type constructors are marked according to their + concrete syntax. Some old translations rules need to be written + for the "type" category, using type constructor application + instead of pseudo-term application of the default category + "logic". + + - 'parse_translation' etc. in ML may use the following + antiquotations: + + @{class_syntax c} -- type class c within parse tree / AST + @{term_syntax c} -- type constructor c within parse tree / AST + @{const_syntax c} -- ML version of "CONST c" above + @{syntax_const c} -- literally c (checked wrt. 'syntax' declarations) + + - Literal types within 'typed_print_translations', i.e. those *not* + represented as pseudo-terms are represented verbatim. Use @{class + c} or @{type_name c} here instead of the above syntax + antiquotations. + +Note that old non-authentic syntax was based on unqualified base +names, so all of the above "constant" names would coincide. Recall +that 'print_syntax' and ML_command "set Syntax.trace_ast" help to +diagnose syntax problems. + +* Type constructors admit general mixfix syntax, not just infix. + +* Concrete syntax may be attached to local entities without a proof +body, too. This works via regular mixfix annotations for 'fix', +'def', 'obtain' etc. or via the explicit 'write' command, which is +similar to the 'notation' command in theory specifications. + +* Discontinued unnamed infix syntax (legacy feature for many years) -- +need to specify constant name and syntax separately. Internal ML +datatype constructors have been renamed from InfixName to Infix etc. +Minor INCOMPATIBILITY. + +* Schematic theorem statements need to be explicitly markup as such, +via commands 'schematic_lemma', 'schematic_theorem', +'schematic_corollary'. Thus the relevance of the proof is made +syntactically clear, which impacts performance in a parallel or +asynchronous interactive environment. Minor INCOMPATIBILITY. + +* Use of cumulative prems via "!" in some proof methods has been +discontinued (old legacy feature). + +* References 'trace_simp' and 'debug_simp' have been replaced by +configuration options stored in the context. Enabling tracing (the +case of debugging is similar) in proofs works via + + using [[trace_simp = true]] + +Tracing is then active for all invocations of the simplifier in +subsequent goal refinement steps. Tracing may also still be enabled or +disabled via the ProofGeneral settings menu. + +* Separate commands 'hide_class', 'hide_type', 'hide_const', +'hide_fact' replace the former 'hide' KIND command. Minor +INCOMPATIBILITY. + +* Improved parallelism of proof term normalization: usedir -p2 -q0 is +more efficient than combinations with -q1 or -q2. + + +*** Pure *** + +* Proofterms record type-class reasoning explicitly, using the +"unconstrain" operation internally. This eliminates all sort +constraints from a theorem and proof, introducing explicit +OFCLASS-premises. On the proof term level, this operation is +automatically applied at theorem boundaries, such that closed proofs +are always free of sort constraints. INCOMPATIBILITY for tools that +inspect proof terms. + +* Local theory specifications may depend on extra type variables that +are not present in the result type -- arguments TYPE('a) :: 'a itself +are added internally. For example: + + definition unitary :: bool where "unitary = (ALL (x::'a) y. x = y)" + +* Predicates of locales introduced by classes carry a mandatory +"class" prefix. INCOMPATIBILITY. + +* Vacuous class specifications observe default sort. INCOMPATIBILITY. + +* Old 'axclass' command has been discontinued. INCOMPATIBILITY, use +'class' instead. + +* Command 'code_reflect' allows to incorporate generated ML code into +runtime environment; replaces immature code_datatype antiquotation. +INCOMPATIBILITY. + +* Code generator: simple concept for abstract datatypes obeying +invariants. + +* Code generator: details of internal data cache have no impact on the +user space functionality any longer. + +* Methods "unfold_locales" and "intro_locales" ignore non-locale +subgoals. This is more appropriate for interpretations with 'where'. +INCOMPATIBILITY. + +* Command 'example_proof' opens an empty proof body. This allows to +experiment with Isar, without producing any persistent result. + +* Commands 'type_notation' and 'no_type_notation' declare type syntax +within a local theory context, with explicit checking of the +constructors involved (in contrast to the raw 'syntax' versions). + +* Commands 'types' and 'typedecl' now work within a local theory +context -- without introducing dependencies on parameters or +assumptions, which is not possible in Isabelle/Pure. + +* Command 'defaultsort' has been renamed to 'default_sort', it works +within a local theory context. Minor INCOMPATIBILITY. + + +*** HOL *** + +* Command 'typedef' now works within a local theory context -- without +introducing dependencies on parameters or assumptions, which is not +possible in Isabelle/Pure/HOL. Note that the logical environment may +contain multiple interpretations of local typedefs (with different +non-emptiness proofs), even in a global theory context. + +* New package for quotient types. Commands 'quotient_type' and +'quotient_definition' may be used for defining types and constants by +quotient constructions. An example is the type of integers created by +quotienting pairs of natural numbers: + + fun + intrel :: "(nat * nat) => (nat * nat) => bool" + where + "intrel (x, y) (u, v) = (x + v = u + y)" + + quotient_type int = "nat * nat" / intrel + by (auto simp add: equivp_def expand_fun_eq) + + quotient_definition + "0::int" is "(0::nat, 0::nat)" + +The method "lifting" can be used to lift of theorems from the +underlying "raw" type to the quotient type. The example +src/HOL/Quotient_Examples/FSet.thy includes such a quotient +construction and provides a reasoning infrastructure for finite sets. + +* Renamed Library/Quotient.thy to Library/Quotient_Type.thy to avoid +clash with new theory Quotient in Main HOL. + +* Moved the SMT binding into the main HOL session, eliminating +separate HOL-SMT session. + +* List membership infix mem operation is only an input abbreviation. +INCOMPATIBILITY. + +* Theory Library/Word.thy has been removed. Use library Word/Word.thy +for future developements; former Library/Word.thy is still present in +the AFP entry RSAPPS. + +* Theorem Int.int_induct renamed to Int.int_of_nat_induct and is no +longer shadowed. INCOMPATIBILITY. + +* Dropped theorem duplicate comp_arith; use semiring_norm instead. +INCOMPATIBILITY. + +* Dropped theorem RealPow.real_sq_order; use power2_le_imp_le instead. +INCOMPATIBILITY. + +* Dropped normalizing_semiring etc; use the facts in semiring classes +instead. INCOMPATIBILITY. + +* Dropped several real-specific versions of lemmas about floor and +ceiling; use the generic lemmas from theory "Archimedean_Field" +instead. INCOMPATIBILITY. + + floor_number_of_eq ~> floor_number_of + le_floor_eq_number_of ~> number_of_le_floor + le_floor_eq_zero ~> zero_le_floor + le_floor_eq_one ~> one_le_floor + floor_less_eq_number_of ~> floor_less_number_of + floor_less_eq_zero ~> floor_less_zero + floor_less_eq_one ~> floor_less_one + less_floor_eq_number_of ~> number_of_less_floor + less_floor_eq_zero ~> zero_less_floor + less_floor_eq_one ~> one_less_floor + floor_le_eq_number_of ~> floor_le_number_of + floor_le_eq_zero ~> floor_le_zero + floor_le_eq_one ~> floor_le_one + floor_subtract_number_of ~> floor_diff_number_of + floor_subtract_one ~> floor_diff_one + ceiling_number_of_eq ~> ceiling_number_of + ceiling_le_eq_number_of ~> ceiling_le_number_of + ceiling_le_zero_eq ~> ceiling_le_zero + ceiling_le_eq_one ~> ceiling_le_one + less_ceiling_eq_number_of ~> number_of_less_ceiling + less_ceiling_eq_zero ~> zero_less_ceiling + less_ceiling_eq_one ~> one_less_ceiling + ceiling_less_eq_number_of ~> ceiling_less_number_of + ceiling_less_eq_zero ~> ceiling_less_zero + ceiling_less_eq_one ~> ceiling_less_one + le_ceiling_eq_number_of ~> number_of_le_ceiling + le_ceiling_eq_zero ~> zero_le_ceiling + le_ceiling_eq_one ~> one_le_ceiling + ceiling_subtract_number_of ~> ceiling_diff_number_of + ceiling_subtract_one ~> ceiling_diff_one + +* Theory "Finite_Set": various folding_XXX locales facilitate the +application of the various fold combinators on finite sets. + +* Library theory "RBT" renamed to "RBT_Impl"; new library theory "RBT" +provides abstract red-black tree type which is backed by "RBT_Impl" as +implementation. INCOMPATIBILITY. + +* Theory Library/Coinductive_List has been removed -- superseded by +AFP/thys/Coinductive. + +* Theory PReal, including the type "preal" and related operations, has +been removed. INCOMPATIBILITY. + +* Real: new development using Cauchy Sequences. + +* Split off theory "Big_Operators" containing setsum, setprod, +Inf_fin, Sup_fin, Min, Max from theory Finite_Set. INCOMPATIBILITY. + +* Theory "Rational" renamed to "Rat", for consistency with "Nat", +"Int" etc. INCOMPATIBILITY. + +* Constant Rat.normalize needs to be qualified. INCOMPATIBILITY. + +* New set of rules "ac_simps" provides combined assoc / commute +rewrites for all interpretations of the appropriate generic locales. + +* Renamed theory "OrderedGroup" to "Groups" and split theory +"Ring_and_Field" into theories "Rings" and "Fields"; for more +appropriate and more consistent names suitable for name prefixes +within the HOL theories. INCOMPATIBILITY. + +* Some generic constants have been put to appropriate theories: + - less_eq, less: Orderings + - zero, one, plus, minus, uminus, times, abs, sgn: Groups + - inverse, divide: Rings +INCOMPATIBILITY. + +* More consistent naming of type classes involving orderings (and +lattices): + + lower_semilattice ~> semilattice_inf + upper_semilattice ~> semilattice_sup + + dense_linear_order ~> dense_linorder + + pordered_ab_group_add ~> ordered_ab_group_add + pordered_ab_group_add_abs ~> ordered_ab_group_add_abs + pordered_ab_semigroup_add ~> ordered_ab_semigroup_add + pordered_ab_semigroup_add_imp_le ~> ordered_ab_semigroup_add_imp_le + pordered_cancel_ab_semigroup_add ~> ordered_cancel_ab_semigroup_add + pordered_cancel_comm_semiring ~> ordered_cancel_comm_semiring + pordered_cancel_semiring ~> ordered_cancel_semiring + pordered_comm_monoid_add ~> ordered_comm_monoid_add + pordered_comm_ring ~> ordered_comm_ring + pordered_comm_semiring ~> ordered_comm_semiring + pordered_ring ~> ordered_ring + pordered_ring_abs ~> ordered_ring_abs + pordered_semiring ~> ordered_semiring + + ordered_ab_group_add ~> linordered_ab_group_add + ordered_ab_semigroup_add ~> linordered_ab_semigroup_add + ordered_cancel_ab_semigroup_add ~> linordered_cancel_ab_semigroup_add + ordered_comm_semiring_strict ~> linordered_comm_semiring_strict + ordered_field ~> linordered_field + ordered_field_no_lb ~> linordered_field_no_lb + ordered_field_no_ub ~> linordered_field_no_ub + ordered_field_dense_linear_order ~> dense_linordered_field + ordered_idom ~> linordered_idom + ordered_ring ~> linordered_ring + ordered_ring_le_cancel_factor ~> linordered_ring_le_cancel_factor + ordered_ring_less_cancel_factor ~> linordered_ring_less_cancel_factor + ordered_ring_strict ~> linordered_ring_strict + ordered_semidom ~> linordered_semidom + ordered_semiring ~> linordered_semiring + ordered_semiring_1 ~> linordered_semiring_1 + ordered_semiring_1_strict ~> linordered_semiring_1_strict + ordered_semiring_strict ~> linordered_semiring_strict + + The following slightly odd type classes have been moved to a + separate theory Library/Lattice_Algebras: + + lordered_ab_group_add ~> lattice_ab_group_add + lordered_ab_group_add_abs ~> lattice_ab_group_add_abs + lordered_ab_group_add_meet ~> semilattice_inf_ab_group_add + lordered_ab_group_add_join ~> semilattice_sup_ab_group_add + lordered_ring ~> lattice_ring + +INCOMPATIBILITY. + +* Refined field classes: + - classes division_ring_inverse_zero, field_inverse_zero, + linordered_field_inverse_zero include rule inverse 0 = 0 -- + subsumes former division_by_zero class; + - numerous lemmas have been ported from field to division_ring. +INCOMPATIBILITY. + +* Refined algebra theorem collections: + - dropped theorem group group_simps, use algebra_simps instead; + - dropped theorem group ring_simps, use field_simps instead; + - proper theorem collection field_simps subsumes former theorem + groups field_eq_simps and field_simps; + - dropped lemma eq_minus_self_iff which is a duplicate for + equal_neg_zero. +INCOMPATIBILITY. + +* Theory Finite_Set and List: some lemmas have been generalized from +sets to lattices: + + fun_left_comm_idem_inter ~> fun_left_comm_idem_inf + fun_left_comm_idem_union ~> fun_left_comm_idem_sup + inter_Inter_fold_inter ~> inf_Inf_fold_inf + union_Union_fold_union ~> sup_Sup_fold_sup + Inter_fold_inter ~> Inf_fold_inf + Union_fold_union ~> Sup_fold_sup + inter_INTER_fold_inter ~> inf_INFI_fold_inf + union_UNION_fold_union ~> sup_SUPR_fold_sup + INTER_fold_inter ~> INFI_fold_inf + UNION_fold_union ~> SUPR_fold_sup + +* Theory "Complete_Lattice": lemmas top_def and bot_def have been +replaced by the more convenient lemmas Inf_empty and Sup_empty. +Dropped lemmas Inf_insert_simp and Sup_insert_simp, which are subsumed +by Inf_insert and Sup_insert. Lemmas Inf_UNIV and Sup_UNIV replace +former Inf_Univ and Sup_Univ. Lemmas inf_top_right and sup_bot_right +subsume inf_top and sup_bot respectively. INCOMPATIBILITY. + +* Reorganized theory Multiset: swapped notation of pointwise and +multiset order: + + - pointwise ordering is instance of class order with standard syntax + <= and <; + - multiset ordering has syntax <=# and <#; partial order properties + are provided by means of interpretation with prefix + multiset_order; + - less duplication, less historical organization of sections, + conversion from associations lists to multisets, rudimentary code + generation; + - use insert_DiffM2 [symmetric] instead of elem_imp_eq_diff_union, + if needed. + +Renamed: + + multiset_eq_conv_count_eq ~> multiset_ext_iff + multi_count_ext ~> multiset_ext + diff_union_inverse2 ~> diff_union_cancelR + +INCOMPATIBILITY. + +* Theory Permutation: replaced local "remove" by List.remove1. + +* Code generation: ML and OCaml code is decorated with signatures. + +* Theory List: added transpose. + +* Library/Nat_Bijection.thy is a collection of bijective functions +between nat and other types, which supersedes the older libraries +Library/Nat_Int_Bij.thy and HOLCF/NatIso.thy. INCOMPATIBILITY. + + Constants: + Nat_Int_Bij.nat2_to_nat ~> prod_encode + Nat_Int_Bij.nat_to_nat2 ~> prod_decode + Nat_Int_Bij.int_to_nat_bij ~> int_encode + Nat_Int_Bij.nat_to_int_bij ~> int_decode + Countable.pair_encode ~> prod_encode + NatIso.prod2nat ~> prod_encode + NatIso.nat2prod ~> prod_decode + NatIso.sum2nat ~> sum_encode + NatIso.nat2sum ~> sum_decode + NatIso.list2nat ~> list_encode + NatIso.nat2list ~> list_decode + NatIso.set2nat ~> set_encode + NatIso.nat2set ~> set_decode + + Lemmas: + Nat_Int_Bij.bij_nat_to_int_bij ~> bij_int_decode + Nat_Int_Bij.nat2_to_nat_inj ~> inj_prod_encode + Nat_Int_Bij.nat2_to_nat_surj ~> surj_prod_encode + Nat_Int_Bij.nat_to_nat2_inj ~> inj_prod_decode + Nat_Int_Bij.nat_to_nat2_surj ~> surj_prod_decode + Nat_Int_Bij.i2n_n2i_id ~> int_encode_inverse + Nat_Int_Bij.n2i_i2n_id ~> int_decode_inverse + Nat_Int_Bij.surj_nat_to_int_bij ~> surj_int_encode + Nat_Int_Bij.surj_int_to_nat_bij ~> surj_int_decode + Nat_Int_Bij.inj_nat_to_int_bij ~> inj_int_encode + Nat_Int_Bij.inj_int_to_nat_bij ~> inj_int_decode + Nat_Int_Bij.bij_nat_to_int_bij ~> bij_int_encode + Nat_Int_Bij.bij_int_to_nat_bij ~> bij_int_decode + +* Sledgehammer: + - Renamed ATP commands: + atp_info ~> sledgehammer running_atps + atp_kill ~> sledgehammer kill_atps + atp_messages ~> sledgehammer messages + atp_minimize ~> sledgehammer minimize + print_atps ~> sledgehammer available_atps + INCOMPATIBILITY. + - Added user's manual ("isabelle doc sledgehammer"). + - Added option syntax and "sledgehammer_params" to customize + Sledgehammer's behavior. See the manual for details. + - Modified the Isar proof reconstruction code so that it produces + direct proofs rather than proofs by contradiction. (This feature + is still experimental.) + - Made Isar proof reconstruction work for SPASS, remote ATPs, and in + full-typed mode. + - Added support for TPTP syntax for SPASS via the "spass_tptp" ATP. + +* Nitpick: + - Added and implemented "binary_ints" and "bits" options. + - Added "std" option and implemented support for nonstandard models. + - Added and implemented "finitize" option to improve the precision + of infinite datatypes based on a monotonicity analysis. + - Added support for quotient types. + - Added support for "specification" and "ax_specification" + constructs. + - Added support for local definitions (for "function" and + "termination" proofs). + - Added support for term postprocessors. + - Optimized "Multiset.multiset" and "FinFun.finfun". + - Improved efficiency of "destroy_constrs" optimization. + - Fixed soundness bugs related to "destroy_constrs" optimization and + record getters. + - Fixed soundness bug related to higher-order constructors. + - Fixed soundness bug when "full_descrs" is enabled. + - Improved precision of set constructs. + - Added "atoms" option. + - Added cache to speed up repeated Kodkod invocations on the same + problems. + - Renamed "MiniSatJNI", "zChaffJNI", "BerkMinAlloy", and + "SAT4JLight" to "MiniSat_JNI", "zChaff_JNI", "BerkMin_Alloy", and + "SAT4J_Light". INCOMPATIBILITY. + - Removed "skolemize", "uncurry", "sym_break", "flatten_prop", + "sharing_depth", and "show_skolems" options. INCOMPATIBILITY. + - Removed "nitpick_intro" attribute. INCOMPATIBILITY. + +* Method "induct" now takes instantiations of the form t, where t is not + a variable, as a shorthand for "x == t", where x is a fresh variable. + If this is not intended, t has to be enclosed in parentheses. + By default, the equalities generated by definitional instantiations + are pre-simplified, which may cause parameters of inductive cases + to disappear, or may even delete some of the inductive cases. + Use "induct (no_simp)" instead of "induct" to restore the old + behaviour. The (no_simp) option is also understood by the "cases" + and "nominal_induct" methods, which now perform pre-simplification, too. + INCOMPATIBILITY. + + +*** HOLCF *** + +* Variable names in lemmas generated by the domain package have +changed; the naming scheme is now consistent with the HOL datatype +package. Some proof scripts may be affected, INCOMPATIBILITY. + +* The domain package no longer defines the function "foo_copy" for +recursive domain "foo". The reach lemma is now stated directly in +terms of "foo_take". Lemmas and proofs that mention "foo_copy" must +be reformulated in terms of "foo_take", INCOMPATIBILITY. + +* Most definedness lemmas generated by the domain package (previously +of the form "x ~= UU ==> foo$x ~= UU") now have an if-and-only-if form +like "foo$x = UU <-> x = UU", which works better as a simp rule. +Proofs that used definedness lemmas as intro rules may break, +potential INCOMPATIBILITY. + +* Induction and casedist rules generated by the domain package now +declare proper case_names (one called "bottom", and one named for each +constructor). INCOMPATIBILITY. + +* For mutually-recursive domains, separate "reach" and "take_lemma" +rules are generated for each domain, INCOMPATIBILITY. + + foo_bar.reach ~> foo.reach bar.reach + foo_bar.take_lemmas ~> foo.take_lemma bar.take_lemma + +* Some lemmas generated by the domain package have been renamed for +consistency with the datatype package, INCOMPATIBILITY. + + foo.ind ~> foo.induct + foo.finite_ind ~> foo.finite_induct + foo.coind ~> foo.coinduct + foo.casedist ~> foo.exhaust + foo.exhaust ~> foo.nchotomy + +* For consistency with other definition packages, the fixrec package +now generates qualified theorem names, INCOMPATIBILITY. + + foo_simps ~> foo.simps + foo_unfold ~> foo.unfold + foo_induct ~> foo.induct + +* The "fixrec_simp" attribute has been removed. The "fixrec_simp" +method and internal fixrec proofs now use the default simpset instead. +INCOMPATIBILITY. + +* The "contlub" predicate has been removed. Proof scripts should use +lemma contI2 in place of monocontlub2cont, INCOMPATIBILITY. + +* The "admw" predicate has been removed, INCOMPATIBILITY. + +* The constants cpair, cfst, and csnd have been removed in favor of +Pair, fst, and snd from Isabelle/HOL, INCOMPATIBILITY. + + +*** ML *** + +* Antiquotations for basic formal entities: + + @{class NAME} -- type class + @{class_syntax NAME} -- syntax representation of the above + + @{type_name NAME} -- logical type + @{type_abbrev NAME} -- type abbreviation + @{nonterminal NAME} -- type of concrete syntactic category + @{type_syntax NAME} -- syntax representation of any of the above + + @{const_name NAME} -- logical constant (INCOMPATIBILITY) + @{const_abbrev NAME} -- abbreviated constant + @{const_syntax NAME} -- syntax representation of any of the above + +* Antiquotation @{syntax_const NAME} ensures that NAME refers to a raw +syntax constant (cf. 'syntax' command). + +* Antiquotation @{make_string} inlines a function to print arbitrary +values similar to the ML toplevel. The result is compiler dependent +and may fall back on "?" in certain situations. + +* Diagnostic commands 'ML_val' and 'ML_command' may refer to +antiquotations @{Isar.state} and @{Isar.goal}. This replaces impure +Isar.state() and Isar.goal(), which belong to the old TTY loop and do +not work with the asynchronous Isar document model. + +* Configuration options now admit dynamic default values, depending on +the context or even global references. + +* SHA1.digest digests strings according to SHA-1 (see RFC 3174). It +uses an efficient external library if available (for Poly/ML). + +* Renamed some important ML structures, while keeping the old names +for some time as aliases within the structure Legacy: + + OuterKeyword ~> Keyword + OuterLex ~> Token + OuterParse ~> Parse + OuterSyntax ~> Outer_Syntax + PrintMode ~> Print_Mode + SpecParse ~> Parse_Spec + ThyInfo ~> Thy_Info + ThyLoad ~> Thy_Load + ThyOutput ~> Thy_Output + TypeInfer ~> Type_Infer + +Note that "open Legacy" simplifies porting of sources, but forgetting +to remove it again will complicate porting again in the future. + +* Most operations that refer to a global context are named +accordingly, e.g. Simplifier.global_context or +ProofContext.init_global. There are some situations where a global +context actually works, but under normal circumstances one needs to +pass the proper local context through the code! + +* Discontinued old TheoryDataFun with its copy/init operation -- data +needs to be pure. Functor Theory_Data_PP retains the traditional +Pretty.pp argument to merge, which is absent in the standard +Theory_Data version. + +* Sorts.certify_sort and derived "cert" operations for types and terms +no longer minimize sorts. Thus certification at the boundary of the +inference kernel becomes invariant under addition of class relations, +which is an important monotonicity principle. Sorts are now minimized +in the syntax layer only, at the boundary between the end-user and the +system. Subtle INCOMPATIBILITY, may have to use Sign.minimize_sort +explicitly in rare situations. + +* Renamed old-style Drule.standard to Drule.export_without_context, to +emphasize that this is in no way a standard operation. +INCOMPATIBILITY. + +* Subgoal.FOCUS (and variants): resulting goal state is normalized as +usual for resolution. Rare INCOMPATIBILITY. + +* Renamed varify/unvarify operations to varify_global/unvarify_global +to emphasize that these only work in a global situation (which is +quite rare). + +* Curried take and drop in library.ML; negative length is interpreted +as infinity (as in chop). Subtle INCOMPATIBILITY. + +* Proof terms: type substitutions on proof constants now use canonical +order of type variables. INCOMPATIBILITY for tools working with proof +terms. + +* Raw axioms/defs may no longer carry sort constraints, and raw defs +may no longer carry premises. User-level specifications are +transformed accordingly by Thm.add_axiom/add_def. + + +*** System *** + +* Discontinued special HOL_USEDIR_OPTIONS for the main HOL image; +ISABELLE_USEDIR_OPTIONS applies uniformly to all sessions. Note that +proof terms are enabled unconditionally in the new HOL-Proofs image. + +* Discontinued old ISABELLE and ISATOOL environment settings (legacy +feature since Isabelle2009). Use ISABELLE_PROCESS and ISABELLE_TOOL, +respectively. + +* Old lib/scripts/polyml-platform is superseded by the +ISABELLE_PLATFORM setting variable, which defaults to the 32 bit +variant, even on a 64 bit machine. The following example setting +prefers 64 bit if available: + + ML_PLATFORM="${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}" + +* The preliminary Isabelle/jEdit application demonstrates the emerging +Isabelle/Scala layer for advanced prover interaction and integration. +See src/Tools/jEdit or "isabelle jedit" provided by the properly built +component. + +* "IsabelleText" is a Unicode font derived from Bitstream Vera Mono +and Bluesky TeX fonts. It provides the usual Isabelle symbols, +similar to the default assignment of the document preparation system +(cf. isabellesym.sty). The Isabelle/Scala class Isabelle_System +provides some operations for direct access to the font without asking +the user for manual installation. + + + +New in Isabelle2009-1 (December 2009) +------------------------------------- + +*** General *** + +* Discontinued old form of "escaped symbols" such as \\. Only +one backslash should be used, even in ML sources. + + +*** Pure *** + +* Locale interpretation propagates mixins along the locale hierarchy. +The currently only available mixins are the equations used to map +local definitions to terms of the target domain of an interpretation. + +* Reactivated diagnostic command 'print_interps'. Use "print_interps +loc" to print all interpretations of locale "loc" in the theory. +Interpretations in proofs are not shown. + +* Thoroughly revised locales tutorial. New section on conditional +interpretation. + +* On instantiation of classes, remaining undefined class parameters +are formally declared. INCOMPATIBILITY. + + +*** Document preparation *** + +* New generalized style concept for printing terms: @{foo (style) ...} +instead of @{foo_style style ...} (old form is still retained for +backward compatibility). Styles can be also applied for +antiquotations prop, term_type and typeof. + + +*** HOL *** + +* New proof method "smt" for a combination of first-order logic with +equality, linear and nonlinear (natural/integer/real) arithmetic, and +fixed-size bitvectors; there is also basic support for higher-order +features (esp. lambda abstractions). It is an incomplete decision +procedure based on external SMT solvers using the oracle mechanism; +for the SMT solver Z3, this method is proof-producing. Certificates +are provided to avoid calling the external solvers solely for +re-checking proofs. Due to a remote SMT service there is no need for +installing SMT solvers locally. See src/HOL/SMT. + +* New commands to load and prove verification conditions generated by +the Boogie program verifier or derived systems (e.g. the Verifying C +Compiler (VCC) or Spec#). See src/HOL/Boogie. + +* New counterexample generator tool 'nitpick' based on the Kodkod +relational model finder. See src/HOL/Tools/Nitpick and +src/HOL/Nitpick_Examples. + +* New commands 'code_pred' and 'values' to invoke the predicate +compiler and to enumerate values of inductive predicates. + +* A tabled implementation of the reflexive transitive closure. + +* New implementation of quickcheck uses generic code generator; +default generators are provided for all suitable HOL types, records +and datatypes. Old quickcheck can be re-activated importing theory +Library/SML_Quickcheck. + +* New testing tool Mirabelle for automated proof tools. Applies +several tools and tactics like sledgehammer, metis, or quickcheck, to +every proof step in a theory. To be used in batch mode via the +"mirabelle" utility. + +* New proof method "sos" (sum of squares) for nonlinear real +arithmetic (originally due to John Harison). It requires theory +Library/Sum_Of_Squares. It is not a complete decision procedure but +works well in practice on quantifier-free real arithmetic with +, -, +*, ^, =, <= and <, i.e. boolean combinations of equalities and +inequalities between polynomials. It makes use of external +semidefinite programming solvers. Method "sos" generates a +certificate that can be pasted into the proof thus avoiding the need +to call an external tool every time the proof is checked. See +src/HOL/Library/Sum_Of_Squares. + +* New method "linarith" invokes existing linear arithmetic decision +procedure only. + +* New command 'atp_minimal' reduces result produced by Sledgehammer. + +* New Sledgehammer option "Full Types" in Proof General settings menu. +Causes full type information to be output to the ATPs. This slows +ATPs down considerably but eliminates a source of unsound "proofs" +that fail later. + +* New method "metisFT": A version of metis that uses full type +information in order to avoid failures of proof reconstruction. + +* New evaluator "approximate" approximates an real valued term using +the same method as the approximation method. + +* Method "approximate" now supports arithmetic expressions as +boundaries of intervals and implements interval splitting and Taylor +series expansion. + +* ML antiquotation @{code_datatype} inserts definition of a datatype +generated by the code generator; e.g. see src/HOL/Predicate.thy. + +* New theory SupInf of the supremum and infimum operators for sets of +reals. + +* New theory Probability, which contains a development of measure +theory, eventually leading to Lebesgue integration and probability. + +* Extended Multivariate Analysis to include derivation and Brouwer's +fixpoint theorem. + +* Reorganization of number theory, INCOMPATIBILITY: + - new number theory development for nat and int, in theories Divides + and GCD as well as in new session Number_Theory + - some constants and facts now suffixed with _nat and _int + accordingly + - former session NumberTheory now named Old_Number_Theory, including + theories Legacy_GCD and Primes (prefer Number_Theory if possible) + - moved theory Pocklington from src/HOL/Library to + src/HOL/Old_Number_Theory + +* Theory GCD includes functions Gcd/GCD and Lcm/LCM for the gcd and +lcm of finite and infinite sets. It is shown that they form a complete +lattice. + +* Class semiring_div requires superclass no_zero_divisors and proof of +div_mult_mult1; theorems div_mult_mult1, div_mult_mult2, +div_mult_mult1_if, div_mult_mult1 and div_mult_mult2 have been +generalized to class semiring_div, subsuming former theorems +zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and +zdiv_zmult_zmult2. div_mult_mult1 is now [simp] by default. +INCOMPATIBILITY. + +* Refinements to lattice classes and sets: + - less default intro/elim rules in locale variant, more default + intro/elim rules in class variant: more uniformity + - lemma ge_sup_conv renamed to le_sup_iff, in accordance with + le_inf_iff + - dropped lemma alias inf_ACI for inf_aci (same for sup_ACI and + sup_aci) + - renamed ACI to inf_sup_aci + - new class "boolean_algebra" + - class "complete_lattice" moved to separate theory + "Complete_Lattice"; corresponding constants (and abbreviations) + renamed and with authentic syntax: + Set.Inf ~> Complete_Lattice.Inf + Set.Sup ~> Complete_Lattice.Sup + Set.INFI ~> Complete_Lattice.INFI + Set.SUPR ~> Complete_Lattice.SUPR + Set.Inter ~> Complete_Lattice.Inter + Set.Union ~> Complete_Lattice.Union + Set.INTER ~> Complete_Lattice.INTER + Set.UNION ~> Complete_Lattice.UNION + - authentic syntax for + Set.Pow + Set.image + - mere abbreviations: + Set.empty (for bot) + Set.UNIV (for top) + Set.inter (for inf, formerly Set.Int) + Set.union (for sup, formerly Set.Un) + Complete_Lattice.Inter (for Inf) + Complete_Lattice.Union (for Sup) + Complete_Lattice.INTER (for INFI) + Complete_Lattice.UNION (for SUPR) + - object-logic definitions as far as appropriate + +INCOMPATIBILITY. Care is required when theorems Int_subset_iff or +Un_subset_iff are explicitly deleted as default simp rules; then also +their lattice counterparts le_inf_iff and le_sup_iff have to be +deleted to achieve the desired effect. + +* Rules inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp +rules by default any longer; the same applies to min_max.inf_absorb1 +etc. INCOMPATIBILITY. + +* Rules sup_Int_eq and sup_Un_eq are no longer declared as +pred_set_conv by default. INCOMPATIBILITY. + +* Power operations on relations and functions are now one dedicated +constant "compow" with infix syntax "^^". Power operation on +multiplicative monoids retains syntax "^" and is now defined generic +in class power. INCOMPATIBILITY. + +* Relation composition "R O S" now has a more standard argument order: +"R O S = {(x, z). EX y. (x, y) : R & (y, z) : S}". INCOMPATIBILITY, +rewrite propositions with "S O R" --> "R O S". Proofs may occasionally +break, since the O_assoc rule was not rewritten like this. Fix using +O_assoc[symmetric]. The same applies to the curried version "R OO S". + +* Function "Inv" is renamed to "inv_into" and function "inv" is now an +abbreviation for "inv_into UNIV". Lemmas are renamed accordingly. +INCOMPATIBILITY. + +* Most rules produced by inductive and datatype package have mandatory +prefixes. INCOMPATIBILITY. + +* Changed "DERIV_intros" to a dynamic fact, which can be augmented by +the attribute of the same name. Each of the theorems in the list +DERIV_intros assumes composition with an additional function and +matches a variable to the derivative, which has to be solved by the +Simplifier. Hence (auto intro!: DERIV_intros) computes the derivative +of most elementary terms. Former Maclauren.DERIV_tac and +Maclauren.deriv_tac should be replaced by (auto intro!: DERIV_intros). +INCOMPATIBILITY. + +* Code generator attributes follow the usual underscore convention: + code_unfold replaces code unfold + code_post replaces code post + etc. + INCOMPATIBILITY. + +* Renamed methods: + sizechange -> size_change + induct_scheme -> induction_schema + INCOMPATIBILITY. + +* Discontinued abbreviation "arbitrary" of constant "undefined". +INCOMPATIBILITY, use "undefined" directly. + +* Renamed theorems: + Suc_eq_add_numeral_1 -> Suc_eq_plus1 + Suc_eq_add_numeral_1_left -> Suc_eq_plus1_left + Suc_plus1 -> Suc_eq_plus1 + *anti_sym -> *antisym* + vector_less_eq_def -> vector_le_def + INCOMPATIBILITY. + +* Added theorem List.map_map as [simp]. Removed List.map_compose. +INCOMPATIBILITY. + +* Removed predicate "M hassize n" (<--> card M = n & finite M). +INCOMPATIBILITY. + + +*** HOLCF *** + +* Theory Representable defines a class "rep" of domains that are +representable (via an ep-pair) in the universal domain type "udom". +Instances are provided for all type constructors defined in HOLCF. + +* The 'new_domain' command is a purely definitional version of the +domain package, for representable domains. Syntax is identical to the +old domain package. The 'new_domain' package also supports indirect +recursion using previously-defined type constructors. See +src/HOLCF/ex/New_Domain.thy for examples. + +* Method "fixrec_simp" unfolds one step of a fixrec-defined constant +on the left-hand side of an equation, and then performs +simplification. Rewriting is done using rules declared with the +"fixrec_simp" attribute. The "fixrec_simp" method is intended as a +replacement for "fixpat"; see src/HOLCF/ex/Fixrec_ex.thy for examples. + +* The pattern-match compiler in 'fixrec' can now handle constructors +with HOL function types. Pattern-match combinators for the Pair +constructor are pre-configured. + +* The 'fixrec' package now produces better fixed-point induction rules +for mutually-recursive definitions: Induction rules have conclusions +of the form "P foo bar" instead of "P ". + +* The constant "sq_le" (with infix syntax "<<" or "\") has +been renamed to "below". The name "below" now replaces "less" in many +theorem names. (Legacy theorem names using "less" are still supported +as well.) + +* The 'fixrec' package now supports "bottom patterns". Bottom +patterns can be used to generate strictness rules, or to make +functions more strict (much like the bang-patterns supported by the +Glasgow Haskell Compiler). See src/HOLCF/ex/Fixrec_ex.thy for +examples. + + +*** ML *** + +* Support for Poly/ML 5.3.0, with improved reporting of compiler +errors and run-time exceptions, including detailed source positions. + +* Structure Name_Space (formerly NameSpace) now manages uniquely +identified entries, with some additional information such as source +position, logical grouping etc. + +* Theory and context data is now introduced by the simplified and +modernized functors Theory_Data, Proof_Data, Generic_Data. Data needs +to be pure, but the old TheoryDataFun for mutable data (with explicit +copy operation) is still available for some time. + +* Structure Synchronized (cf. src/Pure/Concurrent/synchronized.ML) +provides a high-level programming interface to synchronized state +variables with atomic update. This works via pure function +application within a critical section -- its runtime should be as +short as possible; beware of deadlocks if critical code is nested, +either directly or indirectly via other synchronized variables! + +* Structure Unsynchronized (cf. src/Pure/ML-Systems/unsynchronized.ML) +wraps raw ML references, explicitly indicating their non-thread-safe +behaviour. The Isar toplevel keeps this structure open, to +accommodate Proof General as well as quick and dirty interactive +experiments with references. + +* PARALLEL_CHOICE and PARALLEL_GOALS provide basic support for +parallel tactical reasoning. + +* Tacticals Subgoal.FOCUS, Subgoal.FOCUS_PREMS, Subgoal.FOCUS_PARAMS +are similar to SUBPROOF, but are slightly more flexible: only the +specified parts of the subgoal are imported into the context, and the +body tactic may introduce new subgoals and schematic variables. + +* Old tactical METAHYPS, which does not observe the proof context, has +been renamed to Old_Goals.METAHYPS and awaits deletion. Use SUBPROOF +or Subgoal.FOCUS etc. + +* Renamed functor TableFun to Table, and GraphFun to Graph. (Since +functors have their own ML name space there is no point to mark them +separately.) Minor INCOMPATIBILITY. + +* Renamed NamedThmsFun to Named_Thms. INCOMPATIBILITY. + +* Renamed several structures FooBar to Foo_Bar. Occasional, +INCOMPATIBILITY. + +* Operations of structure Skip_Proof no longer require quick_and_dirty +mode, which avoids critical setmp. + +* Eliminated old Attrib.add_attributes, Method.add_methods and related +combinators for "args". INCOMPATIBILITY, need to use simplified +Attrib/Method.setup introduced in Isabelle2009. + +* Proper context for simpset_of, claset_of, clasimpset_of. May fall +back on global_simpset_of, global_claset_of, global_clasimpset_of as +last resort. INCOMPATIBILITY. + +* Display.pretty_thm now requires a proper context (cf. former +ProofContext.pretty_thm). May fall back on Display.pretty_thm_global +or even Display.pretty_thm_without_context as last resort. +INCOMPATIBILITY. + +* Discontinued Display.pretty_ctyp/cterm etc. INCOMPATIBILITY, use +Syntax.pretty_typ/term directly, preferably with proper context +instead of global theory. + + +*** System *** + +* Further fine tuning of parallel proof checking, scales up to 8 cores +(max. speedup factor 5.0). See also Goal.parallel_proofs in ML and +usedir option -q. + +* Support for additional "Isabelle components" via etc/components, see +also the system manual. + +* The isabelle makeall tool now operates on all components with +IsaMakefile, not just hardwired "logics". + +* Removed "compress" option from isabelle-process and isabelle usedir; +this is always enabled. + +* Discontinued support for Poly/ML 4.x versions. + +* Isabelle tool "wwwfind" provides web interface for 'find_theorems' +on a given logic image. This requires the lighttpd webserver and is +currently supported on Linux only. + + + +New in Isabelle2009 (April 2009) +-------------------------------- + +*** General *** + +* Simplified main Isabelle executables, with less surprises on +case-insensitive file-systems (such as Mac OS). + + - The main Isabelle tool wrapper is now called "isabelle" instead of + "isatool." + + - The former "isabelle" alias for "isabelle-process" has been + removed (should rarely occur to regular users). + + - The former "isabelle-interface" and its alias "Isabelle" have been + removed (interfaces are now regular Isabelle tools). + +Within scripts and make files, the Isabelle environment variables +ISABELLE_TOOL and ISABELLE_PROCESS replace old ISATOOL and ISABELLE, +respectively. (The latter are still available as legacy feature.) + +The old isabelle-interface wrapper could react in confusing ways if +the interface was uninstalled or changed otherwise. Individual +interface tool configuration is now more explicit, see also the +Isabelle system manual. In particular, Proof General is now available +via "isabelle emacs". + +INCOMPATIBILITY, need to adapt derivative scripts. Users may need to +purge installed copies of Isabelle executables and re-run "isabelle +install -p ...", or use symlinks. + +* The default for ISABELLE_HOME_USER is now ~/.isabelle instead of the +old ~/isabelle, which was slightly non-standard and apt to cause +surprises on case-insensitive file-systems (such as Mac OS). + +INCOMPATIBILITY, need to move existing ~/isabelle/etc, +~/isabelle/heaps, ~/isabelle/browser_info to the new place. Special +care is required when using older releases of Isabelle. Note that +ISABELLE_HOME_USER can be changed in Isabelle/etc/settings of any +Isabelle distribution, in order to use the new ~/.isabelle uniformly. + +* Proofs of fully specified statements are run in parallel on +multi-core systems. A speedup factor of 2.5 to 3.2 can be expected on +a regular 4-core machine, if the initial heap space is made reasonably +large (cf. Poly/ML option -H). (Requires Poly/ML 5.2.1 or later.) + +* The main reference manuals ("isar-ref", "implementation", and +"system") have been updated and extended. Formally checked references +as hyperlinks are now available uniformly. + + +*** Pure *** + +* Complete re-implementation of locales. INCOMPATIBILITY in several +respects. The most important changes are listed below. See the +Tutorial on Locales ("locales" manual) for details. + +- In locale expressions, instantiation replaces renaming. Parameters +must be declared in a for clause. To aid compatibility with previous +parameter inheritance, in locale declarations, parameters that are not +'touched' (instantiation position "_" or omitted) are implicitly added +with their syntax at the beginning of the for clause. + +- Syntax from abbreviations and definitions in locales is available in +locale expressions and context elements. The latter is particularly +useful in locale declarations. + +- More flexible mechanisms to qualify names generated by locale +expressions. Qualifiers (prefixes) may be specified in locale +expressions, and can be marked as mandatory (syntax: "name!:") or +optional (syntax "name?:"). The default depends for plain "name:" +depends on the situation where a locale expression is used: in +commands 'locale' and 'sublocale' prefixes are optional, in +'interpretation' and 'interpret' prefixes are mandatory. The old +implicit qualifiers derived from the parameter names of a locale are +no longer generated. + +- Command "sublocale l < e" replaces "interpretation l < e". The +instantiation clause in "interpretation" and "interpret" (square +brackets) is no longer available. Use locale expressions. + +- When converting proof scripts, mandatory qualifiers in +'interpretation' and 'interpret' should be retained by default, even +if this is an INCOMPATIBILITY compared to former behavior. In the +worst case, use the "name?:" form for non-mandatory ones. Qualifiers +in locale expressions range over a single locale instance only. + +- Dropped locale element "includes". This is a major INCOMPATIBILITY. +In existing theorem specifications replace the includes element by the +respective context elements of the included locale, omitting those +that are already present in the theorem specification. Multiple +assume elements of a locale should be replaced by a single one +involving the locale predicate. In the proof body, declarations (most +notably theorems) may be regained by interpreting the respective +locales in the proof context as required (command "interpret"). + +If using "includes" in replacement of a target solely because the +parameter types in the theorem are not as general as in the target, +consider declaring a new locale with additional type constraints on +the parameters (context element "constrains"). + +- Discontinued "locale (open)". INCOMPATIBILITY. + +- Locale interpretation commands no longer attempt to simplify goal. +INCOMPATIBILITY: in rare situations the generated goal differs. Use +methods intro_locales and unfold_locales to clarify. + +- Locale interpretation commands no longer accept interpretation +attributes. INCOMPATIBILITY. + +* Class declaration: so-called "base sort" must not be given in import +list any longer, but is inferred from the specification. Particularly +in HOL, write + + class foo = ... + +instead of + + class foo = type + ... + +* Class target: global versions of theorems stemming do not carry a +parameter prefix any longer. INCOMPATIBILITY. + +* Class 'instance' command no longer accepts attached definitions. +INCOMPATIBILITY, use proper 'instantiation' target instead. + +* Recovered hiding of consts, which was accidentally broken in +Isabelle2007. Potential INCOMPATIBILITY, ``hide const c'' really +makes c inaccessible; consider using ``hide (open) const c'' instead. + +* Slightly more coherent Pure syntax, with updated documentation in +isar-ref manual. Removed locales meta_term_syntax and +meta_conjunction_syntax: TERM and &&& (formerly &&) are now permanent, +INCOMPATIBILITY in rare situations. Note that &&& should not be used +directly in regular applications. + +* There is a new syntactic category "float_const" for signed decimal +fractions (e.g. 123.45 or -123.45). + +* Removed exotic 'token_translation' command. INCOMPATIBILITY, use ML +interface with 'setup' command instead. + +* Command 'local_setup' is similar to 'setup', but operates on a local +theory context. + +* The 'axiomatization' command now only works within a global theory +context. INCOMPATIBILITY. + +* Goal-directed proof now enforces strict proof irrelevance wrt. sort +hypotheses. Sorts required in the course of reasoning need to be +covered by the constraints in the initial statement, completed by the +type instance information of the background theory. Non-trivial sort +hypotheses, which rarely occur in practice, may be specified via +vacuous propositions of the form SORT_CONSTRAINT('a::c). For example: + + lemma assumes "SORT_CONSTRAINT('a::empty)" shows False ... + +The result contains an implicit sort hypotheses as before -- +SORT_CONSTRAINT premises are eliminated as part of the canonical rule +normalization. + +* Generalized Isar history, with support for linear undo, direct state +addressing etc. + +* Changed defaults for unify configuration options: + + unify_trace_bound = 50 (formerly 25) + unify_search_bound = 60 (formerly 30) + +* Different bookkeeping for code equations (INCOMPATIBILITY): + + a) On theory merge, the last set of code equations for a particular + constant is taken (in accordance with the policy applied by other + parts of the code generator framework). + + b) Code equations stemming from explicit declarations (e.g. code + attribute) gain priority over default code equations stemming + from definition, primrec, fun etc. + +* Keyword 'code_exception' now named 'code_abort'. INCOMPATIBILITY. + +* Unified theorem tables for both code generators. Thus [code +func] has disappeared and only [code] remains. INCOMPATIBILITY. + +* Command 'find_consts' searches for constants based on type and name +patterns, e.g. + + find_consts "_ => bool" + +By default, matching is against subtypes, but it may be restricted to +the whole type. Searching by name is possible. Multiple queries are +conjunctive and queries may be negated by prefixing them with a +hyphen: + + find_consts strict: "_ => bool" name: "Int" -"int => int" + +* New 'find_theorems' criterion "solves" matches theorems that +directly solve the current goal (modulo higher-order unification). + +* Auto solve feature for main theorem statements: whenever a new goal +is stated, "find_theorems solves" is called; any theorems that could +solve the lemma directly are listed as part of the goal state. +Cf. associated options in Proof General Isabelle settings menu, +enabled by default, with reasonable timeout for pathological cases of +higher-order unification. + + +*** Document preparation *** + +* Antiquotation @{lemma} now imitates a regular terminal proof, +demanding keyword 'by' and supporting the full method expression +syntax just like the Isar command 'by'. + + +*** HOL *** + +* Integrated main parts of former image HOL-Complex with HOL. Entry +points Main and Complex_Main remain as before. + +* Logic image HOL-Plain provides a minimal HOL with the most important +tools available (inductive, datatype, primrec, ...). This facilitates +experimentation and tool development. Note that user applications +(and library theories) should never refer to anything below theory +Main, as before. + +* Logic image HOL-Main stops at theory Main, and thus facilitates +experimentation due to shorter build times. + +* Logic image HOL-NSA contains theories of nonstandard analysis which +were previously part of former HOL-Complex. Entry point Hyperreal +remains valid, but theories formerly using Complex_Main should now use +new entry point Hypercomplex. + +* Generic ATP manager for Sledgehammer, based on ML threads instead of +Posix processes. Avoids potentially expensive forking of the ML +process. New thread-based implementation also works on non-Unix +platforms (Cygwin). Provers are no longer hardwired, but defined +within the theory via plain ML wrapper functions. Basic Sledgehammer +commands are covered in the isar-ref manual. + +* Wrapper scripts for remote SystemOnTPTP service allows to use +sledgehammer without local ATP installation (Vampire etc.). Other +provers may be included via suitable ML wrappers, see also +src/HOL/ATP_Linkup.thy. + +* ATP selection (E/Vampire/Spass) is now via Proof General's settings +menu. + +* The metis method no longer fails because the theorem is too trivial +(contains the empty clause). + +* The metis method now fails in the usual manner, rather than raising +an exception, if it determines that it cannot prove the theorem. + +* Method "coherent" implements a prover for coherent logic (see also +src/Tools/coherent.ML). + +* Constants "undefined" and "default" replace "arbitrary". Usually +"undefined" is the right choice to replace "arbitrary", though +logically there is no difference. INCOMPATIBILITY. + +* Command "value" now integrates different evaluation mechanisms. The +result of the first successful evaluation mechanism is printed. In +square brackets a particular named evaluation mechanisms may be +specified (currently, [SML], [code] or [nbe]). See further +src/HOL/ex/Eval_Examples.thy. + +* Normalization by evaluation now allows non-leftlinear equations. +Declare with attribute [code nbe]. + +* Methods "case_tac" and "induct_tac" now refer to the very same rules +as the structured Isar versions "cases" and "induct", cf. the +corresponding "cases" and "induct" attributes. Mutual induction rules +are now presented as a list of individual projections +(e.g. foo_bar.inducts for types foo and bar); the old format with +explicit HOL conjunction is no longer supported. INCOMPATIBILITY, in +rare situations a different rule is selected --- notably nested tuple +elimination instead of former prod.exhaust: use explicit (case_tac t +rule: prod.exhaust) here. + +* Attributes "cases", "induct", "coinduct" support "del" option. + +* Removed fact "case_split_thm", which duplicates "case_split". + +* The option datatype has been moved to a new theory Option. Renamed +option_map to Option.map, and o2s to Option.set, INCOMPATIBILITY. + +* New predicate "strict_mono" classifies strict functions on partial +orders. With strict functions on linear orders, reasoning about +(in)equalities is facilitated by theorems "strict_mono_eq", +"strict_mono_less_eq" and "strict_mono_less". + +* Some set operations are now proper qualified constants with +authentic syntax. INCOMPATIBILITY: + + op Int ~> Set.Int + op Un ~> Set.Un + INTER ~> Set.INTER + UNION ~> Set.UNION + Inter ~> Set.Inter + Union ~> Set.Union + {} ~> Set.empty + UNIV ~> Set.UNIV + +* Class complete_lattice with operations Inf, Sup, INFI, SUPR now in +theory Set. + +* Auxiliary class "itself" has disappeared -- classes without any +parameter are treated as expected by the 'class' command. + +* Leibnitz's Series for Pi and the arcus tangens and logarithm series. + +* Common decision procedures (Cooper, MIR, Ferrack, Approximation, +Dense_Linear_Order) are now in directory HOL/Decision_Procs. + +* Theory src/HOL/Decision_Procs/Approximation provides the new proof +method "approximation". It proves formulas on real values by using +interval arithmetic. In the formulas are also the transcendental +functions sin, cos, tan, atan, ln, exp and the constant pi are +allowed. For examples see +src/HOL/Descision_Procs/ex/Approximation_Ex.thy. + +* Theory "Reflection" now resides in HOL/Library. + +* Entry point to Word library now simply named "Word". +INCOMPATIBILITY. + +* Made source layout more coherent with logical distribution +structure: + + src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy + src/HOL/Library/Code_Message.thy ~> src/HOL/ + src/HOL/Library/GCD.thy ~> src/HOL/ + src/HOL/Library/Order_Relation.thy ~> src/HOL/ + src/HOL/Library/Parity.thy ~> src/HOL/ + src/HOL/Library/Univ_Poly.thy ~> src/HOL/ + src/HOL/Real/ContNotDenum.thy ~> src/HOL/Library/ + src/HOL/Real/Lubs.thy ~> src/HOL/ + src/HOL/Real/PReal.thy ~> src/HOL/ + src/HOL/Real/Rational.thy ~> src/HOL/ + src/HOL/Real/RComplete.thy ~> src/HOL/ + src/HOL/Real/RealDef.thy ~> src/HOL/ + src/HOL/Real/RealPow.thy ~> src/HOL/ + src/HOL/Real/Real.thy ~> src/HOL/ + src/HOL/Complex/Complex_Main.thy ~> src/HOL/ + src/HOL/Complex/Complex.thy ~> src/HOL/ + src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/Library/ + src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/Library/ + src/HOL/Hyperreal/Deriv.thy ~> src/HOL/ + src/HOL/Hyperreal/Fact.thy ~> src/HOL/ + src/HOL/Hyperreal/Integration.thy ~> src/HOL/ + src/HOL/Hyperreal/Lim.thy ~> src/HOL/ + src/HOL/Hyperreal/Ln.thy ~> src/HOL/ + src/HOL/Hyperreal/Log.thy ~> src/HOL/ + src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/ + src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/ + src/HOL/Hyperreal/Series.thy ~> src/HOL/ + src/HOL/Hyperreal/SEQ.thy ~> src/HOL/ + src/HOL/Hyperreal/Taylor.thy ~> src/HOL/ + src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/ + src/HOL/Real/Float ~> src/HOL/Library/ + src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach + src/HOL/Real/RealVector.thy ~> src/HOL/ + + src/HOL/arith_data.ML ~> src/HOL/Tools + src/HOL/hologic.ML ~> src/HOL/Tools + src/HOL/simpdata.ML ~> src/HOL/Tools + src/HOL/int_arith1.ML ~> src/HOL/Tools/int_arith.ML + src/HOL/int_factor_simprocs.ML ~> src/HOL/Tools + src/HOL/nat_simprocs.ML ~> src/HOL/Tools + src/HOL/Real/float_arith.ML ~> src/HOL/Tools + src/HOL/Real/float_syntax.ML ~> src/HOL/Tools + src/HOL/Real/rat_arith.ML ~> src/HOL/Tools + src/HOL/Real/real_arith.ML ~> src/HOL/Tools + + src/HOL/Library/Array.thy ~> src/HOL/Imperative_HOL + src/HOL/Library/Heap_Monad.thy ~> src/HOL/Imperative_HOL + src/HOL/Library/Heap.thy ~> src/HOL/Imperative_HOL + src/HOL/Library/Imperative_HOL.thy ~> src/HOL/Imperative_HOL + src/HOL/Library/Ref.thy ~> src/HOL/Imperative_HOL + src/HOL/Library/Relational.thy ~> src/HOL/Imperative_HOL + +* If methods "eval" and "evaluation" encounter a structured proof +state with !!/==>, only the conclusion is evaluated to True (if +possible), avoiding strange error messages. + +* Method "sizechange" automates termination proofs using (a +modification of) the size-change principle. Requires SAT solver. See +src/HOL/ex/Termination.thy for examples. + +* Simplifier: simproc for let expressions now unfolds if bound +variable occurs at most once in let expression body. INCOMPATIBILITY. + +* Method "arith": Linear arithmetic now ignores all inequalities when +fast_arith_neq_limit is exceeded, instead of giving up entirely. + +* New attribute "arith" for facts that should always be used +automatically by arithmetic. It is intended to be used locally in +proofs, e.g. + + assumes [arith]: "x > 0" + +Global usage is discouraged because of possible performance impact. + +* New classes "top" and "bot" with corresponding operations "top" and +"bot" in theory Orderings; instantiation of class "complete_lattice" +requires instantiation of classes "top" and "bot". INCOMPATIBILITY. + +* Changed definition lemma "less_fun_def" in order to provide an +instance for preorders on functions; use lemma "less_le" instead. +INCOMPATIBILITY. + +* Theory Orderings: class "wellorder" moved here, with explicit +induction rule "less_induct" as assumption. For instantiation of +"wellorder" by means of predicate "wf", use rule wf_wellorderI. +INCOMPATIBILITY. + +* Theory Orderings: added class "preorder" as superclass of "order". +INCOMPATIBILITY: Instantiation proofs for order, linorder +etc. slightly changed. Some theorems named order_class.* now named +preorder_class.*. + +* Theory Relation: renamed "refl" to "refl_on", "reflexive" to "refl, +"diag" to "Id_on". + +* Theory Finite_Set: added a new fold combinator of type + + ('a => 'b => 'b) => 'b => 'a set => 'b + +Occasionally this is more convenient than the old fold combinator +which is now defined in terms of the new one and renamed to +fold_image. + +* Theories Ring_and_Field and OrderedGroup: The lemmas "group_simps" +and "ring_simps" have been replaced by "algebra_simps" (which can be +extended with further lemmas!). At the moment both still exist but +the former will disappear at some point. + +* Theory Power: Lemma power_Suc is now declared as a simp rule in +class recpower. Type-specific simp rules for various recpower types +have been removed. INCOMPATIBILITY, rename old lemmas as follows: + +rat_power_0 -> power_0 +rat_power_Suc -> power_Suc +realpow_0 -> power_0 +realpow_Suc -> power_Suc +complexpow_0 -> power_0 +complexpow_Suc -> power_Suc +power_poly_0 -> power_0 +power_poly_Suc -> power_Suc + +* Theories Ring_and_Field and Divides: Definition of "op dvd" has been +moved to separate class dvd in Ring_and_Field; a couple of lemmas on +dvd has been generalized to class comm_semiring_1. Likewise a bunch +of lemmas from Divides has been generalized from nat to class +semiring_div. INCOMPATIBILITY. This involves the following theorem +renames resulting from duplicate elimination: + + dvd_def_mod ~> dvd_eq_mod_eq_0 + zero_dvd_iff ~> dvd_0_left_iff + dvd_0 ~> dvd_0_right + DIVISION_BY_ZERO_DIV ~> div_by_0 + DIVISION_BY_ZERO_MOD ~> mod_by_0 + mult_div ~> div_mult_self2_is_id + mult_mod ~> mod_mult_self2_is_0 + +* Theory IntDiv: removed many lemmas that are instances of class-based +generalizations (from Divides and Ring_and_Field). INCOMPATIBILITY, +rename old lemmas as follows: + +dvd_diff -> nat_dvd_diff +dvd_zminus_iff -> dvd_minus_iff +mod_add1_eq -> mod_add_eq +mod_mult1_eq -> mod_mult_right_eq +mod_mult1_eq' -> mod_mult_left_eq +mod_mult_distrib_mod -> mod_mult_eq +nat_mod_add_left_eq -> mod_add_left_eq +nat_mod_add_right_eq -> mod_add_right_eq +nat_mod_div_trivial -> mod_div_trivial +nat_mod_mod_trivial -> mod_mod_trivial +zdiv_zadd_self1 -> div_add_self1 +zdiv_zadd_self2 -> div_add_self2 +zdiv_zmult_self1 -> div_mult_self2_is_id +zdiv_zmult_self2 -> div_mult_self1_is_id +zdvd_triv_left -> dvd_triv_left +zdvd_triv_right -> dvd_triv_right +zdvd_zmult_cancel_disj -> dvd_mult_cancel_left +zmod_eq0_zdvd_iff -> dvd_eq_mod_eq_0[symmetric] +zmod_zadd_left_eq -> mod_add_left_eq +zmod_zadd_right_eq -> mod_add_right_eq +zmod_zadd_self1 -> mod_add_self1 +zmod_zadd_self2 -> mod_add_self2 +zmod_zadd1_eq -> mod_add_eq +zmod_zdiff1_eq -> mod_diff_eq +zmod_zdvd_zmod -> mod_mod_cancel +zmod_zmod_cancel -> mod_mod_cancel +zmod_zmult_self1 -> mod_mult_self2_is_0 +zmod_zmult_self2 -> mod_mult_self1_is_0 +zmod_1 -> mod_by_1 +zdiv_1 -> div_by_1 +zdvd_abs1 -> abs_dvd_iff +zdvd_abs2 -> dvd_abs_iff +zdvd_refl -> dvd_refl +zdvd_trans -> dvd_trans +zdvd_zadd -> dvd_add +zdvd_zdiff -> dvd_diff +zdvd_zminus_iff -> dvd_minus_iff +zdvd_zminus2_iff -> minus_dvd_iff +zdvd_zmultD -> dvd_mult_right +zdvd_zmultD2 -> dvd_mult_left +zdvd_zmult_mono -> mult_dvd_mono +zdvd_0_right -> dvd_0_right +zdvd_0_left -> dvd_0_left_iff +zdvd_1_left -> one_dvd +zminus_dvd_iff -> minus_dvd_iff + +* Theory Rational: 'Fract k 0' now equals '0'. INCOMPATIBILITY. + +* The real numbers offer decimal input syntax: 12.34 is translated +into 1234/10^2. This translation is not reversed upon output. + +* Theory Library/Polynomial defines an abstract type 'a poly of +univariate polynomials with coefficients of type 'a. In addition to +the standard ring operations, it also supports div and mod. Code +generation is also supported, using list-style constructors. + +* Theory Library/Inner_Product defines a class of real_inner for real +inner product spaces, with an overloaded operation inner :: 'a => 'a +=> real. Class real_inner is a subclass of real_normed_vector from +theory RealVector. + +* Theory Library/Product_Vector provides instances for the product +type 'a * 'b of several classes from RealVector and Inner_Product. +Definitions of addition, subtraction, scalar multiplication, norms, +and inner products are included. + +* Theory Library/Bit defines the field "bit" of integers modulo 2. In +addition to the field operations, numerals and case syntax are also +supported. + +* Theory Library/Diagonalize provides constructive version of Cantor's +first diagonalization argument. + +* Theory Library/GCD: Curried operations gcd, lcm (for nat) and zgcd, +zlcm (for int); carried together from various gcd/lcm developements in +the HOL Distribution. Constants zgcd and zlcm replace former igcd and +ilcm; corresponding theorems renamed accordingly. INCOMPATIBILITY, +may recover tupled syntax as follows: + + hide (open) const gcd + abbreviation gcd where + "gcd == (%(a, b). GCD.gcd a b)" + notation (output) + GCD.gcd ("gcd '(_, _')") + +The same works for lcm, zgcd, zlcm. + +* Theory Library/Nat_Infinity: added addition, numeral syntax and more +instantiations for algebraic structures. Removed some duplicate +theorems. Changes in simp rules. INCOMPATIBILITY. + +* ML antiquotation @{code} takes a constant as argument and generates +corresponding code in background and inserts name of the corresponding +resulting ML value/function/datatype constructor binding in place. +All occurrences of @{code} with a single ML block are generated +simultaneously. Provides a generic and safe interface for +instrumentalizing code generation. See +src/HOL/Decision_Procs/Ferrack.thy for a more ambitious application. +In future you ought to refrain from ad-hoc compiling generated SML +code on the ML toplevel. Note that (for technical reasons) @{code} +cannot refer to constants for which user-defined serializations are +set. Refer to the corresponding ML counterpart directly in that +cases. + +* Command 'rep_datatype': instead of theorem names the command now +takes a list of terms denoting the constructors of the type to be +represented as datatype. The characteristic theorems have to be +proven. INCOMPATIBILITY. Also observe that the following theorems +have disappeared in favour of existing ones: + + unit_induct ~> unit.induct + prod_induct ~> prod.induct + sum_induct ~> sum.induct + Suc_Suc_eq ~> nat.inject + Suc_not_Zero Zero_not_Suc ~> nat.distinct + + +*** HOL-Algebra *** + +* New locales for orders and lattices where the equivalence relation +is not restricted to equality. INCOMPATIBILITY: all order and lattice +locales use a record structure with field eq for the equivalence. + +* New theory of factorial domains. + +* Units_l_inv and Units_r_inv are now simp rules by default. +INCOMPATIBILITY. Simplifier proof that require deletion of l_inv +and/or r_inv will now also require deletion of these lemmas. + +* Renamed the following theorems, INCOMPATIBILITY: + +UpperD ~> Upper_memD +LowerD ~> Lower_memD +least_carrier ~> least_closed +greatest_carrier ~> greatest_closed +greatest_Lower_above ~> greatest_Lower_below +one_zero ~> carrier_one_zero +one_not_zero ~> carrier_one_not_zero (collision with assumption) + + +*** HOL-Nominal *** + +* Nominal datatypes can now contain type-variables. + +* Commands 'nominal_inductive' and 'equivariance' work with local +theory targets. + +* Nominal primrec can now works with local theory targets and its +specification syntax now conforms to the general format as seen in +'inductive' etc. + +* Method "perm_simp" honours the standard simplifier attributes +(no_asm), (no_asm_use) etc. + +* The new predicate #* is defined like freshness, except that on the +left hand side can be a set or list of atoms. + +* Experimental command 'nominal_inductive2' derives strong induction +principles for inductive definitions. In contrast to +'nominal_inductive', which can only deal with a fixed number of +binders, it can deal with arbitrary expressions standing for sets of +atoms to be avoided. The only inductive definition we have at the +moment that needs this generalisation is the typing rule for Lets in +the algorithm W: + + Gamma |- t1 : T1 (x,close Gamma T1)::Gamma |- t2 : T2 x#Gamma + ----------------------------------------------------------------- + Gamma |- Let x be t1 in t2 : T2 + +In this rule one wants to avoid all the binders that are introduced by +"close Gamma T1". We are looking for other examples where this +feature might be useful. Please let us know. + + +*** HOLCF *** + +* Reimplemented the simplification procedure for proving continuity +subgoals. The new simproc is extensible; users can declare additional +continuity introduction rules with the attribute [cont2cont]. + +* The continuity simproc now uses a different introduction rule for +solving continuity subgoals on terms with lambda abstractions. In +some rare cases the new simproc may fail to solve subgoals that the +old one could solve, and "simp add: cont2cont_LAM" may be necessary. +Potential INCOMPATIBILITY. + +* Command 'fixrec': specification syntax now conforms to the general +format as seen in 'inductive' etc. See src/HOLCF/ex/Fixrec_ex.thy for +examples. INCOMPATIBILITY. + + +*** ZF *** + +* Proof of Zorn's Lemma for partial orders. + + +*** ML *** + +* Multithreading for Poly/ML 5.1/5.2 is no longer supported, only for +Poly/ML 5.2.1 or later. Important note: the TimeLimit facility +depends on multithreading, so timouts will not work before Poly/ML +5.2.1! + +* High-level support for concurrent ML programming, see +src/Pure/Cuncurrent. The data-oriented model of "future values" is +particularly convenient to organize independent functional +computations. The concept of "synchronized variables" provides a +higher-order interface for components with shared state, avoiding the +delicate details of mutexes and condition variables. (Requires +Poly/ML 5.2.1 or later.) + +* ML bindings produced via Isar commands are stored within the Isar +context (theory or proof). Consequently, commands like 'use' and 'ML' +become thread-safe and work with undo as expected (concerning +top-level bindings, not side-effects on global references). +INCOMPATIBILITY, need to provide proper Isar context when invoking the +compiler at runtime; really global bindings need to be given outside a +theory. (Requires Poly/ML 5.2 or later.) + +* Command 'ML_prf' is analogous to 'ML' but works within a proof +context. Top-level ML bindings are stored within the proof context in +a purely sequential fashion, disregarding the nested proof structure. +ML bindings introduced by 'ML_prf' are discarded at the end of the +proof. (Requires Poly/ML 5.2 or later.) + +* Simplified ML attribute and method setup, cf. functions Attrib.setup +and Method.setup, as well as Isar commands 'attribute_setup' and +'method_setup'. INCOMPATIBILITY for 'method_setup', need to simplify +existing code accordingly, or use plain 'setup' together with old +Method.add_method. + +* Simplified ML oracle interface Thm.add_oracle promotes 'a -> cterm +to 'a -> thm, while results are always tagged with an authentic oracle +name. The Isar command 'oracle' is now polymorphic, no argument type +is specified. INCOMPATIBILITY, need to simplify existing oracle code +accordingly. Note that extra performance may be gained by producing +the cterm carefully, avoiding slow Thm.cterm_of. + +* Simplified interface for defining document antiquotations via +ThyOutput.antiquotation, ThyOutput.output, and optionally +ThyOutput.maybe_pretty_source. INCOMPATIBILITY, need to simplify user +antiquotations accordingly, see src/Pure/Thy/thy_output.ML for common +examples. + +* More systematic treatment of long names, abstract name bindings, and +name space operations. Basic operations on qualified names have been +move from structure NameSpace to Long_Name, e.g. Long_Name.base_name, +Long_Name.append. Old type bstring has been mostly replaced by +abstract type binding (see structure Binding), which supports precise +qualification by packages and local theory targets, as well as proper +tracking of source positions. INCOMPATIBILITY, need to wrap old +bstring values into Binding.name, or better pass through abstract +bindings everywhere. See further src/Pure/General/long_name.ML, +src/Pure/General/binding.ML and src/Pure/General/name_space.ML + +* Result facts (from PureThy.note_thms, ProofContext.note_thms, +LocalTheory.note etc.) now refer to the *full* internal name, not the +bstring as before. INCOMPATIBILITY, not detected by ML type-checking! + +* Disposed old type and term read functions (Sign.read_def_typ, +Sign.read_typ, Sign.read_def_terms, Sign.read_term, +Thm.read_def_cterms, Thm.read_cterm etc.). INCOMPATIBILITY, should +use regular Syntax.read_typ, Syntax.read_term, Syntax.read_typ_global, +Syntax.read_term_global etc.; see also OldGoals.read_term as last +resort for legacy applications. + +* Disposed old declarations, tactics, tactic combinators that refer to +the simpset or claset of an implicit theory (such as Addsimps, +Simp_tac, SIMPSET). INCOMPATIBILITY, should use @{simpset} etc. in +embedded ML text, or local_simpset_of with a proper context passed as +explicit runtime argument. + +* Rules and tactics that read instantiations (read_instantiate, +res_inst_tac, thin_tac, subgoal_tac etc.) now demand a proper proof +context, which is required for parsing and type-checking. Moreover, +the variables are specified as plain indexnames, not string encodings +thereof. INCOMPATIBILITY. + +* Generic Toplevel.add_hook interface allows to analyze the result of +transactions. E.g. see src/Pure/ProofGeneral/proof_general_pgip.ML +for theorem dependency output of transactions resulting in a new +theory state. + +* ML antiquotations: block-structured compilation context indicated by +\ ... \; additional antiquotation forms: + + @{binding name} - basic name binding + @{let ?pat = term} - term abbreviation (HO matching) + @{note name = fact} - fact abbreviation + @{thm fact} - singleton fact (with attributes) + @{thms fact} - general fact (with attributes) + @{lemma prop by method} - singleton goal + @{lemma prop by meth1 meth2} - singleton goal + @{lemma prop1 ... propN by method} - general goal + @{lemma prop1 ... propN by meth1 meth2} - general goal + @{lemma (open) ...} - open derivation + + +*** System *** + +* The Isabelle "emacs" tool provides a specific interface to invoke +Proof General / Emacs, with more explicit failure if that is not +installed (the old isabelle-interface script silently falls back on +isabelle-process). The PROOFGENERAL_HOME setting determines the +installation location of the Proof General distribution. + +* Isabelle/lib/classes/Pure.jar provides basic support to integrate +the Isabelle process into a JVM/Scala application. See +Isabelle/lib/jedit/plugin for a minimal example. (The obsolete Java +process wrapper has been discontinued.) + +* Added homegrown Isabelle font with unicode layout, see lib/fonts. + +* Various status messages (with exact source position information) are +emitted, if proper markup print mode is enabled. This allows +user-interface components to provide detailed feedback on internal +prover operations. + + + +New in Isabelle2008 (June 2008) +------------------------------- + +*** General *** + +* The Isabelle/Isar Reference Manual (isar-ref) has been reorganized +and updated, with formally checked references as hyperlinks. + +* Theory loader: use_thy (and similar operations) no longer set the +implicit ML context, which was occasionally hard to predict and in +conflict with concurrency. INCOMPATIBILITY, use ML within Isar which +provides a proper context already. + +* Theory loader: old-style ML proof scripts being *attached* to a thy +file are no longer supported. INCOMPATIBILITY, regular 'uses' and +'use' within a theory file will do the job. + +* Name space merge now observes canonical order, i.e. the second space +is inserted into the first one, while existing entries in the first +space take precedence. INCOMPATIBILITY in rare situations, may try to +swap theory imports. + +* Syntax: symbol \ is now considered a letter. Potential +INCOMPATIBILITY in identifier syntax etc. + +* Outer syntax: string tokens no longer admit escaped white space, +which was an accidental (undocumented) feature. INCOMPATIBILITY, use +white space without escapes. + +* Outer syntax: string tokens may contain arbitrary character codes +specified via 3 decimal digits (as in SML). E.g. "foo\095bar" for +"foo_bar". + + +*** Pure *** + +* Context-dependent token translations. Default setup reverts locally +fixed variables, and adds hilite markup for undeclared frees. + +* Unused theorems can be found using the new command 'unused_thms'. +There are three ways of invoking it: + +(1) unused_thms + Only finds unused theorems in the current theory. + +(2) unused_thms thy_1 ... thy_n - + Finds unused theorems in the current theory and all of its ancestors, + excluding the theories thy_1 ... thy_n and all of their ancestors. + +(3) unused_thms thy_1 ... thy_n - thy'_1 ... thy'_m + Finds unused theorems in the theories thy'_1 ... thy'_m and all of + their ancestors, excluding the theories thy_1 ... thy_n and all of + their ancestors. + +In order to increase the readability of the list produced by +unused_thms, theorems that have been created by a particular instance +of a theory command such as 'inductive' or 'function' are considered +to belong to the same "group", meaning that if at least one theorem in +this group is used, the other theorems in the same group are no longer +reported as unused. Moreover, if all theorems in the group are +unused, only one theorem in the group is displayed. + +Note that proof objects have to be switched on in order for +unused_thms to work properly (i.e. !proofs must be >= 1, which is +usually the case when using Proof General with the default settings). + +* Authentic naming of facts disallows ad-hoc overwriting of previous +theorems within the same name space. INCOMPATIBILITY, need to remove +duplicate fact bindings, or even accidental fact duplications. Note +that tools may maintain dynamically scoped facts systematically, using +PureThy.add_thms_dynamic. + +* Command 'hide' now allows to hide from "fact" name space as well. + +* Eliminated destructive theorem database, simpset, claset, and +clasimpset. Potential INCOMPATIBILITY, really need to observe linear +update of theories within ML code. + +* Eliminated theory ProtoPure and CPure, leaving just one Pure theory. +INCOMPATIBILITY, object-logics depending on former Pure require +additional setup PureThy.old_appl_syntax_setup; object-logics +depending on former CPure need to refer to Pure. + +* Commands 'use' and 'ML' are now purely functional, operating on +theory/local_theory. Removed former 'ML_setup' (on theory), use 'ML' +instead. Added 'ML_val' as mere diagnostic replacement for 'ML'. +INCOMPATIBILITY. + +* Command 'setup': discontinued implicit version with ML reference. + +* Instantiation target allows for simultaneous specification of class +instance operations together with an instantiation proof. +Type-checking phase allows to refer to class operations uniformly. +See src/HOL/Complex/Complex.thy for an Isar example and +src/HOL/Library/Eval.thy for an ML example. + +* Indexing of literal facts: be more serious about including only +facts from the visible specification/proof context, but not the +background context (locale etc.). Affects `prop` notation and method +"fact". INCOMPATIBILITY: need to name facts explicitly in rare +situations. + +* Method "cases", "induct", "coinduct": removed obsolete/undocumented +"(open)" option, which used to expose internal bound variables to the +proof text. + +* Isar statements: removed obsolete case "rule_context". +INCOMPATIBILITY, better use explicit fixes/assumes. + +* Locale proofs: default proof step now includes 'unfold_locales'; +hence 'proof' without argument may be used to unfold locale +predicates. + + +*** Document preparation *** + +* Simplified pdfsetup.sty: color/hyperref is used unconditionally for +both pdf and dvi (hyperlinks usually work in xdvi as well); removed +obsolete thumbpdf setup (contemporary PDF viewers do this on the +spot); renamed link color from "darkblue" to "linkcolor" (default +value unchanged, can be redefined via \definecolor); no longer sets +"a4paper" option (unnecessary or even intrusive). + +* Antiquotation @{lemma A method} proves proposition A by the given +method (either a method name or a method name plus (optional) method +arguments in parentheses) and prints A just like @{prop A}. + + +*** HOL *** + +* New primrec package. Specification syntax conforms in style to +definition/function/.... No separate induction rule is provided. The +"primrec" command distinguishes old-style and new-style specifications +by syntax. The former primrec package is now named OldPrimrecPackage. +When adjusting theories, beware: constants stemming from new-style +primrec specifications have authentic syntax. + +* Metis prover is now an order of magnitude faster, and also works +with multithreading. + +* Metis: the maximum number of clauses that can be produced from a +theorem is now given by the attribute max_clauses. Theorems that +exceed this number are ignored, with a warning printed. + +* Sledgehammer no longer produces structured proofs by default. To +enable, declare [[sledgehammer_full = true]]. Attributes +reconstruction_modulus, reconstruction_sorts renamed +sledgehammer_modulus, sledgehammer_sorts. INCOMPATIBILITY. + +* Method "induct_scheme" derives user-specified induction rules +from well-founded induction and completeness of patterns. This factors +out some operations that are done internally by the function package +and makes them available separately. See +src/HOL/ex/Induction_Scheme.thy for examples. + +* More flexible generation of measure functions for termination +proofs: Measure functions can be declared by proving a rule of the +form "is_measure f" and giving it the [measure_function] attribute. +The "is_measure" predicate is logically meaningless (always true), and +just guides the heuristic. To find suitable measure functions, the +termination prover sets up the goal "is_measure ?f" of the appropriate +type and generates all solutions by Prolog-style backward proof using +the declared rules. + +This setup also deals with rules like + + "is_measure f ==> is_measure (list_size f)" + +which accommodates nested datatypes that recurse through lists. +Similar rules are predeclared for products and option types. + +* Turned the type of sets "'a set" into an abbreviation for "'a => bool" + + INCOMPATIBILITIES: + + - Definitions of overloaded constants on sets have to be replaced by + definitions on => and bool. + + - Some definitions of overloaded operators on sets can now be proved + using the definitions of the operators on => and bool. Therefore, + the following theorems have been renamed: + + subset_def -> subset_eq + psubset_def -> psubset_eq + set_diff_def -> set_diff_eq + Compl_def -> Compl_eq + Sup_set_def -> Sup_set_eq + Inf_set_def -> Inf_set_eq + sup_set_def -> sup_set_eq + inf_set_def -> inf_set_eq + + - Due to the incompleteness of the HO unification algorithm, some + rules such as subst may require manual instantiation, if some of + the unknowns in the rule is a set. + + - Higher order unification and forward proofs: + The proof pattern + + have "P (S::'a set)" <...> + then have "EX S. P S" .. + + no longer works (due to the incompleteness of the HO unification + algorithm) and must be replaced by the pattern + + have "EX S. P S" + proof + show "P S" <...> + qed + + - Calculational reasoning with subst (or similar rules): + The proof pattern + + have "P (S::'a set)" <...> + also have "S = T" <...> + finally have "P T" . + + no longer works (for similar reasons as the previous example) and + must be replaced by something like + + have "P (S::'a set)" <...> + moreover have "S = T" <...> + ultimately have "P T" by simp + + - Tactics or packages written in ML code: + Code performing pattern matching on types via + + Type ("set", [T]) => ... + + must be rewritten. Moreover, functions like strip_type or + binder_types no longer return the right value when applied to a + type of the form + + T1 => ... => Tn => U => bool + + rather than + + T1 => ... => Tn => U set + +* Merged theories Wellfounded_Recursion, Accessible_Part and +Wellfounded_Relations to theory Wellfounded. + +* Explicit class "eq" for executable equality. INCOMPATIBILITY. + +* Class finite no longer treats UNIV as class parameter. Use class +enum from theory Library/Enum instead to achieve a similar effect. +INCOMPATIBILITY. + +* Theory List: rule list_induct2 now has explicitly named cases "Nil" +and "Cons". INCOMPATIBILITY. + +* HOL (and FOL): renamed variables in rules imp_elim and swap. +Potential INCOMPATIBILITY. + +* Theory Product_Type: duplicated lemmas split_Pair_apply and +injective_fst_snd removed, use split_eta and prod_eqI instead. +Renamed upd_fst to apfst and upd_snd to apsnd. INCOMPATIBILITY. + +* Theory Nat: removed redundant lemmas that merely duplicate lemmas of +the same name in theory Orderings: + + less_trans + less_linear + le_imp_less_or_eq + le_less_trans + less_le_trans + less_not_sym + less_asym + +Renamed less_imp_le to less_imp_le_nat, and less_irrefl to +less_irrefl_nat. Potential INCOMPATIBILITY due to more general types +and different variable names. + +* Library/Option_ord.thy: Canonical order on option type. + +* Library/RBT.thy: Red-black trees, an efficient implementation of +finite maps. + +* Library/Countable.thy: Type class for countable types. + +* Theory Int: The representation of numerals has changed. The infix +operator BIT and the bit datatype with constructors B0 and B1 have +disappeared. INCOMPATIBILITY, use "Int.Bit0 x" and "Int.Bit1 y" in +place of "x BIT bit.B0" and "y BIT bit.B1", respectively. Theorems +involving BIT, B0, or B1 have been renamed with "Bit0" or "Bit1" +accordingly. + +* Theory Nat: definition of <= and < on natural numbers no longer +depend on well-founded relations. INCOMPATIBILITY. Definitions +le_def and less_def have disappeared. Consider lemmas not_less +[symmetric, where ?'a = nat] and less_eq [symmetric] instead. + +* Theory Finite_Set: locales ACf, ACe, ACIf, ACIfSL and ACIfSLlin +(whose purpose mainly is for various fold_set functionals) have been +abandoned in favor of the existing algebraic classes +ab_semigroup_mult, comm_monoid_mult, ab_semigroup_idem_mult, +lower_semilattice (resp. upper_semilattice) and linorder. +INCOMPATIBILITY. + +* Theory Transitive_Closure: induct and cases rules now declare proper +case_names ("base" and "step"). INCOMPATIBILITY. + +* Theorem Inductive.lfp_ordinal_induct generalized to complete +lattices. The form set-specific version is available as +Inductive.lfp_ordinal_induct_set. + +* Renamed theorems "power.simps" to "power_int.simps". +INCOMPATIBILITY. + +* Class semiring_div provides basic abstract properties of semirings +with division and modulo operations. Subsumes former class dvd_mod. + +* Merged theories IntDef, Numeral and IntArith into unified theory +Int. INCOMPATIBILITY. + +* Theory Library/Code_Index: type "index" now represents natural +numbers rather than integers. INCOMPATIBILITY. + +* New class "uminus" with operation "uminus" (split of from class +"minus" which now only has operation "minus", binary). +INCOMPATIBILITY. + +* Constants "card", "internal_split", "option_map" now with authentic +syntax. INCOMPATIBILITY. + +* Definitions subset_def, psubset_def, set_diff_def, Compl_def, +le_bool_def, less_bool_def, le_fun_def, less_fun_def, inf_bool_def, +sup_bool_def, Inf_bool_def, Sup_bool_def, inf_fun_def, sup_fun_def, +Inf_fun_def, Sup_fun_def, inf_set_def, sup_set_def, Inf_set_def, +Sup_set_def, le_def, less_def, option_map_def now with object +equality. INCOMPATIBILITY. + +* Records. Removed K_record, and replaced it by pure lambda term +%x. c. The simplifier setup is now more robust against eta expansion. +INCOMPATIBILITY: in cases explicitly referring to K_record. + +* Library/Multiset: {#a, b, c#} abbreviates {#a#} + {#b#} + {#c#}. + +* Library/ListVector: new theory of arithmetic vector operations. + +* Library/Order_Relation: new theory of various orderings as sets of +pairs. Defines preorders, partial orders, linear orders and +well-orders on sets and on types. + + +*** ZF *** + +* Renamed some theories to allow to loading both ZF and HOL in the +same session: + + Datatype -> Datatype_ZF + Inductive -> Inductive_ZF + Int -> Int_ZF + IntDiv -> IntDiv_ZF + Nat -> Nat_ZF + List -> List_ZF + Main -> Main_ZF + +INCOMPATIBILITY: ZF theories that import individual theories below +Main might need to be adapted. Regular theory Main is still +available, as trivial extension of Main_ZF. + + +*** ML *** + +* ML within Isar: antiquotation @{const name} or @{const +name(typargs)} produces statically-checked Const term. + +* Functor NamedThmsFun: data is available to the user as dynamic fact +(of the same name). Removed obsolete print command. + +* Removed obsolete "use_legacy_bindings" function. + +* The ``print mode'' is now a thread-local value derived from a global +template (the former print_mode reference), thus access becomes +non-critical. The global print_mode reference is for session +management only; user-code should use print_mode_value, +print_mode_active, PrintMode.setmp etc. INCOMPATIBILITY. + +* Functions system/system_out provide a robust way to invoke external +shell commands, with propagation of interrupts (requires Poly/ML +5.2.1). Do not use OS.Process.system etc. from the basis library! + + +*** System *** + +* Default settings: PROOFGENERAL_OPTIONS no longer impose xemacs --- +in accordance with Proof General 3.7, which prefers GNU emacs. + +* isatool tty runs Isabelle process with plain tty interaction; +optional line editor may be specified via ISABELLE_LINE_EDITOR +setting, the default settings attempt to locate "ledit" and "rlwrap". + +* isatool browser now works with Cygwin as well, using general +"javapath" function defined in Isabelle process environment. + +* YXML notation provides a simple and efficient alternative to +standard XML transfer syntax. See src/Pure/General/yxml.ML and +isatool yxml as described in the Isabelle system manual. + +* JVM class isabelle.IsabelleProcess (located in Isabelle/lib/classes) +provides general wrapper for managing an Isabelle process in a robust +fashion, with ``cooked'' output from stdin/stderr. + +* Rudimentary Isabelle plugin for jEdit (see Isabelle/lib/jedit), +based on Isabelle/JVM process wrapper (see Isabelle/lib/classes). + +* Removed obsolete THIS_IS_ISABELLE_BUILD feature. NB: the documented +way of changing the user's settings is via +ISABELLE_HOME_USER/etc/settings, which is a fully featured bash +script. + +* Multithreading.max_threads := 0 refers to the number of actual CPU +cores of the underlying machine, which is a good starting point for +optimal performance tuning. The corresponding usedir option -M allows +"max" as an alias for "0". WARNING: does not work on certain versions +of Mac OS (with Poly/ML 5.1). + +* isabelle-process: non-ML sessions are run with "nice", to reduce the +adverse effect of Isabelle flooding interactive front-ends (notably +ProofGeneral / XEmacs). + + + +New in Isabelle2007 (November 2007) +----------------------------------- + +*** General *** + +* More uniform information about legacy features, notably a +warning/error of "Legacy feature: ...", depending on the state of the +tolerate_legacy_features flag (default true). FUTURE INCOMPATIBILITY: +legacy features will disappear eventually. + +* Theory syntax: the header format ``theory A = B + C:'' has been +discontinued in favour of ``theory A imports B C begin''. Use isatool +fixheaders to convert existing theory files. INCOMPATIBILITY. + +* Theory syntax: the old non-Isar theory file format has been +discontinued altogether. Note that ML proof scripts may still be used +with Isar theories; migration is usually quite simple with the ML +function use_legacy_bindings. INCOMPATIBILITY. + +* Theory syntax: some popular names (e.g. 'class', 'declaration', +'fun', 'help', 'if') are now keywords. INCOMPATIBILITY, use double +quotes. + +* Theory loader: be more serious about observing the static theory +header specifications (including optional directories), but not the +accidental file locations of previously successful loads. The strict +update policy of former update_thy is now already performed by +use_thy, so the former has been removed; use_thys updates several +theories simultaneously, just as 'imports' within a theory header +specification, but without merging the results. Potential +INCOMPATIBILITY: may need to refine theory headers and commands +ROOT.ML which depend on load order. + +* Theory loader: optional support for content-based file +identification, instead of the traditional scheme of full physical +path plus date stamp; configured by the ISABELLE_FILE_IDENT setting +(cf. the system manual). The new scheme allows to work with +non-finished theories in persistent session images, such that source +files may be moved later on without requiring reloads. + +* Theory loader: old-style ML proof scripts being *attached* to a thy +file (with the same base name as the theory) are considered a legacy +feature, which will disappear eventually. Even now, the theory loader +no longer maintains dependencies on such files. + +* Syntax: the scope for resolving ambiguities via type-inference is +now limited to individual terms, instead of whole simultaneous +specifications as before. This greatly reduces the complexity of the +syntax module and improves flexibility by separating parsing and +type-checking. INCOMPATIBILITY: additional type-constraints (explicit +'fixes' etc.) are required in rare situations. + +* Syntax: constants introduced by new-style packages ('definition', +'abbreviation' etc.) are passed through the syntax module in +``authentic mode''. This means that associated mixfix annotations +really stick to such constants, independently of potential name space +ambiguities introduced later on. INCOMPATIBILITY: constants in parse +trees are represented slightly differently, may need to adapt syntax +translations accordingly. Use CONST marker in 'translations' and +@{const_syntax} antiquotation in 'parse_translation' etc. + +* Legacy goal package: reduced interface to the bare minimum required +to keep existing proof scripts running. Most other user-level +functions are now part of the OldGoals structure, which is *not* open +by default (consider isatool expandshort before open OldGoals). +Removed top_sg, prin, printyp, pprint_term/typ altogether, because +these tend to cause confusion about the actual goal (!) context being +used here, which is not necessarily the same as the_context(). + +* Command 'find_theorems': supports "*" wild-card in "name:" +criterion; "with_dups" option. Certain ProofGeneral versions might +support a specific search form (see ProofGeneral/CHANGES). + +* The ``prems limit'' option (cf. ProofContext.prems_limit) is now -1 +by default, which means that "prems" (and also "fixed variables") are +suppressed from proof state output. Note that the ProofGeneral +settings mechanism allows to change and save options persistently, but +older versions of Isabelle will fail to start up if a negative prems +limit is imposed. + +* Local theory targets may be specified by non-nested blocks of +``context/locale/class ... begin'' followed by ``end''. The body may +contain definitions, theorems etc., including any derived mechanism +that has been implemented on top of these primitives. This concept +generalizes the existing ``theorem (in ...)'' towards more versatility +and scalability. + +* Proof General interface: proper undo of final 'end' command; +discontinued Isabelle/classic mode (ML proof scripts). + + +*** Document preparation *** + +* Added antiquotation @{theory name} which prints the given name, +after checking that it refers to a valid ancestor theory in the +current context. + +* Added antiquotations @{ML_type text} and @{ML_struct text} which +check the given source text as ML type/structure, printing verbatim. + +* Added antiquotation @{abbrev "c args"} which prints the abbreviation +"c args == rhs" given in the current context. (Any number of +arguments may be given on the LHS.) + + +*** Pure *** + +* The 'class' package offers a combination of axclass and locale to +achieve Haskell-like type classes in Isabelle. Definitions and +theorems within a class context produce both relative results (with +implicit parameters according to the locale context), and polymorphic +constants with qualified polymorphism (according to the class +context). Within the body context of a 'class' target, a separate +syntax layer ("user space type system") takes care of converting +between global polymorphic consts and internal locale representation. +See src/HOL/ex/Classpackage.thy for examples (as well as main HOL). +"isatool doc classes" provides a tutorial. + +* Generic code generator framework allows to generate executable +code for ML and Haskell (including Isabelle classes). A short usage +sketch: + + internal compilation: + export_code in SML + writing SML code to a file: + export_code in SML + writing OCaml code to a file: + export_code in OCaml + writing Haskell code to a bunch of files: + export_code in Haskell + + evaluating closed propositions to True/False using code generation: + method ``eval'' + +Reasonable default setup of framework in HOL. + +Theorem attributs for selecting and transforming function equations theorems: + + [code fun]: select a theorem as function equation for a specific constant + [code fun del]: deselect a theorem as function equation for a specific constant + [code inline]: select an equation theorem for unfolding (inlining) in place + [code inline del]: deselect an equation theorem for unfolding (inlining) in place + +User-defined serializations (target in {SML, OCaml, Haskell}): + + code_const + {(target) }+ + + code_type + {(target) }+ + + code_instance + {(target)}+ + where instance ::= :: + + code_class + {(target) }+ + where class target syntax ::= {where { == }+}? + +code_instance and code_class only are effective to target Haskell. + +For example usage see src/HOL/ex/Codegenerator.thy and +src/HOL/ex/Codegenerator_Pretty.thy. A separate tutorial on code +generation from Isabelle/HOL theories is available via "isatool doc +codegen". + +* Code generator: consts in 'consts_code' Isar commands are now +referred to by usual term syntax (including optional type +annotations). + +* Command 'no_translations' removes translation rules from theory +syntax. + +* Overloaded definitions are now actually checked for acyclic +dependencies. The overloading scheme is slightly more general than +that of Haskell98, although Isabelle does not demand an exact +correspondence to type class and instance declarations. +INCOMPATIBILITY, use ``defs (unchecked overloaded)'' to admit more +exotic versions of overloading -- at the discretion of the user! + +Polymorphic constants are represented via type arguments, i.e. the +instantiation that matches an instance against the most general +declaration given in the signature. For example, with the declaration +c :: 'a => 'a => 'a, an instance c :: nat => nat => nat is represented +as c(nat). Overloading is essentially simultaneous structural +recursion over such type arguments. Incomplete specification patterns +impose global constraints on all occurrences, e.g. c('a * 'a) on the +LHS means that more general c('a * 'b) will be disallowed on any RHS. +Command 'print_theory' outputs the normalized system of recursive +equations, see section "definitions". + +* Configuration options are maintained within the theory or proof +context (with name and type bool/int/string), providing a very simple +interface to a poor-man's version of general context data. Tools may +declare options in ML (e.g. using Attrib.config_int) and then refer to +these values using Config.get etc. Users may change options via an +associated attribute of the same name. This form of context +declaration works particularly well with commands 'declare' or +'using', for example ``declare [[foo = 42]]''. Thus it has become +very easy to avoid global references, which would not observe Isar +toplevel undo/redo and fail to work with multithreading. + +Various global ML references of Pure and HOL have been turned into +configuration options: + + Unify.search_bound unify_search_bound + Unify.trace_bound unify_trace_bound + Unify.trace_simp unify_trace_simp + Unify.trace_types unify_trace_types + Simplifier.simp_depth_limit simp_depth_limit + Blast.depth_limit blast_depth_limit + DatatypeProp.dtK datatype_distinctness_limit + fast_arith_neq_limit fast_arith_neq_limit + fast_arith_split_limit fast_arith_split_limit + +* Named collections of theorems may be easily installed as context +data using the functor NamedThmsFun (see also +src/Pure/Tools/named_thms.ML). The user may add or delete facts via +attributes; there is also a toplevel print command. This facility is +just a common case of general context data, which is the preferred way +for anything more complex than just a list of facts in canonical +order. + +* Isar: command 'declaration' augments a local theory by generic +declaration functions written in ML. This enables arbitrary content +being added to the context, depending on a morphism that tells the +difference of the original declaration context wrt. the application +context encountered later on. + +* Isar: proper interfaces for simplification procedures. Command +'simproc_setup' declares named simprocs (with match patterns, and body +text in ML). Attribute "simproc" adds/deletes simprocs in the current +context. ML antiquotation @{simproc name} retrieves named simprocs. + +* Isar: an extra pair of brackets around attribute declarations +abbreviates a theorem reference involving an internal dummy fact, +which will be ignored later --- only the effect of the attribute on +the background context will persist. This form of in-place +declarations is particularly useful with commands like 'declare' and +'using', for example ``have A using [[simproc a]] by simp''. + +* Isar: method "assumption" (and implicit closing of subproofs) now +takes simple non-atomic goal assumptions into account: after applying +an assumption as a rule the resulting subgoals are solved by atomic +assumption steps. This is particularly useful to finish 'obtain' +goals, such as "!!x. (!!x. P x ==> thesis) ==> P x ==> thesis", +without referring to the original premise "!!x. P x ==> thesis" in the +Isar proof context. POTENTIAL INCOMPATIBILITY: method "assumption" is +more permissive. + +* Isar: implicit use of prems from the Isar proof context is +considered a legacy feature. Common applications like ``have A .'' +may be replaced by ``have A by fact'' or ``note `A`''. In general, +referencing facts explicitly here improves readability and +maintainability of proof texts. + +* Isar: improper proof element 'guess' is like 'obtain', but derives +the obtained context from the course of reasoning! For example: + + assume "EX x y. A x & B y" -- "any previous fact" + then guess x and y by clarify + +This technique is potentially adventurous, depending on the facts and +proof tools being involved here. + +* Isar: known facts from the proof context may be specified as literal +propositions, using ASCII back-quote syntax. This works wherever +named facts used to be allowed so far, in proof commands, proof +methods, attributes etc. Literal facts are retrieved from the context +according to unification of type and term parameters. For example, +provided that "A" and "A ==> B" and "!!x. P x ==> Q x" are known +theorems in the current context, then these are valid literal facts: +`A` and `A ==> B` and `!!x. P x ==> Q x" as well as `P a ==> Q a` etc. + +There is also a proof method "fact" which does the same composition +for explicit goal states, e.g. the following proof texts coincide with +certain special cases of literal facts: + + have "A" by fact == note `A` + have "A ==> B" by fact == note `A ==> B` + have "!!x. P x ==> Q x" by fact == note `!!x. P x ==> Q x` + have "P a ==> Q a" by fact == note `P a ==> Q a` + +* Isar: ":" (colon) is no longer a symbolic identifier character in +outer syntax. Thus symbolic identifiers may be used without +additional white space in declarations like this: ``assume *: A''. + +* Isar: 'print_facts' prints all local facts of the current context, +both named and unnamed ones. + +* Isar: 'def' now admits simultaneous definitions, e.g.: + + def x == "t" and y == "u" + +* Isar: added command 'unfolding', which is structurally similar to +'using', but affects both the goal state and facts by unfolding given +rewrite rules. Thus many occurrences of the 'unfold' method or +'unfolded' attribute may be replaced by first-class proof text. + +* Isar: methods 'unfold' / 'fold', attributes 'unfolded' / 'folded', +and command 'unfolding' now all support object-level equalities +(potentially conditional). The underlying notion of rewrite rule is +analogous to the 'rule_format' attribute, but *not* that of the +Simplifier (which is usually more generous). + +* Isar: the new attribute [rotated n] (default n = 1) rotates the +premises of a theorem by n. Useful in conjunction with drule. + +* Isar: the goal restriction operator [N] (default N = 1) evaluates a +method expression within a sandbox consisting of the first N +sub-goals, which need to exist. For example, ``simp_all [3]'' +simplifies the first three sub-goals, while (rule foo, simp_all)[] +simplifies all new goals that emerge from applying rule foo to the +originally first one. + +* Isar: schematic goals are no longer restricted to higher-order +patterns; e.g. ``lemma "?P(?x)" by (rule TrueI)'' now works as +expected. + +* Isar: the conclusion of a long theorem statement is now either +'shows' (a simultaneous conjunction, as before), or 'obtains' +(essentially a disjunction of cases with local parameters and +assumptions). The latter allows to express general elimination rules +adequately; in this notation common elimination rules look like this: + + lemma exE: -- "EX x. P x ==> (!!x. P x ==> thesis) ==> thesis" + assumes "EX x. P x" + obtains x where "P x" + + lemma conjE: -- "A & B ==> (A ==> B ==> thesis) ==> thesis" + assumes "A & B" + obtains A and B + + lemma disjE: -- "A | B ==> (A ==> thesis) ==> (B ==> thesis) ==> thesis" + assumes "A | B" + obtains + A + | B + +The subsequent classical rules even refer to the formal "thesis" +explicitly: + + lemma classical: -- "(~ thesis ==> thesis) ==> thesis" + obtains "~ thesis" + + lemma Peirce's_Law: -- "((thesis ==> something) ==> thesis) ==> thesis" + obtains "thesis ==> something" + +The actual proof of an 'obtains' statement is analogous to that of the +Isar proof element 'obtain', only that there may be several cases. +Optional case names may be specified in parentheses; these will be +available both in the present proof and as annotations in the +resulting rule, for later use with the 'cases' method (cf. attribute +case_names). + +* Isar: the assumptions of a long theorem statement are available as +"assms" fact in the proof context. This is more appropriate than the +(historical) "prems", which refers to all assumptions of the current +context, including those from the target locale, proof body etc. + +* Isar: 'print_statement' prints theorems from the current theory or +proof context in long statement form, according to the syntax of a +top-level lemma. + +* Isar: 'obtain' takes an optional case name for the local context +introduction rule (default "that"). + +* Isar: removed obsolete 'concl is' patterns. INCOMPATIBILITY, use +explicit (is "_ ==> ?foo") in the rare cases where this still happens +to occur. + +* Pure: syntax "CONST name" produces a fully internalized constant +according to the current context. This is particularly useful for +syntax translations that should refer to internal constant +representations independently of name spaces. + +* Pure: syntax constant for foo (binder "FOO ") is called "foo_binder" +instead of "FOO ". This allows multiple binder declarations to coexist +in the same context. INCOMPATIBILITY. + +* Isar/locales: 'notation' provides a robust interface to the 'syntax' +primitive that also works in a locale context (both for constants and +fixed variables). Type declaration and internal syntactic representation +of given constants retrieved from the context. Likewise, the +'no_notation' command allows to remove given syntax annotations from the +current context. + +* Isar/locales: new derived specification elements 'axiomatization', +'definition', 'abbreviation', which support type-inference, admit +object-level specifications (equality, equivalence). See also the +isar-ref manual. Examples: + + axiomatization + eq (infix "===" 50) where + eq_refl: "x === x" and eq_subst: "x === y ==> P x ==> P y" + + definition "f x y = x + y + 1" + definition g where "g x = f x x" + + abbreviation + neq (infix "=!=" 50) where + "x =!= y == ~ (x === y)" + +These specifications may be also used in a locale context. Then the +constants being introduced depend on certain fixed parameters, and the +constant name is qualified by the locale base name. An internal +abbreviation takes care for convenient input and output, making the +parameters implicit and using the original short name. See also +src/HOL/ex/Abstract_NAT.thy for an example of deriving polymorphic +entities from a monomorphic theory. + +Presently, abbreviations are only available 'in' a target locale, but +not inherited by general import expressions. Also note that +'abbreviation' may be used as a type-safe replacement for 'syntax' + +'translations' in common applications. The "no_abbrevs" print mode +prevents folding of abbreviations in term output. + +Concrete syntax is attached to specified constants in internal form, +independently of name spaces. The parse tree representation is +slightly different -- use 'notation' instead of raw 'syntax', and +'translations' with explicit "CONST" markup to accommodate this. + +* Pure/Isar: unified syntax for new-style specification mechanisms +(e.g. 'definition', 'abbreviation', or 'inductive' in HOL) admits +full type inference and dummy patterns ("_"). For example: + + definition "K x _ = x" + + inductive conj for A B + where "A ==> B ==> conj A B" + +* Pure: command 'print_abbrevs' prints all constant abbreviations of +the current context. Print mode "no_abbrevs" prevents inversion of +abbreviations on output. + +* Isar/locales: improved parameter handling: use of locales "var" and +"struct" no longer necessary; - parameter renamings are no longer +required to be injective. For example, this allows to define +endomorphisms as locale endom = homom mult mult h. + +* Isar/locales: changed the way locales with predicates are defined. +Instead of accumulating the specification, the imported expression is +now an interpretation. INCOMPATIBILITY: different normal form of +locale expressions. In particular, in interpretations of locales with +predicates, goals repesenting already interpreted fragments are not +removed automatically. Use methods `intro_locales' and +`unfold_locales'; see below. + +* Isar/locales: new methods `intro_locales' and `unfold_locales' +provide backward reasoning on locales predicates. The methods are +aware of interpretations and discharge corresponding goals. +`intro_locales' is less aggressive then `unfold_locales' and does not +unfold predicates to assumptions. + +* Isar/locales: the order in which locale fragments are accumulated +has changed. This enables to override declarations from fragments due +to interpretations -- for example, unwanted simp rules. + +* Isar/locales: interpretation in theories and proof contexts has been +extended. One may now specify (and prove) equations, which are +unfolded in interpreted theorems. This is useful for replacing +defined concepts (constants depending on locale parameters) by +concepts already existing in the target context. Example: + + interpretation partial_order ["op <= :: [int, int] => bool"] + where "partial_order.less (op <=) (x::int) y = (x < y)" + +Typically, the constant `partial_order.less' is created by a +definition specification element in the context of locale +partial_order. + +* Method "induct": improved internal context management to support +local fixes and defines on-the-fly. Thus explicit meta-level +connectives !! and ==> are rarely required anymore in inductive goals +(using object-logic connectives for this purpose has been long +obsolete anyway). Common proof patterns are explained in +src/HOL/Induct/Common_Patterns.thy, see also +src/HOL/Isar_examples/Puzzle.thy and src/HOL/Lambda for realistic +examples. + +* Method "induct": improved handling of simultaneous goals. Instead of +introducing object-level conjunction, the statement is now split into +several conclusions, while the corresponding symbolic cases are nested +accordingly. INCOMPATIBILITY, proofs need to be structured explicitly, +see src/HOL/Induct/Common_Patterns.thy, for example. + +* Method "induct": mutual induction rules are now specified as a list +of rule sharing the same induction cases. HOL packages usually provide +foo_bar.inducts for mutually defined items foo and bar (e.g. inductive +predicates/sets or datatypes). INCOMPATIBILITY, users need to specify +mutual induction rules differently, i.e. like this: + + (induct rule: foo_bar.inducts) + (induct set: foo bar) + (induct pred: foo bar) + (induct type: foo bar) + +The ML function ProjectRule.projections turns old-style rules into the +new format. + +* Method "coinduct": dual of induction, see +src/HOL/Library/Coinductive_List.thy for various examples. + +* Method "cases", "induct", "coinduct": the ``(open)'' option is +considered a legacy feature. + +* Attribute "symmetric" produces result with standardized schematic +variables (index 0). Potential INCOMPATIBILITY. + +* Simplifier: by default the simplifier trace only shows top level +rewrites now. That is, trace_simp_depth_limit is set to 1 by +default. Thus there is less danger of being flooded by the trace. The +trace indicates where parts have been suppressed. + +* Provers/classical: removed obsolete classical version of elim_format +attribute; classical elim/dest rules are now treated uniformly when +manipulating the claset. + +* Provers/classical: stricter checks to ensure that supplied intro, +dest and elim rules are well-formed; dest and elim rules must have at +least one premise. + +* Provers/classical: attributes dest/elim/intro take an optional +weight argument for the rule (just as the Pure versions). Weights are +ignored by automated tools, but determine the search order of single +rule steps. + +* Syntax: input syntax now supports dummy variable binding "%_. b", +where the body does not mention the bound variable. Note that dummy +patterns implicitly depend on their context of bounds, which makes +"{_. _}" match any set comprehension as expected. Potential +INCOMPATIBILITY -- parse translations need to cope with syntactic +constant "_idtdummy" in the binding position. + +* Syntax: removed obsolete syntactic constant "_K" and its associated +parse translation. INCOMPATIBILITY -- use dummy abstraction instead, +for example "A -> B" => "Pi A (%_. B)". + +* Pure: 'class_deps' command visualizes the subclass relation, using +the graph browser tool. + +* Pure: 'print_theory' now suppresses certain internal declarations by +default; use '!' option for full details. + + +*** HOL *** + +* Method "metis" proves goals by applying the Metis general-purpose +resolution prover (see also http://gilith.com/software/metis/). +Examples are in the directory MetisExamples. WARNING: the +Isabelle/HOL-Metis integration does not yet work properly with +multi-threading. + +* Command 'sledgehammer' invokes external automatic theorem provers as +background processes. It generates calls to the "metis" method if +successful. These can be pasted into the proof. Users do not have to +wait for the automatic provers to return. WARNING: does not really +work with multi-threading. + +* New "auto_quickcheck" feature tests outermost goal statements for +potential counter-examples. Controlled by ML references +auto_quickcheck (default true) and auto_quickcheck_time_limit (default +5000 milliseconds). Fails silently if statements is outside of +executable fragment, or any other codgenerator problem occurs. + +* New constant "undefined" with axiom "undefined x = undefined". + +* Added class "HOL.eq", allowing for code generation with polymorphic +equality. + +* Some renaming of class constants due to canonical name prefixing in +the new 'class' package: + + HOL.abs ~> HOL.abs_class.abs + HOL.divide ~> HOL.divide_class.divide + 0 ~> HOL.zero_class.zero + 1 ~> HOL.one_class.one + op + ~> HOL.plus_class.plus + op - ~> HOL.minus_class.minus + uminus ~> HOL.minus_class.uminus + op * ~> HOL.times_class.times + op < ~> HOL.ord_class.less + op <= > HOL.ord_class.less_eq + Nat.power ~> Power.power_class.power + Nat.size ~> Nat.size_class.size + Numeral.number_of ~> Numeral.number_class.number_of + FixedPoint.Inf ~> Lattices.complete_lattice_class.Inf + FixedPoint.Sup ~> Lattices.complete_lattice_class.Sup + Orderings.min ~> Orderings.ord_class.min + Orderings.max ~> Orderings.ord_class.max + Divides.op div ~> Divides.div_class.div + Divides.op mod ~> Divides.div_class.mod + Divides.op dvd ~> Divides.div_class.dvd + +INCOMPATIBILITY. Adaptions may be required in the following cases: + +a) User-defined constants using any of the names "plus", "minus", +"times", "less" or "less_eq". The standard syntax translations for +"+", "-" and "*" may go wrong. INCOMPATIBILITY: use more specific +names. + +b) Variables named "plus", "minus", "times", "less", "less_eq" +INCOMPATIBILITY: use more specific names. + +c) Permutative equations (e.g. "a + b = b + a") +Since the change of names also changes the order of terms, permutative +rewrite rules may get applied in a different order. Experience shows +that this is rarely the case (only two adaptions in the whole Isabelle +distribution). INCOMPATIBILITY: rewrite proofs + +d) ML code directly refering to constant names +This in general only affects hand-written proof tactics, simprocs and +so on. INCOMPATIBILITY: grep your sourcecode and replace names. +Consider using @{const_name} antiquotation. + +* New class "default" with associated constant "default". + +* Function "sgn" is now overloaded and available on int, real, complex +(and other numeric types), using class "sgn". Two possible defs of +sgn are given as equational assumptions in the classes sgn_if and +sgn_div_norm; ordered_idom now also inherits from sgn_if. +INCOMPATIBILITY. + +* Locale "partial_order" now unified with class "order" (cf. theory +Orderings), added parameter "less". INCOMPATIBILITY. + +* Renamings in classes "order" and "linorder": facts "refl", "trans" and +"cases" to "order_refl", "order_trans" and "linorder_cases", to avoid +clashes with HOL "refl" and "trans". INCOMPATIBILITY. + +* Classes "order" and "linorder": potential INCOMPATIBILITY due to +changed order of proof goals in instance proofs. + +* The transitivity reasoner for partial and linear orders is set up +for classes "order" and "linorder". Instances of the reasoner are available +in all contexts importing or interpreting the corresponding locales. +Method "order" invokes the reasoner separately; the reasoner +is also integrated with the Simplifier as a solver. Diagnostic +command 'print_orders' shows the available instances of the reasoner +in the current context. + +* Localized monotonicity predicate in theory "Orderings"; integrated +lemmas max_of_mono and min_of_mono with this predicate. +INCOMPATIBILITY. + +* Formulation of theorem "dense" changed slightly due to integration +with new class dense_linear_order. + +* Uniform lattice theory development in HOL. + + constants "meet" and "join" now named "inf" and "sup" + constant "Meet" now named "Inf" + + classes "meet_semilorder" and "join_semilorder" now named + "lower_semilattice" and "upper_semilattice" + class "lorder" now named "lattice" + class "comp_lat" now named "complete_lattice" + + Instantiation of lattice classes allows explicit definitions + for "inf" and "sup" operations (or "Inf" and "Sup" for complete lattices). + + INCOMPATIBILITY. Theorem renames: + + meet_left_le ~> inf_le1 + meet_right_le ~> inf_le2 + join_left_le ~> sup_ge1 + join_right_le ~> sup_ge2 + meet_join_le ~> inf_sup_ord + le_meetI ~> le_infI + join_leI ~> le_supI + le_meet ~> le_inf_iff + le_join ~> ge_sup_conv + meet_idempotent ~> inf_idem + join_idempotent ~> sup_idem + meet_comm ~> inf_commute + join_comm ~> sup_commute + meet_leI1 ~> le_infI1 + meet_leI2 ~> le_infI2 + le_joinI1 ~> le_supI1 + le_joinI2 ~> le_supI2 + meet_assoc ~> inf_assoc + join_assoc ~> sup_assoc + meet_left_comm ~> inf_left_commute + meet_left_idempotent ~> inf_left_idem + join_left_comm ~> sup_left_commute + join_left_idempotent ~> sup_left_idem + meet_aci ~> inf_aci + join_aci ~> sup_aci + le_def_meet ~> le_iff_inf + le_def_join ~> le_iff_sup + join_absorp2 ~> sup_absorb2 + join_absorp1 ~> sup_absorb1 + meet_absorp1 ~> inf_absorb1 + meet_absorp2 ~> inf_absorb2 + meet_join_absorp ~> inf_sup_absorb + join_meet_absorp ~> sup_inf_absorb + distrib_join_le ~> distrib_sup_le + distrib_meet_le ~> distrib_inf_le + + add_meet_distrib_left ~> add_inf_distrib_left + add_join_distrib_left ~> add_sup_distrib_left + is_join_neg_meet ~> is_join_neg_inf + is_meet_neg_join ~> is_meet_neg_sup + add_meet_distrib_right ~> add_inf_distrib_right + add_join_distrib_right ~> add_sup_distrib_right + add_meet_join_distribs ~> add_sup_inf_distribs + join_eq_neg_meet ~> sup_eq_neg_inf + meet_eq_neg_join ~> inf_eq_neg_sup + add_eq_meet_join ~> add_eq_inf_sup + meet_0_imp_0 ~> inf_0_imp_0 + join_0_imp_0 ~> sup_0_imp_0 + meet_0_eq_0 ~> inf_0_eq_0 + join_0_eq_0 ~> sup_0_eq_0 + neg_meet_eq_join ~> neg_inf_eq_sup + neg_join_eq_meet ~> neg_sup_eq_inf + join_eq_if ~> sup_eq_if + + mono_meet ~> mono_inf + mono_join ~> mono_sup + meet_bool_eq ~> inf_bool_eq + join_bool_eq ~> sup_bool_eq + meet_fun_eq ~> inf_fun_eq + join_fun_eq ~> sup_fun_eq + meet_set_eq ~> inf_set_eq + join_set_eq ~> sup_set_eq + meet1_iff ~> inf1_iff + meet2_iff ~> inf2_iff + meet1I ~> inf1I + meet2I ~> inf2I + meet1D1 ~> inf1D1 + meet2D1 ~> inf2D1 + meet1D2 ~> inf1D2 + meet2D2 ~> inf2D2 + meet1E ~> inf1E + meet2E ~> inf2E + join1_iff ~> sup1_iff + join2_iff ~> sup2_iff + join1I1 ~> sup1I1 + join2I1 ~> sup2I1 + join1I1 ~> sup1I1 + join2I2 ~> sup1I2 + join1CI ~> sup1CI + join2CI ~> sup2CI + join1E ~> sup1E + join2E ~> sup2E + + is_meet_Meet ~> is_meet_Inf + Meet_bool_def ~> Inf_bool_def + Meet_fun_def ~> Inf_fun_def + Meet_greatest ~> Inf_greatest + Meet_lower ~> Inf_lower + Meet_set_def ~> Inf_set_def + + Sup_def ~> Sup_Inf + Sup_bool_eq ~> Sup_bool_def + Sup_fun_eq ~> Sup_fun_def + Sup_set_eq ~> Sup_set_def + + listsp_meetI ~> listsp_infI + listsp_meet_eq ~> listsp_inf_eq + + meet_min ~> inf_min + join_max ~> sup_max + +* Added syntactic class "size"; overloaded constant "size" now has +type "'a::size ==> bool" + +* Internal reorganisation of `size' of datatypes: size theorems +"foo.size" are no longer subsumed by "foo.simps" (but are still +simplification rules by default!); theorems "prod.size" now named +"*.size". + +* Class "div" now inherits from class "times" rather than "type". +INCOMPATIBILITY. + +* HOL/Finite_Set: "name-space" locales Lattice, Distrib_lattice, +Linorder etc. have disappeared; operations defined in terms of +fold_set now are named Inf_fin, Sup_fin. INCOMPATIBILITY. + +* HOL/Nat: neq0_conv no longer declared as iff. INCOMPATIBILITY. + +* HOL-Word: New extensive library and type for generic, fixed size +machine words, with arithmetic, bit-wise, shifting and rotating +operations, reflection into int, nat, and bool lists, automation for +linear arithmetic (by automatic reflection into nat or int), including +lemmas on overflow and monotonicity. Instantiated to all appropriate +arithmetic type classes, supporting automatic simplification of +numerals on all operations. + +* Library/Boolean_Algebra: locales for abstract boolean algebras. + +* Library/Numeral_Type: numbers as types, e.g. TYPE(32). + +* Code generator library theories: + - Code_Integer represents HOL integers by big integer literals in target + languages. + - Code_Char represents HOL characters by character literals in target + languages. + - Code_Char_chr like Code_Char, but also offers treatment of character + codes; includes Code_Integer. + - Executable_Set allows to generate code for finite sets using lists. + - Executable_Rat implements rational numbers as triples (sign, enumerator, + denominator). + - Executable_Real implements a subset of real numbers, namly those + representable by rational numbers. + - Efficient_Nat implements natural numbers by integers, which in general will + result in higher efficency; pattern matching with 0/Suc is eliminated; + includes Code_Integer. + - Code_Index provides an additional datatype index which is mapped to + target-language built-in integers. + - Code_Message provides an additional datatype message_string which is isomorphic to + strings; messages are mapped to target-language strings. + +* New package for inductive predicates + + An n-ary predicate p with m parameters z_1, ..., z_m can now be defined via + + inductive + p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" + for z_1 :: U_1 and ... and z_n :: U_m + where + rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n" + | ... + + with full support for type-inference, rather than + + consts s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" + + abbreviation p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" + where "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m" + + inductive "s z_1 ... z_m" + intros + rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m" + ... + + For backward compatibility, there is a wrapper allowing inductive + sets to be defined with the new package via + + inductive_set + s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" + for z_1 :: U_1 and ... and z_n :: U_m + where + rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m" + | ... + + or + + inductive_set + s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" + and p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" + for z_1 :: U_1 and ... and z_n :: U_m + where + "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m" + | rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n" + | ... + + if the additional syntax "p ..." is required. + + Numerous examples can be found in the subdirectories src/HOL/Auth, + src/HOL/Bali, src/HOL/Induct, and src/HOL/MicroJava. + + INCOMPATIBILITIES: + + - Since declaration and definition of inductive sets or predicates + is no longer separated, abbreviations involving the newly + introduced sets or predicates must be specified together with the + introduction rules after the 'where' keyword (see above), rather + than before the actual inductive definition. + + - The variables in induction and elimination rules are now + quantified in the order of their occurrence in the introduction + rules, rather than in alphabetical order. Since this may break + some proofs, these proofs either have to be repaired, e.g. by + reordering the variables a_i_1 ... a_i_{k_i} in Isar 'case' + statements of the form + + case (rule_i a_i_1 ... a_i_{k_i}) + + or the old order of quantification has to be restored by explicitly adding + meta-level quantifiers in the introduction rules, i.e. + + | rule_i: "!!a_i_1 ... a_i_{k_i}. ... ==> p z_1 ... z_m t_i_1 ... t_i_n" + + - The format of the elimination rules is now + + p z_1 ... z_m x_1 ... x_n ==> + (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P) + ==> ... ==> P + + for predicates and + + (x_1, ..., x_n) : s z_1 ... z_m ==> + (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P) + ==> ... ==> P + + for sets rather than + + x : s z_1 ... z_m ==> + (!!a_1_1 ... a_1_{k_1}. x = (t_1_1, ..., t_1_n) ==> ... ==> P) + ==> ... ==> P + + This may require terms in goals to be expanded to n-tuples + (e.g. using case_tac or simplification with the split_paired_all + rule) before the above elimination rule is applicable. + + - The elimination or case analysis rules for (mutually) inductive + sets or predicates are now called "p_1.cases" ... "p_k.cases". The + list of rules "p_1_..._p_k.elims" is no longer available. + +* New package "function"/"fun" for general recursive functions, +supporting mutual and nested recursion, definitions in local contexts, +more general pattern matching and partiality. See HOL/ex/Fundefs.thy +for small examples, and the separate tutorial on the function +package. The old recdef "package" is still available as before, but +users are encouraged to use the new package. + +* Method "lexicographic_order" automatically synthesizes termination +relations as lexicographic combinations of size measures. + +* Case-expressions allow arbitrary constructor-patterns (including +"_") and take their order into account, like in functional +programming. Internally, this is translated into nested +case-expressions; missing cases are added and mapped to the predefined +constant "undefined". In complicated cases printing may no longer show +the original input but the internal form. Lambda-abstractions allow +the same form of pattern matching: "% pat1 => e1 | ..." is an +abbreviation for "%x. case x of pat1 => e1 | ..." where x is a new +variable. + +* IntDef: The constant "int :: nat => int" has been removed; now "int" +is an abbreviation for "of_nat :: nat => int". The simplification +rules for "of_nat" have been changed to work like "int" did +previously. Potential INCOMPATIBILITY: + - "of_nat (Suc m)" simplifies to "1 + of_nat m" instead of "of_nat m + 1" + - of_nat_diff and of_nat_mult are no longer default simp rules + +* Method "algebra" solves polynomial equations over (semi)rings using +Groebner bases. The (semi)ring structure is defined by locales and the +tool setup depends on that generic context. Installing the method for +a specific type involves instantiating the locale and possibly adding +declarations for computation on the coefficients. The method is +already instantiated for natural numbers and for the axiomatic class +of idoms with numerals. See also the paper by Chaieb and Wenzel at +CALCULEMUS 2007 for the general principles underlying this +architecture of context-aware proof-tools. + +* Method "ferrack" implements quantifier elimination over +special-purpose dense linear orders using locales (analogous to +"algebra"). The method is already installed for class +{ordered_field,recpower,number_ring} which subsumes real, hyperreal, +rat, etc. + +* Former constant "List.op @" now named "List.append". Use ML +antiquotations @{const_name List.append} or @{term " ... @ ... "} to +circumvent possible incompatibilities when working on ML level. + +* primrec: missing cases mapped to "undefined" instead of "arbitrary". + +* New function listsum :: 'a list => 'a for arbitrary monoids. +Special syntax: "SUM x <- xs. f x" (and latex variants) + +* New syntax for Haskell-like list comprehension (input only), eg. +[(x,y). x <- xs, y <- ys, x ~= y], see also src/HOL/List.thy. + +* The special syntax for function "filter" has changed from [x : +xs. P] to [x <- xs. P] to avoid an ambiguity caused by list +comprehension syntax, and for uniformity. INCOMPATIBILITY. + +* [a..b] is now defined for arbitrary linear orders. It used to be +defined on nat only, as an abbreviation for [a.. B" for equality on bool (with priority +25 like -->); output depends on the "iff" print_mode, the default is +"A = B" (with priority 50). + +* Relations less (<) and less_eq (<=) are also available on type bool. +Modified syntax to disallow nesting without explicit parentheses, +e.g. "(x < y) < z" or "x < (y < z)", but NOT "x < y < z". Potential +INCOMPATIBILITY. + +* "LEAST x:A. P" expands to "LEAST x. x:A & P" (input only). + +* Relation composition operator "op O" now has precedence 75 and binds +stronger than union and intersection. INCOMPATIBILITY. + +* The old set interval syntax "{m..n(}" (and relatives) has been +removed. Use "{m.. ==> False", equivalences +(i.e. "=" on type bool) are handled, variable names of the form +"lit_" are no longer reserved, significant speedup. + +* Methods "sat" and "satx" can now replay MiniSat proof traces. +zChaff is still supported as well. + +* 'inductive' and 'datatype': provide projections of mutual rules, +bundled as foo_bar.inducts; + +* Library: moved theories Parity, GCD, Binomial, Infinite_Set to +Library. + +* Library: moved theory Accessible_Part to main HOL. + +* Library: added theory Coinductive_List of potentially infinite lists +as greatest fixed-point. + +* Library: added theory AssocList which implements (finite) maps as +association lists. + +* Method "evaluation" solves goals (i.e. a boolean expression) +efficiently by compiling it to ML. The goal is "proved" (via an +oracle) if it evaluates to True. + +* Linear arithmetic now splits certain operators (e.g. min, max, abs) +also when invoked by the simplifier. This results in the Simplifier +being more powerful on arithmetic goals. INCOMPATIBILITY. +Configuration option fast_arith_split_limit=0 recovers the old +behavior. + +* Support for hex (0x20) and binary (0b1001) numerals. + +* New method: reify eqs (t), where eqs are equations for an +interpretation I :: 'a list => 'b => 'c and t::'c is an optional +parameter, computes a term s::'b and a list xs::'a list and proves the +theorem I xs s = t. This is also known as reification or quoting. The +resulting theorem is applied to the subgoal to substitute t with I xs +s. If t is omitted, the subgoal itself is reified. + +* New method: reflection corr_thm eqs (t). The parameters eqs and (t) +are as explained above. corr_thm is a theorem for I vs (f t) = I vs t, +where f is supposed to be a computable function (in the sense of code +generattion). The method uses reify to compute s and xs as above then +applies corr_thm and uses normalization by evaluation to "prove" f s = +r and finally gets the theorem t = r, which is again applied to the +subgoal. An Example is available in src/HOL/ex/ReflectionEx.thy. + +* Reflection: Automatic reification now handels binding, an example is +available in src/HOL/ex/ReflectionEx.thy + +* HOL-Statespace: ``State Spaces: The Locale Way'' introduces a +command 'statespace' that is similar to 'record', but introduces an +abstract specification based on the locale infrastructure instead of +HOL types. This leads to extra flexibility in composing state spaces, +in particular multiple inheritance and renaming of components. + + +*** HOL-Complex *** + +* Hyperreal: Functions root and sqrt are now defined on negative real +inputs so that root n (- x) = - root n x and sqrt (- x) = - sqrt x. +Nonnegativity side conditions have been removed from many lemmas, so +that more subgoals may now be solved by simplification; potential +INCOMPATIBILITY. + +* Real: new type classes formalize real normed vector spaces and +algebras, using new overloaded constants scaleR :: real => 'a => 'a +and norm :: 'a => real. + +* Real: constant of_real :: real => 'a::real_algebra_1 injects from +reals into other types. The overloaded constant Reals :: 'a set is now +defined as range of_real; potential INCOMPATIBILITY. + +* Real: proper support for ML code generation, including 'quickcheck'. +Reals are implemented as arbitrary precision rationals. + +* Hyperreal: Several constants that previously worked only for the +reals have been generalized, so they now work over arbitrary vector +spaces. Type annotations may need to be added in some cases; potential +INCOMPATIBILITY. + + Infinitesimal :: ('a::real_normed_vector) star set + HFinite :: ('a::real_normed_vector) star set + HInfinite :: ('a::real_normed_vector) star set + approx :: ('a::real_normed_vector) star => 'a star => bool + monad :: ('a::real_normed_vector) star => 'a star set + galaxy :: ('a::real_normed_vector) star => 'a star set + (NS)LIMSEQ :: [nat => 'a::real_normed_vector, 'a] => bool + (NS)convergent :: (nat => 'a::real_normed_vector) => bool + (NS)Bseq :: (nat => 'a::real_normed_vector) => bool + (NS)Cauchy :: (nat => 'a::real_normed_vector) => bool + (NS)LIM :: ['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool + is(NS)Cont :: ['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool + deriv :: ['a::real_normed_field => 'a, 'a, 'a] => bool + sgn :: 'a::real_normed_vector => 'a + exp :: 'a::{recpower,real_normed_field,banach} => 'a + +* Complex: Some complex-specific constants are now abbreviations for +overloaded ones: complex_of_real = of_real, cmod = norm, hcmod = +hnorm. Other constants have been entirely removed in favor of the +polymorphic versions (INCOMPATIBILITY): + + approx <-- capprox + HFinite <-- CFinite + HInfinite <-- CInfinite + Infinitesimal <-- CInfinitesimal + monad <-- cmonad + galaxy <-- cgalaxy + (NS)LIM <-- (NS)CLIM, (NS)CRLIM + is(NS)Cont <-- is(NS)Contc, is(NS)contCR + (ns)deriv <-- (ns)cderiv + + +*** HOL-Algebra *** + +* Formalisation of ideals and the quotient construction over rings. + +* Order and lattice theory no longer based on records. +INCOMPATIBILITY. + +* Renamed lemmas least_carrier -> least_closed and greatest_carrier -> +greatest_closed. INCOMPATIBILITY. + +* Method algebra is now set up via an attribute. For examples see +Ring.thy. INCOMPATIBILITY: the method is now weaker on combinations +of algebraic structures. + +* Renamed theory CRing to Ring. + + +*** HOL-Nominal *** + +* Substantial, yet incomplete support for nominal datatypes (binding +structures) based on HOL-Nominal logic. See src/HOL/Nominal and +src/HOL/Nominal/Examples. Prospective users should consult +http://isabelle.in.tum.de/nominal/ + + +*** ML *** + +* ML basics: just one true type int, which coincides with IntInf.int +(even on SML/NJ). + +* ML within Isar: antiquotations allow to embed statically-checked +formal entities in the source, referring to the context available at +compile-time. For example: + +ML {* @{sort "{zero,one}"} *} +ML {* @{typ "'a => 'b"} *} +ML {* @{term "%x. x"} *} +ML {* @{prop "x == y"} *} +ML {* @{ctyp "'a => 'b"} *} +ML {* @{cterm "%x. x"} *} +ML {* @{cprop "x == y"} *} +ML {* @{thm asm_rl} *} +ML {* @{thms asm_rl} *} +ML {* @{type_name c} *} +ML {* @{type_syntax c} *} +ML {* @{const_name c} *} +ML {* @{const_syntax c} *} +ML {* @{context} *} +ML {* @{theory} *} +ML {* @{theory Pure} *} +ML {* @{theory_ref} *} +ML {* @{theory_ref Pure} *} +ML {* @{simpset} *} +ML {* @{claset} *} +ML {* @{clasimpset} *} + +The same works for sources being ``used'' within an Isar context. + +* ML in Isar: improved error reporting; extra verbosity with +ML_Context.trace enabled. + +* Pure/General/table.ML: the join operations now works via exceptions +DUP/SAME instead of type option. This is simpler in simple cases, and +admits slightly more efficient complex applications. + +* Pure: 'advanced' translation functions (parse_translation etc.) now +use Context.generic instead of just theory. + +* Pure: datatype Context.generic joins theory/Proof.context and +provides some facilities for code that works in either kind of +context, notably GenericDataFun for uniform theory and proof data. + +* Pure: simplified internal attribute type, which is now always +Context.generic * thm -> Context.generic * thm. Global (theory) vs. +local (Proof.context) attributes have been discontinued, while +minimizing code duplication. Thm.rule_attribute and +Thm.declaration_attribute build canonical attributes; see also structure +Context for further operations on Context.generic, notably +GenericDataFun. INCOMPATIBILITY, need to adapt attribute type +declarations and definitions. + +* Context data interfaces (Theory/Proof/GenericDataFun): removed +name/print, uninitialized data defaults to ad-hoc copy of empty value, +init only required for impure data. INCOMPATIBILITY: empty really need +to be empty (no dependencies on theory content!) + +* Pure/kernel: consts certification ignores sort constraints given in +signature declarations. (This information is not relevant to the +logic, but only for type inference.) SIGNIFICANT INTERNAL CHANGE, +potential INCOMPATIBILITY. + +* Pure: axiomatic type classes are now purely definitional, with +explicit proofs of class axioms and super class relations performed +internally. See Pure/axclass.ML for the main internal interfaces -- +notably AxClass.define_class supercedes AxClass.add_axclass, and +AxClass.axiomatize_class/classrel/arity supersede +Sign.add_classes/classrel/arities. + +* Pure/Isar: Args/Attrib parsers operate on Context.generic -- +global/local versions on theory vs. Proof.context have been +discontinued; Attrib.syntax and Method.syntax have been adapted +accordingly. INCOMPATIBILITY, need to adapt parser expressions for +attributes, methods, etc. + +* Pure: several functions of signature "... -> theory -> theory * ..." +have been reoriented to "... -> theory -> ... * theory" in order to +allow natural usage in combination with the ||>, ||>>, |-> and +fold_map combinators. + +* Pure: official theorem names (closed derivations) and additional +comments (tags) are now strictly separate. Name hints -- which are +maintained as tags -- may be attached any time without affecting the +derivation. + +* Pure: primitive rule lift_rule now takes goal cterm instead of an +actual goal state (thm). Use Thm.lift_rule (Thm.cprem_of st i) to +achieve the old behaviour. + +* Pure: the "Goal" constant is now called "prop", supporting a +slightly more general idea of ``protecting'' meta-level rule +statements. + +* Pure: Logic.(un)varify only works in a global context, which is now +enforced instead of silently assumed. INCOMPATIBILITY, may use +Logic.legacy_(un)varify as temporary workaround. + +* Pure: structure Name provides scalable operations for generating +internal variable names, notably Name.variants etc. This replaces +some popular functions from term.ML: + + Term.variant -> Name.variant + Term.variantlist -> Name.variant_list + Term.invent_names -> Name.invent_list + +Note that low-level renaming rarely occurs in new code -- operations +from structure Variable are used instead (see below). + +* Pure: structure Variable provides fundamental operations for proper +treatment of fixed/schematic variables in a context. For example, +Variable.import introduces fixes for schematics of given facts and +Variable.export reverses the effect (up to renaming) -- this replaces +various freeze_thaw operations. + +* Pure: structure Goal provides simple interfaces for +init/conclude/finish and tactical prove operations (replacing former +Tactic.prove). Goal.prove is the canonical way to prove results +within a given context; Goal.prove_global is a degraded version for +theory level goals, including a global Drule.standard. Note that +OldGoals.prove_goalw_cterm has long been obsolete, since it is +ill-behaved in a local proof context (e.g. with local fixes/assumes or +in a locale context). + +* Pure/Syntax: generic interfaces for parsing (Syntax.parse_term etc.) +and type checking (Syntax.check_term etc.), with common combinations +(Syntax.read_term etc.). These supersede former Sign.read_term etc. +which are considered legacy and await removal. + +* Pure/Syntax: generic interfaces for type unchecking +(Syntax.uncheck_terms etc.) and unparsing (Syntax.unparse_term etc.), +with common combinations (Syntax.pretty_term, Syntax.string_of_term +etc.). Former Sign.pretty_term, Sign.string_of_term etc. are still +available for convenience, but refer to the very same operations using +a mere theory instead of a full context. + +* Isar: simplified treatment of user-level errors, using exception +ERROR of string uniformly. Function error now merely raises ERROR, +without any side effect on output channels. The Isar toplevel takes +care of proper display of ERROR exceptions. ML code may use plain +handle/can/try; cat_error may be used to concatenate errors like this: + + ... handle ERROR msg => cat_error msg "..." + +Toplevel ML code (run directly or through the Isar toplevel) may be +embedded into the Isar toplevel with exception display/debug like +this: + + Isar.toplevel (fn () => ...) + +INCOMPATIBILITY, removed special transform_error facilities, removed +obsolete variants of user-level exceptions (ERROR_MESSAGE, +Context.PROOF, ProofContext.CONTEXT, Proof.STATE, ProofHistory.FAIL) +-- use plain ERROR instead. + +* Isar: theory setup now has type (theory -> theory), instead of a +list. INCOMPATIBILITY, may use #> to compose setup functions. + +* Isar: ML toplevel pretty printer for type Proof.context, subject to +ProofContext.debug/verbose flags. + +* Isar: Toplevel.theory_to_proof admits transactions that modify the +theory before entering a proof state. Transactions now always see a +quasi-functional intermediate checkpoint, both in interactive and +batch mode. + +* Isar: simplified interfaces for outer syntax. Renamed +OuterSyntax.add_keywords to OuterSyntax.keywords. Removed +OuterSyntax.add_parsers -- this functionality is now included in +OuterSyntax.command etc. INCOMPATIBILITY. + +* Simplifier: the simpset of a running simplification process now +contains a proof context (cf. Simplifier.the_context), which is the +very context that the initial simpset has been retrieved from (by +simpset_of/local_simpset_of). Consequently, all plug-in components +(solver, looper etc.) may depend on arbitrary proof data. + +* Simplifier.inherit_context inherits the proof context (plus the +local bounds) of the current simplification process; any simproc +etc. that calls the Simplifier recursively should do this! Removed +former Simplifier.inherit_bounds, which is already included here -- +INCOMPATIBILITY. Tools based on low-level rewriting may even have to +specify an explicit context using Simplifier.context/theory_context. + +* Simplifier/Classical Reasoner: more abstract interfaces +change_simpset/claset for modifying the simpset/claset reference of a +theory; raw versions simpset/claset_ref etc. have been discontinued -- +INCOMPATIBILITY. + +* Provers: more generic wrt. syntax of object-logics, avoid hardwired +"Trueprop" etc. + + +*** System *** + +* settings: the default heap location within ISABELLE_HOME_USER now +includes ISABELLE_IDENTIFIER. This simplifies use of multiple +Isabelle installations. + +* isabelle-process: option -S (secure mode) disables some critical +operations, notably runtime compilation and evaluation of ML source +code. + +* Basic Isabelle mode for jEdit, see Isabelle/lib/jedit/. + +* Support for parallel execution, using native multicore support of +Poly/ML 5.1. The theory loader exploits parallelism when processing +independent theories, according to the given theory header +specifications. The maximum number of worker threads is specified via +usedir option -M or the "max-threads" setting in Proof General. A +speedup factor of 1.5--3.5 can be expected on a 4-core machine, and up +to 6 on a 8-core machine. User-code needs to observe certain +guidelines for thread-safe programming, see appendix A in the Isar +Implementation manual. + + + +New in Isabelle2005 (October 2005) +---------------------------------- + +*** General *** + +* Theory headers: the new header syntax for Isar theories is + + theory + imports ... + uses ... + begin + +where the 'uses' part is optional. The previous syntax + + theory = + ... + : + +will disappear in the next release. Use isatool fixheaders to convert +existing theory files. Note that there is no change in ancient +non-Isar theories now, but these will disappear soon. + +* Theory loader: parent theories can now also be referred to via +relative and absolute paths. + +* Command 'find_theorems' searches for a list of criteria instead of a +list of constants. Known criteria are: intro, elim, dest, name:string, +simp:term, and any term. Criteria can be preceded by '-' to select +theorems that do not match. Intro, elim, dest select theorems that +match the current goal, name:s selects theorems whose fully qualified +name contain s, and simp:term selects all simplification rules whose +lhs match term. Any other term is interpreted as pattern and selects +all theorems matching the pattern. Available in ProofGeneral under +'ProofGeneral -> Find Theorems' or C-c C-f. Example: + + C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL." + +prints the last 100 theorems matching the pattern "(_::nat) + _ + _", +matching the current goal as introduction rule and not having "HOL." +in their name (i.e. not being defined in theory HOL). + +* Command 'thms_containing' has been discontinued in favour of +'find_theorems'; INCOMPATIBILITY. + +* Communication with Proof General is now 8bit clean, which means that +Unicode text in UTF-8 encoding may be used within theory texts (both +formal and informal parts). Cf. option -U of the Isabelle Proof +General interface. Here are some simple examples (cf. src/HOL/ex): + + http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html + http://isabelle.in.tum.de/library/HOL/ex/Chinese.html + +* Improved efficiency of the Simplifier and, to a lesser degree, the +Classical Reasoner. Typical big applications run around 2 times +faster. + + +*** Document preparation *** + +* Commands 'display_drafts' and 'print_drafts' perform simple output +of raw sources. Only those symbols that do not require additional +LaTeX packages (depending on comments in isabellesym.sty) are +displayed properly, everything else is left verbatim. isatool display +and isatool print are used as front ends (these are subject to the +DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively). + +* Command tags control specific markup of certain regions of text, +notably folding and hiding. Predefined tags include "theory" (for +theory begin and end), "proof" for proof commands, and "ML" for +commands involving ML code; the additional tags "visible" and +"invisible" are unused by default. Users may give explicit tag +specifications in the text, e.g. ''by %invisible (auto)''. The +interpretation of tags is determined by the LaTeX job during document +preparation: see option -V of isatool usedir, or options -n and -t of +isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag, +\isadroptag. + +Several document versions may be produced at the same time via isatool +usedir (the generated index.html will link all of them). Typical +specifications include ''-V document=theory,proof,ML'' to present +theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold +proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit +these parts without any formal replacement text. The Isabelle site +default settings produce ''document'' and ''outline'' versions as +specified above. + +* Several new antiquotations: + + @{term_type term} prints a term with its type annotated; + + @{typeof term} prints the type of a term; + + @{const const} is the same as @{term const}, but checks that the + argument is a known logical constant; + + @{term_style style term} and @{thm_style style thm} print a term or + theorem applying a "style" to it + + @{ML text} + +Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of +definitions, equations, inequations etc., 'concl' printing only the +conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19' +to print the specified premise. TermStyle.add_style provides an ML +interface for introducing further styles. See also the "LaTeX Sugar" +document practical applications. The ML antiquotation prints +type-checked ML expressions verbatim. + +* Markup commands 'chapter', 'section', 'subsection', 'subsubsection', +and 'text' support optional locale specification '(in loc)', which +specifies the default context for interpreting antiquotations. For +example: 'text (in lattice) {* @{thm inf_assoc}*}'. + +* Option 'locale=NAME' of antiquotations specifies an alternative +context interpreting the subsequent argument. For example: @{thm +[locale=lattice] inf_assoc}. + +* Proper output of proof terms (@{prf ...} and @{full_prf ...}) within +a proof context. + +* Proper output of antiquotations for theory commands involving a +proof context (such as 'locale' or 'theorem (in loc) ...'). + +* Delimiters of outer tokens (string etc.) now produce separate LaTeX +macros (\isachardoublequoteopen, isachardoublequoteclose etc.). + +* isatool usedir: new option -C (default true) controls whether option +-D should include a copy of the original document directory; -C false +prevents unwanted effects such as copying of administrative CVS data. + + +*** Pure *** + +* Considerably improved version of 'constdefs' command. Now performs +automatic type-inference of declared constants; additional support for +local structure declarations (cf. locales and HOL records), see also +isar-ref manual. Potential INCOMPATIBILITY: need to observe strictly +sequential dependencies of definitions within a single 'constdefs' +section; moreover, the declared name needs to be an identifier. If +all fails, consider to fall back on 'consts' and 'defs' separately. + +* Improved indexed syntax and implicit structures. First of all, +indexed syntax provides a notational device for subscripted +application, using the new syntax \<^bsub>term\<^esub> for arbitrary +expressions. Secondly, in a local context with structure +declarations, number indexes \<^sub>n or the empty index (default +number 1) refer to a certain fixed variable implicitly; option +show_structs controls printing of implicit structures. Typical +applications of these concepts involve record types and locales. + +* New command 'no_syntax' removes grammar declarations (and +translations) resulting from the given syntax specification, which is +interpreted in the same manner as for the 'syntax' command. + +* 'Advanced' translation functions (parse_translation etc.) may depend +on the signature of the theory context being presently used for +parsing/printing, see also isar-ref manual. + +* Improved 'oracle' command provides a type-safe interface to turn an +ML expression of type theory -> T -> term into a primitive rule of +type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle +is already included here); see also FOL/ex/IffExample.thy; +INCOMPATIBILITY. + +* axclass: name space prefix for class "c" is now "c_class" (was "c" +before); "cI" is no longer bound, use "c.intro" instead. +INCOMPATIBILITY. This change avoids clashes of fact bindings for +axclasses vs. locales. + +* Improved internal renaming of symbolic identifiers -- attach primes +instead of base 26 numbers. + +* New flag show_question_marks controls printing of leading question +marks in schematic variable names. + +* In schematic variable names, *any* symbol following \<^isub> or +\<^isup> is now treated as part of the base name. For example, the +following works without printing of awkward ".0" indexes: + + lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1" + by simp + +* Inner syntax includes (*(*nested*) comments*). + +* Pretty printer now supports unbreakable blocks, specified in mixfix +annotations as "(00...)". + +* Clear separation of logical types and nonterminals, where the latter +may only occur in 'syntax' specifications or type abbreviations. +Before that distinction was only partially implemented via type class +"logic" vs. "{}". Potential INCOMPATIBILITY in rare cases of improper +use of 'types'/'consts' instead of 'nonterminals'/'syntax'. Some very +exotic syntax specifications may require further adaption +(e.g. Cube/Cube.thy). + +* Removed obsolete type class "logic", use the top sort {} instead. +Note that non-logical types should be declared as 'nonterminals' +rather than 'types'. INCOMPATIBILITY for new object-logic +specifications. + +* Attributes 'induct' and 'cases': type or set names may now be +locally fixed variables as well. + +* Simplifier: can now control the depth to which conditional rewriting +is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth +Limit. + +* Simplifier: simplification procedures may now take the current +simpset into account (cf. Simplifier.simproc(_i) / mk_simproc +interface), which is very useful for calling the Simplifier +recursively. Minor INCOMPATIBILITY: the 'prems' argument of simprocs +is gone -- use prems_of_ss on the simpset instead. Moreover, the +low-level mk_simproc no longer applies Logic.varify internally, to +allow for use in a context of fixed variables. + +* thin_tac now works even if the assumption being deleted contains !! +or ==>. More generally, erule now works even if the major premise of +the elimination rule contains !! or ==>. + +* Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY. + +* Reorganized bootstrapping of the Pure theories; CPure is now derived +from Pure, which contains all common declarations already. Both +theories are defined via plain Isabelle/Isar .thy files. +INCOMPATIBILITY: elements of CPure (such as the CPure.intro / +CPure.elim / CPure.dest attributes) now appear in the Pure name space; +use isatool fixcpure to adapt your theory and ML sources. + +* New syntax 'name(i-j, i-, i, ...)' for referring to specific +selections of theorems in named facts via index ranges. + +* 'print_theorems': in theory mode, really print the difference +wrt. the last state (works for interactive theory development only), +in proof mode print all local facts (cf. 'print_facts'); + +* 'hide': option '(open)' hides only base names. + +* More efficient treatment of intermediate checkpoints in interactive +theory development. + +* Code generator is now invoked via code_module (incremental code +generation) and code_library (modular code generation, ML structures +for each theory). INCOMPATIBILITY: new keywords 'file' and 'contains' +must be quoted when used as identifiers. + +* New 'value' command for reading, evaluating and printing terms using +the code generator. INCOMPATIBILITY: command keyword 'value' must be +quoted when used as identifier. + + +*** Locales *** + +* New commands for the interpretation of locale expressions in +theories (1), locales (2) and proof contexts (3). These generate +proof obligations from the expression specification. After the +obligations have been discharged, theorems of the expression are added +to the theory, target locale or proof context. The synopsis of the +commands is a follows: + + (1) interpretation expr inst + (2) interpretation target < expr + (3) interpret expr inst + +Interpretation in theories and proof contexts require a parameter +instantiation of terms from the current context. This is applied to +specifications and theorems of the interpreted expression. +Interpretation in locales only permits parameter renaming through the +locale expression. Interpretation is smart in that interpretations +that are active already do not occur in proof obligations, neither are +instantiated theorems stored in duplicate. Use 'print_interps' to +inspect active interpretations of a particular locale. For details, +see the Isar Reference manual. Examples can be found in +HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy. + +INCOMPATIBILITY: former 'instantiate' has been withdrawn, use +'interpret' instead. + +* New context element 'constrains' for adding type constraints to +parameters. + +* Context expressions: renaming of parameters with syntax +redeclaration. + +* Locale declaration: 'includes' disallowed. + +* Proper static binding of attribute syntax -- i.e. types / terms / +facts mentioned as arguments are always those of the locale definition +context, independently of the context of later invocations. Moreover, +locale operations (renaming and type / term instantiation) are applied +to attribute arguments as expected. + +INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of +actual attributes; rare situations may require Attrib.attribute to +embed those attributes into Attrib.src that lack concrete syntax. +Attribute implementations need to cooperate properly with the static +binding mechanism. Basic parsers Args.XXX_typ/term/prop and +Attrib.XXX_thm etc. already do the right thing without further +intervention. Only unusual applications -- such as "where" or "of" +(cf. src/Pure/Isar/attrib.ML), which process arguments depending both +on the context and the facts involved -- may have to assign parsed +values to argument tokens explicitly. + +* Changed parameter management in theorem generation for long goal +statements with 'includes'. INCOMPATIBILITY: produces a different +theorem statement in rare situations. + +* Locale inspection command 'print_locale' omits notes elements. Use +'print_locale!' to have them included in the output. + + +*** Provers *** + +* Provers/hypsubst.ML: improved version of the subst method, for +single-step rewriting: it now works in bound variable contexts. New is +'subst (asm)', for rewriting an assumption. INCOMPATIBILITY: may +rewrite a different subterm than the original subst method, which is +still available as 'simplesubst'. + +* Provers/quasi.ML: new transitivity reasoners for transitivity only +and quasi orders. + +* Provers/trancl.ML: new transitivity reasoner for transitive and +reflexive-transitive closure of relations. + +* Provers/blast.ML: new reference depth_limit to make blast's depth +limit (previously hard-coded with a value of 20) user-definable. + +* Provers/simplifier.ML has been moved to Pure, where Simplifier.setup +is peformed already. Object-logics merely need to finish their +initial simpset configuration as before. INCOMPATIBILITY. + + +*** HOL *** + +* Symbolic syntax of Hilbert Choice Operator is now as follows: + + syntax (epsilon) + "_Eps" :: "[pttrn, bool] => 'a" ("(3\_./ _)" [0, 10] 10) + +The symbol \ is displayed as the alternative epsilon of LaTeX +and x-symbol; use option '-m epsilon' to get it actually printed. +Moreover, the mathematically important symbolic identifier \ +becomes available as variable, constant etc. INCOMPATIBILITY, + +* "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x". +Similarly for all quantifiers: "ALL x > y" etc. The x-symbol for >= +is \. New transitivity rules have been added to HOL/Orderings.thy to +support corresponding Isar calculations. + +* "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\" +instead of ":". + +* theory SetInterval: changed the syntax for open intervals: + + Old New + {..n(} {.. {\1<\.\.} + \.\.\([^(}]*\)(} -> \.\.<\1} + +* Theory Commutative_Ring (in Library): method comm_ring for proving +equalities in commutative rings; method 'algebra' provides a generic +interface. + +* Theory Finite_Set: changed the syntax for 'setsum', summation over +finite sets: "setsum (%x. e) A", which used to be "\x:A. e", is +now either "SUM x:A. e" or "\x \ A. e". The bound variable can +be a tuple pattern. + +Some new syntax forms are available: + + "\x | P. e" for "setsum (%x. e) {x. P}" + "\x = a..b. e" for "setsum (%x. e) {a..b}" + "\x = a..x < k. e" for "setsum (%x. e) {..x < k. e" used to be based on a separate +function "Summation", which has been discontinued. + +* theory Finite_Set: in structured induction proofs, the insert case +is now 'case (insert x F)' instead of the old counterintuitive 'case +(insert F x)'. + +* The 'refute' command has been extended to support a much larger +fragment of HOL, including axiomatic type classes, constdefs and +typedefs, inductive datatypes and recursion. + +* New tactics 'sat' and 'satx' to prove propositional tautologies. +Requires zChaff with proof generation to be installed. See +HOL/ex/SAT_Examples.thy for examples. + +* Datatype induction via method 'induct' now preserves the name of the +induction variable. For example, when proving P(xs::'a list) by +induction on xs, the induction step is now P(xs) ==> P(a#xs) rather +than P(list) ==> P(a#list) as previously. Potential INCOMPATIBILITY +in unstructured proof scripts. + +* Reworked implementation of records. Improved scalability for +records with many fields, avoiding performance problems for type +inference. Records are no longer composed of nested field types, but +of nested extension types. Therefore the record type only grows linear +in the number of extensions and not in the number of fields. The +top-level (users) view on records is preserved. Potential +INCOMPATIBILITY only in strange cases, where the theory depends on the +old record representation. The type generated for a record is called +_ext_type. + +Flag record_quick_and_dirty_sensitive can be enabled to skip the +proofs triggered by a record definition or a simproc (if +quick_and_dirty is enabled). Definitions of large records can take +quite long. + +New simproc record_upd_simproc for simplification of multiple record +updates enabled by default. Moreover, trivial updates are also +removed: r(|x := x r|) = r. INCOMPATIBILITY: old proofs break +occasionally, since simplification is more powerful by default. + +* typedef: proper support for polymorphic sets, which contain extra +type-variables in the term. + +* Simplifier: automatically reasons about transitivity chains +involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics +provided by Provers/trancl.ML as additional solvers. INCOMPATIBILITY: +old proofs break occasionally as simplification may now solve more +goals than previously. + +* Simplifier: converts x <= y into x = y if assumption y <= x is +present. Works for all partial orders (class "order"), in particular +numbers and sets. For linear orders (e.g. numbers) it treats ~ x < y +just like y <= x. + +* Simplifier: new simproc for "let x = a in f x". If a is a free or +bound variable or a constant then the let is unfolded. Otherwise +first a is simplified to b, and then f b is simplified to g. If +possible we abstract b from g arriving at "let x = b in h x", +otherwise we unfold the let and arrive at g. The simproc can be +enabled/disabled by the reference use_let_simproc. Potential +INCOMPATIBILITY since simplification is more powerful by default. + +* Classical reasoning: the meson method now accepts theorems as arguments. + +* Prover support: pre-release of the Isabelle-ATP linkup, which runs background +jobs to provide advice on the provability of subgoals. + +* Theory OrderedGroup and Ring_and_Field: various additions and +improvements to faciliate calculations involving equalities and +inequalities. + +The following theorems have been eliminated or modified +(INCOMPATIBILITY): + + abs_eq now named abs_of_nonneg + abs_of_ge_0 now named abs_of_nonneg + abs_minus_eq now named abs_of_nonpos + imp_abs_id now named abs_of_nonneg + imp_abs_neg_id now named abs_of_nonpos + mult_pos now named mult_pos_pos + mult_pos_le now named mult_nonneg_nonneg + mult_pos_neg_le now named mult_nonneg_nonpos + mult_pos_neg2_le now named mult_nonneg_nonpos2 + mult_neg now named mult_neg_neg + mult_neg_le now named mult_nonpos_nonpos + +* The following lemmas in Ring_and_Field have been added to the simplifier: + + zero_le_square + not_square_less_zero + + The following lemmas have been deleted from Real/RealPow: + + realpow_zero_zero + realpow_two + realpow_less + zero_le_power + realpow_two_le + abs_realpow_two + realpow_two_abs + +* Theory Parity: added rules for simplifying exponents. + +* Theory List: + +The following theorems have been eliminated or modified +(INCOMPATIBILITY): + + list_all_Nil now named list_all.simps(1) + list_all_Cons now named list_all.simps(2) + list_all_conv now named list_all_iff + set_mem_eq now named mem_iff + +* Theories SetsAndFunctions and BigO (see HOL/Library) support +asymptotic "big O" calculations. See the notes in BigO.thy. + + +*** HOL-Complex *** + +* Theory RealDef: better support for embedding natural numbers and +integers in the reals. + +The following theorems have been eliminated or modified +(INCOMPATIBILITY): + + exp_ge_add_one_self now requires no hypotheses + real_of_int_add reversed direction of equality (use [symmetric]) + real_of_int_minus reversed direction of equality (use [symmetric]) + real_of_int_diff reversed direction of equality (use [symmetric]) + real_of_int_mult reversed direction of equality (use [symmetric]) + +* Theory RComplete: expanded support for floor and ceiling functions. + +* Theory Ln is new, with properties of the natural logarithm + +* Hyperreal: There is a new type constructor "star" for making +nonstandard types. The old type names are now type synonyms: + + hypreal = real star + hypnat = nat star + hcomplex = complex star + +* Hyperreal: Many groups of similarly-defined constants have been +replaced by polymorphic versions (INCOMPATIBILITY): + + star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex + + starset <-- starsetNat, starsetC + *s* <-- *sNat*, *sc* + starset_n <-- starsetNat_n, starsetC_n + *sn* <-- *sNatn*, *scn* + InternalSets <-- InternalNatSets, InternalCSets + + starfun <-- starfun{Nat,Nat2,C,RC,CR} + *f* <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR* + starfun_n <-- starfun{Nat,Nat2,C,RC,CR}_n + *fn* <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn* + InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs + +* Hyperreal: Many type-specific theorems have been removed in favor of +theorems specific to various axiomatic type classes (INCOMPATIBILITY): + + add_commute <-- {hypreal,hypnat,hcomplex}_add_commute + add_assoc <-- {hypreal,hypnat,hcomplex}_add_assocs + OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left + OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right + right_minus <-- hypreal_add_minus + left_minus <-- {hypreal,hcomplex}_add_minus_left + mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute + mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc + mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left + mult_1_right <-- hcomplex_mult_one_right + mult_zero_left <-- hcomplex_mult_zero_left + left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib + right_distrib <-- hypnat_add_mult_distrib2 + zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one + right_inverse <-- hypreal_mult_inverse + left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left + order_refl <-- {hypreal,hypnat}_le_refl + order_trans <-- {hypreal,hypnat}_le_trans + order_antisym <-- {hypreal,hypnat}_le_anti_sym + order_less_le <-- {hypreal,hypnat}_less_le + linorder_linear <-- {hypreal,hypnat}_le_linear + add_left_mono <-- {hypreal,hypnat}_add_left_mono + mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2 + add_nonneg_nonneg <-- hypreal_le_add_order + +* Hyperreal: Separate theorems having to do with type-specific +versions of constants have been merged into theorems that apply to the +new polymorphic constants (INCOMPATIBILITY): + + STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set + STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set + STAR_Un <-- {STAR,NatStar,STARC}_Un + STAR_Int <-- {STAR,NatStar,STARC}_Int + STAR_Compl <-- {STAR,NatStar,STARC}_Compl + STAR_subset <-- {STAR,NatStar,STARC}_subset + STAR_mem <-- {STAR,NatStar,STARC}_mem + STAR_mem_Compl <-- {STAR,STARC}_mem_Compl + STAR_diff <-- {STAR,STARC}_diff + STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real, + STARC_hcomplex_of_complex}_image_subset + starset_n_Un <-- starset{Nat,C}_n_Un + starset_n_Int <-- starset{Nat,C}_n_Int + starset_n_Compl <-- starset{Nat,C}_n_Compl + starset_n_diff <-- starset{Nat,C}_n_diff + InternalSets_Un <-- Internal{Nat,C}Sets_Un + InternalSets_Int <-- Internal{Nat,C}Sets_Int + InternalSets_Compl <-- Internal{Nat,C}Sets_Compl + InternalSets_diff <-- Internal{Nat,C}Sets_diff + InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff + InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n + starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq + starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C} + starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR} + starfun <-- starfun{Nat,Nat2,C,RC,CR} + starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult + starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add + starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus + starfun_diff <-- starfun{C,RC,CR}_diff + starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o + starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2 + starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun + starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse + starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq + starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff + starfun_Id <-- starfunC_Id + starfun_approx <-- starfun{Nat,CR}_approx + starfun_capprox <-- starfun{C,RC}_capprox + starfun_abs <-- starfunNat_rabs + starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel + starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2 + starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox + starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox + starfun_add_capprox <-- starfun{C,RC}_add_capprox + starfun_add_approx <-- starfunCR_add_approx + starfun_inverse_inverse <-- starfunC_inverse_inverse + starfun_divide <-- starfun{C,CR,RC}_divide + starfun_n <-- starfun{Nat,C}_n + starfun_n_mult <-- starfun{Nat,C}_n_mult + starfun_n_add <-- starfun{Nat,C}_n_add + starfun_n_add_minus <-- starfunNat_n_add_minus + starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun + starfun_n_minus <-- starfun{Nat,C}_n_minus + starfun_n_eq <-- starfun{Nat,C}_n_eq + + star_n_add <-- {hypreal,hypnat,hcomplex}_add + star_n_minus <-- {hypreal,hcomplex}_minus + star_n_diff <-- {hypreal,hcomplex}_diff + star_n_mult <-- {hypreal,hcomplex}_mult + star_n_inverse <-- {hypreal,hcomplex}_inverse + star_n_le <-- {hypreal,hypnat}_le + star_n_less <-- {hypreal,hypnat}_less + star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num + star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num + star_n_abs <-- hypreal_hrabs + star_n_divide <-- hcomplex_divide + + star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add + star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus + star_of_diff <-- hypreal_of_real_diff + star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult + star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one + star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero + star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff + star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff + star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff + star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse + star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide + star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat + star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int + star_of_number_of <-- {hypreal,hcomplex}_number_of + star_of_number_less <-- number_of_less_hypreal_of_real_iff + star_of_number_le <-- number_of_le_hypreal_of_real_iff + star_of_eq_number <-- hypreal_of_real_eq_number_of_iff + star_of_less_number <-- hypreal_of_real_less_number_of_iff + star_of_le_number <-- hypreal_of_real_le_number_of_iff + star_of_power <-- hypreal_of_real_power + star_of_eq_0 <-- hcomplex_of_complex_zero_iff + +* Hyperreal: new method "transfer" that implements the transfer +principle of nonstandard analysis. With a subgoal that mentions +nonstandard types like "'a star", the command "apply transfer" +replaces it with an equivalent one that mentions only standard types. +To be successful, all free variables must have standard types; non- +standard variables must have explicit universal quantifiers. + +* Hyperreal: A theory of Taylor series. + + +*** HOLCF *** + +* Discontinued special version of 'constdefs' (which used to support +continuous functions) in favor of the general Pure one with full +type-inference. + +* New simplification procedure for solving continuity conditions; it +is much faster on terms with many nested lambda abstractions (cubic +instead of exponential time). + +* New syntax for domain package: selector names are now optional. +Parentheses should be omitted unless argument is lazy, for example: + + domain 'a stream = cons "'a" (lazy "'a stream") + +* New command 'fixrec' for defining recursive functions with pattern +matching; defining multiple functions with mutual recursion is also +supported. Patterns may include the constants cpair, spair, up, sinl, +sinr, or any data constructor defined by the domain package. The given +equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for +syntax and examples. + +* New commands 'cpodef' and 'pcpodef' for defining predicate subtypes +of cpo and pcpo types. Syntax is exactly like the 'typedef' command, +but the proof obligation additionally includes an admissibility +requirement. The packages generate instances of class cpo or pcpo, +with continuity and strictness theorems for Rep and Abs. + +* HOLCF: Many theorems have been renamed according to a more standard naming +scheme (INCOMPATIBILITY): + + foo_inject: "foo$x = foo$y ==> x = y" + foo_eq: "(foo$x = foo$y) = (x = y)" + foo_less: "(foo$x << foo$y) = (x << y)" + foo_strict: "foo$UU = UU" + foo_defined: "... ==> foo$x ~= UU" + foo_defined_iff: "(foo$x = UU) = (x = UU)" + + +*** ZF *** + +* ZF/ex: theories Group and Ring provide examples in abstract algebra, +including the First Isomorphism Theorem (on quotienting by the kernel +of a homomorphism). + +* ZF/Simplifier: install second copy of type solver that actually +makes use of TC rules declared to Isar proof contexts (or locales); +the old version is still required for ML proof scripts. + + +*** Cube *** + +* Converted to Isar theory format; use locales instead of axiomatic +theories. + + +*** ML *** + +* Pure/library.ML: added ##>, ##>>, #>> -- higher-order counterparts +for ||>, ||>>, |>>, + +* Pure/library.ML no longer defines its own option datatype, but uses +that of the SML basis, which has constructors NONE and SOME instead of +None and Some, as well as exception Option.Option instead of OPTION. +The functions the, if_none, is_some, is_none have been adapted +accordingly, while Option.map replaces apsome. + +* Pure/library.ML: the exception LIST has been given up in favour of +the standard exceptions Empty and Subscript, as well as +Library.UnequalLengths. Function like Library.hd and Library.tl are +superceded by the standard hd and tl functions etc. + +A number of basic list functions are no longer exported to the ML +toplevel, as they are variants of predefined functions. The following +suggests how one can translate existing code: + + rev_append xs ys = List.revAppend (xs, ys) + nth_elem (i, xs) = List.nth (xs, i) + last_elem xs = List.last xs + flat xss = List.concat xss + seq fs = List.app fs + partition P xs = List.partition P xs + mapfilter f xs = List.mapPartial f xs + +* Pure/library.ML: several combinators for linear functional +transformations, notably reverse application and composition: + + x |> f f #> g + (x, y) |-> f f #-> g + +* Pure/library.ML: introduced/changed precedence of infix operators: + + infix 1 |> |-> ||> ||>> |>> |>>> #> #->; + infix 2 ?; + infix 3 o oo ooo oooo; + infix 4 ~~ upto downto; + +Maybe INCOMPATIBILITY when any of those is used in conjunction with other +infix operators. + +* Pure/library.ML: natural list combinators fold, fold_rev, and +fold_map support linear functional transformations and nesting. For +example: + + fold f [x1, ..., xN] y = + y |> f x1 |> ... |> f xN + + (fold o fold) f [xs1, ..., xsN] y = + y |> fold f xs1 |> ... |> fold f xsN + + fold f [x1, ..., xN] = + f x1 #> ... #> f xN + + (fold o fold) f [xs1, ..., xsN] = + fold f xs1 #> ... #> fold f xsN + +* Pure/library.ML: the following selectors on type 'a option are +available: + + the: 'a option -> 'a (*partial*) + these: 'a option -> 'a where 'a = 'b list + the_default: 'a -> 'a option -> 'a + the_list: 'a option -> 'a list + +* Pure/General: structure AList (cf. Pure/General/alist.ML) provides +basic operations for association lists, following natural argument +order; moreover the explicit equality predicate passed here avoids +potentially expensive polymorphic runtime equality checks. +The old functions may be expressed as follows: + + assoc = uncurry (AList.lookup (op =)) + assocs = these oo AList.lookup (op =) + overwrite = uncurry (AList.update (op =)) o swap + +* Pure/General: structure AList (cf. Pure/General/alist.ML) provides + + val make: ('a -> 'b) -> 'a list -> ('a * 'b) list + val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list + +replacing make_keylist and keyfilter (occassionally used) +Naive rewrites: + + make_keylist = AList.make + keyfilter = AList.find (op =) + +* eq_fst and eq_snd now take explicit equality parameter, thus + avoiding eqtypes. Naive rewrites: + + eq_fst = eq_fst (op =) + eq_snd = eq_snd (op =) + +* Removed deprecated apl and apr (rarely used). + Naive rewrites: + + apl (n, op) =>>= curry op n + apr (op, m) =>>= fn n => op (n, m) + +* Pure/General: structure OrdList (cf. Pure/General/ord_list.ML) +provides a reasonably efficient light-weight implementation of sets as +lists. + +* Pure/General: generic tables (cf. Pure/General/table.ML) provide a +few new operations; existing lookup and update are now curried to +follow natural argument order (for use with fold etc.); +INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort. + +* Pure/General: output via the Isabelle channels of +writeln/warning/error etc. is now passed through Output.output, with a +hook for arbitrary transformations depending on the print_mode +(cf. Output.add_mode -- the first active mode that provides a output +function wins). Already formatted output may be embedded into further +text via Output.raw; the result of Pretty.string_of/str_of and derived +functions (string_of_term/cterm/thm etc.) is already marked raw to +accommodate easy composition of diagnostic messages etc. Programmers +rarely need to care about Output.output or Output.raw at all, with +some notable exceptions: Output.output is required when bypassing the +standard channels (writeln etc.), or in token translations to produce +properly formatted results; Output.raw is required when capturing +already output material that will eventually be presented to the user +a second time. For the default print mode, both Output.output and +Output.raw have no effect. + +* Pure/General: Output.time_accumulator NAME creates an operator ('a +-> 'b) -> 'a -> 'b to measure runtime and count invocations; the +cumulative results are displayed at the end of a batch session. + +* Pure/General: File.sysify_path and File.quote_sysify path have been +replaced by File.platform_path and File.shell_path (with appropriate +hooks). This provides a clean interface for unusual systems where the +internal and external process view of file names are different. + +* Pure: more efficient orders for basic syntactic entities: added +fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord +and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is +NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast +orders now -- potential INCOMPATIBILITY for code that depends on a +particular order for Symtab.keys, Symtab.dest, etc. (consider using +Library.sort_strings on result). + +* Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types, +fold_types traverse types/terms from left to right, observing natural +argument order. Supercedes previous foldl_XXX versions, add_frees, +add_vars etc. have been adapted as well: INCOMPATIBILITY. + +* Pure: name spaces have been refined, with significant changes of the +internal interfaces -- INCOMPATIBILITY. Renamed cond_extern(_table) +to extern(_table). The plain name entry path is superceded by a +general 'naming' context, which also includes the 'policy' to produce +a fully qualified name and external accesses of a fully qualified +name; NameSpace.extend is superceded by context dependent +Sign.declare_name. Several theory and proof context operations modify +the naming context. Especially note Theory.restore_naming and +ProofContext.restore_naming to get back to a sane state; note that +Theory.add_path is no longer sufficient to recover from +Theory.absolute_path in particular. + +* Pure: new flags short_names (default false) and unique_names +(default true) for controlling output of qualified names. If +short_names is set, names are printed unqualified. If unique_names is +reset, the name prefix is reduced to the minimum required to achieve +the original result when interning again, even if there is an overlap +with earlier declarations. + +* Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is +now 'extend', and 'merge' gets an additional Pretty.pp argument +(useful for printing error messages). INCOMPATIBILITY. + +* Pure: major reorganization of the theory context. Type Sign.sg and +Theory.theory are now identified, referring to the universal +Context.theory (see Pure/context.ML). Actual signature and theory +content is managed as theory data. The old code and interfaces were +spread over many files and structures; the new arrangement introduces +considerable INCOMPATIBILITY to gain more clarity: + + Context -- theory management operations (name, identity, inclusion, + parents, ancestors, merge, etc.), plus generic theory data; + + Sign -- logical signature and syntax operations (declaring consts, + types, etc.), plus certify/read for common entities; + + Theory -- logical theory operations (stating axioms, definitions, + oracles), plus a copy of logical signature operations (consts, + types, etc.); also a few basic management operations (Theory.copy, + Theory.merge, etc.) + +The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm +etc.) as well as the sign field in Thm.rep_thm etc. have been retained +for convenience -- they merely return the theory. + +* Pure: type Type.tsig is superceded by theory in most interfaces. + +* Pure: the Isar proof context type is already defined early in Pure +as Context.proof (note that ProofContext.context and Proof.context are +aliases, where the latter is the preferred name). This enables other +Isabelle components to refer to that type even before Isar is present. + +* Pure/sign/theory: discontinued named name spaces (i.e. classK, +typeK, constK, axiomK, oracleK), but provide explicit operations for +any of these kinds. For example, Sign.intern typeK is now +Sign.intern_type, Theory.hide_space Sign.typeK is now +Theory.hide_types. Also note that former +Theory.hide_classes/types/consts are now +Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions +internalize their arguments! INCOMPATIBILITY. + +* Pure: get_thm interface (of PureThy and ProofContext) expects +datatype thmref (with constructors Name and NameSelection) instead of +plain string -- INCOMPATIBILITY; + +* Pure: cases produced by proof methods specify options, where NONE +means to remove case bindings -- INCOMPATIBILITY in +(RAW_)METHOD_CASES. + +* Pure: the following operations retrieve axioms or theorems from a +theory node or theory hierarchy, respectively: + + Theory.axioms_of: theory -> (string * term) list + Theory.all_axioms_of: theory -> (string * term) list + PureThy.thms_of: theory -> (string * thm) list + PureThy.all_thms_of: theory -> (string * thm) list + +* Pure: print_tac now outputs the goal through the trace channel. + +* Isar toplevel: improved diagnostics, mostly for Poly/ML only. +Reference Toplevel.debug (default false) controls detailed printing +and tracing of low-level exceptions; Toplevel.profiling (default 0) +controls execution profiling -- set to 1 for time and 2 for space +(both increase the runtime). + +* Isar session: The initial use of ROOT.ML is now always timed, +i.e. the log will show the actual process times, in contrast to the +elapsed wall-clock time that the outer shell wrapper produces. + +* Simplifier: improved handling of bound variables (nameless +representation, avoid allocating new strings). Simprocs that invoke +the Simplifier recursively should use Simplifier.inherit_bounds to +avoid local name clashes. Failure to do so produces warnings +"Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds +for further details. + +* ML functions legacy_bindings and use_legacy_bindings produce ML fact +bindings for all theorems stored within a given theory; this may help +in porting non-Isar theories to Isar ones, while keeping ML proof +scripts for the time being. + +* ML operator HTML.with_charset specifies the charset begin used for +generated HTML files. For example: + + HTML.with_charset "utf-8" use_thy "Hebrew"; + HTML.with_charset "utf-8" use_thy "Chinese"; + + +*** System *** + +* Allow symlinks to all proper Isabelle executables (Isabelle, +isabelle, isatool etc.). + +* ISABELLE_DOC_FORMAT setting specifies preferred document format (for +isatool doc, isatool mkdir, display_drafts etc.). + +* isatool usedir: option -f allows specification of the ML file to be +used by Isabelle; default is ROOT.ML. + +* New isatool version outputs the version identifier of the Isabelle +distribution being used. + +* HOL: new isatool dimacs2hol converts files in DIMACS CNF format +(containing Boolean satisfiability problems) into Isabelle/HOL +theories. + + + +New in Isabelle2004 (April 2004) +-------------------------------- + +*** General *** + +* Provers/order.ML: new efficient reasoner for partial and linear orders. + Replaces linorder.ML. + +* Pure: Greek letters (except small lambda, \), as well as Gothic + (\...\\...\), calligraphic (\...\), and Euler + (\...\), are now considered normal letters, and can therefore + be used anywhere where an ASCII letter (a...zA...Z) has until + now. COMPATIBILITY: This obviously changes the parsing of some + terms, especially where a symbol has been used as a binder, say + '\x. ...', which is now a type error since \x will be parsed + as an identifier. Fix it by inserting a space around former + symbols. Call 'isatool fixgreek' to try to fix parsing errors in + existing theory and ML files. + +* Pure: Macintosh and Windows line-breaks are now allowed in theory files. + +* Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now + allowed in identifiers. Similar to Greek letters \<^isub> is now considered + a normal (but invisible) letter. For multiple letter subscripts repeat + \<^isub> like this: x\<^isub>1\<^isub>2. + +* Pure: There are now sub-/superscripts that can span more than one + character. Text between \<^bsub> and \<^esub> is set in subscript in + ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in + superscript. The new control characters are not identifier parts. + +* Pure: Control-symbols of the form \<^raw:...> will literally print the + content of "..." to the latex file instead of \isacntrl... . The "..." + may consist of any printable characters excluding the end bracket >. + +* Pure: Using new Isar command "finalconsts" (or the ML functions + Theory.add_finals or Theory.add_finals_i) it is now possible to + declare constants "final", which prevents their being given a definition + later. It is useful for constants whose behaviour is fixed axiomatically + rather than definitionally, such as the meta-logic connectives. + +* Pure: 'instance' now handles general arities with general sorts + (i.e. intersections of classes), + +* Presentation: generated HTML now uses a CSS style sheet to make layout + (somewhat) independent of content. It is copied from lib/html/isabelle.css. + It can be changed to alter the colors/layout of generated pages. + + +*** Isar *** + +* Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac, + cut_tac, subgoal_tac and thin_tac: + - Now understand static (Isar) contexts. As a consequence, users of Isar + locales are no longer forced to write Isar proof scripts. + For details see Isar Reference Manual, paragraph 4.3.2: Further tactic + emulations. + - INCOMPATIBILITY: names of variables to be instantiated may no + longer be enclosed in quotes. Instead, precede variable name with `?'. + This is consistent with the instantiation attribute "where". + +* Attributes "where" and "of": + - Now take type variables of instantiated theorem into account when reading + the instantiation string. This fixes a bug that caused instantiated + theorems to have too special types in some circumstances. + - "where" permits explicit instantiations of type variables. + +* Calculation commands "moreover" and "also" no longer interfere with + current facts ("this"), admitting arbitrary combinations with "then" + and derived forms. + +* Locales: + - Goal statements involving the context element "includes" no longer + generate theorems with internal delta predicates (those ending on + "_axioms") in the premise. + Resolve particular premise with .intro to obtain old form. + - Fixed bug in type inference ("unify_frozen") that prevented mix of target + specification and "includes" elements in goal statement. + - Rule sets .intro and .axioms no longer declared as + [intro?] and [elim?] (respectively) by default. + - Experimental command for instantiation of locales in proof contexts: + instantiate