1 Isabelle NEWS -- history of user-relevant changes
2 =================================================
4 (Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.)
7 New in Isabelle2019 (June 2019)
8 -------------------------------
12 * The font collection "Isabelle DejaVu" is systematically derived from
13 the existing "DejaVu" fonts, with variants "Sans Mono", "Sans", "Serif"
14 and styles "Normal", "Bold", "Italic/Oblique", "Bold-Italic/Oblique".
15 The DejaVu base fonts are retricted to well-defined Unicode ranges and
16 augmented by special Isabelle symbols, taken from the former
17 "IsabelleText" font (which is no longer provided separately). The line
18 metrics and overall rendering quality is closer to original DejaVu.
19 INCOMPATIBILITY with display configuration expecting the old
20 "IsabelleText" font: use e.g. "Isabelle DejaVu Sans Mono" instead.
22 * The Isabelle fonts render "\<inverse>" properly as superscript "-1".
24 * Old-style inner comments (* ... *) within the term language are no
25 longer supported (legacy feature in Isabelle2018).
27 * Old-style {* verbatim *} tokens are explicitly marked as legacy
28 feature and will be removed soon. Use \<open>cartouche\<close> syntax instead, e.g.
29 via "isabelle update_cartouches -t" (available since Isabelle2015).
31 * Infix operators that begin or end with a "*" can now be paranthesized
32 without additional spaces, eg "(*)" instead of "( * )". Minor
35 * Mixfix annotations may use cartouches instead of old-style double
36 quotes, e.g. (infixl \<open>+\<close> 60). The command-line tool "isabelle update -u
37 mixfix_cartouches" allows to update existing theory sources
40 * ML setup commands (e.g. 'setup', 'method_setup', 'parse_translation')
41 need to provide a closed expression -- without trailing semicolon. Minor
44 * Commands 'generate_file', 'export_generated_files', and
45 'compile_generated_files' support a stateless (PIDE-conformant) model
46 for generated sources and compiled binaries of other languages. The
47 compilation process is managed in Isabelle/ML, and results exported to
48 the session database for further use (e.g. with "isabelle export" or
52 *** Isabelle/jEdit Prover IDE ***
54 * Fonts for the text area, gutter, GUI elements etc. use the "Isabelle
55 DejaVu" collection by default, which provides uniform rendering quality
56 with the usual Isabelle symbols. Line spacing no longer needs to be
57 adjusted: properties for the old IsabelleText font had "Global Options /
58 Text Area / Extra vertical line spacing (in pixels): -2", it now
59 defaults to 1, but 0 works as well.
61 * The jEdit File Browser is more prominent in the default GUI layout of
62 Isabelle/jEdit: various virtual file-systems provide access to Isabelle
63 resources, notably via "favorites:" (or "Edit Favorites").
65 * Further markup and rendering for "plain text" (e.g. informal prose)
66 and "raw text" (e.g. verbatim sources). This improves the visual
67 appearance of formal comments inside the term language, or in general
68 for repeated alternation of formal and informal text.
70 * Action "isabelle-export-browser" points the File Browser to the theory
71 exports of the current buffer, based on the "isabelle-export:" virtual
72 file-system. The directory view needs to be reloaded manually to follow
73 ongoing document processing.
75 * Action "isabelle-session-browser" points the File Browser to session
76 information, based on the "isabelle-session:" virtual file-system. Its
77 entries are structured according to chapter / session names, the open
78 operation is redirected to the session ROOT file.
80 * Support for user-defined file-formats via class isabelle.File_Format
81 in Isabelle/Scala (e.g. see isabelle.Bibtex.File_Format), configured via
82 the shell function "isabelle_file_format" in etc/settings (e.g. of an
85 * System option "jedit_text_overview" allows to disable the text
88 * Command-line options "-s" and "-u" of "isabelle jedit" override the
89 default for system option "system_heaps" that determines the heap
90 storage directory for "isabelle build". Option "-n" is now clearly
91 separated from option "-s".
93 * The Isabelle/jEdit desktop application uses the same options as
94 "isabelle jedit" for its internal "isabelle build" process: the implicit
95 option "-o system_heaps" (or "-s") has been discontinued. This reduces
96 the potential for surprise wrt. command-line tools.
98 * The official download of the Isabelle/jEdit application already
99 contains heap images for Isabelle/HOL within its main directory: thus
100 the first encounter becomes faster and more robust (e.g. when run from a
101 read-only directory).
103 * Isabelle DejaVu fonts are available with hinting by default, which is
104 relevant for low-resolution displays. This may be disabled via system
105 option "isabelle_fonts_hinted = false" in
106 $ISABELLE_HOME_USER/etc/preferences -- it occasionally yields better
109 * OpenJDK 11 has quite different font rendering, with better glyph
110 shapes and improved sub-pixel anti-aliasing. In some situations results
111 might be *worse* than Oracle Java 8, though -- a proper HiDPI / UHD
112 display is recommended.
115 *** Document preparation ***
117 * Document markers are formal comments of the form \<^marker>\<open>marker_body\<close> that
118 are stripped from document output: the effect is to modify the semantic
119 presentation context or to emit markup to the PIDE document. Some
120 predefined markers are taken from the Dublin Core Metadata Initiative,
121 e.g. \<^marker>\<open>contributor arg\<close> or \<^marker>\<open>license arg\<close> and produce PIDE markup that
122 can retrieved from the document database.
124 * Old-style command tags %name are re-interpreted as markers with
125 proof-scope \<^marker>\<open>tag (proof) name\<close> and produce LaTeX environments as
126 before. Potential INCOMPATIBILITY: multiple markers are composed in
127 canonical order, resulting in a reversed list of tags in the
128 presentation context.
130 * Marker \<^marker>\<open>tag name\<close> does not apply to the proof of a top-level goal
131 statement by default (e.g. 'theorem', 'lemma'). This is a subtle change
132 of semantics wrt. old-style %name.
134 * In Isabelle/jEdit, the string "\tag" may be completed to a "\<^marker>\<open>tag \<close>"
137 * Document antiquotation option "cartouche" indicates if the output
138 should be delimited as cartouche; this takes precedence over the
139 analogous option "quotes".
141 * Many document antiquotations are internally categorized as "embedded"
142 and expect one cartouche argument, which is typically used with the
143 \<^control>\<open>cartouche\<close> notation (e.g. \<^term>\<open>\<lambda>x y. x\<close>). The cartouche
144 delimiters are stripped in output of the source (antiquotation option
145 "source"), but it is possible to enforce delimiters via option
146 "source_cartouche", e.g. @{term [source_cartouche] \<open>\<lambda>x y. x\<close>}.
151 * Implicit cases goal1, goal2, goal3, etc. have been discontinued
152 (legacy feature since Isabelle2016).
154 * More robust treatment of structural errors: begin/end blocks take
155 precedence over goal/proof. This is particularly relevant for the
156 headless PIDE session and server.
158 * Command keywords of kind thy_decl / thy_goal may be more specifically
159 fit into the traditional document model of "definition-statement-proof"
160 via thy_defn / thy_stmt / thy_goal_defn / thy_goal_stmt.
165 * Command 'export_code' produces output as logical files within the
166 theory context, as well as formal session exports that can be
167 materialized via command-line tools "isabelle export" or "isabelle build
168 -e" (with 'export_files' in the session ROOT). Isabelle/jEdit also
169 provides a virtual file-system "isabelle-export:" that can be explored
170 in the regular file-browser. A 'file_prefix' argument allows to specify
171 an explicit name prefix for the target file (SML, OCaml, Scala) or
172 directory (Haskell); the default is "export" with a consecutive number
175 * Command 'export_code': the 'file' argument is now legacy and will be
176 removed soon: writing to the physical file-system is not well-defined in
177 a reactive/parallel application like Isabelle. The empty 'file' argument
178 has been discontinued already: it is superseded by the file-browser in
179 Isabelle/jEdit on "isabelle-export:". Minor INCOMPATIBILITY.
181 * Command 'code_reflect' no longer supports the 'file' argument: it has
182 been superseded by 'file_prefix' for stateless file management as in
183 'export_code'. Minor INCOMPATIBILITY.
185 * Code generation for OCaml: proper strings are used for literals.
186 Minor INCOMPATIBILITY.
188 * Code generation for OCaml: Zarith supersedes Nums as library for
189 proper integer arithmetic. The library is located via standard
190 invocations of "ocamlfind" (via ISABELLE_OCAMLFIND settings variable).
191 The environment provided by "isabelle ocaml_setup" already contains this
192 tool and the required packages. Minor INCOMPATIBILITY.
194 * Code generation for Haskell: code includes for Haskell must contain
195 proper module frame, nothing is added magically any longer.
198 * Code generation: slightly more conventional syntax for 'code_stmts'
199 antiquotation. Minor INCOMPATIBILITY.
201 * Theory List: the precedence of the list_update operator has changed:
202 "f a [n := x]" now needs to be written "(f a)[n := x]".
204 * The functions \<Union>, \<Inter>, \<Squnion>, \<Sqinter> (not the corresponding binding operators)
205 now have the same precedence as any other prefix function symbol. Minor
208 * Simplified syntax setup for big operators under image. In rare
209 situations, type conversions are not inserted implicitly any longer
210 and need to be given explicitly. Auxiliary abbreviations INFIMUM,
211 SUPREMUM, UNION, INTER should now rarely occur in output and are just
212 retained as migration auxiliary. INCOMPATIBILITY.
214 * The simplifier uses image_cong_simp as a congruence rule. The historic
215 and not really well-formed congruence rules INF_cong*, SUP_cong*, are
216 not used by default any longer. INCOMPATIBILITY; consider using declare
217 image_cong_simp [cong del] in extreme situations.
219 * INF_image and SUP_image are no default simp rules any longer.
220 INCOMPATIBILITY, prefer image_comp as simp rule if needed.
222 * Strong congruence rules (with =simp=> in the premises) for constant f
223 are now uniformly called f_cong_simp, in accordance with congruence
224 rules produced for mappers by the datatype package. INCOMPATIBILITY.
226 * Retired lemma card_Union_image; use the simpler card_UN_disjoint
227 instead. INCOMPATIBILITY.
229 * Facts sum_mset.commute and prod_mset.commute have been renamed to
230 sum_mset.swap and prod_mset.swap, similarly to sum.swap and prod.swap.
233 * ML structure Inductive: slightly more conventional naming schema.
234 Minor INCOMPATIBILITY.
236 * ML: Various _global variants of specification tools have been removed.
237 Minor INCOMPATIBILITY, prefer combinators
238 Named_Target.theory_map[_result] to lift specifications to the global
241 * Theory HOL-Library.Simps_Case_Conv: 'case_of_simps' now supports
242 overlapping and non-exhaustive patterns and handles arbitrarily nested
243 patterns. It uses on the same algorithm as HOL-Library.Code_Lazy, which
244 assumes sequential left-to-right pattern matching. The generated
245 equation no longer tuples the arguments on the right-hand side.
248 * Theory HOL-Library.Multiset: the <Union># operator now has the same
249 precedence as any other prefix function symbol.
251 * Theory HOL-Library.Cardinal_Notations has been discontinued in favor
252 of the bundle cardinal_syntax (available in theory Main). Minor
255 * Session HOL-Library and HOL-Number_Theory: Exponentiation by squaring,
256 used for computing powers in class "monoid_mult" and modular
259 * Session HOL-Computational_Algebra: Formal Laurent series and overhaul
260 of Formal power series.
262 * Session HOL-Number_Theory: More material on residue rings in
263 Carmichael's function, primitive roots, more properties for "ord".
265 * Session HOL-Analysis: Better organization and much more material
266 at the level of abstract topological spaces.
268 * Session HOL-Algebra: Free abelian groups, etc., ported from HOL Light;
269 proofs towards algebraic closure by de Vilhena and Baillon.
271 * Session HOL-Homology has been added. It is a port of HOL Light's
272 homology library, with new proofs of "invariance of domain" and related
275 * Session HOL-SPARK: .prv files are no longer written to the
276 file-system, but exported to the session database. Results may be
277 retrieved via "isabelle build -e HOL-SPARK-Examples" on the
281 - The URL for SystemOnTPTP, which is used by remote provers, has been
283 - The machine-learning-based filter MaSh has been optimized to take
284 less time (in most cases).
286 * SMT: reconstruction is now possible using the SMT solver veriT.
289 * New theory More_Word as comprehensive entrance point.
290 * Merged type class bitss into type class bits.
296 * Command 'generate_file' allows to produce sources for other languages,
297 with antiquotations in the Isabelle context (only the control-cartouche
298 form). The default "cartouche" antiquotation evaluates an ML expression
299 of type string and inlines the result as a string literal of the target
300 language. For example, this works for Haskell as follows:
302 generate_file "Pure.hs" = \<open>
303 module Isabelle.Pure where
304 allConst, impConst, eqConst :: String
305 allConst = \<open>\<^const_name>\<open>Pure.all\<close>\<close>
306 impConst = \<open>\<^const_name>\<open>Pure.imp\<close>\<close>
307 eqConst = \<open>\<^const_name>\<open>Pure.eq\<close>\<close>
310 See also commands 'export_generated_files' and 'compile_generated_files'
313 * ML evaluation (notably via commands 'ML' and 'ML_file') is subject to
314 option ML_environment to select a named environment, such as "Isabelle"
315 for Isabelle/ML, or "SML" for official Standard ML. It is also possible
316 to move toplevel bindings between environments, using a notation with
317 ">" as separator. For example:
319 declare [[ML_environment = "Isabelle>SML"]]
320 ML \<open>val println = writeln\<close>
322 declare [[ML_environment = "SML"]]
323 ML \<open>println "test"\<close>
325 declare [[ML_environment = "Isabelle"]]
326 ML \<open>println\<close> \<comment> \<open>not present\<close>
328 The Isabelle/ML function ML_Env.setup defines new ML environments. This
329 is useful to incorporate big SML projects in an isolated name space, and
330 potentially with variations on ML syntax (existing ML_Env.SML_operations
331 observes the official standard).
333 * ML antiquotation @{master_dir} refers to the master directory of the
334 underlying theory, i.e. the directory of the theory file.
336 * ML antiquotation @{verbatim} inlines its argument as string literal,
337 preserving newlines literally. The short form \<^verbatim>\<open>abc\<close> is particularly
340 * Local_Theory.reset is no longer available in user space. Regular
341 definitional packages should use balanced blocks of
342 Local_Theory.open_target versus Local_Theory.close_target instead, or
343 the Local_Theory.subtarget(_result) combinator. Rare INCOMPATIBILITY.
345 * Original PolyML.pointerEq is retained as a convenience for tools that
346 don't use Isabelle/ML (where this is called "pointer_eq").
351 * Update to OpenJDK 11: the current long-term support version of Java.
353 * Update to Poly/ML 5.8 allows to use the native x86_64 platform without
354 the full overhead of 64-bit values everywhere. This special x86_64_32
355 mode provides up to 16GB ML heap, while program code and stacks are
356 allocated elsewhere. Thus approx. 5 times more memory is available for
357 applications compared to old x86 mode (which is no longer used by
358 Isabelle). The switch to the x86_64 CPU architecture also avoids
359 compatibility problems with Linux and macOS, where 32-bit applications
360 are gradually phased out.
362 * System option "checkpoint" has been discontinued: obsolete thanks to
363 improved memory management in Poly/ML.
365 * System option "system_heaps" determines where to store the session
366 image of "isabelle build" (and other tools using that internally).
367 Former option "-s" is superseded by option "-o system_heaps".
368 INCOMPATIBILITY in command-line syntax.
370 * Session directory $ISABELLE_HOME/src/Tools/Haskell provides some
371 source modules for Isabelle tools implemented in Haskell, notably for
374 * The command-line tool "isabelle build -e" retrieves theory exports
375 from the session build database, using 'export_files' in session ROOT
378 * The command-line tool "isabelle update" uses Isabelle/PIDE in
379 batch-mode to update theory sources based on semantic markup produced in
380 Isabelle/ML. Actual updates depend on system options that may be enabled
381 via "-u OPT" (for "update_OPT"), see also $ISABELLE_HOME/etc/options
382 section "Theory update". Theory sessions are specified as in "isabelle
385 * The command-line tool "isabelle update -u control_cartouches" changes
386 antiquotations into control-symbol format (where possible): @{NAME}
387 becomes \<^NAME> and @{NAME ARG} becomes \<^NAME>\<open>ARG\<close>.
389 * Support for Isabelle command-line tools defined in Isabelle/Scala.
390 Instances of class Isabelle_Scala_Tools may be configured via the shell
391 function "isabelle_scala_tools" in etc/settings (e.g. of an Isabelle
394 * Isabelle Server command "use_theories" supports "nodes_status_delay"
395 for continuous output of node status information. The time interval is
396 specified in seconds; a negative value means it is disabled (default).
398 * Isabelle Server command "use_theories" terminates more robustly in the
399 presence of structurally broken sources: full consolidation of theories
400 is no longer required.
402 * OCaml tools and libraries are now accesed via ISABELLE_OCAMLFIND,
403 which needs to point to a suitable version of "ocamlfind" (e.g. via
404 OPAM, see below). INCOMPATIBILITY: settings variables ISABELLE_OCAML and
405 ISABELLE_OCAMLC are no longer supported.
407 * Support for managed installations of Glasgow Haskell Compiler and
408 OCaml via the following command-line tools:
416 The global installation state is determined by the following settings
417 (and corresponding directory contents):
420 ISABELLE_STACK_RESOLVER
424 ISABELLE_OCAML_VERSION
426 After setup, the following Isabelle settings are automatically
427 redirected (overriding existing user settings):
433 The old meaning of these settings as locally installed executables may
434 be recovered by purging the directories ISABELLE_STACK_ROOT /
435 ISABELLE_OPAM_ROOT, or by resetting these variables in
436 $ISABELLE_HOME_USER/etc/settings.
440 New in Isabelle2018 (August 2018)
441 ---------------------------------
445 * Session-qualified theory names are mandatory: it is no longer possible
446 to refer to unqualified theories from the parent session.
447 INCOMPATIBILITY for old developments that have not been updated to
448 Isabelle2017 yet (using the "isabelle imports" tool).
450 * Only the most fundamental theory names are global, usually the entry
451 points to major logic sessions: Pure, Main, Complex_Main, HOLCF, IFOL,
452 FOL, ZF, ZFC etc. INCOMPATIBILITY, need to use qualified names for
453 formerly global "HOL-Probability.Probability" and "HOL-SPARK.SPARK".
455 * Global facts need to be closed: no free variables and no hypotheses.
456 Rare INCOMPATIBILITY.
458 * Facts stemming from locale interpretation are subject to lazy
459 evaluation for improved performance. Rare INCOMPATIBILITY: errors
460 stemming from interpretation morphisms might be deferred and thus
461 difficult to locate; enable system option "strict_facts" temporarily to
464 * Marginal comments need to be written exclusively in the new-style form
465 "\<comment> \<open>text\<close>", old ASCII variants like "-- {* ... *}" are no longer
466 supported. INCOMPATIBILITY, use the command-line tool "isabelle
467 update_comments" to update existing theory files.
469 * Old-style inner comments (* ... *) within the term language are legacy
470 and will be discontinued soon: use formal comments "\<comment> \<open>...\<close>" or "\<^cancel>\<open>...\<close>"
473 * The "op <infix-op>" syntax for infix operators has been replaced by
474 "(<infix-op>)". If <infix-op> begins or ends with a "*", there needs to
475 be a space between the "*" and the corresponding parenthesis.
476 INCOMPATIBILITY, use the command-line tool "isabelle update_op" to
477 convert theory and ML files to the new syntax. Because it is based on
478 regular expression matching, the result may need a bit of manual
479 postprocessing. Invoking "isabelle update_op" converts all files in the
480 current directory (recursively). In case you want to exclude conversion
481 of ML files (because the tool frequently also converts ML's "op"
482 syntax), use option "-m".
484 * Theory header 'abbrevs' specifications need to be separated by 'and'.
487 * Command 'external_file' declares the formal dependency on the given
488 file name, such that the Isabelle build process knows about it, but
489 without specific Prover IDE management.
491 * Session ROOT entries no longer allow specification of 'files'. Rare
492 INCOMPATIBILITY, use command 'external_file' within a proper theory
495 * Session root directories may be specified multiple times: each
496 accessible ROOT file is processed only once. This facilitates
497 specification of $ISABELLE_HOME_USER/ROOTS or command-line options like
498 -d or -D for "isabelle build" and "isabelle jedit". Example:
500 isabelle build -D '~~/src/ZF'
502 * The command 'display_drafts' has been discontinued. INCOMPATIBILITY,
503 use action "isabelle.draft" (or "print") in Isabelle/jEdit instead.
505 * In HTML output, the Isabelle symbol "\<hyphen>" is rendered as explicit
506 Unicode hyphen U+2010, to avoid unclear meaning of the old "soft hyphen"
507 U+00AD. Rare INCOMPATIBILITY, e.g. copy-paste of historic Isabelle HTML
511 *** Isabelle/jEdit Prover IDE ***
513 * The command-line tool "isabelle jedit" provides more flexible options
514 for session management:
516 - option -R builds an auxiliary logic image with all theories from
517 other sessions that are not already present in its parent
519 - option -S is like -R, with a focus on the selected session and its
520 descendants (this reduces startup time for big projects like AFP)
522 - option -A specifies an alternative ancestor session for options -R
525 - option -i includes additional sessions into the name-space of
529 isabelle jedit -R HOL-Number_Theory
530 isabelle jedit -R HOL-Number_Theory -A HOL
531 isabelle jedit -d '$AFP' -S Formal_SSA -A HOL
532 isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis
533 isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis -i CryptHOL
535 * PIDE markup for session ROOT files: allows to complete session names,
536 follow links to theories and document files etc.
538 * Completion supports theory header imports, using theory base name.
539 E.g. "Prob" may be completed to "HOL-Probability.Probability".
541 * Named control symbols (without special Unicode rendering) are shown as
542 bold-italic keyword. This is particularly useful for the short form of
543 antiquotations with control symbol: \<^name>\<open>argument\<close>. The action
544 "isabelle.antiquoted_cartouche" turns an antiquotation with 0 or 1
545 arguments into this format.
547 * Completion provides templates for named symbols with arguments,
548 e.g. "\<comment> \<open>ARGUMENT\<close>" or "\<^emph>\<open>ARGUMENT\<close>".
550 * Slightly more parallel checking, notably for high priority print
551 functions (e.g. State output).
553 * The view title is set dynamically, according to the Isabelle
554 distribution and the logic session name. The user can override this via
555 set-view-title (stored persistently in $JEDIT_SETTINGS/perspective.xml).
557 * System options "spell_checker_include" and "spell_checker_exclude"
558 supersede former "spell_checker_elements" to determine regions of text
559 that are subject to spell-checking. Minor INCOMPATIBILITY.
561 * Action "isabelle.preview" is able to present more file formats,
562 notably bibtex database files and ML files.
564 * Action "isabelle.draft" is similar to "isabelle.preview", but shows a
565 plain-text document draft. Both are available via the menu "Plugins /
568 * When loading text files, the Isabelle symbols encoding UTF-8-Isabelle
569 is only used if there is no conflict with existing Unicode sequences in
570 the file. Otherwise, the fallback encoding is plain UTF-8 and Isabelle
571 symbols remain in literal \<symbol> form. This avoids accidental loss of
572 Unicode content when saving the file.
574 * Bibtex database files (.bib) are semantically checked.
576 * Update to jedit-5.5.0, the latest release.
579 *** Isabelle/VSCode Prover IDE ***
581 * HTML preview of theories and other file-formats similar to
584 * Command-line tool "isabelle vscode_server" accepts the same options
585 -A, -R, -S, -i for session selection as "isabelle jedit". This is
586 relevant for isabelle.args configuration settings in VSCode. The former
587 option -A (explore all known session files) has been discontinued: it is
588 enabled by default, unless option -S is used to focus on a particular
589 spot in the session structure. INCOMPATIBILITY.
592 *** Document preparation ***
594 * Formal comments work uniformly in outer syntax, inner syntax (term
595 language), Isabelle/ML and some other embedded languages of Isabelle.
596 See also "Document comments" in the isar-ref manual. The following forms
599 - marginal text comment: \<comment> \<open>\<dots>\<close>
600 - canceled source: \<^cancel>\<open>\<dots>\<close>
601 - raw LaTeX: \<^latex>\<open>\<dots>\<close>
603 * Outside of the inner theory body, the default presentation context is
604 theory Pure. Thus elementary antiquotations may be used in markup
605 commands (e.g. 'chapter', 'section', 'text') and formal comments.
607 * System option "document_tags" specifies alternative command tags. This
608 is occasionally useful to control the global visibility of commands via
609 session options (e.g. in ROOT).
611 * Document markup commands ('section', 'text' etc.) are implicitly
612 tagged as "document" and visible by default. This avoids the application
613 of option "document_tags" to these commands.
615 * Isabelle names are mangled into LaTeX macro names to allow the full
616 identifier syntax with underscore, prime, digits. This is relevant for
617 antiquotations in control symbol notation, e.g. \<^const_name> becomes
618 \isactrlconstUNDERSCOREname.
620 * Document preparation with skip_proofs option now preserves the content
621 more accurately: only terminal proof steps ('by' etc.) are skipped.
623 * Document antiquotation @{theory name} requires the long
624 session-qualified theory name: this is what users reading the text
625 normally need to import.
627 * Document antiquotation @{session name} checks and prints the given
628 session name verbatim.
630 * Document antiquotation @{cite} now checks the given Bibtex entries
631 against the Bibtex database files -- only in batch-mode session builds.
633 * Command-line tool "isabelle document" has been re-implemented in
634 Isabelle/Scala, with simplified arguments and explicit errors from the
635 latex and bibtex process. Minor INCOMPATIBILITY.
637 * Session ROOT entry: empty 'document_files' means there is no document
638 for this session. There is no need to specify options [document = false]
644 * Command 'interpret' no longer exposes resulting theorems as literal
645 facts, notably for the \<open>prop\<close> notation or the "fact" proof method. This
646 improves modularity of proofs and scalability of locale interpretation.
647 Rare INCOMPATIBILITY, need to refer to explicitly named facts instead
648 (e.g. use 'find_theorems' or 'try' to figure this out).
650 * The old 'def' command has been discontinued (legacy since
651 Isbelle2016-1). INCOMPATIBILITY, use 'define' instead -- usually with
652 object-logic equality or equivalence.
657 * The inner syntax category "sort" now includes notation "_" for the
658 dummy sort: it is effectively ignored in type-inference.
660 * Rewrites clauses (keyword 'rewrites') were moved into the locale
661 expression syntax, where they are part of locale instances. In
662 interpretation commands rewrites clauses now need to occur before 'for'
663 and 'defines'. Rare INCOMPATIBILITY; definitions immediately subject to
664 rewriting may need to be pulled up into the surrounding theory.
666 * For 'rewrites' clauses, if activating a locale instance fails, fall
667 back to reading the clause first. This helps avoid qualification of
668 locale instances where the qualifier's sole purpose is avoiding
669 duplicate constant declarations.
671 * Proof method "simp" now supports a new modifier "flip:" followed by a
672 list of theorems. Each of these theorems is removed from the simpset
673 (without warning if it is not there) and the symmetric version of the
674 theorem (i.e. lhs and rhs exchanged) is added to the simpset. For "auto"
675 and friends the modifier is "simp flip:".
680 * Sledgehammer: bundled version of "vampire" (for non-commercial users)
681 helps to avoid fragility of "remote_vampire" service.
683 * Clarified relationship of characters, strings and code generation:
685 - Type "char" is now a proper datatype of 8-bit values.
687 - Conversions "nat_of_char" and "char_of_nat" are gone; use more
688 general conversions "of_char" and "char_of" with suitable type
691 - The zero character is just written "CHR 0x00", not "0" any longer.
693 - Type "String.literal" (for code generation) is now isomorphic to
694 lists of 7-bit (ASCII) values; concrete values can be written as
695 "STR ''...''" for sequences of printable characters and "STR 0x..."
696 for one single ASCII code point given as hexadecimal numeral.
698 - Type "String.literal" supports concatenation "... + ..." for all
699 standard target languages.
701 - Theory HOL-Library.Code_Char is gone; study the explanations
702 concerning "String.literal" in the tutorial on code generation to
703 get an idea how target-language string literals can be converted to
704 HOL string values and vice versa.
706 - Session Imperative-HOL: operation "raise" directly takes a value of
707 type "String.literal" as argument, not type "string".
711 * Code generation: Code generation takes an explicit option
712 "case_insensitive" to accomodate case-insensitive file systems.
714 * Abstract bit operations as part of Main: push_bit, take_bit, drop_bit.
716 * New, more general, axiomatization of complete_distrib_lattice. The
719 "sup x (Inf X) = Inf (sup x ` X)" and "inf x (Sup X) = Sup (inf x ` X)"
723 "Inf (Sup ` A) <= Sup (Inf ` {f ` A | f . (! Y \<in> A . f Y \<in> Y)})"
725 The instantiations of sets and functions as complete_distrib_lattice are
726 moved to Hilbert_Choice.thy because their proofs need the Hilbert choice
727 operator. The dual of this property is also proved in theory
730 * New syntax for the minimum/maximum of a function over a finite set:
731 MIN x\<in>A. B and even MIN x. B (only useful for finite types), also MAX.
733 * Clarifed theorem names:
735 Min.antimono ~> Min.subset_imp
736 Max.antimono ~> Max.subset_imp
738 Minor INCOMPATIBILITY.
742 - The 'smt_oracle' option is now necessary when using the 'smt' method
743 with a solver other than Z3. INCOMPATIBILITY.
745 - The encoding to first-order logic is now more complete in the
746 presence of higher-order quantifiers. An 'smt_explicit_application'
747 option has been added to control this. INCOMPATIBILITY.
749 * Facts sum.commute(_restrict) and prod.commute(_restrict) renamed to
750 sum.swap(_restrict) and prod.swap(_restrict), to avoid name clashes on
751 interpretation of abstract locales. INCOMPATIBILITY.
753 * Predicate coprime is now a real definition, not a mere abbreviation.
756 * Predicate pairwise_coprime abolished, use "pairwise coprime" instead.
759 * The relator rel_filter on filters has been strengthened to its
760 canonical categorical definition with better properties.
763 * Generalized linear algebra involving linear, span, dependent, dim
764 from type class real_vector to locales module and vector_space.
767 span_inc ~> span_superset
768 span_superset ~> span_base
769 span_eq ~> span_eq_iff
773 * Class linordered_semiring_1 covers zero_less_one also, ruling out
774 pathologic instances. Minor INCOMPATIBILITY.
776 * Theory HOL.List: functions "sorted_wrt" and "sorted" now compare every
777 element in a list to all following elements, not just the next one.
779 * Theory HOL.List syntax:
781 - filter-syntax "[x <- xs. P]" is no longer output syntax, but only
784 - list comprehension syntax now supports tuple patterns in "pat <- xs"
786 * Theory Map: "empty" must now be qualified as "Map.empty".
788 * Removed nat-int transfer machinery. Rare INCOMPATIBILITY.
790 * Fact mod_mult_self4 (on nat) renamed to Suc_mod_mult_self3, to avoid
791 clash with fact mod_mult_self4 (on more generic semirings).
794 * Eliminated some theorem aliasses:
795 even_times_iff ~> even_mult_iff
796 mod_2_not_eq_zero_eq_one_nat ~> not_mod_2_eq_0_eq_1
797 even_of_nat ~> even_int_iff
801 * Eliminated some theorem duplicate variations:
803 - dvd_eq_mod_eq_0_numeral can be replaced by dvd_eq_mod_eq_0
804 - mod_Suc_eq_Suc_mod can be replaced by mod_Suc
805 - mod_Suc_eq_Suc_mod [symmetrict] can be replaced by mod_simps
806 - mod_eq_0_iff can be replaced by mod_eq_0_iff_dvd and dvd_def
807 - the witness of mod_eqD can be given directly as "_ div _"
811 * Classical setup: Assumption "m mod d = 0" (for m d :: nat) is no
812 longer aggresively destroyed to "\<exists>q. m = d * q". INCOMPATIBILITY, adding
813 "elim!: dvd" to classical proof methods in most situations restores
816 * Theory HOL-Library.Conditional_Parametricity provides command
817 'parametric_constant' for proving parametricity of non-recursive
818 definitions. For constants that are not fully parametric the command
819 will infer conditions on relations (e.g., bi_unique, bi_total, or type
820 class conditions such as "respects 0") sufficient for parametricity. See
821 theory HOL-ex.Conditional_Parametricity_Examples for some examples.
823 * Theory HOL-Library.Code_Lazy provides a new preprocessor for the code
824 generator to generate code for algebraic types with lazy evaluation
825 semantics even in call-by-value target languages. See the theories
826 HOL-ex.Code_Lazy_Demo and HOL-Codegenerator_Test.Code_Lazy_Test for some
829 * Theory HOL-Library.Landau_Symbols has been moved here from AFP.
831 * Theory HOL-Library.Old_Datatype no longer provides the legacy command
832 'old_datatype'. INCOMPATIBILITY.
834 * Theory HOL-Computational_Algebra.Polynomial_Factorial does not provide
835 instances of rat, real, complex as factorial rings etc. Import
836 HOL-Computational_Algebra.Field_as_Ring explicitly in case of need.
839 * Session HOL-Algebra: renamed (^) to [^] to avoid conflict with new
840 infix/prefix notation.
842 * Session HOL-Algebra: revamped with much new material. The set of
843 isomorphisms between two groups is now denoted iso rather than iso_set.
846 * Session HOL-Analysis: the Arg function now respects the same interval
847 as Ln, namely (-pi,pi]; the old Arg function has been renamed Arg2pi.
850 * Session HOL-Analysis: the functions zorder, zer_poly, porder and
851 pol_poly have been redefined. All related lemmas have been reworked.
854 * Session HOL-Analysis: infinite products, Moebius functions, the
855 Riemann mapping theorem, the Vitali covering theorem,
856 change-of-variables results for integration and measures.
858 * Session HOL-Real_Asymp: proof method "real_asymp" proves asymptotics
859 or real-valued functions (limits, "Big-O", etc.) automatically.
860 See also ~~/src/HOL/Real_Asymp/Manual for some documentation.
862 * Session HOL-Types_To_Sets: more tool support (unoverload_type combines
863 internalize_sorts and unoverload) and larger experimental application
864 (type based linear algebra transferred to linear algebra on subspaces).
869 * Operation Export.export emits theory exports (arbitrary blobs), which
870 are stored persistently in the session build database.
872 * Command 'ML_export' exports ML toplevel bindings to the global
873 bootstrap environment of the ML process. This allows ML evaluation
874 without a formal theory context, e.g. in command-line tools like
880 * Mac OS X 10.10 Yosemite is now the baseline version; Mavericks is no
883 * Linux and Windows/Cygwin is for x86_64 only, old 32bit platform
884 support has been discontinued.
886 * Java runtime is for x86_64 only. Corresponding Isabelle settings have
887 been renamed to ISABELLE_TOOL_JAVA_OPTIONS and JEDIT_JAVA_OPTIONS,
888 instead of former 32/64 variants. INCOMPATIBILITY.
890 * Old settings ISABELLE_PLATFORM and ISABELLE_WINDOWS_PLATFORM should be
891 phased out due to unclear preference of 32bit vs. 64bit architecture.
892 Explicit GNU bash expressions are now preferred, for example (with
895 #Posix executables (Unix or Cygwin), with preference for 64bit
896 "${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}"
898 #native Windows or Unix executables, with preference for 64bit
899 "${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"
901 #native Windows (32bit) or Unix executables (preference for 64bit)
902 "${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}"
904 * Command-line tool "isabelle build" supports new options:
905 - option -B NAME: include session NAME and all descendants
906 - option -S: only observe changes of sources, not heap images
907 - option -f: forces a fresh build
909 * Command-line tool "isabelle build" options -c -x -B refer to
910 descendants wrt. the session parent or import graph. Subtle
911 INCOMPATIBILITY: options -c -x used to refer to the session parent graph
914 * Command-line tool "isabelle build" takes "condition" options with the
915 corresponding environment values into account, when determining the
916 up-to-date status of a session.
918 * The command-line tool "dump" dumps information from the cumulative
919 PIDE session database: many sessions may be loaded into a given logic
920 image, results from all loaded theories are written to the output
923 * Command-line tool "isabelle imports -I" also reports actual session
924 imports. This helps to minimize the session dependency graph.
926 * The command-line tool "export" and 'export_files' in session ROOT
927 entries retrieve theory exports from the session build database.
929 * The command-line tools "isabelle server" and "isabelle client" provide
930 access to the Isabelle Server: it supports responsive session management
931 and concurrent use of theories, based on Isabelle/PIDE infrastructure.
932 See also the "system" manual.
934 * The command-line tool "isabelle update_comments" normalizes formal
935 comments in outer syntax as follows: \<comment> \<open>text\<close> (whith a single space to
936 approximate the appearance in document output). This is more specific
937 than former "isabelle update_cartouches -c": the latter tool option has
940 * The command-line tool "isabelle mkroot" now always produces a document
941 outline: its options have been adapted accordingly. INCOMPATIBILITY.
943 * The command-line tool "isabelle mkroot -I" initializes a Mercurial
944 repository for the generated session files.
946 * Settings ISABELLE_HEAPS + ISABELLE_BROWSER_INFO (or
947 ISABELLE_HEAPS_SYSTEM + ISABELLE_BROWSER_INFO_SYSTEM in "system build
948 mode") determine the directory locations of the main build artefacts --
949 instead of hard-wired directories in ISABELLE_HOME_USER (or
952 * Settings ISABELLE_PATH and ISABELLE_OUTPUT have been discontinued:
953 heap images and session databases are always stored in
954 $ISABELLE_HEAPS/$ML_IDENTIFIER (command-line default) or
955 $ISABELLE_HEAPS_SYSTEM/$ML_IDENTIFIER (main Isabelle application or
956 "isabelle jedit -s" or "isabelle build -s").
958 * ISABELLE_LATEX and ISABELLE_PDFLATEX now include platform-specific
959 options for improved error reporting. Potential INCOMPATIBILITY with
960 unusual LaTeX installations, may have to adapt these settings.
962 * Update to Poly/ML 5.7.1 with slightly improved performance and PIDE
963 markup for identifier bindings. It now uses The GNU Multiple Precision
964 Arithmetic Library (libgmp) on all platforms, notably Mac OS X with
969 New in Isabelle2017 (October 2017)
970 ----------------------------------
974 * Experimental support for Visual Studio Code (VSCode) as alternative
975 Isabelle/PIDE front-end, see also
976 https://marketplace.visualstudio.com/items?itemName=makarius.Isabelle2017
978 VSCode is a new type of application that continues the concepts of
979 "programmer's editor" and "integrated development environment" towards
980 fully semantic editing and debugging -- in a relatively light-weight
981 manner. Thus it fits nicely on top of the Isabelle/PIDE infrastructure.
982 Technically, VSCode is based on the Electron application framework
983 (Node.js + Chromium browser + V8), which is implemented in JavaScript
984 and TypeScript, while Isabelle/VSCode mainly consists of Isabelle/Scala
985 modules around a Language Server implementation.
987 * Theory names are qualified by the session name that they belong to.
988 This affects imports, but not the theory name space prefix (which is
989 just the theory base name as before).
991 In order to import theories from other sessions, the ROOT file format
992 provides a new 'sessions' keyword. In contrast, a theory that is
993 imported in the old-fashioned manner via an explicit file-system path
994 belongs to the current session, and might cause theory name conflicts
995 later on. Theories that are imported from other sessions are excluded
996 from the current session document. The command-line tool "isabelle
997 imports" helps to update theory imports.
999 * The main theory entry points for some non-HOL sessions have changed,
1000 to avoid confusion with the global name "Main" of the session HOL. This
1001 leads to the follow renamings:
1003 CTT/Main.thy ~> CTT/CTT.thy
1004 ZF/Main.thy ~> ZF/ZF.thy
1005 ZF/Main_ZF.thy ~> ZF/ZF.thy
1006 ZF/Main_ZFC.thy ~> ZF/ZFC.thy
1007 ZF/ZF.thy ~> ZF/ZF_Base.thy
1011 * Commands 'alias' and 'type_alias' introduce aliases for constants and
1012 type constructors, respectively. This allows adhoc changes to name-space
1013 accesses within global or local theory contexts, e.g. within a 'bundle'.
1015 * Document antiquotations @{prf} and @{full_prf} output proof terms
1016 (again) in the same way as commands 'prf' and 'full_prf'.
1018 * Computations generated by the code generator can be embedded directly
1019 into ML, alongside with @{code} antiquotations, using the following
1022 @{computation ... terms: ... datatypes: ...} :
1023 ((term -> term) -> 'ml option -> 'a) -> Proof.context -> term -> 'a
1024 @{computation_conv ... terms: ... datatypes: ...} :
1025 (Proof.context -> 'ml -> conv) -> Proof.context -> conv
1026 @{computation_check terms: ... datatypes: ...} : Proof.context -> conv
1028 See src/HOL/ex/Computations.thy,
1029 src/HOL/Decision_Procs/Commutative_Ring.thy and
1030 src/HOL/Decision_Procs/Reflective_Field.thy for examples and the
1031 tutorial on code generation.
1034 *** Prover IDE -- Isabelle/Scala/jEdit ***
1036 * Session-qualified theory imports allow the Prover IDE to process
1037 arbitrary theory hierarchies independently of the underlying logic
1038 session image (e.g. option "isabelle jedit -l"), but the directory
1039 structure needs to be known in advance (e.g. option "isabelle jedit -d"
1040 or a line in the file $ISABELLE_HOME_USER/ROOTS).
1042 * The PIDE document model maintains file content independently of the
1043 status of jEdit editor buffers. Reloading jEdit buffers no longer causes
1044 changes of formal document content. Theory dependencies are always
1045 resolved internally, without the need for corresponding editor buffers.
1046 The system option "jedit_auto_load" has been discontinued: it is
1047 effectively always enabled.
1049 * The Theories dockable provides a "Purge" button, in order to restrict
1050 the document model to theories that are required for open editor
1053 * The Theories dockable indicates the overall status of checking of each
1054 entry. When all forked tasks of a theory are finished, the border is
1055 painted with thick lines; remaining errors in this situation are
1056 represented by a different border color.
1058 * Automatic indentation is more careful to avoid redundant spaces in
1059 intermediate situations. Keywords are indented after input (via typed
1060 characters or completion); see also option "jedit_indent_input".
1062 * Action "isabelle.preview" opens an HTML preview of the current theory
1063 document in the default web browser.
1065 * Command-line invocation "isabelle jedit -R -l LOGIC" opens the ROOT
1066 entry of the specified logic session in the editor, while its parent is
1067 used for formal checking.
1069 * The main Isabelle/jEdit plugin may be restarted manually (using the
1070 jEdit Plugin Manager), as long as the "Isabelle Base" plugin remains
1071 enabled at all times.
1073 * Update to current jedit-5.4.0.
1078 * Deleting the last code equations for a particular function using
1079 [code del] results in function with no equations (runtime abort) rather
1080 than an unimplemented function (generation time abort). Use explicit
1081 [[code drop:]] to enforce the latter. Minor INCOMPATIBILITY.
1083 * Proper concept of code declarations in code.ML:
1084 - Regular code declarations act only on the global theory level, being
1085 ignored with warnings if syntactically malformed.
1086 - Explicitly global code declarations yield errors if syntactically
1088 - Default code declarations are silently ignored if syntactically
1090 Minor INCOMPATIBILITY.
1092 * Clarified and standardized internal data bookkeeping of code
1093 declarations: history of serials allows to track potentially
1094 non-monotonous declarations appropriately. Minor INCOMPATIBILITY.
1099 * The Nunchaku model finder is now part of "Main".
1102 - A new option, 'smt_nat_as_int', has been added to translate 'nat' to
1103 'int' and benefit from the SMT solver's theory reasoning. It is
1104 disabled by default.
1105 - The legacy module "src/HOL/Library/Old_SMT.thy" has been removed.
1106 - Several small issues have been rectified in the 'smt' command.
1108 * (Co)datatype package: The 'size_gen_o_map' lemma is no longer
1109 generated for datatypes with type class annotations. As a result, the
1110 tactic that derives it no longer fails on nested datatypes. Slight
1113 * Command and antiquotation "value" with modified default strategy:
1114 terms without free variables are always evaluated using plain evaluation
1115 only, with no fallback on normalization by evaluation. Minor
1118 * Theories "GCD" and "Binomial" are already included in "Main" (instead
1121 * Constant "surj" is a full input/output abbreviation (again).
1122 Minor INCOMPATIBILITY.
1124 * Dropped aliasses RangeP, DomainP for Rangep, Domainp respectively.
1127 * Renamed ii to imaginary_unit in order to free up ii as a variable
1128 name. The syntax \<i> remains available. INCOMPATIBILITY.
1130 * Dropped abbreviations transP, antisymP, single_valuedP; use constants
1131 transp, antisymp, single_valuedp instead. INCOMPATIBILITY.
1133 * Constant "subseq" in Topological_Spaces has been removed -- it is
1134 subsumed by "strict_mono". Some basic lemmas specific to "subseq" have
1135 been renamed accordingly, e.g. "subseq_o" -> "strict_mono_o" etc.
1137 * Theory List: "sublist" renamed to "nths" in analogy with "nth", and
1138 "sublisteq" renamed to "subseq". Minor INCOMPATIBILITY.
1140 * Theory List: new generic function "sorted_wrt".
1142 * Named theorems mod_simps covers various congruence rules concerning
1143 mod, replacing former zmod_simps. INCOMPATIBILITY.
1145 * Swapped orientation of congruence rules mod_add_left_eq,
1146 mod_add_right_eq, mod_add_eq, mod_mult_left_eq, mod_mult_right_eq,
1147 mod_mult_eq, mod_minus_eq, mod_diff_left_eq, mod_diff_right_eq,
1148 mod_diff_eq. INCOMPATIBILITY.
1150 * Generalized some facts:
1153 zminus_zmod ~> mod_minus_eq
1154 zdiff_zmod_left ~> mod_diff_left_eq
1155 zdiff_zmod_right ~> mod_diff_right_eq
1156 zmod_eq_dvd_iff ~> mod_eq_dvd_iff
1159 * Algebraic type class hierarchy of euclidean (semi)rings in HOL:
1160 euclidean_(semi)ring, euclidean_(semi)ring_cancel,
1161 unique_euclidean_(semi)ring; instantiation requires provision of a
1164 * Theory "HOL-Number_Theory.Euclidean_Algorithm" has been reworked:
1165 - Euclidean induction is available as rule eucl_induct.
1166 - Constants Euclidean_Algorithm.gcd, Euclidean_Algorithm.lcm,
1167 Euclidean_Algorithm.Gcd and Euclidean_Algorithm.Lcm allow
1168 easy instantiation of euclidean (semi)rings as GCD (semi)rings.
1169 - Coefficients obtained by extended euclidean algorithm are
1170 available as "bezout_coefficients".
1173 * Theory "Number_Theory.Totient" introduces basic notions about Euler's
1174 totient function previously hidden as solitary example in theory
1175 Residues. Definition changed so that "totient 1 = 1" in agreement with
1176 the literature. Minor INCOMPATIBILITY.
1178 * New styles in theory "HOL-Library.LaTeXsugar":
1179 - "dummy_pats" for printing equations with "_" on the lhs;
1180 - "eta_expand" for printing eta-expanded terms.
1182 * Theory "HOL-Library.Permutations": theorem bij_swap_ompose_bij has
1183 been renamed to bij_swap_compose_bij. INCOMPATIBILITY.
1185 * New theory "HOL-Library.Going_To_Filter" providing the "f going_to F"
1186 filter for describing points x such that f(x) is in the filter F.
1188 * Theory "HOL-Library.Formal_Power_Series": constants X/E/L/F have been
1189 renamed to fps_X/fps_exp/fps_ln/fps_hypergeo to avoid polluting the name
1190 space. INCOMPATIBILITY.
1192 * Theory "HOL-Library.FinFun" has been moved to AFP (again).
1195 * Theory "HOL-Library.FuncSet": some old and rarely used ASCII
1196 replacement syntax has been removed. INCOMPATIBILITY, standard syntax
1197 with symbols should be used instead. The subsequent commands help to
1198 reproduce the old forms, e.g. to simplify porting old theories:
1201 "_PiE" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(3PIE _:_./ _)" 10)
1202 "_Pi" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(3PI _:_./ _)" 10)
1203 "_lam" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("(3%_:_./ _)" [0,0,3] 3)
1205 * Theory "HOL-Library.Multiset": the simprocs on subsets operators of
1206 multisets have been renamed:
1208 msetless_cancel_numerals ~> msetsubset_cancel
1209 msetle_cancel_numerals ~> msetsubset_eq_cancel
1213 * Theory "HOL-Library.Pattern_Aliases" provides input and output syntax
1214 for pattern aliases as known from Haskell, Scala and ML.
1216 * Theory "HOL-Library.Uprod" formalizes the type of unordered pairs.
1218 * Session HOL-Analysis: more material involving arcs, paths, covering
1219 spaces, innessential maps, retracts, infinite products, simplicial
1220 complexes. Baire Category theorem. Major results include the Jordan
1221 Curve Theorem and the Great Picard Theorem.
1223 * Session HOL-Algebra has been extended by additional lattice theory:
1224 the Knaster-Tarski fixed point theorem and Galois Connections.
1226 * Sessions HOL-Computational_Algebra and HOL-Number_Theory: new notions
1227 of squarefreeness, n-th powers, and prime powers.
1229 * Session "HOL-Computional_Algebra" covers many previously scattered
1230 theories, notably Euclidean_Algorithm, Factorial_Ring,
1231 Formal_Power_Series, Fraction_Field, Fundamental_Theorem_Algebra,
1232 Normalized_Fraction, Polynomial_FPS, Polynomial, Primes. Minor
1238 * Isabelle/Scala: the SQL module supports access to relational
1239 databases, either as plain file (SQLite) or full-scale server
1240 (PostgreSQL via local port or remote ssh connection).
1242 * Results of "isabelle build" are recorded as SQLite database (i.e.
1243 "Application File Format" in the sense of
1244 https://www.sqlite.org/appfileformat.html). This allows systematic
1245 access via operations from module Sessions.Store in Isabelle/Scala.
1247 * System option "parallel_proofs" is 1 by default (instead of more
1248 aggressive 2). This requires less heap space and avoids burning parallel
1249 CPU cycles, while full subproof parallelization is enabled for repeated
1250 builds (according to parallel_subproofs_threshold).
1252 * System option "record_proofs" allows to change the global
1253 Proofterm.proofs variable for a session. Regular values are are 0, 1, 2;
1254 a negative value means the current state in the ML heap image remains
1257 * Isabelle settings variable ISABELLE_SCALA_BUILD_OPTIONS has been
1258 renamed to ISABELLE_SCALAC_OPTIONS. Rare INCOMPATIBILITY.
1260 * Isabelle settings variables ISABELLE_WINDOWS_PLATFORM,
1261 ISABELLE_WINDOWS_PLATFORM32, ISABELLE_WINDOWS_PLATFORM64 indicate the
1262 native Windows platform (independently of the Cygwin installation). This
1263 is analogous to ISABELLE_PLATFORM, ISABELLE_PLATFORM32,
1264 ISABELLE_PLATFORM64.
1266 * Command-line tool "isabelle build_docker" builds a Docker image from
1267 the Isabelle application bundle for Linux. See also
1268 https://hub.docker.com/r/makarius/isabelle
1270 * Command-line tool "isabelle vscode_server" provides a Language Server
1271 Protocol implementation, e.g. for the Visual Studio Code editor. It
1272 serves as example for alternative PIDE front-ends.
1274 * Command-line tool "isabelle imports" helps to maintain theory imports
1275 wrt. session structure. Examples for the main Isabelle distribution:
1277 isabelle imports -I -a
1278 isabelle imports -U -a
1279 isabelle imports -U -i -a
1280 isabelle imports -M -a -d '~~/src/Benchmarks'
1284 New in Isabelle2016-1 (December 2016)
1285 -------------------------------------
1289 * Splitter in proof methods "simp", "auto" and friends:
1290 - The syntax "split add" has been discontinued, use plain "split",
1292 - For situations with many conditional or case expressions, there is
1293 an alternative splitting strategy that can be much faster. It is
1294 selected by writing "split!" instead of "split". It applies safe
1295 introduction and elimination rules after each split rule. As a
1296 result the subgoal may be split into several subgoals.
1298 * Command 'bundle' provides a local theory target to define a bundle
1299 from the body of specification commands (such as 'declare',
1300 'declaration', 'notation', 'lemmas', 'lemma'). For example:
1308 * Command 'unbundle' is like 'include', but works within a local theory
1309 context. Unlike "context includes ... begin", the effect of 'unbundle'
1310 on the target context persists, until different declarations are given.
1312 * Simplified outer syntax: uniform category "name" includes long
1313 identifiers. Former "xname" / "nameref" / "name reference" has been
1316 * Embedded content (e.g. the inner syntax of types, terms, props) may be
1317 delimited uniformly via cartouches. This works better than old-fashioned
1318 quotes when sub-languages are nested.
1320 * Mixfix annotations support general block properties, with syntax
1321 "(\<open>x=a y=b z \<dots>\<close>". Notable property names are "indent", "consistent",
1322 "unbreakable", "markup". The existing notation "(DIGITS" is equivalent
1323 to "(\<open>indent=DIGITS\<close>". The former notation "(00" for unbreakable blocks
1324 is superseded by "(\<open>unbreabable\<close>" --- rare INCOMPATIBILITY.
1326 * Proof method "blast" is more robust wrt. corner cases of Pure
1327 statements without object-logic judgment.
1329 * Commands 'prf' and 'full_prf' are somewhat more informative (again):
1330 proof terms are reconstructed and cleaned from administrative thm nodes.
1332 * Code generator: config option "code_timing" triggers measurements of
1333 different phases of code generation. See src/HOL/ex/Code_Timing.thy for
1336 * Code generator: implicits in Scala (stemming from type class
1337 instances) are generated into companion object of corresponding type
1338 class, to resolve some situations where ambiguities may occur.
1340 * Solve direct: option "solve_direct_strict_warnings" gives explicit
1341 warnings for lemma statements with trivial proofs.
1344 *** Prover IDE -- Isabelle/Scala/jEdit ***
1346 * More aggressive flushing of machine-generated input, according to
1347 system option editor_generated_input_delay (in addition to existing
1348 editor_input_delay for regular user edits). This may affect overall PIDE
1349 reactivity and CPU usage.
1351 * Syntactic indentation according to Isabelle outer syntax. Action
1352 "indent-lines" (shortcut C+i) indents the current line according to
1353 command keywords and some command substructure. Action
1354 "isabelle.newline" (shortcut ENTER) indents the old and the new line
1355 according to command keywords only; see also option
1356 "jedit_indent_newline".
1358 * Semantic indentation for unstructured proof scripts ('apply' etc.) via
1359 number of subgoals. This requires information of ongoing document
1360 processing and may thus lag behind, when the user is editing too
1361 quickly; see also option "jedit_script_indent" and
1362 "jedit_script_indent_limit".
1364 * Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed'
1365 are treated as delimiters for fold structure; 'begin' and 'end'
1366 structure of theory specifications is treated as well.
1368 * Command 'proof' provides information about proof outline with cases,
1369 e.g. for proof methods "cases", "induct", "goal_cases".
1371 * Completion templates for commands involving "begin ... end" blocks,
1372 e.g. 'context', 'notepad'.
1374 * Sidekick parser "isabelle-context" shows nesting of context blocks
1375 according to 'begin' and 'end' structure.
1377 * Highlighting of entity def/ref positions wrt. cursor.
1379 * Action "isabelle.select-entity" (shortcut CS+ENTER) selects all
1380 occurrences of the formal entity at the caret position. This facilitates
1381 systematic renaming.
1383 * PIDE document markup works across multiple Isar commands, e.g. the
1384 results established at the end of a proof are properly identified in the
1387 * Cartouche abbreviations work both for " and ` to accomodate typical
1388 situations where old ASCII notation may be updated.
1390 * Dockable window "Symbols" also provides access to 'abbrevs' from the
1391 outer syntax of the current theory buffer. This provides clickable
1392 syntax templates, including entries with empty abbrevs name (which are
1393 inaccessible via keyboard completion).
1395 * IDE support for the Isabelle/Pure bootstrap process, with the
1396 following independent stages:
1401 src/Pure/ML_Bootstrap.thy
1403 The ML ROOT files act like quasi-theories in the context of theory
1404 ML_Bootstrap: this allows continuous checking of all loaded ML files.
1405 The theory files are presented with a modified header to import Pure
1406 from the running Isabelle instance. Results from changed versions of
1407 each stage are *not* propagated to the next stage, and isolated from the
1408 actual Isabelle/Pure that runs the IDE itself. The sequential
1409 dependencies of the above files are only observed for batch build.
1411 * Isabelle/ML and Standard ML files are presented in Sidekick with the
1412 tree structure of section headings: this special comment format is
1413 described in "implementation" chapter 0, e.g. (*** section ***).
1415 * Additional abbreviations for syntactic completion may be specified
1416 within the theory header as 'abbrevs'. The theory syntax for 'keywords'
1417 has been simplified accordingly: optional abbrevs need to go into the
1418 new 'abbrevs' section.
1420 * Global abbreviations via $ISABELLE_HOME/etc/abbrevs and
1421 $ISABELLE_HOME_USER/etc/abbrevs are no longer supported. Minor
1422 INCOMPATIBILITY, use 'abbrevs' within theory header instead.
1424 * Action "isabelle.keymap-merge" asks the user to resolve pending
1425 Isabelle keymap changes that are in conflict with the current jEdit
1426 keymap; non-conflicting changes are always applied implicitly. This
1427 action is automatically invoked on Isabelle/jEdit startup and thus
1428 increases chances that users see new keyboard shortcuts when re-using
1431 * ML and document antiquotations for file-systems paths are more uniform
1434 @{path NAME} -- no file-system check
1435 @{file NAME} -- check for plain file
1436 @{dir NAME} -- check for directory
1438 Minor INCOMPATIBILITY, former uses of @{file} and @{file_unchecked} may
1442 *** Document preparation ***
1444 * New symbol \<circle>, e.g. for temporal operator.
1446 * New document and ML antiquotation @{locale} for locales, similar to
1447 existing antiquotation @{class}.
1449 * Mixfix annotations support delimiters like \<^control>\<open>cartouche\<close> --
1450 this allows special forms of document output.
1452 * Raw LaTeX output now works via \<^latex>\<open>...\<close> instead of raw control
1453 symbol \<^raw:...>. INCOMPATIBILITY, notably for LaTeXsugar.thy and its
1456 * \<^raw:...> symbols are no longer supported.
1458 * Old 'header' command is no longer supported (legacy since
1464 * Many specification elements support structured statements with 'if' /
1465 'for' eigen-context, e.g. 'axiomatization', 'abbreviation',
1466 'definition', 'inductive', 'function'.
1468 * Toplevel theorem statements support eigen-context notation with 'if' /
1469 'for' (in postfix), which corresponds to 'assumes' / 'fixes' in the
1470 traditional long statement form (in prefix). Local premises are called
1471 "that" or "assms", respectively. Empty premises are *not* bound in the
1472 context: INCOMPATIBILITY.
1474 * Command 'define' introduces a local (non-polymorphic) definition, with
1475 optional abstraction over local parameters. The syntax resembles
1476 'definition' and 'obtain'. It fits better into the Isar language than
1477 old 'def', which is now a legacy feature.
1479 * Command 'obtain' supports structured statements with 'if' / 'for'
1482 * Command '\<proof>' is an alias for 'sorry', with different
1483 typesetting. E.g. to produce proof holes in examples and documentation.
1485 * The defining position of a literal fact \<open>prop\<close> is maintained more
1486 carefully, and made accessible as hyperlink in the Prover IDE.
1488 * Commands 'finally' and 'ultimately' used to expose the result as
1489 literal fact: this accidental behaviour has been discontinued. Rare
1490 INCOMPATIBILITY, use more explicit means to refer to facts in Isar.
1492 * Command 'axiomatization' has become more restrictive to correspond
1493 better to internal axioms as singleton facts with mandatory name. Minor
1496 * Proof methods may refer to the main facts via the dynamic fact
1497 "method_facts". This is particularly useful for Eisbach method
1500 * Proof method "use" allows to modify the main facts of a given method
1504 (use facts in \<open>simp add: ...\<close>)
1506 * The old proof method "default" has been removed (legacy since
1507 Isabelle2016). INCOMPATIBILITY, use "standard" instead.
1512 * Pure provides basic versions of proof methods "simp" and "simp_all"
1513 that only know about meta-equality (==). Potential INCOMPATIBILITY in
1514 theory imports that merge Pure with e.g. Main of Isabelle/HOL: the order
1515 is relevant to avoid confusion of Pure.simp vs. HOL.simp.
1517 * The command 'unfolding' and proof method "unfold" include a second
1518 stage where given equations are passed through the attribute "abs_def"
1519 before rewriting. This ensures that definitions are fully expanded,
1520 regardless of the actual parameters that are provided. Rare
1521 INCOMPATIBILITY in some corner cases: use proof method (simp only:)
1522 instead, or declare [[unfold_abs_def = false]] in the proof context.
1524 * Type-inference improves sorts of newly introduced type variables for
1525 the object-logic, using its base sort (i.e. HOL.type for Isabelle/HOL).
1526 Thus terms like "f x" or "\<And>x. P x" without any further syntactic context
1527 produce x::'a::type in HOL instead of x::'a::{} in Pure. Rare
1528 INCOMPATIBILITY, need to provide explicit type constraints for Pure
1529 types where this is really intended.
1534 * New proof method "argo" using the built-in Argo solver based on SMT
1535 technology. The method can be used to prove goals of quantifier-free
1536 propositional logic, goals based on a combination of quantifier-free
1537 propositional logic with equality, and goals based on a combination of
1538 quantifier-free propositional logic with linear real arithmetic
1539 including min/max/abs. See HOL/ex/Argo_Examples.thy for examples.
1541 * The new "nunchaku" command integrates the Nunchaku model finder. The
1542 tool is experimental. See ~~/src/HOL/Nunchaku/Nunchaku.thy for details.
1544 * Metis: The problem encoding has changed very slightly. This might
1545 break existing proofs. INCOMPATIBILITY.
1548 - The MaSh relevance filter is now faster than before.
1549 - Produce syntactically correct Vampire 4.0 problem files.
1551 * (Co)datatype package:
1552 - New commands for defining corecursive functions and reasoning about
1553 them in "~~/src/HOL/Library/BNF_Corec.thy": 'corec', 'corecursive',
1554 'friend_of_corec', and 'corecursion_upto'; and 'corec_unique' proof
1555 method. See 'isabelle doc corec'.
1556 - The predicator :: ('a \<Rightarrow> bool) \<Rightarrow> 'a F \<Rightarrow> bool is now a first-class
1557 citizen in bounded natural functors.
1558 - 'primrec' now allows nested calls through the predicator in addition
1559 to the map function.
1560 - 'bnf' automatically discharges reflexive proof obligations.
1561 - 'bnf' outputs a slightly modified proof obligation expressing rel in
1562 terms of map and set
1563 (not giving a specification for rel makes this one reflexive).
1564 - 'bnf' outputs a new proof obligation expressing pred in terms of set
1565 (not giving a specification for pred makes this one reflexive).
1566 INCOMPATIBILITY: manual 'bnf' declarations may need adjustment.
1568 rel_prod_apply ~> rel_prod_inject
1569 pred_prod_apply ~> pred_prod_inject
1571 - The "size" plugin has been made compatible again with locales.
1572 - The theorems about "rel" and "set" may have a slightly different (but
1576 * The 'coinductive' command produces a proper coinduction rule for
1577 mutual coinductive predicates. This new rule replaces the old rule,
1578 which exposed details of the internal fixpoint construction and was
1579 hard to use. INCOMPATIBILITY.
1581 * New abbreviations for negated existence (but not bounded existence):
1583 \<nexists>x. P x \<equiv> \<not> (\<exists>x. P x)
1584 \<nexists>!x. P x \<equiv> \<not> (\<exists>!x. P x)
1586 * The print mode "HOL" for ASCII syntax of binders "!", "?", "?!", "@"
1587 has been removed for output. It is retained for input only, until it is
1588 eliminated altogether.
1590 * The unique existence quantifier no longer provides 'binder' syntax,
1591 but uses syntax translations (as for bounded unique existence). Thus
1592 iterated quantification \<exists>!x y. P x y with its slightly confusing
1593 sequential meaning \<exists>!x. \<exists>!y. P x y is no longer possible. Instead,
1594 pattern abstraction admits simultaneous unique existence \<exists>!(x, y). P x y
1595 (analogous to existing notation \<exists>!(x, y)\<in>A. P x y). Potential
1596 INCOMPATIBILITY in rare situations.
1598 * Conventional syntax "%(). t" for unit abstractions. Slight syntactic
1601 * Renamed constants and corresponding theorems:
1606 listprod ~> prod_list
1610 * Sligthly more standardized theorem names:
1611 sgn_times ~> sgn_mult
1612 sgn_mult' ~> Real_Vector_Spaces.sgn_mult
1613 divide_zero_left ~> div_0
1614 zero_mod_left ~> mod_0
1615 divide_zero ~> div_by_0
1616 divide_1 ~> div_by_1
1617 nonzero_mult_divide_cancel_left ~> nonzero_mult_div_cancel_left
1618 div_mult_self1_is_id ~> nonzero_mult_div_cancel_left
1619 nonzero_mult_divide_cancel_right ~> nonzero_mult_div_cancel_right
1620 div_mult_self2_is_id ~> nonzero_mult_div_cancel_right
1621 is_unit_divide_mult_cancel_left ~> is_unit_div_mult_cancel_left
1622 is_unit_divide_mult_cancel_right ~> is_unit_div_mult_cancel_right
1623 mod_div_equality ~> div_mult_mod_eq
1624 mod_div_equality2 ~> mult_div_mod_eq
1625 mod_div_equality3 ~> mod_div_mult_eq
1626 mod_div_equality4 ~> mod_mult_div_eq
1627 minus_div_eq_mod ~> minus_div_mult_eq_mod
1628 minus_div_eq_mod2 ~> minus_mult_div_eq_mod
1629 minus_mod_eq_div ~> minus_mod_eq_div_mult
1630 minus_mod_eq_div2 ~> minus_mod_eq_mult_div
1631 div_mod_equality' ~> minus_mod_eq_div_mult [symmetric]
1632 mod_div_equality' ~> minus_div_mult_eq_mod [symmetric]
1633 zmod_zdiv_equality ~> mult_div_mod_eq [symmetric]
1634 zmod_zdiv_equality' ~> minus_div_mult_eq_mod [symmetric]
1635 Divides.mult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
1636 mult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
1637 zmult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
1638 div_1 ~> div_by_Suc_0
1639 mod_1 ~> mod_by_Suc_0
1642 * New type class "idom_abs_sgn" specifies algebraic properties
1643 of sign and absolute value functions. Type class "sgn_if" has
1644 disappeared. Slight INCOMPATIBILITY.
1646 * Dedicated syntax LENGTH('a) for length of types.
1648 * Characters (type char) are modelled as finite algebraic type
1649 corresponding to {0..255}.
1651 - Logical representation:
1652 * 0 is instantiated to the ASCII zero character.
1653 * All other characters are represented as "Char n"
1654 with n being a raw numeral expression less than 256.
1655 * Expressions of the form "Char n" with n greater than 255
1657 - Printing and parsing:
1658 * Printable characters are printed and parsed as "CHR ''\<dots>''"
1660 * The ASCII zero character is printed and parsed as "0".
1661 * All other canonical characters are printed as "CHR 0xXX"
1662 with XX being the hexadecimal character code. "CHR n"
1663 is parsable for every numeral expression n.
1664 * Non-canonical characters have no special syntax and are
1665 printed as their logical representation.
1666 - Explicit conversions from and to the natural numbers are
1667 provided as char_of_nat, nat_of_char (as before).
1668 - The auxiliary nibble type has been discontinued.
1672 * Type class "div" with operation "mod" renamed to type class "modulo"
1673 with operation "modulo", analogously to type class "divide". This
1674 eliminates the need to qualify any of those names in the presence of
1675 infix "mod" syntax. INCOMPATIBILITY.
1677 * Statements and proofs of Knaster-Tarski fixpoint combinators lfp/gfp
1678 have been clarified. The fixpoint properties are lfp_fixpoint, its
1679 symmetric lfp_unfold (as before), and the duals for gfp. Auxiliary items
1680 for the proof (lfp_lemma2 etc.) are no longer exported, but can be
1681 easily recovered by composition with eq_refl. Minor INCOMPATIBILITY.
1683 * Constant "surj" is a mere input abbreviation, to avoid hiding an
1684 equation in term output. Minor INCOMPATIBILITY.
1686 * Command 'code_reflect' accepts empty constructor lists for datatypes,
1687 which renders those abstract effectively.
1689 * Command 'export_code' checks given constants for abstraction
1690 violations: a small guarantee that given constants specify a safe
1691 interface for the generated code.
1693 * Code generation for Scala: ambiguous implicts in class diagrams are
1694 spelt out explicitly.
1696 * Static evaluators (Code_Evaluation.static_* in Isabelle/ML) rely on
1697 explicitly provided auxiliary definitions for required type class
1698 dictionaries rather than half-working magic. INCOMPATIBILITY, see the
1699 tutorial on code generation for details.
1701 * Theory Set_Interval: substantial new theorems on indexed sums and
1704 * Locale bijection establishes convenient default simp rules such as
1705 "inv f (f a) = a" for total bijections.
1707 * Abstract locales semigroup, abel_semigroup, semilattice,
1708 semilattice_neutr, ordering, ordering_top, semilattice_order,
1709 semilattice_neutr_order, comm_monoid_set, semilattice_set,
1710 semilattice_neutr_set, semilattice_order_set,
1711 semilattice_order_neutr_set monoid_list, comm_monoid_list,
1712 comm_monoid_list_set, comm_monoid_mset, comm_monoid_fun use boldified
1713 syntax uniformly that does not clash with corresponding global syntax.
1716 * Former locale lifting_syntax is now a bundle, which is easier to
1717 include in a local context or theorem statement, e.g. "context includes
1718 lifting_syntax begin ... end". Minor INCOMPATIBILITY.
1720 * Some old / obsolete theorems have been renamed / removed, potential
1723 nat_less_cases -- removed, use linorder_cases instead
1724 inv_image_comp -- removed, use image_inv_f_f instead
1725 image_surj_f_inv_f ~> image_f_inv_f
1727 * Some theorems about groups and orders have been generalised from
1728 groups to semi-groups that are also monoids:
1731 less_add_same_cancel1
1732 less_add_same_cancel2
1735 add_less_same_cancel1
1736 add_less_same_cancel2
1738 * Some simplifications theorems about rings have been removed, since
1739 superseeded by a more general version:
1740 less_add_cancel_left_greater_zero ~> less_add_same_cancel1
1741 less_add_cancel_right_greater_zero ~> less_add_same_cancel2
1742 less_eq_add_cancel_left_greater_eq_zero ~> le_add_same_cancel1
1743 less_eq_add_cancel_right_greater_eq_zero ~> le_add_same_cancel2
1744 less_eq_add_cancel_left_less_eq_zero ~> add_le_same_cancel1
1745 less_eq_add_cancel_right_less_eq_zero ~> add_le_same_cancel2
1746 less_add_cancel_left_less_zero ~> add_less_same_cancel1
1747 less_add_cancel_right_less_zero ~> add_less_same_cancel2
1750 * Renamed split_if -> if_split and split_if_asm -> if_split_asm to
1751 resemble the f.split naming convention, INCOMPATIBILITY.
1753 * Added class topological_monoid.
1755 * The following theorems have been renamed:
1757 setsum_left_distrib ~> sum_distrib_right
1758 setsum_right_distrib ~> sum_distrib_left
1762 * Compound constants INFIMUM and SUPREMUM are mere abbreviations now.
1765 * "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional
1766 comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f `
1769 * Class semiring_Lcd merged into semiring_Gcd. INCOMPATIBILITY.
1771 * The type class ordered_comm_monoid_add is now called
1772 ordered_cancel_comm_monoid_add. A new type class ordered_comm_monoid_add
1773 is introduced as the combination of ordered_ab_semigroup_add +
1774 comm_monoid_add. INCOMPATIBILITY.
1776 * Introduced the type classes canonically_ordered_comm_monoid_add and
1779 * Introduced the type class ordered_ab_semigroup_monoid_add_imp_le. When
1780 instantiating linordered_semiring_strict and ordered_ab_group_add, an
1781 explicit instantiation of ordered_ab_semigroup_monoid_add_imp_le might
1782 be required. INCOMPATIBILITY.
1784 * Dropped various legacy fact bindings, whose replacements are often
1785 of a more general type also:
1786 lcm_left_commute_nat ~> lcm.left_commute
1787 lcm_left_commute_int ~> lcm.left_commute
1788 gcd_left_commute_nat ~> gcd.left_commute
1789 gcd_left_commute_int ~> gcd.left_commute
1790 gcd_greatest_iff_nat ~> gcd_greatest_iff
1791 gcd_greatest_iff_int ~> gcd_greatest_iff
1792 coprime_dvd_mult_nat ~> coprime_dvd_mult
1793 coprime_dvd_mult_int ~> coprime_dvd_mult
1794 zpower_numeral_even ~> power_numeral_even
1795 gcd_mult_cancel_nat ~> gcd_mult_cancel
1796 gcd_mult_cancel_int ~> gcd_mult_cancel
1797 div_gcd_coprime_nat ~> div_gcd_coprime
1798 div_gcd_coprime_int ~> div_gcd_coprime
1799 zpower_numeral_odd ~> power_numeral_odd
1800 zero_less_int_conv ~> of_nat_0_less_iff
1801 gcd_greatest_nat ~> gcd_greatest
1802 gcd_greatest_int ~> gcd_greatest
1803 coprime_mult_nat ~> coprime_mult
1804 coprime_mult_int ~> coprime_mult
1805 lcm_commute_nat ~> lcm.commute
1806 lcm_commute_int ~> lcm.commute
1807 int_less_0_conv ~> of_nat_less_0_iff
1808 gcd_commute_nat ~> gcd.commute
1809 gcd_commute_int ~> gcd.commute
1810 Gcd_insert_nat ~> Gcd_insert
1811 Gcd_insert_int ~> Gcd_insert
1812 of_int_int_eq ~> of_int_of_nat_eq
1813 lcm_least_nat ~> lcm_least
1814 lcm_least_int ~> lcm_least
1815 lcm_assoc_nat ~> lcm.assoc
1816 lcm_assoc_int ~> lcm.assoc
1817 int_le_0_conv ~> of_nat_le_0_iff
1818 int_eq_0_conv ~> of_nat_eq_0_iff
1819 Gcd_empty_nat ~> Gcd_empty
1820 Gcd_empty_int ~> Gcd_empty
1821 gcd_assoc_nat ~> gcd.assoc
1822 gcd_assoc_int ~> gcd.assoc
1823 zero_zle_int ~> of_nat_0_le_iff
1824 lcm_dvd2_nat ~> dvd_lcm2
1825 lcm_dvd2_int ~> dvd_lcm2
1826 lcm_dvd1_nat ~> dvd_lcm1
1827 lcm_dvd1_int ~> dvd_lcm1
1828 gcd_zero_nat ~> gcd_eq_0_iff
1829 gcd_zero_int ~> gcd_eq_0_iff
1830 gcd_dvd2_nat ~> gcd_dvd2
1831 gcd_dvd2_int ~> gcd_dvd2
1832 gcd_dvd1_nat ~> gcd_dvd1
1833 gcd_dvd1_int ~> gcd_dvd1
1834 int_numeral ~> of_nat_numeral
1835 lcm_ac_nat ~> ac_simps
1836 lcm_ac_int ~> ac_simps
1837 gcd_ac_nat ~> ac_simps
1838 gcd_ac_int ~> ac_simps
1839 abs_int_eq ~> abs_of_nat
1840 zless_int ~> of_nat_less_iff
1841 zdiff_int ~> of_nat_diff
1842 zadd_int ~> of_nat_add
1843 int_mult ~> of_nat_mult
1844 int_Suc ~> of_nat_Suc
1845 inj_int ~> inj_of_nat
1848 Lcm_empty_nat ~> Lcm_empty
1849 Lcm_empty_int ~> Lcm_empty
1850 Lcm_insert_nat ~> Lcm_insert
1851 Lcm_insert_int ~> Lcm_insert
1852 comp_fun_idem_gcd_nat ~> comp_fun_idem_gcd
1853 comp_fun_idem_gcd_int ~> comp_fun_idem_gcd
1854 comp_fun_idem_lcm_nat ~> comp_fun_idem_lcm
1855 comp_fun_idem_lcm_int ~> comp_fun_idem_lcm
1856 Lcm_eq_0 ~> Lcm_eq_0_I
1857 Lcm0_iff ~> Lcm_0_iff
1858 Lcm_dvd_int ~> Lcm_least
1859 divides_mult_nat ~> divides_mult
1860 divides_mult_int ~> divides_mult
1861 lcm_0_nat ~> lcm_0_right
1862 lcm_0_int ~> lcm_0_right
1863 lcm_0_left_nat ~> lcm_0_left
1864 lcm_0_left_int ~> lcm_0_left
1865 dvd_gcd_D1_nat ~> dvd_gcdD1
1866 dvd_gcd_D1_int ~> dvd_gcdD1
1867 dvd_gcd_D2_nat ~> dvd_gcdD2
1868 dvd_gcd_D2_int ~> dvd_gcdD2
1869 coprime_dvd_mult_iff_nat ~> coprime_dvd_mult_iff
1870 coprime_dvd_mult_iff_int ~> coprime_dvd_mult_iff
1871 realpow_minus_mult ~> power_minus_mult
1872 realpow_Suc_le_self ~> power_Suc_le_self
1873 dvd_Gcd, dvd_Gcd_nat, dvd_Gcd_int removed in favour of Gcd_greatest
1876 * Renamed HOL/Quotient_Examples/FSet.thy to
1877 HOL/Quotient_Examples/Quotient_FSet.thy INCOMPATIBILITY.
1879 * Session HOL-Library: theory FinFun bundles "finfun_syntax" and
1880 "no_finfun_syntax" allow to control optional syntax in local contexts;
1881 this supersedes former theory FinFun_Syntax. INCOMPATIBILITY, e.g. use
1882 "unbundle finfun_syntax" to imitate import of
1883 "~~/src/HOL/Library/FinFun_Syntax".
1885 * Session HOL-Library: theory Multiset_Permutations (executably) defines
1886 the set of permutations of a given set or multiset, i.e. the set of all
1887 lists that contain every element of the carrier (multi-)set exactly
1890 * Session HOL-Library: multiset membership is now expressed using
1891 set_mset rather than count.
1893 - Expressions "count M a > 0" and similar simplify to membership
1896 - Converting between "count M a = 0" and non-membership happens using
1897 equations count_eq_zero_iff and not_in_iff.
1899 - Rules count_inI and in_countE obtain facts of the form
1900 "count M a = n" from membership.
1902 - Rules count_in_diffI and in_diff_countE obtain facts of the form
1903 "count M a = n + count N a" from membership on difference sets.
1907 * Session HOL-Library: theory LaTeXsugar uses new-style "dummy_pats" for
1908 displaying equations in functional programming style --- variables
1909 present on the left-hand but not on the righ-hand side are replaced by
1912 * Session HOL-Library: theory Combinator_PER provides combinator to
1913 build partial equivalence relations from a predicate and an equivalence
1916 * Session HOL-Library: theory Perm provides basic facts about almost
1917 everywhere fix bijections.
1919 * Session HOL-Library: theory Normalized_Fraction allows viewing an
1920 element of a field of fractions as a normalized fraction (i.e. a pair of
1921 numerator and denominator such that the two are coprime and the
1922 denominator is normalized wrt. unit factors).
1924 * Session HOL-NSA has been renamed to HOL-Nonstandard_Analysis.
1926 * Session HOL-Multivariate_Analysis has been renamed to HOL-Analysis.
1928 * Session HOL-Analysis: measure theory has been moved here from
1929 HOL-Probability. When importing HOL-Analysis some theorems need
1930 additional name spaces prefixes due to name clashes. INCOMPATIBILITY.
1932 * Session HOL-Analysis: more complex analysis including Cauchy's
1933 inequality, Liouville theorem, open mapping theorem, maximum modulus
1934 principle, Residue theorem, Schwarz Lemma.
1936 * Session HOL-Analysis: Theory of polyhedra: faces, extreme points,
1937 polytopes, and the Krein–Milman Minkowski theorem.
1939 * Session HOL-Analysis: Numerous results ported from the HOL Light
1940 libraries: homeomorphisms, continuous function extensions, invariance of
1943 * Session HOL-Probability: the type of emeasure and nn_integral was
1944 changed from ereal to ennreal, INCOMPATIBILITY.
1946 emeasure :: 'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal
1947 nn_integral :: 'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal
1949 * Session HOL-Probability: Code generation and QuickCheck for
1950 Probability Mass Functions.
1952 * Session HOL-Probability: theory Random_Permutations contains some
1953 theory about choosing a permutation of a set uniformly at random and
1954 folding over a list in random order.
1956 * Session HOL-Probability: theory SPMF formalises discrete
1957 subprobability distributions.
1959 * Session HOL-Library: the names of multiset theorems have been
1960 normalised to distinguish which ordering the theorems are about
1962 mset_less_eqI ~> mset_subset_eqI
1963 mset_less_insertD ~> mset_subset_insertD
1964 mset_less_eq_count ~> mset_subset_eq_count
1965 mset_less_diff_self ~> mset_subset_diff_self
1966 mset_le_exists_conv ~> mset_subset_eq_exists_conv
1967 mset_le_mono_add_right_cancel ~> mset_subset_eq_mono_add_right_cancel
1968 mset_le_mono_add_left_cancel ~> mset_subset_eq_mono_add_left_cancel
1969 mset_le_mono_add ~> mset_subset_eq_mono_add
1970 mset_le_add_left ~> mset_subset_eq_add_left
1971 mset_le_add_right ~> mset_subset_eq_add_right
1972 mset_le_single ~> mset_subset_eq_single
1973 mset_le_multiset_union_diff_commute ~> mset_subset_eq_multiset_union_diff_commute
1974 diff_le_self ~> diff_subset_eq_self
1975 mset_leD ~> mset_subset_eqD
1976 mset_lessD ~> mset_subsetD
1977 mset_le_insertD ~> mset_subset_eq_insertD
1978 mset_less_of_empty ~> mset_subset_of_empty
1979 mset_less_size ~> mset_subset_size
1980 wf_less_mset_rel ~> wf_subset_mset_rel
1981 count_le_replicate_mset_le ~> count_le_replicate_mset_subset_eq
1982 mset_remdups_le ~> mset_remdups_subset_eq
1983 ms_lesseq_impl ~> subset_eq_mset_impl
1985 Some functions have been renamed:
1986 ms_lesseq_impl -> subset_eq_mset_impl
1988 * HOL-Library: multisets are now ordered with the multiset ordering
1989 #\<subseteq># ~> \<le>
1991 le_multiset ~> less_eq_multiset
1992 less_multiset ~> le_multiset
1995 * Session HOL-Library: the prefix multiset_order has been discontinued:
1996 the theorems can be directly accessed. As a consequence, the lemmas
1997 "order_multiset" and "linorder_multiset" have been discontinued, and the
1998 interpretations "multiset_linorder" and "multiset_wellorder" have been
1999 replaced by instantiations. INCOMPATIBILITY.
2001 * Session HOL-Library: some theorems about the multiset ordering have
2004 le_multiset_def ~> less_eq_multiset_def
2005 less_multiset_def ~> le_multiset_def
2006 less_eq_imp_le_multiset ~> subset_eq_imp_le_multiset
2007 mult_less_not_refl ~> mset_le_not_refl
2008 mult_less_trans ~> mset_le_trans
2009 mult_less_not_sym ~> mset_le_not_sym
2010 mult_less_asym ~> mset_le_asym
2011 mult_less_irrefl ~> mset_le_irrefl
2012 union_less_mono2{,1,2} ~> union_le_mono2{,1,2}
2014 le_multiset\<^sub>H\<^sub>O ~> less_eq_multiset\<^sub>H\<^sub>O
2015 le_multiset_total ~> less_eq_multiset_total
2016 less_multiset_right_total ~> subset_eq_imp_le_multiset
2017 le_multiset_empty_left ~> less_eq_multiset_empty_left
2018 le_multiset_empty_right ~> less_eq_multiset_empty_right
2019 less_multiset_empty_right ~> le_multiset_empty_left
2020 less_multiset_empty_left ~> le_multiset_empty_right
2021 union_less_diff_plus ~> union_le_diff_plus
2022 ex_gt_count_imp_less_multiset ~> ex_gt_count_imp_le_multiset
2023 less_multiset_plus_left_nonempty ~> le_multiset_plus_left_nonempty
2024 le_multiset_plus_right_nonempty ~> le_multiset_plus_right_nonempty
2027 * Session HOL-Library: the lemma mset_map has now the attribute [simp].
2030 * Session HOL-Library: some theorems about multisets have been removed.
2031 INCOMPATIBILITY, use the following replacements:
2033 le_multiset_plus_plus_left_iff ~> add_less_cancel_right
2034 less_multiset_plus_plus_left_iff ~> add_less_cancel_right
2035 le_multiset_plus_plus_right_iff ~> add_less_cancel_left
2036 less_multiset_plus_plus_right_iff ~> add_less_cancel_left
2037 add_eq_self_empty_iff ~> add_cancel_left_right
2038 mset_subset_add_bothsides ~> subset_mset.add_less_cancel_right
2039 mset_less_add_bothsides ~> subset_mset.add_less_cancel_right
2040 mset_le_add_bothsides ~> subset_mset.add_less_cancel_right
2041 empty_inter ~> subset_mset.inf_bot_left
2042 inter_empty ~> subset_mset.inf_bot_right
2043 empty_sup ~> subset_mset.sup_bot_left
2044 sup_empty ~> subset_mset.sup_bot_right
2045 bdd_below_multiset ~> subset_mset.bdd_above_bot
2046 subset_eq_empty ~> subset_mset.le_zero_eq
2047 le_empty ~> subset_mset.le_zero_eq
2048 mset_subset_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero
2049 mset_less_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero
2051 * Session HOL-Library: some typeclass constraints about multisets have
2052 been reduced from ordered or linordered to preorder. Multisets have the
2053 additional typeclasses order_bot, no_top,
2054 ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add,
2055 linordered_cancel_ab_semigroup_add, and
2056 ordered_ab_semigroup_monoid_add_imp_le. INCOMPATIBILITY.
2058 * Session HOL-Library: there are some new simplification rules about
2059 multisets, the multiset ordering, and the subset ordering on multisets.
2062 * Session HOL-Library: the subset ordering on multisets has now the
2063 interpretations ordered_ab_semigroup_monoid_add_imp_le and
2064 bounded_lattice_bot. INCOMPATIBILITY.
2066 * Session HOL-Library, theory Multiset: single has been removed in favor
2067 of add_mset that roughly corresponds to Set.insert. Some theorems have
2070 single_not_empty ~> add_mset_not_empty or empty_not_add_mset
2071 fold_mset_insert ~> fold_mset_add_mset
2072 image_mset_insert ~> image_mset_add_mset
2073 union_single_eq_diff
2074 multi_self_add_other_not_self
2075 diff_single_eq_union
2078 * Session HOL-Library, theory Multiset: some theorems have been changed
2079 to use add_mset instead of single:
2082 multi_self_add_other_not_self
2083 diff_single_eq_union
2084 union_single_eq_diff
2085 union_single_eq_member
2090 multiset_add_sub_el_shuffle
2091 mset_subset_eq_insertD
2093 insert_subset_eq_iff
2094 insert_union_subset_iff
2095 multi_psub_of_add_self
2104 size_eq_Suc_imp_eq_union
2105 multi_nonempty_split
2110 mset_zip_take_Cons_drop_twice
2116 le_multiset_right_total
2118 multiset_induct2_size
2122 * Session HOL-Library, theory Multiset: the definitions of some
2123 constants have changed to use add_mset instead of adding a single
2136 * Session HOL-Library, theory Multiset: due to the above changes, the
2137 attributes of some multiset theorems have been changed:
2139 insert_DiffM [] ~> [simp]
2140 insert_DiffM2 [simp] ~> []
2141 diff_add_mset_swap [simp]
2142 fold_mset_add_mset [simp]
2143 diff_diff_add [simp] (for multisets only)
2144 diff_cancel [simp] ~> []
2145 count_single [simp] ~> []
2146 set_mset_single [simp] ~> []
2147 size_multiset_single [simp] ~> []
2148 size_single [simp] ~> []
2149 image_mset_single [simp] ~> []
2150 mset_subset_eq_mono_add_right_cancel [simp] ~> []
2151 mset_subset_eq_mono_add_left_cancel [simp] ~> []
2152 fold_mset_single [simp] ~> []
2153 subset_eq_empty [simp] ~> []
2154 empty_sup [simp] ~> []
2155 sup_empty [simp] ~> []
2156 inter_empty [simp] ~> []
2157 empty_inter [simp] ~> []
2160 * Session HOL-Library, theory Multiset: the order of the variables in
2161 the second cases of multiset_induct, multiset_induct2_size,
2162 multiset_induct2 has been changed (e.g. Add A a ~> Add a A).
2165 * Session HOL-Library, theory Multiset: there is now a simplification
2166 procedure on multisets. It mimics the behavior of the procedure on
2167 natural numbers. INCOMPATIBILITY.
2169 * Session HOL-Library, theory Multiset: renamed sums and products of
2173 msetprod ~> prod_mset
2175 * Session HOL-Library, theory Multiset: the notation for intersection
2176 and union of multisets have been changed:
2178 #\<inter> ~> \<inter>#
2179 #\<union> ~> \<union>#
2183 * Session HOL-Library, theory Multiset: the lemma
2184 one_step_implies_mult_aux on multisets has been removed, use
2185 one_step_implies_mult instead. INCOMPATIBILITY.
2187 * Session HOL-Library: theory Complete_Partial_Order2 provides reasoning
2188 support for monotonicity and continuity in chain-complete partial orders
2189 and about admissibility conditions for fixpoint inductions.
2191 * Session HOL-Library: theory Library/Polynomial contains also
2192 derivation of polynomials (formerly in Library/Poly_Deriv) but not
2193 gcd/lcm on polynomials over fields. This has been moved to a separate
2194 theory Library/Polynomial_GCD_euclidean.thy, to pave way for a possible
2195 future different type class instantiation for polynomials over factorial
2196 rings. INCOMPATIBILITY.
2198 * Session HOL-Library: theory Sublist provides function "prefixes" with
2199 the following renaming
2202 prefix -> strict_prefix
2204 suffix -> strict_suffix
2206 Added theory of longest common prefixes.
2208 * Session HOL-Number_Theory: algebraic foundation for primes:
2209 Generalisation of predicate "prime" and introduction of predicates
2210 "prime_elem", "irreducible", a "prime_factorization" function, and the
2211 "factorial_ring" typeclass with instance proofs for nat, int, poly. Some
2212 theorems now have different names, most notably "prime_def" is now
2213 "prime_nat_iff". INCOMPATIBILITY.
2215 * Session Old_Number_Theory has been removed, after porting remaining
2218 * Session HOL-Types_To_Sets provides an experimental extension of
2219 Higher-Order Logic to allow translation of types to sets.
2224 * Integer.gcd and Integer.lcm use efficient operations from the Poly/ML
2225 library (notably for big integers). Subtle change of semantics:
2226 Integer.gcd and Integer.lcm both normalize the sign, results are never
2227 negative. This coincides with the definitions in HOL/GCD.thy.
2230 * Structure Rat for rational numbers is now an integral part of
2231 Isabelle/ML, with special notation @int/nat or @int for numerals (an
2232 abbreviation for antiquotation @{Pure.rat argument}) and ML pretty
2233 printing. Standard operations on type Rat.rat are provided via ad-hoc
2234 overloading of + - * / < <= > >= ~ abs. INCOMPATIBILITY, need to
2235 use + instead of +/ etc. Moreover, exception Rat.DIVZERO has been
2236 superseded by General.Div.
2238 * ML antiquotation @{path} is superseded by @{file}, which ensures that
2239 the argument is a plain file. Minor INCOMPATIBILITY.
2241 * Antiquotation @{make_string} is available during Pure bootstrap --
2242 with approximative output quality.
2244 * Low-level ML system structures (like PolyML and RunCall) are no longer
2245 exposed to Isabelle/ML user-space. Potential INCOMPATIBILITY.
2247 * The ML function "ML" provides easy access to run-time compilation.
2248 This is particularly useful for conditional compilation, without
2249 requiring separate files.
2251 * Option ML_exception_debugger controls detailed exception trace via the
2252 Poly/ML debugger. Relevant ML modules need to be compiled beforehand
2253 with ML_file_debug, or with ML_file and option ML_debugger enabled. Note
2254 debugger information requires consirable time and space: main
2255 Isabelle/HOL with full debugger support may need ML_system_64.
2257 * Local_Theory.restore has been renamed to Local_Theory.reset to
2258 emphasize its disruptive impact on the cumulative context, notably the
2259 scope of 'private' or 'qualified' names. Note that Local_Theory.reset is
2260 only appropriate when targets are managed, e.g. starting from a global
2261 theory and returning to it. Regular definitional packages should use
2262 balanced blocks of Local_Theory.open_target versus
2263 Local_Theory.close_target instead. Rare INCOMPATIBILITY.
2265 * Structure TimeLimit (originally from the SML/NJ library) has been
2266 replaced by structure Timeout, with slightly different signature.
2269 * Discontinued cd and pwd operations, which are not well-defined in a
2270 multi-threaded environment. Note that files are usually located
2271 relatively to the master directory of a theory (see also
2272 File.full_path). Potential INCOMPATIBILITY.
2274 * Binding.empty_atts supersedes Thm.empty_binding and
2275 Attrib.empty_binding. Minor INCOMPATIBILITY.
2280 * SML/NJ and old versions of Poly/ML are no longer supported.
2282 * Poly/ML heaps now follow the hierarchy of sessions, and thus require
2283 much less disk space.
2285 * The Isabelle ML process is now managed directly by Isabelle/Scala, and
2286 shell scripts merely provide optional command-line access. In
2289 . Scala module ML_Process to connect to the raw ML process,
2290 with interaction via stdin/stdout/stderr or in batch mode;
2291 . command-line tool "isabelle console" as interactive wrapper;
2292 . command-line tool "isabelle process" as batch mode wrapper.
2294 * The executable "isabelle_process" has been discontinued. Tools and
2295 prover front-ends should use ML_Process or Isabelle_Process in
2296 Isabelle/Scala. INCOMPATIBILITY.
2298 * New command-line tool "isabelle process" supports ML evaluation of
2299 literal expressions (option -e) or files (option -f) in the context of a
2300 given heap image. Errors lead to premature exit of the ML process with
2303 * The command-line tool "isabelle build" supports option -N for cyclic
2304 shuffling of NUMA CPU nodes. This may help performance tuning on Linux
2305 servers with separate CPU/memory modules.
2307 * System option "threads" (for the size of the Isabelle/ML thread farm)
2308 is also passed to the underlying ML runtime system as --gcthreads,
2309 unless there is already a default provided via ML_OPTIONS settings.
2311 * System option "checkpoint" helps to fine-tune the global heap space
2312 management of isabelle build. This is relevant for big sessions that may
2313 exhaust the small 32-bit address space of the ML process (which is used
2316 * System option "profiling" specifies the mode for global ML profiling
2317 in "isabelle build". Possible values are "time", "allocations". The
2318 command-line tool "isabelle profiling_report" helps to digest the
2319 resulting log files.
2321 * System option "ML_process_policy" specifies an optional command prefix
2322 for the underlying ML process, e.g. to control CPU affinity on
2323 multiprocessor systems. The "isabelle jedit" tool allows to override the
2324 implicit default via option -p.
2326 * Command-line tool "isabelle console" provides option -r to help to
2327 bootstrapping Isabelle/Pure interactively.
2329 * Command-line tool "isabelle yxml" has been discontinued.
2330 INCOMPATIBILITY, use operations from the modules "XML" and "YXML" in
2331 Isabelle/ML or Isabelle/Scala.
2333 * Many Isabelle tools that require a Java runtime system refer to the
2334 settings ISABELLE_TOOL_JAVA_OPTIONS32 / ISABELLE_TOOL_JAVA_OPTIONS64,
2335 depending on the underlying platform. The settings for "isabelle build"
2336 ISABELLE_BUILD_JAVA_OPTIONS32 / ISABELLE_BUILD_JAVA_OPTIONS64 have been
2337 discontinued. Potential INCOMPATIBILITY.
2339 * The Isabelle system environment always ensures that the main
2340 executables are found within the shell search $PATH: "isabelle" and
2341 "isabelle_scala_script".
2343 * Isabelle tools may consist of .scala files: the Scala compiler is
2344 invoked on the spot. The source needs to define some object that extends
2347 * File.bash_string, File.bash_path etc. represent Isabelle/ML and
2348 Isabelle/Scala strings authentically within GNU bash. This is useful to
2349 produce robust shell scripts under program control, without worrying
2350 about spaces or special characters. Note that user output works via
2351 Path.print (ML) or Path.toString (Scala). INCOMPATIBILITY, the old (and
2352 less versatile) operations File.shell_quote, File.shell_path etc. have
2355 * The isabelle_java executable allows to run a Java process within the
2356 name space of Java and Scala components that are bundled with Isabelle,
2357 but without the Isabelle settings environment.
2359 * Isabelle/Scala: the SSH module supports ssh and sftp connections, for
2360 remote command-execution and file-system access. This resembles
2361 operations from module File and Isabelle_System to some extent. Note
2362 that Path specifications need to be resolved remotely via
2363 ssh.remote_path instead of File.standard_path: the implicit process
2364 environment is different, Isabelle settings are not available remotely.
2366 * Isabelle/Scala: the Mercurial module supports repositories via the
2367 regular hg command-line interface. The repositroy clone and working
2368 directory may reside on a local or remote file-system (via ssh
2373 New in Isabelle2016 (February 2016)
2374 -----------------------------------
2378 * Eisbach is now based on Pure instead of HOL. Objects-logics may import
2379 either the theory ~~/src/HOL/Eisbach/Eisbach (for HOL etc.) or
2380 ~~/src/HOL/Eisbach/Eisbach_Old_Appl_Syntax (for FOL, ZF etc.). Note that
2381 the HOL-Eisbach session located in ~~/src/HOL/Eisbach/ contains further
2382 examples that do require HOL.
2384 * Better resource usage on all platforms (Linux, Windows, Mac OS X) for
2385 both Isabelle/ML and Isabelle/Scala. Slightly reduced heap space usage.
2387 * Former "xsymbols" syntax with Isabelle symbols is used by default,
2388 without any special print mode. Important ASCII replacement syntax
2389 remains available under print mode "ASCII", but less important syntax
2390 has been removed (see below).
2392 * Support for more arrow symbols, with rendering in LaTeX and Isabelle
2393 fonts: \<Lleftarrow> \<Rrightarrow> \<longlongleftarrow> \<longlongrightarrow> \<longlonglongleftarrow> \<longlonglongrightarrow>.
2395 * Special notation \<struct> for the first implicit 'structure' in the
2396 context has been discontinued. Rare INCOMPATIBILITY, use explicit
2397 structure name instead, notably in indexed notation with block-subscript
2398 (e.g. \<odot>\<^bsub>A\<^esub>).
2400 * The glyph for \<diamond> in the IsabelleText font now corresponds better to its
2401 counterpart \<box> as quantifier-like symbol. A small diamond is available as
2402 \<diamondop>; the old symbol \<struct> loses this rendering and any special
2405 * Syntax for formal comments "-- text" now also supports the symbolic
2406 form "\<comment> text". Command-line tool "isabelle update_cartouches -c" helps
2407 to update old sources.
2409 * Toplevel theorem statements have been simplified as follows:
2412 schematic_lemma ~> schematic_goal
2413 schematic_theorem ~> schematic_goal
2414 schematic_corollary ~> schematic_goal
2416 Command-line tool "isabelle update_theorems" updates theory sources
2419 * Toplevel theorem statement 'proposition' is another alias for
2422 * The old 'defs' command has been removed (legacy since Isabelle2014).
2423 INCOMPATIBILITY, use regular 'definition' instead. Overloaded and/or
2424 deferred definitions require a surrounding 'overloading' block.
2427 *** Prover IDE -- Isabelle/Scala/jEdit ***
2429 * IDE support for the source-level debugger of Poly/ML, to work with
2430 Isabelle/ML and official Standard ML. Option "ML_debugger" and commands
2431 'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug',
2432 'SML_file_no_debug' control compilation of sources with or without
2433 debugging information. The Debugger panel allows to set breakpoints (via
2434 context menu), step through stopped threads, evaluate local ML
2435 expressions etc. At least one Debugger view needs to be active to have
2436 any effect on the running ML program.
2438 * The State panel manages explicit proof state output, with dynamic
2439 auto-update according to cursor movement. Alternatively, the jEdit
2440 action "isabelle.update-state" (shortcut S+ENTER) triggers manual
2443 * The Output panel no longer shows proof state output by default, to
2444 avoid GUI overcrowding. INCOMPATIBILITY, use the State panel instead or
2445 enable option "editor_output_state".
2447 * The text overview column (status of errors, warnings etc.) is updated
2448 asynchronously, leading to much better editor reactivity. Moreover, the
2449 full document node content is taken into account. The width of the
2450 column is scaled according to the main text area font, for improved
2453 * The main text area no longer changes its color hue in outdated
2454 situations. The text overview column takes over the role to indicate
2455 unfinished edits in the PIDE pipeline. This avoids flashing text display
2456 due to ad-hoc updates by auxiliary GUI components, such as the State
2459 * Slightly improved scheduling for urgent print tasks (e.g. command
2460 state output, interactive queries) wrt. long-running background tasks.
2462 * Completion of symbols via prefix of \<name> or \<^name> or \name is
2463 always possible, independently of the language context. It is never
2464 implicit: a popup will show up unconditionally.
2466 * Additional abbreviations for syntactic completion may be specified in
2467 $ISABELLE_HOME/etc/abbrevs and $ISABELLE_HOME_USER/etc/abbrevs, with
2468 support for simple templates using ASCII 007 (bell) as placeholder.
2470 * Symbols \<oplus>, \<Oplus>, \<otimes>, \<Otimes>, \<odot>, \<Odot>, \<ominus>, \<oslash> no longer provide abbreviations for
2471 completion like "+o", "*o", ".o" etc. -- due to conflicts with other
2472 ASCII syntax. INCOMPATIBILITY, use plain backslash-completion or define
2473 suitable abbreviations in $ISABELLE_HOME_USER/etc/abbrevs.
2475 * Action "isabelle-emph" (with keyboard shortcut C+e LEFT) controls
2476 emphasized text style; the effect is visible in document output, not in
2479 * Action "isabelle-reset" now uses keyboard shortcut C+e BACK_SPACE,
2480 instead of former C+e LEFT.
2482 * The command-line tool "isabelle jedit" and the isabelle.Main
2483 application wrapper treat the default $USER_HOME/Scratch.thy more
2484 uniformly, and allow the dummy file argument ":" to open an empty buffer
2487 * New command-line tool "isabelle jedit_client" allows to connect to an
2488 already running Isabelle/jEdit process. This achieves the effect of
2489 single-instance applications seen on common GUI desktops.
2491 * The default look-and-feel for Linux is the traditional "Metal", which
2492 works better with GUI scaling for very high-resolution displays (e.g.
2493 4K). Moreover, it is generally more robust than "Nimbus".
2495 * Update to jedit-5.3.0, with improved GUI scaling and support of
2496 high-resolution displays (e.g. 4K).
2498 * The main Isabelle executable is managed as single-instance Desktop
2499 application uniformly on all platforms: Linux, Windows, Mac OS X.
2502 *** Document preparation ***
2504 * Commands 'paragraph' and 'subparagraph' provide additional section
2505 headings. Thus there are 6 levels of standard headings, as in HTML.
2507 * Command 'text_raw' has been clarified: input text is processed as in
2508 'text' (with antiquotations and control symbols). The key difference is
2509 the lack of the surrounding isabelle markup environment in output.
2511 * Text is structured in paragraphs and nested lists, using notation that
2512 is similar to Markdown. The control symbols for list items are as
2517 \<^descr> description
2519 * There is a new short form for antiquotations with a single argument
2520 that is a cartouche: \<^name>\<open>...\<close> is equivalent to @{name \<open>...\<close>} and
2521 \<open>...\<close> without control symbol is equivalent to @{cartouche \<open>...\<close>}.
2522 \<^name> without following cartouche is equivalent to @{name}. The
2523 standard Isabelle fonts provide glyphs to render important control
2524 symbols, e.g. "\<^verbatim>", "\<^emph>", "\<^bold>".
2526 * Antiquotations @{noindent}, @{smallskip}, @{medskip}, @{bigskip} with
2527 corresponding control symbols \<^noindent>, \<^smallskip>, \<^medskip>, \<^bigskip> specify spacing formally, using
2528 standard LaTeX macros of the same names.
2530 * Antiquotation @{cartouche} in Isabelle/Pure is the same as @{text}.
2531 Consequently, \<open>...\<close> without any decoration prints literal quasi-formal
2532 text. Command-line tool "isabelle update_cartouches -t" helps to update
2533 old sources, by approximative patching of the content of string and
2534 cartouche tokens seen in theory sources.
2536 * The @{text} antiquotation now ignores the antiquotation option
2537 "source". The given text content is output unconditionally, without any
2538 surrounding quotes etc. Subtle INCOMPATIBILITY, put quotes into the
2539 argument where they are really intended, e.g. @{text \<open>"foo"\<close>}. Initial
2540 or terminal spaces are ignored.
2542 * Antiquotations @{emph} and @{bold} output LaTeX source recursively,
2543 adding appropriate text style markup. These may be used in the short
2544 form \<^emph>\<open>...\<close> and \<^bold>\<open>...\<close>.
2546 * Document antiquotation @{footnote} outputs LaTeX source recursively,
2547 marked as \footnote{}. This may be used in the short form \<^footnote>\<open>...\<close>.
2549 * Antiquotation @{verbatim [display]} supports option "indent".
2551 * Antiquotation @{theory_text} prints uninterpreted theory source text
2552 (Isar outer syntax with command keywords etc.). This may be used in the
2553 short form \<^theory_text>\<open>...\<close>. @{theory_text [display]} supports option "indent".
2555 * Antiquotation @{doc ENTRY} provides a reference to the given
2556 documentation, with a hyperlink in the Prover IDE.
2558 * Antiquotations @{command}, @{method}, @{attribute} print checked
2559 entities of the Isar language.
2561 * HTML presentation uses the standard IsabelleText font and Unicode
2562 rendering of Isabelle symbols like Isabelle/Scala/jEdit. The former
2563 print mode "HTML" loses its special meaning.
2568 * Local goals ('have', 'show', 'hence', 'thus') allow structured rule
2569 statements like fixes/assumes/shows in theorem specifications, but the
2570 notation is postfix with keywords 'if' (or 'when') and 'for'. For
2573 have result: "C x y"
2575 for x :: 'a and y :: 'a
2578 The local assumptions are bound to the name "that". The result is
2579 exported from context of the statement as usual. The above roughly
2580 corresponds to a raw proof block like this:
2583 fix x :: 'a and y :: 'a
2584 assume that: "A x" "B y"
2585 have "C x y" <proof>
2589 The keyword 'when' may be used instead of 'if', to indicate 'presume'
2590 instead of 'assume' above.
2592 * Assumptions ('assume', 'presume') allow structured rule statements
2593 using 'if' and 'for', similar to 'have' etc. above. For example:
2595 assume result: "C x y"
2597 for x :: 'a and y :: 'a
2599 This assumes "\<And>x y::'a. A x \<Longrightarrow> B y \<Longrightarrow> C x y" and produces a general
2600 result as usual: "A ?x \<Longrightarrow> B ?y \<Longrightarrow> C ?x ?y".
2602 Vacuous quantification in assumptions is omitted, i.e. a for-context
2603 only effects propositions according to actual use of variables. For
2606 assume "A x" and "B y" for x and y
2610 assume "\<And>x. A x" and "\<And>y. B y"
2612 * The meaning of 'show' with Pure rule statements has changed: premises
2613 are treated in the sense of 'assume', instead of 'presume'. This means,
2614 a goal like "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x" can be solved completely as
2617 show "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
2621 show "C x" if "A x" "B x" for x
2623 Rare INCOMPATIBILITY, the old behaviour may be recovered as follows:
2625 show "C x" when "A x" "B x" for x
2627 * New command 'consider' states rules for generalized elimination and
2628 case splitting. This is like a toplevel statement "theorem obtains" used
2629 within a proof body; or like a multi-branch 'obtain' without activation
2630 of the local context elements yet.
2632 * Proof method "cases" allows to specify the rule as first entry of
2633 chained facts. This is particularly useful with 'consider':
2635 consider (a) A | (b) B | (c) C <proof>
2639 then show ?thesis <proof>
2642 then show ?thesis <proof>
2645 then show ?thesis <proof>
2648 * Command 'case' allows fact name and attribute specification like this:
2651 case a [attributes]: (c xs)
2653 Facts that are introduced by invoking the case context are uniformly
2654 qualified by "a"; the same name is used for the cumulative fact. The old
2655 form "case (c xs) [attributes]" is no longer supported. Rare
2656 INCOMPATIBILITY, need to adapt uses of case facts in exotic situations,
2657 and always put attributes in front.
2659 * The standard proof method of commands 'proof' and '..' is now called
2660 "standard" to make semantically clear what it is; the old name "default"
2661 is still available as legacy for some time. Documentation now explains
2662 '..' more accurately as "by standard" instead of "by rule".
2664 * Nesting of Isar goal structure has been clarified: the context after
2665 the initial backwards refinement is retained for the whole proof, within
2666 all its context sections (as indicated via 'next'). This is e.g.
2667 relevant for 'using', 'including', 'supply':
2669 have "A \<and> A" if a: A for A
2677 * Command 'obtain' binds term abbreviations (via 'is' patterns) in the
2678 proof body as well, abstracted over relevant parameters.
2680 * Improved type-inference for theorem statement 'obtains': separate
2681 parameter scope for of each clause.
2683 * Term abbreviations via 'is' patterns also work for schematic
2684 statements: result is abstracted over unknowns.
2686 * Command 'subgoal' allows to impose some structure on backward
2687 refinements, to avoid proof scripts degenerating into long of 'apply'
2688 sequences. Further explanations and examples are given in the isar-ref
2691 * Command 'supply' supports fact definitions during goal refinement
2694 * Proof method "goal_cases" turns the current subgoals into cases within
2695 the context; the conclusion is bound to variable ?case in each case. For
2698 lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
2699 and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
2702 then show ?case using \<open>A x\<close> \<open>B x\<close> sorry
2705 then show ?case using \<open>U y\<close> \<open>V z\<close> sorry
2708 lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
2709 and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
2712 then show ?case using prems sorry
2715 then show ?case using prems sorry
2718 * The undocumented feature of implicit cases goal1, goal2, goal3, etc.
2719 is marked as legacy, and will be removed eventually. The proof method
2720 "goals" achieves a similar effect within regular Isar; often it can be
2721 done more adequately by other means (e.g. 'consider').
2723 * The vacuous fact "TERM x" may be established "by fact" or as `TERM x`
2724 as well, not just "by this" or "." as before.
2726 * Method "sleep" succeeds after a real-time delay (in seconds). This is
2727 occasionally useful for demonstration and testing purposes.
2732 * Qualifiers in locale expressions default to mandatory ('!') regardless
2733 of the command. Previously, for 'locale' and 'sublocale' the default was
2734 optional ('?'). The old synatx '!' has been discontinued.
2735 INCOMPATIBILITY, remove '!' and add '?' as required.
2737 * Keyword 'rewrites' identifies rewrite morphisms in interpretation
2738 commands. Previously, the keyword was 'where'. INCOMPATIBILITY.
2740 * More gentle suppression of syntax along locale morphisms while
2741 printing terms. Previously 'abbreviation' and 'notation' declarations
2742 would be suppressed for morphisms except term identity. Now
2743 'abbreviation' is also kept for morphims that only change the involved
2744 parameters, and only 'notation' is suppressed. This can be of great help
2745 when working with complex locale hierarchies, because proof states are
2746 displayed much more succinctly. It also means that only notation needs
2747 to be redeclared if desired, as illustrated by this example:
2749 locale struct = fixes composition :: "'a => 'a => 'a" (infixl "\<cdot>" 65)
2751 definition derived (infixl "\<odot>" 65) where ...
2755 left: struct composition + right: struct composition'
2756 for composition (infix "\<cdot>" 65) and composition' (infix "\<cdot>''" 65)
2758 notation right.derived ("\<odot>''")
2761 * Command 'global_interpretation' issues interpretations into global
2762 theories, with optional rewrite definitions following keyword 'defines'.
2764 * Command 'sublocale' accepts optional rewrite definitions after keyword
2767 * Command 'permanent_interpretation' has been discontinued. Use
2768 'global_interpretation' or 'sublocale' instead. INCOMPATIBILITY.
2770 * Command 'print_definitions' prints dependencies of definitional
2771 specifications. This functionality used to be part of 'print_theory'.
2773 * Configuration option rule_insts_schematic has been discontinued
2774 (intermediate legacy feature in Isabelle2015). INCOMPATIBILITY.
2776 * Abbreviations in type classes now carry proper sort constraint. Rare
2777 INCOMPATIBILITY in situations where the previous misbehaviour has been
2780 * Refinement of user-space type system in type classes: pseudo-local
2781 operations behave more similar to abbreviations. Potential
2782 INCOMPATIBILITY in exotic situations.
2787 * The 'typedef' command has been upgraded from a partially checked
2788 "axiomatization", to a full definitional specification that takes the
2789 global collection of overloaded constant / type definitions into
2790 account. Type definitions with open dependencies on overloaded
2791 definitions need to be specified as "typedef (overloaded)". This
2792 provides extra robustness in theory construction. Rare INCOMPATIBILITY.
2794 * Qualification of various formal entities in the libraries is done more
2795 uniformly via "context begin qualified definition ... end" instead of
2796 old-style "hide_const (open) ...". Consequently, both the defined
2797 constant and its defining fact become qualified, e.g. Option.is_none and
2798 Option.is_none_def. Occasional INCOMPATIBILITY in applications.
2800 * Some old and rarely used ASCII replacement syntax has been removed.
2801 INCOMPATIBILITY, standard syntax with symbols should be used instead.
2802 The subsequent commands help to reproduce the old forms, e.g. to
2803 simplify porting old theories:
2805 notation iff (infixr "<->" 25)
2807 notation Times (infixr "<*>" 80)
2809 type_notation Map.map (infixr "~=>" 0)
2810 notation Map.map_comp (infixl "o'_m" 55)
2812 type_notation FinFun.finfun ("(_ =>f /_)" [22, 21] 21)
2814 notation FuncSet.funcset (infixr "->" 60)
2815 notation FuncSet.extensional_funcset (infixr "->\<^sub>E" 60)
2817 notation Omega_Words_Fun.conc (infixr "conc" 65)
2819 notation Preorder.equiv ("op ~~")
2820 and Preorder.equiv ("(_/ ~~ _)" [51, 51] 50)
2822 notation (in topological_space) tendsto (infixr "--->" 55)
2823 notation (in topological_space) LIMSEQ ("((_)/ ----> (_))" [60, 60] 60)
2824 notation LIM ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
2826 notation NSA.approx (infixl "@=" 50)
2827 notation NSLIMSEQ ("((_)/ ----NS> (_))" [60, 60] 60)
2828 notation NSLIM ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
2830 * The alternative notation "\<Colon>" for type and sort constraints has been
2831 removed: in LaTeX document output it looks the same as "::".
2832 INCOMPATIBILITY, use plain "::" instead.
2834 * Commands 'inductive' and 'inductive_set' work better when names for
2835 intro rules are omitted: the "cases" and "induct" rules no longer
2836 declare empty case_names, but no case_names at all. This allows to use
2837 numbered cases in proofs, without requiring method "goal_cases".
2839 * Inductive definitions ('inductive', 'coinductive', etc.) expose
2840 low-level facts of the internal construction only if the option
2841 "inductive_internals" is enabled. This refers to the internal predicate
2842 definition and its monotonicity result. Rare INCOMPATIBILITY.
2844 * Recursive function definitions ('fun', 'function', 'partial_function')
2845 expose low-level facts of the internal construction only if the option
2846 "function_internals" is enabled. Its internal inductive definition is
2847 also subject to "inductive_internals". Rare INCOMPATIBILITY.
2849 * BNF datatypes ('datatype', 'codatatype', etc.) expose low-level facts
2850 of the internal construction only if the option "bnf_internals" is
2851 enabled. This supersedes the former option "bnf_note_all". Rare
2854 * Combinator to represent case distinction on products is named
2855 "case_prod", uniformly, discontinuing any input aliasses. Very popular
2856 theorem aliasses have been retained.
2859 PairE ~> prod.exhaust
2860 Pair_eq ~> prod.inject
2861 pair_collapse ~> prod.collapse
2862 Pair_fst_snd_eq ~> prod_eq_iff
2863 split_twice ~> prod.case_distrib
2864 split_weak_cong ~> prod.case_cong_weak
2865 split_split ~> prod.split
2866 split_split_asm ~> prod.split_asm
2867 splitI ~> case_prodI
2868 splitD ~> case_prodD
2869 splitI2 ~> case_prodI2
2870 splitI2' ~> case_prodI2'
2871 splitE ~> case_prodE
2872 splitE' ~> case_prodE'
2873 split_pair ~> case_prod_Pair
2874 split_eta ~> case_prod_eta
2875 split_comp ~> case_prod_comp
2876 mem_splitI ~> mem_case_prodI
2877 mem_splitI2 ~> mem_case_prodI2
2878 mem_splitE ~> mem_case_prodE
2879 The_split ~> The_case_prod
2880 cond_split_eta ~> cond_case_prod_eta
2881 Collect_split_in_rel_leE ~> Collect_case_prod_in_rel_leE
2882 Collect_split_in_rel_leI ~> Collect_case_prod_in_rel_leI
2883 in_rel_Collect_split_eq ~> in_rel_Collect_case_prod_eq
2884 Collect_split_Grp_eqD ~> Collect_case_prod_Grp_eqD
2885 Collect_split_Grp_inD ~> Collect_case_prod_Grp_in
2886 Domain_Collect_split ~> Domain_Collect_case_prod
2887 Image_Collect_split ~> Image_Collect_case_prod
2888 Range_Collect_split ~> Range_Collect_case_prod
2889 Eps_split ~> Eps_case_prod
2890 Eps_split_eq ~> Eps_case_prod_eq
2891 split_rsp ~> case_prod_rsp
2892 curry_split ~> curry_case_prod
2893 split_curry ~> case_prod_curry
2895 Changes in structure HOLogic:
2896 split_const ~> case_prod_const
2897 mk_split ~> mk_case_prod
2898 mk_psplits ~> mk_ptupleabs
2899 strip_psplits ~> strip_ptupleabs
2903 * The coercions to type 'real' have been reorganised. The function
2904 'real' is no longer overloaded, but has type 'nat => real' and
2905 abbreviates of_nat for that type. Also 'real_of_int :: int => real'
2906 abbreviates of_int for that type. Other overloaded instances of 'real'
2907 have been replaced by 'real_of_ereal' and 'real_of_float'.
2909 Consolidated facts (among others):
2910 real_of_nat_le_iff -> of_nat_le_iff
2911 real_of_nat_numeral of_nat_numeral
2912 real_of_int_zero of_int_0
2913 real_of_nat_zero of_nat_0
2914 real_of_one of_int_1
2915 real_of_int_add of_int_add
2916 real_of_nat_add of_nat_add
2917 real_of_int_diff of_int_diff
2918 real_of_nat_diff of_nat_diff
2919 floor_subtract floor_diff_of_int
2920 real_of_int_inject of_int_eq_iff
2921 real_of_int_gt_zero_cancel_iff of_int_0_less_iff
2922 real_of_int_ge_zero_cancel_iff of_int_0_le_iff
2923 real_of_nat_ge_zero of_nat_0_le_iff
2924 real_of_int_ceiling_ge le_of_int_ceiling
2925 ceiling_less_eq ceiling_less_iff
2926 ceiling_le_eq ceiling_le_iff
2927 less_floor_eq less_floor_iff
2928 floor_less_eq floor_less_iff
2929 floor_divide_eq_div floor_divide_of_int_eq
2930 real_of_int_zero_cancel of_nat_eq_0_iff
2931 ceiling_real_of_int ceiling_of_int
2935 * Theory Map: lemma map_of_is_SomeD was a clone of map_of_SomeD and has
2936 been removed. INCOMPATIBILITY.
2938 * Quickcheck setup for finite sets.
2940 * Discontinued simp_legacy_precond. Potential INCOMPATIBILITY.
2943 - The MaSh relevance filter has been sped up.
2944 - Proof reconstruction has been improved, to minimize the incidence of
2945 cases where Sledgehammer gives a proof that does not work.
2946 - Auto Sledgehammer now minimizes and preplays the results.
2947 - Handle Vampire 4.0 proof output without raising exception.
2948 - Eliminated "MASH" environment variable. Use the "MaSh" option in
2949 Isabelle/jEdit instead. INCOMPATIBILITY.
2950 - Eliminated obsolete "blocking" option and related subcommands.
2953 - Fixed soundness bug in translation of "finite" predicate.
2954 - Fixed soundness bug in "destroy_constrs" optimization.
2955 - Fixed soundness bug in translation of "rat" type.
2956 - Removed "check_potential" and "check_genuine" options.
2957 - Eliminated obsolete "blocking" option.
2959 * (Co)datatype package:
2960 - New commands "lift_bnf" and "copy_bnf" for lifting (copying) a BNF
2961 structure on the raw type to an abstract type defined using typedef.
2962 - Always generate "case_transfer" theorem.
2963 - For mutual types, generate slightly stronger "rel_induct",
2964 "rel_coinduct", and "coinduct" theorems. INCOMPATIBILITY.
2965 - Allow discriminators and selectors with the same name as the type
2967 - Avoid various internal name clashes (e.g., 'datatype f = f').
2969 * Transfer: new methods for interactive debugging of 'transfer' and
2970 'transfer_prover': 'transfer_start', 'transfer_step', 'transfer_end',
2971 'transfer_prover_start' and 'transfer_prover_end'.
2973 * New diagnostic command print_record for displaying record definitions.
2975 * Division on integers is bootstrapped directly from division on
2976 naturals and uses generic numeral algorithm for computations. Slight
2977 INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes former
2978 simprocs binary_int_div and binary_int_mod
2980 * Tightened specification of class semiring_no_zero_divisors. Minor
2983 * Class algebraic_semidom introduces common algebraic notions of
2984 integral (semi)domains, particularly units. Although logically subsumed
2985 by fields, is is not a super class of these in order not to burden
2986 fields with notions that are trivial there.
2988 * Class normalization_semidom specifies canonical representants for
2989 equivalence classes of associated elements in an integral (semi)domain.
2990 This formalizes associated elements as well.
2992 * Abstract specification of gcd/lcm operations in classes semiring_gcd,
2993 semiring_Gcd, semiring_Lcd. Minor INCOMPATIBILITY: facts gcd_nat.commute
2994 and gcd_int.commute are subsumed by gcd.commute, as well as
2995 gcd_nat.assoc and gcd_int.assoc by gcd.assoc.
2997 * Former constants Fields.divide (_ / _) and Divides.div (_ div _) are
2998 logically unified to Rings.divide in syntactic type class Rings.divide,
2999 with infix syntax (_ div _). Infix syntax (_ / _) for field division is
3000 added later as abbreviation in class Fields.inverse. INCOMPATIBILITY,
3001 instantiations must refer to Rings.divide rather than the former
3002 separate constants, hence infix syntax (_ / _) is usually not available
3003 during instantiation.
3005 * New cancellation simprocs for boolean algebras to cancel complementary
3006 terms for sup and inf. For example, "sup x (sup y (- x))" simplifies to
3007 "top". INCOMPATIBILITY.
3009 * Class uniform_space introduces uniform spaces btw topological spaces
3010 and metric spaces. Minor INCOMPATIBILITY: open_<type>_def needs to be
3011 introduced in the form of an uniformity. Some constants are more general
3012 now, it may be necessary to add type class constraints.
3014 open_real_def \<leadsto> open_dist
3015 open_complex_def \<leadsto> open_dist
3017 * Library/Monad_Syntax: notation uses symbols \<bind> and \<then>. INCOMPATIBILITY.
3020 - Renamed multiset inclusion operators:
3028 - Added multiset inclusion operator syntax:
3033 - "'a multiset" is no longer an instance of the "order",
3034 "ordered_ab_semigroup_add_imp_le", "ordered_cancel_comm_monoid_diff",
3035 "semilattice_inf", and "semilattice_sup" type classes. The theorems
3036 previously provided by these type classes (directly or indirectly)
3037 are now available through the "subset_mset" interpretation
3038 (e.g. add_mono ~> subset_mset.add_mono).
3040 - Renamed conversions:
3042 multiset_of_set ~> mset_set
3046 mset_le_def ~> subseteq_mset_def
3047 mset_less_def ~> subset_mset_def
3048 less_eq_multiset.rep_eq ~> subseteq_mset_def
3050 - Removed lemmas generated by lift_definition:
3051 less_eq_multiset.abs_eq, less_eq_multiset.rsp,
3052 less_eq_multiset.transfer, less_eq_multiset_def
3055 * Library/Omega_Words_Fun: Infinite words modeled as functions nat \<Rightarrow> 'a.
3057 * Library/Bourbaki_Witt_Fixpoint: Added formalisation of the
3058 Bourbaki-Witt fixpoint theorem for increasing functions in
3059 chain-complete partial orders.
3061 * Library/Old_Recdef: discontinued obsolete 'defer_recdef' command.
3062 Minor INCOMPATIBILITY, use 'function' instead.
3064 * Library/Periodic_Fun: a locale that provides convenient lemmas for
3067 * Library/Formal_Power_Series: proper definition of division (with
3068 remainder) for formal power series; instances for Euclidean Ring and
3071 * HOL-Imperative_HOL: obsolete theory Legacy_Mrec has been removed.
3073 * HOL-Statespace: command 'statespace' uses mandatory qualifier for
3074 import of parent, as for general 'locale' expressions. INCOMPATIBILITY,
3075 remove '!' and add '?' as required.
3077 * HOL-Decision_Procs: The "approximation" method works with "powr"
3078 (exponentiation on real numbers) again.
3080 * HOL-Multivariate_Analysis: theory Cauchy_Integral_Thm with Contour
3081 integrals (= complex path integrals), Cauchy's integral theorem, winding
3082 numbers and Cauchy's integral formula, Liouville theorem, Fundamental
3083 Theorem of Algebra. Ported from HOL Light.
3085 * HOL-Multivariate_Analysis: topological concepts such as connected
3086 components, homotopic paths and the inside or outside of a set.
3088 * HOL-Multivariate_Analysis: radius of convergence of power series and
3089 various summability tests; Harmonic numbers and the Euler–Mascheroni
3090 constant; the Generalised Binomial Theorem; the complex and real
3091 Gamma/log-Gamma/Digamma/ Polygamma functions and their most important
3094 * HOL-Probability: The central limit theorem based on Levy's uniqueness
3095 and continuity theorems, weak convergence, and characterisitc functions.
3097 * HOL-Data_Structures: new and growing session of standard data
3103 * The following combinators for low-level profiling of the ML runtime
3104 system are available:
3106 profile_time (*CPU time*)
3107 profile_time_thread (*CPU time on this thread*)
3108 profile_allocations (*overall heap allocations*)
3110 * Antiquotation @{undefined} or \<^undefined> inlines (raise Match).
3112 * Antiquotation @{method NAME} inlines the (checked) name of the given
3115 * Pretty printing of Poly/ML compiler output in Isabelle has been
3116 improved: proper treatment of break offsets and blocks with consistent
3119 * The auxiliary module Pure/display.ML has been eliminated. Its
3120 elementary thm print operations are now in Pure/more_thm.ML and thus
3121 called Thm.pretty_thm, Thm.string_of_thm etc. INCOMPATIBILITY.
3123 * Simproc programming interfaces have been simplified:
3124 Simplifier.make_simproc and Simplifier.define_simproc supersede various
3125 forms of Simplifier.mk_simproc, Simplifier.simproc_global etc. Note that
3126 term patterns for the left-hand sides are specified with implicitly
3127 fixed variables, like top-level theorem statements. INCOMPATIBILITY.
3129 * Instantiation rules have been re-organized as follows:
3131 Thm.instantiate (*low-level instantiation with named arguments*)
3132 Thm.instantiate' (*version with positional arguments*)
3134 Drule.infer_instantiate (*instantiation with type inference*)
3135 Drule.infer_instantiate' (*version with positional arguments*)
3137 The LHS only requires variable specifications, instead of full terms.
3138 Old cterm_instantiate is superseded by infer_instantiate.
3139 INCOMPATIBILITY, need to re-adjust some ML names and types accordingly.
3141 * Old tactic shorthands atac, rtac, etac, dtac, ftac have been
3142 discontinued. INCOMPATIBILITY, use regular assume_tac, resolve_tac etc.
3143 instead (with proper context).
3145 * Thm.instantiate (and derivatives) no longer require the LHS of the
3146 instantiation to be certified: plain variables are given directly.
3148 * Subgoal.SUBPROOF and Subgoal.FOCUS combinators use anonymous
3149 quasi-bound variables (like the Simplifier), instead of accidentally
3150 named local fixes. This has the potential to improve stability of proof
3151 tools, but can also cause INCOMPATIBILITY for tools that don't observe
3152 the proof context discipline.
3154 * Isar proof methods are based on a slightly more general type
3155 context_tactic, which allows to change the proof context dynamically
3156 (e.g. to update cases) and indicate explicit Seq.Error results. Former
3157 METHOD_CASES is superseded by CONTEXT_METHOD; further combinators are
3158 provided in src/Pure/Isar/method.ML for convenience. INCOMPATIBILITY.
3163 * Command-line tool "isabelle console" enables print mode "ASCII".
3165 * Command-line tool "isabelle update_then" expands old Isar command
3171 This syntax is more orthogonal and improves readability and
3172 maintainability of proofs.
3174 * Global session timeout is multiplied by timeout_scale factor. This
3175 allows to adjust large-scale tests (e.g. AFP) to overall hardware
3178 * Property values in etc/symbols may contain spaces, if written with the
3179 replacement character "␣" (Unicode point 0x2324). For example:
3181 \<star> code: 0x0022c6 group: operator font: Deja␣Vu␣Sans␣Mono
3183 * Java runtime environment for x86_64-windows allows to use larger heap
3186 * Java runtime options are determined separately for 32bit vs. 64bit
3187 platforms as follows.
3189 - Isabelle desktop application: platform-specific files that are
3190 associated with the main app bundle
3192 - isabelle jedit: settings
3193 JEDIT_JAVA_SYSTEM_OPTIONS
3194 JEDIT_JAVA_OPTIONS32 vs. JEDIT_JAVA_OPTIONS64
3196 - isabelle build: settings
3197 ISABELLE_BUILD_JAVA_OPTIONS32 vs. ISABELLE_BUILD_JAVA_OPTIONS64
3199 * Bash shell function "jvmpath" has been renamed to "platform_path": it
3200 is relevant both for Poly/ML and JVM processes.
3202 * Poly/ML default platform architecture may be changed from 32bit to
3203 64bit via system option ML_system_64. A system restart (and rebuild) is
3204 required after change.
3206 * Poly/ML 5.6 runs natively on x86-windows and x86_64-windows, which
3207 both allow larger heap space than former x86-cygwin.
3209 * Heap images are 10-15% smaller due to less wasteful persistent theory
3210 content (using ML type theory_id instead of theory);
3214 New in Isabelle2015 (May 2015)
3215 ------------------------------
3219 * Local theory specification commands may have a 'private' or
3220 'qualified' modifier to restrict name space accesses to the local scope,
3221 as provided by some "context begin ... end" block. For example:
3226 private definition ...
3229 qualified definition ...
3237 * Command 'experiment' opens an anonymous locale context with private
3240 * Command 'notepad' requires proper nesting of begin/end and its proof
3241 structure in the body: 'oops' is no longer supported here. Minor
3242 INCOMPATIBILITY, use 'sorry' instead.
3244 * Command 'named_theorems' declares a dynamic fact within the context,
3245 together with an attribute to maintain the content incrementally. This
3246 supersedes functor Named_Thms in Isabelle/ML, but with a subtle change
3247 of semantics due to external visual order vs. internal reverse order.
3249 * 'find_theorems': search patterns which are abstractions are
3250 schematically expanded before search. Search results match the naive
3251 expectation more closely, particularly wrt. abbreviations.
3254 * Commands 'method_setup' and 'attribute_setup' now work within a local
3257 * Outer syntax commands are managed authentically within the theory
3258 context, without implicit global state. Potential for accidental
3259 INCOMPATIBILITY, make sure that required theories are really imported.
3261 * Historical command-line terminator ";" is no longer accepted (and
3262 already used differently in Isar). Minor INCOMPATIBILITY, use "isabelle
3263 update_semicolons" to remove obsolete semicolons from old theory
3266 * Structural composition of proof methods (meth1; meth2) in Isar
3267 corresponds to (tac1 THEN_ALL_NEW tac2) in ML.
3269 * The Eisbach proof method language allows to define new proof methods
3270 by combining existing ones with their usual syntax. The "match" proof
3271 method provides basic fact/term matching in addition to
3272 premise/conclusion matching through Subgoal.focus, and binds fact names
3273 from matches as well as term patterns within matches. The Isabelle
3274 documentation provides an entry "eisbach" for the Eisbach User Manual.
3275 Sources and various examples are in ~~/src/HOL/Eisbach/.
3278 *** Prover IDE -- Isabelle/Scala/jEdit ***
3280 * Improved folding mode "isabelle" based on Isar syntax. Alternatively,
3281 the "sidekick" mode may be used for document structure.
3283 * Extended bracket matching based on Isar language structure. System
3284 option jedit_structure_limit determines maximum number of lines to scan
3287 * Support for BibTeX files: context menu, context-sensitive token
3288 marker, SideKick parser.
3290 * Document antiquotation @{cite} provides formal markup, which is
3291 interpreted semi-formally based on .bib files that happen to be open in
3292 the editor (hyperlinks, completion etc.).
3294 * Less waste of vertical space via negative line spacing (see Global
3295 Options / Text Area).
3297 * Improved graphview panel with optional output of PNG or PDF, for
3298 display of 'thy_deps', 'class_deps' etc.
3300 * The commands 'thy_deps' and 'class_deps' allow optional bounds to
3301 restrict the visualized hierarchy.
3303 * Improved scheduling for asynchronous print commands (e.g. provers
3304 managed by the Sledgehammer panel) wrt. ongoing document processing.
3307 *** Document preparation ***
3309 * Document markup commands 'chapter', 'section', 'subsection',
3310 'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any
3311 context, even before the initial 'theory' command. Obsolete proof
3312 commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been
3313 discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw'
3314 instead. The old 'header' command is still retained for some time, but
3315 should be replaced by 'chapter', 'section' etc. (using "isabelle
3316 update_header"). Minor INCOMPATIBILITY.
3318 * Official support for "tt" style variants, via \isatt{...} or
3319 \begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or
3320 verbatim environment of LaTeX is no longer used. This allows @{ML} etc.
3321 as argument to other macros (such as footnotes).
3323 * Document antiquotation @{verbatim} prints ASCII text literally in "tt"
3326 * Discontinued obsolete option "document_graph": session_graph.pdf is
3327 produced unconditionally for HTML browser_info and PDF-LaTeX document.
3329 * Diagnostic commands and document markup commands within a proof do not
3330 affect the command tag for output. Thus commands like 'thm' are subject
3331 to proof document structure, and no longer "stick out" accidentally.
3332 Commands 'text' and 'txt' merely differ in the LaTeX style, not their
3333 tags. Potential INCOMPATIBILITY in exotic situations.
3335 * System option "pretty_margin" is superseded by "thy_output_margin",
3336 which is also accessible via document antiquotation option "margin".
3337 Only the margin for document output may be changed, but not the global
3338 pretty printing: that is 76 for plain console output, and adapted
3339 dynamically in GUI front-ends. Implementations of document
3340 antiquotations need to observe the margin explicitly according to
3341 Thy_Output.string_of_margin. Minor INCOMPATIBILITY.
3343 * Specification of 'document_files' in the session ROOT file is
3344 mandatory for document preparation. The legacy mode with implicit
3345 copying of the document/ directory is no longer supported. Minor
3351 * Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
3352 etc.) allow an optional context of local variables ('for' declaration):
3353 these variables become schematic in the instantiated theorem; this
3354 behaviour is analogous to 'for' in attributes "where" and "of".
3355 Configuration option rule_insts_schematic (default false) controls use
3356 of schematic variables outside the context. Minor INCOMPATIBILITY,
3357 declare rule_insts_schematic = true temporarily and update to use local
3358 variable declarations or dummy patterns instead.
3360 * Explicit instantiation via attributes "where", "of", and proof methods
3361 "rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns
3362 ("_") that stand for anonymous local variables.
3364 * Generated schematic variables in standard format of exported facts are
3365 incremented to avoid material in the proof context. Rare
3366 INCOMPATIBILITY, explicit instantiation sometimes needs to refer to
3369 * Lexical separation of signed and unsigned numerals: categories "num"
3370 and "float" are unsigned. INCOMPATIBILITY: subtle change in precedence
3371 of numeral signs, particularly in expressions involving infix syntax
3374 * Old inner token category "xnum" has been discontinued. Potential
3375 INCOMPATIBILITY for exotic syntax: may use mixfix grammar with "num"
3376 token category instead.
3381 * New (co)datatype package:
3382 - The 'datatype_new' command has been renamed 'datatype'. The old
3383 command of that name is now called 'old_datatype' and is provided
3384 by "~~/src/HOL/Library/Old_Datatype.thy". See
3385 'isabelle doc datatypes' for information on porting.
3388 disc_corec ~> corec_disc
3389 disc_corec_iff ~> corec_disc_iff
3390 disc_exclude ~> distinct_disc
3391 disc_exhaust ~> exhaust_disc
3392 disc_map_iff ~> map_disc_iff
3393 sel_corec ~> corec_sel
3394 sel_exhaust ~> exhaust_sel
3397 sel_split ~> split_sel
3398 sel_split_asm ~> split_sel_asm
3399 strong_coinduct ~> coinduct_strong
3400 weak_case_cong ~> case_cong_weak
3402 - The "no_code" option to "free_constructors", "datatype_new", and
3403 "codatatype" has been renamed "plugins del: code".
3405 - The rules "set_empty" have been removed. They are easy
3406 consequences of other set rules "by auto".
3408 - The rule "set_cases" is now registered with the "[cases set]"
3409 attribute. This can influence the behavior of the "cases" proof
3410 method when more than one case rule is applicable (e.g., an
3411 assumption is of the form "w : set ws" and the method "cases w"
3412 is invoked). The solution is to specify the case rule explicitly
3413 (e.g. "cases w rule: widget.exhaust").
3416 BNF_Comp ~> BNF_Composition
3417 BNF_FP_Base ~> BNF_Fixpoint_Base
3418 BNF_GFP ~> BNF_Greatest_Fixpoint
3419 BNF_LFP ~> BNF_Least_Fixpoint
3420 BNF_Constructions_on_Wellorders ~> BNF_Wellorder_Constructions
3421 Cardinals/Constructions_on_Wellorders ~> Cardinals/Wellorder_Constructions
3423 - Lifting and Transfer setup for basic HOL types sum and prod (also
3424 option) is now performed by the BNF package. Theories Lifting_Sum,
3425 Lifting_Product and Lifting_Option from Main became obsolete and
3426 were removed. Changed definitions of the relators rel_prod and
3427 rel_sum (using inductive).
3428 INCOMPATIBILITY: use rel_prod.simps and rel_sum.simps instead
3429 of rel_prod_def and rel_sum_def.
3430 Minor INCOMPATIBILITY: (rarely used by name) transfer theorem names
3431 changed (e.g. map_prod_transfer ~> prod.map_transfer).
3432 - Parametricity theorems for map functions, relators, set functions,
3433 constructors, case combinators, discriminators, selectors and
3434 (co)recursors are automatically proved and registered as transfer
3437 * Old datatype package:
3438 - The old 'datatype' command has been renamed 'old_datatype', and
3439 'rep_datatype' has been renamed 'old_rep_datatype'. They are
3440 provided by "~~/src/HOL/Library/Old_Datatype.thy". See
3441 'isabelle doc datatypes' for information on porting.
3444 weak_case_cong ~> case_cong_weak
3447 ~~/src/HOL/Datatype.thy ~> ~~/src/HOL/Library/Old_Datatype.thy
3451 - Fixed soundness bug related to the strict and non-strict subset
3455 - CVC4 is now included with Isabelle instead of CVC3 and run by
3457 - Z3 is now always enabled by default, now that it is fully open
3458 source. The "z3_non_commercial" option is discontinued.
3459 - Minimization is now always enabled by default.
3460 Removed sub-command:
3462 - Proof reconstruction, both one-liners and Isar, has been
3463 dramatically improved.
3464 - Improved support for CVC4 and veriT.
3466 * Old and new SMT modules:
3467 - The old 'smt' method has been renamed 'old_smt' and moved to
3468 'src/HOL/Library/Old_SMT.thy'. It is provided for compatibility,
3469 until applications have been ported to use the new 'smt' method. For
3470 the method to work, an older version of Z3 (e.g. Z3 3.2 or 4.0) must
3471 be installed, and the environment variable "OLD_Z3_SOLVER" must
3474 - The 'smt2' method has been renamed 'smt'.
3476 - New option 'smt_reconstruction_step_timeout' to limit the
3477 reconstruction time of Z3 proof steps in the new 'smt' method.
3478 - New option 'smt_statistics' to display statistics of the new 'smt'
3479 method, especially runtime statistics of Z3 proof reconstruction.
3481 * Lifting: command 'lift_definition' allows to execute lifted constants
3482 that have as a return type a datatype containing a subtype. This
3483 overcomes long-time limitations in the area of code generation and
3484 lifting, and avoids tedious workarounds.
3486 * Command and antiquotation "value" provide different evaluation slots
3487 (again), where the previous strategy (NBE after ML) serves as default.
3488 Minor INCOMPATIBILITY.
3490 * Add NO_MATCH-simproc, allows to check for syntactic non-equality.
3492 * field_simps: Use NO_MATCH-simproc for distribution rules, to avoid
3493 non-termination in case of distributing a division. With this change
3494 field_simps is in some cases slightly less powerful, if it fails try to
3495 add algebra_simps, or use divide_simps. Minor INCOMPATIBILITY.
3497 * Separate class no_zero_divisors has been given up in favour of fully
3498 algebraic semiring_no_zero_divisors. INCOMPATIBILITY.
3500 * Class linordered_semidom really requires no zero divisors.
3503 * Classes division_ring, field and linordered_field always demand
3504 "inverse 0 = 0". Given up separate classes division_ring_inverse_zero,
3505 field_inverse_zero and linordered_field_inverse_zero. INCOMPATIBILITY.
3507 * Classes cancel_ab_semigroup_add / cancel_monoid_add specify explicit
3508 additive inverse operation. INCOMPATIBILITY.
3510 * Complex powers and square roots. The functions "ln" and "powr" are now
3511 overloaded for types real and complex, and 0 powr y = 0 by definition.
3512 INCOMPATIBILITY: type constraints may be necessary.
3514 * The functions "sin" and "cos" are now defined for any type of sort
3515 "{real_normed_algebra_1,banach}" type, so in particular on "real" and
3516 "complex" uniformly. Minor INCOMPATIBILITY: type constraints may be
3519 * New library of properties of the complex transcendental functions sin,
3520 cos, tan, exp, Ln, Arctan, Arcsin, Arccos. Ported from HOL Light.
3522 * The factorial function, "fact", now has type "nat => 'a" (of a sort
3523 that admits numeric types including nat, int, real and complex.
3524 INCOMPATIBILITY: an expression such as "fact 3 = 6" may require a type
3525 constraint, and the combination "real (fact k)" is likely to be
3526 unsatisfactory. If a type conversion is still necessary, then use
3527 "of_nat (fact k)" or "real_of_nat (fact k)".
3529 * Removed functions "natfloor" and "natceiling", use "nat o floor" and
3530 "nat o ceiling" instead. A few of the lemmas have been retained and
3531 adapted: in their names "natfloor"/"natceiling" has been replaced by
3532 "nat_floor"/"nat_ceiling".
3534 * Qualified some duplicated fact names required for boostrapping the
3535 type class hierarchy:
3536 ab_add_uminus_conv_diff ~> diff_conv_add_uminus
3537 field_inverse_zero ~> inverse_zero
3538 field_divide_inverse ~> divide_inverse
3539 field_inverse ~> left_inverse
3540 Minor INCOMPATIBILITY.
3542 * Eliminated fact duplicates:
3543 mult_less_imp_less_right ~> mult_right_less_imp_less
3544 mult_less_imp_less_left ~> mult_left_less_imp_less
3545 Minor INCOMPATIBILITY.
3547 * Fact consolidation: even_less_0_iff is subsumed by
3548 double_add_less_zero_iff_single_add_less_zero (simp by default anyway).
3550 * Generalized and consolidated some theorems concerning divsibility:
3551 dvd_reduce ~> dvd_add_triv_right_iff
3552 dvd_plus_eq_right ~> dvd_add_right_iff
3553 dvd_plus_eq_left ~> dvd_add_left_iff
3554 Minor INCOMPATIBILITY.
3556 * "even" and "odd" are mere abbreviations for "2 dvd _" and "~ 2 dvd _"
3557 and part of theory Main.
3558 even_def ~> even_iff_mod_2_eq_zero
3561 * Lemma name consolidation: divide_Numeral1 ~> divide_numeral_1. Minor
3564 * Bootstrap of listsum as special case of abstract product over lists.
3566 listsum_def ~> listsum.eq_foldr
3569 * Product over lists via constant "listprod".
3571 * Theory List: renamed drop_Suc_conv_tl and nth_drop' to
3574 * New infrastructure for compiling, running, evaluating and testing
3575 generated code in target languages in HOL/Library/Code_Test. See
3576 HOL/Codegenerator_Test/Code_Test* for examples.
3579 - Introduced "replicate_mset" operation.
3580 - Introduced alternative characterizations of the multiset ordering in
3581 "Library/Multiset_Order".
3582 - Renamed multiset ordering:
3585 \<subset># ~> #\<subset>#
3586 \<subseteq># ~> #\<subseteq>#
3588 - Introduced abbreviations for ill-named multiset operations:
3589 <#, \<subset># abbreviate < (strict subset)
3590 <=#, \<le>#, \<subseteq># abbreviate <= (subset or equal)
3593 in_multiset_of ~> in_multiset_in_set
3594 Multiset.fold ~> fold_mset
3595 Multiset.filter ~> filter_mset
3597 - Removed mcard, is equal to size.
3599 image_mset.id [simp]
3600 image_mset_id [simp]
3601 elem_multiset_of_set [simp, intro]
3602 comp_fun_commute_plus_mset [simp]
3603 comp_fun_commute.fold_mset_insert [OF comp_fun_commute_plus_mset, simp]
3604 in_mset_fold_plus_iff [iff]
3605 set_of_Union_mset [simp]
3606 in_Union_mset_iff [iff]
3609 * Library/Sum_of_Squares: simplified and improved "sos" method. Always
3610 use local CSDP executable, which is much faster than the NEOS server.
3611 The "sos_cert" functionality is invoked as "sos" with additional
3612 argument. Minor INCOMPATIBILITY.
3614 * HOL-Decision_Procs: New counterexample generator quickcheck
3615 [approximation] for inequalities of transcendental functions. Uses
3616 hardware floating point arithmetic to randomly discover potential
3617 counterexamples. Counterexamples are certified with the "approximation"
3618 method. See HOL/Decision_Procs/ex/Approximation_Quickcheck_Ex.thy for
3621 * HOL-Probability: Reworked measurability prover
3622 - applies destructor rules repeatedly
3623 - removed application splitting (replaced by destructor rule)
3624 - added congruence rules to rewrite measure spaces under the sets
3627 * New proof method "rewrite" (in theory ~~/src/HOL/Library/Rewrite) for
3628 single-step rewriting with subterm selection based on patterns.
3633 * Subtle change of name space policy: undeclared entries are now
3634 considered inaccessible, instead of accessible via the fully-qualified
3635 internal name. This mainly affects Name_Space.intern (and derivatives),
3636 which may produce an unexpected Long_Name.hidden prefix. Note that
3637 contemporary applications use the strict Name_Space.check (and
3638 derivatives) instead, which is not affected by the change. Potential
3639 INCOMPATIBILITY in rare applications of Name_Space.intern.
3641 * Subtle change of error semantics of Toplevel.proof_of: regular user
3642 ERROR instead of internal Toplevel.UNDEF.
3644 * Basic combinators map, fold, fold_map, split_list, apply are available
3645 as parameterized antiquotations, e.g. @{map 4} for lists of quadruples.
3647 * Renamed "pairself" to "apply2", in accordance to @{apply 2}.
3650 * Former combinators NAMED_CRITICAL and CRITICAL for central critical
3651 sections have been discontinued, in favour of the more elementary
3652 Multithreading.synchronized and its high-level derivative
3653 Synchronized.var (which is usually sufficient in applications). Subtle
3654 INCOMPATIBILITY: synchronized access needs to be atomic and cannot be
3657 * Synchronized.value (ML) is actually synchronized (as in Scala): subtle
3658 change of semantics with minimal potential for INCOMPATIBILITY.
3660 * The main operations to certify logical entities are Thm.ctyp_of and
3661 Thm.cterm_of with a local context; old-style global theory variants are
3662 available as Thm.global_ctyp_of and Thm.global_cterm_of.
3665 * Elementary operations in module Thm are no longer pervasive.
3666 INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of,
3669 * Proper context for various elementary tactics: assume_tac,
3670 resolve_tac, eresolve_tac, dresolve_tac, forward_tac, match_tac,
3671 compose_tac, Splitter.split_tac etc. INCOMPATIBILITY.
3673 * Tactical PARALLEL_ALLGOALS is the most common way to refer to
3676 * Goal.prove_multi is superseded by the fully general Goal.prove_common,
3677 which also allows to specify a fork priority.
3679 * Antiquotation @{command_spec "COMMAND"} is superseded by
3680 @{command_keyword COMMAND} (usually without quotes and with PIDE
3681 markup). Minor INCOMPATIBILITY.
3683 * Cartouches within ML sources are turned into values of type
3684 Input.source (with formal position information).
3689 * The Isabelle tool "update_cartouches" changes theory files to use
3690 cartouches instead of old-style {* verbatim *} or `alt_string` tokens.
3692 * The Isabelle tool "build" provides new options -X, -k, -x.
3694 * Discontinued old-fashioned "codegen" tool. Code generation can always
3695 be externally triggered using an appropriate ROOT file plus a
3696 corresponding theory. Parametrization is possible using environment
3697 variables, or ML snippets in the most extreme cases. Minor
3700 * JVM system property "isabelle.threads" determines size of Scala thread
3701 pool, like Isabelle system option "threads" for ML.
3703 * JVM system property "isabelle.laf" determines the default Swing
3704 look-and-feel, via internal class name or symbolic name as in the jEdit
3705 menu Global Options / Appearance.
3707 * Support for Proof General and Isar TTY loop has been discontinued.
3708 Minor INCOMPATIBILITY, use standard PIDE infrastructure instead.
3712 New in Isabelle2014 (August 2014)
3713 ---------------------------------
3717 * Support for official Standard ML within the Isabelle context.
3718 Command 'SML_file' reads and evaluates the given Standard ML file.
3719 Toplevel bindings are stored within the theory context; the initial
3720 environment is restricted to the Standard ML implementation of
3721 Poly/ML, without the add-ons of Isabelle/ML. Commands 'SML_import'
3722 and 'SML_export' allow to exchange toplevel bindings between the two
3723 separate environments. See also ~~/src/Tools/SML/Examples.thy for
3726 * Standard tactics and proof methods such as "clarsimp", "auto" and
3727 "safe" now preserve equality hypotheses "x = expr" where x is a free
3728 variable. Locale assumptions and chained facts containing "x"
3729 continue to be useful. The new method "hypsubst_thin" and the
3730 configuration option "hypsubst_thin" (within the attribute name space)
3731 restore the previous behavior. INCOMPATIBILITY, especially where
3732 induction is done after these methods or when the names of free and
3733 bound variables clash. As first approximation, old proofs may be
3734 repaired by "using [[hypsubst_thin = true]]" in the critical spot.
3736 * More static checking of proof methods, which allows the system to
3737 form a closure over the concrete syntax. Method arguments should be
3738 processed in the original proof context as far as possible, before
3739 operating on the goal state. In any case, the standard discipline for
3740 subgoal-addressing needs to be observed: no subgoals or a subgoal
3741 number that is out of range produces an empty result sequence, not an
3742 exception. Potential INCOMPATIBILITY for non-conformant tactical
3745 * Lexical syntax (inner and outer) supports text cartouches with
3746 arbitrary nesting, and without escapes of quotes etc. The Prover IDE
3747 supports input via ` (backquote).
3749 * The outer syntax categories "text" (for formal comments and document
3750 markup commands) and "altstring" (for literal fact references) allow
3751 cartouches as well, in addition to the traditional mix of quotations.
3753 * Syntax of document antiquotation @{rail} now uses \<newline> instead
3754 of "\\", to avoid the optical illusion of escaped backslash within
3755 string token. General renovation of its syntax using text cartouches.
3756 Minor INCOMPATIBILITY.
3758 * Discontinued legacy_isub_isup, which was a temporary workaround for
3759 Isabelle/ML in Isabelle2013-1. The prover process no longer accepts
3760 old identifier syntax with \<^isub> or \<^isup>. Potential
3763 * Document antiquotation @{url} produces markup for the given URL,
3764 which results in an active hyperlink within the text.
3766 * Document antiquotation @{file_unchecked} is like @{file}, but does
3767 not check existence within the file-system.
3769 * Updated and extended manuals: codegen, datatypes, implementation,
3770 isar-ref, jedit, system.
3773 *** Prover IDE -- Isabelle/Scala/jEdit ***
3775 * Improved Document panel: simplified interaction where every single
3776 mouse click (re)opens document via desktop environment or as jEdit
3779 * Support for Navigator plugin (with toolbar buttons), with connection
3782 * Auxiliary files ('ML_file' etc.) are managed by the Prover IDE.
3783 Open text buffers take precedence over copies within the file-system.
3785 * Improved support for Isabelle/ML, with jEdit mode "isabelle-ml" for
3788 * Improved syntactic and semantic completion mechanism, with simple
3789 templates, completion language context, name-space completion,
3790 file-name completion, spell-checker completion.
3792 * Refined GUI popup for completion: more robust key/mouse event
3793 handling and propagation to enclosing text area -- avoid loosing
3794 keystrokes with slow / remote graphics displays.
3796 * Completion popup supports both ENTER and TAB (default) to select an
3797 item, depending on Isabelle options.
3799 * Refined insertion of completion items wrt. jEdit text: multiple
3800 selections, rectangular selections, rectangular selection as "tall
3803 * Integrated spell-checker for document text, comments etc. with
3804 completion popup and context-menu.
3806 * More general "Query" panel supersedes "Find" panel, with GUI access
3807 to commands 'find_theorems' and 'find_consts', as well as print
3808 operations for the context. Minor incompatibility in keyboard
3809 shortcuts etc.: replace action isabelle-find by isabelle-query.
3811 * Search field for all output panels ("Output", "Query", "Info" etc.)
3812 to highlight text via regular expression.
3814 * Option "jedit_print_mode" (see also "Plugin Options / Isabelle /
3815 General") allows to specify additional print modes for the prover
3816 process, without requiring old-fashioned command-line invocation of
3817 "isabelle jedit -m MODE".
3819 * More support for remote files (e.g. http) using standard Java
3820 networking operations instead of jEdit virtual file-systems.
3822 * Empty editors buffers that are no longer required (e.g.\ via theory
3823 imports) are automatically removed from the document model.
3825 * Improved monitor panel.
3827 * Improved Console/Scala plugin: more uniform scala.Console output,
3828 more robust treatment of threads and interrupts.
3830 * Improved management of dockable windows: clarified keyboard focus
3831 and window placement wrt. main editor view; optional menu item to
3832 "Detach" a copy where this makes sense.
3834 * New Simplifier Trace panel provides an interactive view of the
3835 simplification process, enabled by the "simp_trace_new" attribute
3841 * Low-level type-class commands 'classes', 'classrel', 'arities' have
3842 been discontinued to avoid the danger of non-trivial axiomatization
3843 that is not immediately visible. INCOMPATIBILITY, use regular
3844 'instance' command with proof. The required OFCLASS(...) theorem
3845 might be postulated via 'axiomatization' beforehand, or the proof
3846 finished trivially if the underlying class definition is made vacuous
3847 (without any assumptions). See also Isabelle/ML operations
3848 Axclass.class_axiomatization, Axclass.classrel_axiomatization,
3849 Axclass.arity_axiomatization.
3851 * Basic constants of Pure use more conventional names and are always
3852 qualified. Rare INCOMPATIBILITY, but with potentially serious
3853 consequences, notably for tools in Isabelle/ML. The following
3854 renaming needs to be applied:
3860 dummy_pattern ~> Pure.dummy_pattern
3862 Systematic porting works by using the following theory setup on a
3863 *previous* Isabelle version to introduce the new name accesses for the
3869 |> Sign.const_alias (Binding.qualify true "Pure" @{binding eq}) "=="
3870 |> Sign.const_alias (Binding.qualify true "Pure" @{binding imp}) "==>"
3871 |> Sign.const_alias (Binding.qualify true "Pure" @{binding all}) "all"
3872 |> Sign.restore_naming thy
3875 Thus ML antiquotations like @{const_name Pure.eq} may be used already.
3876 Later the application is moved to the current Isabelle version, and
3877 the auxiliary aliases are deleted.
3879 * Attributes "where" and "of" allow an optional context of local
3880 variables ('for' declaration): these variables become schematic in the
3881 instantiated theorem.
3883 * Obsolete attribute "standard" has been discontinued (legacy since
3884 Isabelle2012). Potential INCOMPATIBILITY, use explicit 'for' context
3885 where instantiations with schematic variables are intended (for
3886 declaration commands like 'lemmas' or attributes like "of"). The
3887 following temporary definition may help to port old applications:
3889 attribute_setup standard =
3890 "Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))"
3892 * More thorough check of proof context for goal statements and
3893 attributed fact expressions (concerning background theory, declared
3894 hyps). Potential INCOMPATIBILITY, tools need to observe standard
3895 context discipline. See also Assumption.add_assumes and the more
3896 primitive Thm.assume_hyps.
3898 * Inner syntax token language allows regular quoted strings "..."
3899 (only makes sense in practice, if outer syntax is delimited
3900 differently, e.g. via cartouches).
3902 * Command 'print_term_bindings' supersedes 'print_binds' for clarity,
3903 but the latter is retained some time as Proof General legacy.
3905 * Code generator preprocessor: explicit control of simp tracing on a
3906 per-constant basis. See attribute "code_preproc".
3911 * Code generator: enforce case of identifiers only for strict target
3912 language requirements. INCOMPATIBILITY.
3914 * Code generator: explicit proof contexts in many ML interfaces.
3917 * Code generator: minimize exported identifiers by default. Minor
3920 * Code generation for SML and OCaml: dropped arcane "no_signatures"
3921 option. Minor INCOMPATIBILITY.
3923 * "declare [[code abort: ...]]" replaces "code_abort ...".
3926 * "declare [[code drop: ...]]" drops all code equations associated
3927 with the given constants.
3929 * Code generations are provided for make, fields, extend and truncate
3930 operations on records.
3932 * Command and antiquotation "value" are now hardcoded against nbe and
3933 ML. Minor INCOMPATIBILITY.
3935 * Renamed command 'enriched_type' to 'functor'. INCOMPATIBILITY.
3937 * The symbol "\<newline>" may be used within char or string literals
3938 to represent (Char Nibble0 NibbleA), i.e. ASCII newline.
3940 * Qualified String.implode and String.explode. INCOMPATIBILITY.
3942 * Simplifier: Enhanced solver of preconditions of rewrite rules can
3943 now deal with conjunctions. For help with converting proofs, the old
3944 behaviour of the simplifier can be restored like this: declare/using
3945 [[simp_legacy_precond]]. This configuration option will disappear
3946 again in the future. INCOMPATIBILITY.
3948 * Simproc "finite_Collect" is no longer enabled by default, due to
3949 spurious crashes and other surprises. Potential INCOMPATIBILITY.
3951 * Moved new (co)datatype package and its dependencies from session
3952 "HOL-BNF" to "HOL". The commands 'bnf', 'wrap_free_constructors',
3953 'datatype_new', 'codatatype', 'primcorec', 'primcorecursive' are now
3954 part of theory "Main".
3957 FunDef.thy ~> Fun_Def.thy (and Fun_Def_Base.thy)
3958 Library/Wfrec.thy ~> Wfrec.thy
3959 Library/Zorn.thy ~> Zorn.thy
3960 Cardinals/Order_Relation.thy ~> Order_Relation.thy
3961 Library/Order_Union.thy ~> Cardinals/Order_Union.thy
3962 Cardinals/Cardinal_Arithmetic_Base.thy ~> BNF_Cardinal_Arithmetic.thy
3963 Cardinals/Cardinal_Order_Relation_Base.thy ~> BNF_Cardinal_Order_Relation.thy
3964 Cardinals/Constructions_on_Wellorders_Base.thy ~> BNF_Constructions_on_Wellorders.thy
3965 Cardinals/Wellorder_Embedding_Base.thy ~> BNF_Wellorder_Embedding.thy
3966 Cardinals/Wellorder_Relation_Base.thy ~> BNF_Wellorder_Relation.thy
3967 BNF/Ctr_Sugar.thy ~> Ctr_Sugar.thy
3968 BNF/Basic_BNFs.thy ~> Basic_BNFs.thy
3969 BNF/BNF_Comp.thy ~> BNF_Comp.thy
3970 BNF/BNF_Def.thy ~> BNF_Def.thy
3971 BNF/BNF_FP_Base.thy ~> BNF_FP_Base.thy
3972 BNF/BNF_GFP.thy ~> BNF_GFP.thy
3973 BNF/BNF_LFP.thy ~> BNF_LFP.thy
3974 BNF/BNF_Util.thy ~> BNF_Util.thy
3975 BNF/Coinduction.thy ~> Coinduction.thy
3976 BNF/More_BNFs.thy ~> Library/More_BNFs.thy
3977 BNF/Countable_Type.thy ~> Library/Countable_Set_Type.thy
3978 BNF/Examples/* ~> BNF_Examples/*
3981 Wellorder_Extension.thy (split from Zorn.thy)
3982 Library/Cardinal_Notations.thy
3983 Library/BNF_Axomatization.thy
3984 BNF_Examples/Misc_Primcorec.thy
3985 BNF_Examples/Stream_Processor.thy
3987 Discontinued theories:
3989 BNF/Equiv_Relations_More.thy
3993 * New (co)datatype package:
3994 - Command 'primcorec' is fully implemented.
3995 - Command 'datatype_new' generates size functions ("size_xxx" and
3996 "size") as required by 'fun'.
3997 - BNFs are integrated with the Lifting tool and new-style
3998 (co)datatypes with Transfer.
4000 datatype_new_compat ~> datatype_compat
4001 primrec_new ~> primrec
4002 wrap_free_constructors ~> free_constructors
4004 - The generated constants "xxx_case" and "xxx_rec" have been renamed
4005 "case_xxx" and "rec_xxx" (e.g., "prod_case" ~> "case_prod").
4007 - The constant "xxx_(un)fold" and related theorems are no longer
4008 generated. Use "xxx_(co)rec" or define "xxx_(un)fold" manually
4009 using "prim(co)rec".
4011 - No discriminators are generated for nullary constructors by
4012 default, eliminating the need for the odd "=:" syntax.
4014 - No discriminators or selectors are generated by default by
4015 "datatype_new", unless custom names are specified or the new
4016 "discs_sels" option is passed.
4019 * Old datatype package:
4020 - The generated theorems "xxx.cases" and "xxx.recs" have been
4021 renamed "xxx.case" and "xxx.rec" (e.g., "sum.cases" ->
4022 "sum.case"). INCOMPATIBILITY.
4023 - The generated constants "xxx_case", "xxx_rec", and "xxx_size" have
4024 been renamed "case_xxx", "rec_xxx", and "size_xxx" (e.g.,
4025 "prod_case" ~> "case_prod"). INCOMPATIBILITY.
4027 * The types "'a list" and "'a option", their set and map functions,
4028 their relators, and their selectors are now produced using the new
4029 BNF-based datatype package.
4032 Option.set ~> set_option
4033 Option.map ~> map_option
4034 option_rel ~> rel_option
4037 set_def ~> set_rec[abs_def]
4038 map_def ~> map_rec[abs_def]
4039 Option.map_def ~> map_option_case[abs_def] (with "case_option" instead of "rec_option")
4040 option.recs ~> option.rec
4041 list_all2_def ~> list_all2_iff
4042 set.simps ~> set_simps (or the slightly different "list.set")
4043 map.simps ~> list.map
4044 hd.simps ~> list.sel(1)
4045 tl.simps ~> list.sel(2-3)
4046 the.simps ~> option.sel
4050 * The following map functions and relators have been renamed:
4052 map_pair ~> map_prod
4053 prod_rel ~> rel_prod
4057 filter_rel ~> rel_filter
4058 fset_rel ~> rel_fset (in "src/HOL/Library/FSet.thy")
4059 cset_rel ~> rel_cset (in "src/HOL/Library/Countable_Set_Type.thy")
4060 vset ~> rel_vset (in "src/HOL/Library/Quotient_Set.thy")
4064 * Lifting and Transfer:
4065 - a type variable as a raw type is supported
4066 - stronger reflexivity prover
4067 - rep_eq is always generated by lift_definition
4068 - setup for Lifting/Transfer is now automated for BNFs
4069 + holds for BNFs that do not contain a dead variable
4070 + relator_eq, relator_mono, relator_distr, relator_domain,
4071 relator_eq_onp, quot_map, transfer rules for bi_unique, bi_total,
4072 right_unique, right_total, left_unique, left_total are proved
4074 + definition of a predicator is generated automatically
4075 + simplification rules for a predicator definition are proved
4076 automatically for datatypes
4077 - consolidation of the setup of Lifting/Transfer
4078 + property that a relator preservers reflexivity is not needed any
4080 Minor INCOMPATIBILITY.
4081 + left_total and left_unique rules are now transfer rules
4082 (reflexivity_rule attribute not needed anymore)
4084 + Domainp does not have to be a separate assumption in
4085 relator_domain theorems (=> more natural statement)
4087 - registration of code equations is more robust
4088 Potential INCOMPATIBILITY.
4089 - respectfulness proof obligation is preprocessed to a more readable
4091 Potential INCOMPATIBILITY.
4092 - eq_onp is always unfolded in respectfulness proof obligation
4093 Potential INCOMPATIBILITY.
4094 - unregister lifting setup for Code_Numeral.integer and
4095 Code_Numeral.natural
4096 Potential INCOMPATIBILITY.
4097 - Lifting.invariant -> eq_onp
4100 * New internal SAT solver "cdclite" that produces models and proof
4101 traces. This solver replaces the internal SAT solvers "enumerate" and
4102 "dpll". Applications that explicitly used one of these two SAT
4103 solvers should use "cdclite" instead. In addition, "cdclite" is now
4104 the default SAT solver for the "sat" and "satx" proof methods and
4105 corresponding tactics; the old default can be restored using "declare
4106 [[sat_solver = zchaff_with_proofs]]". Minor INCOMPATIBILITY.
4108 * SMT module: A new version of the SMT module, temporarily called
4109 "SMT2", uses SMT-LIB 2 and supports recent versions of Z3 (e.g.,
4110 4.3). The new proof method is called "smt2". CVC3 and CVC4 are also
4111 supported as oracles. Yices is no longer supported, because no version
4112 of the solver can handle both SMT-LIB 2 and quantifiers.
4114 * Activation of Z3 now works via "z3_non_commercial" system option
4115 (without requiring restart), instead of former settings variable
4116 "Z3_NON_COMMERCIAL". The option can be edited in Isabelle/jEdit menu
4117 Plugin Options / Isabelle / General.
4120 - Z3 can now produce Isar proofs.
4122 . New SML-based learning algorithms eliminate the dependency on
4123 Python and increase performance and reliability.
4124 . MaSh and MeSh are now used by default together with the
4125 traditional MePo (Meng-Paulson) relevance filter. To disable
4126 MaSh, set the "MaSh" system option in Isabelle/jEdit Plugin
4127 Options / Isabelle / General to "none".
4131 isar_compress ~> compress
4136 * Removed solvers remote_cvc3 and remote_z3. Use cvc3 and z3 instead.
4139 - Fixed soundness bug whereby mutually recursive datatypes could
4140 take infinite values.
4141 - Fixed soundness bug with low-level number functions such as
4142 "Abs_Integ" and "Rep_Integ".
4143 - Removed "std" option.
4144 - Renamed "show_datatypes" to "show_types" and "hide_datatypes" to
4147 * Metis: Removed legacy proof method 'metisFT'. Use 'metis
4148 (full_types)' instead. INCOMPATIBILITY.
4150 * Try0: Added 'algebra' and 'meson' to the set of proof methods.
4152 * Adjustion of INF and SUP operations:
4153 - Elongated constants INFI and SUPR to INFIMUM and SUPREMUM.
4154 - Consolidated theorem names containing INFI and SUPR: have INF and
4155 SUP instead uniformly.
4156 - More aggressive normalization of expressions involving INF and Inf
4158 - INF_image and SUP_image do not unfold composition.
4159 - Dropped facts INF_comp, SUP_comp.
4160 - Default congruence rules strong_INF_cong and strong_SUP_cong, with
4161 simplifier implication in premises. Generalize and replace former
4166 * SUP and INF generalized to conditionally_complete_lattice.
4168 * Swapped orientation of facts image_comp and vimage_comp:
4170 image_compose ~> image_comp [symmetric]
4171 image_comp ~> image_comp [symmetric]
4172 vimage_compose ~> vimage_comp [symmetric]
4173 vimage_comp ~> vimage_comp [symmetric]
4177 * Theory reorganization: split of Big_Operators.thy into
4178 Groups_Big.thy and Lattices_Big.thy.
4180 * Consolidated some facts about big group operators:
4182 setsum_0' ~> setsum.neutral
4183 setsum_0 ~> setsum.neutral_const
4184 setsum_addf ~> setsum.distrib
4185 setsum_cartesian_product ~> setsum.cartesian_product
4186 setsum_cases ~> setsum.If_cases
4187 setsum_commute ~> setsum.commute
4188 setsum_cong ~> setsum.cong
4189 setsum_delta ~> setsum.delta
4190 setsum_delta' ~> setsum.delta'
4191 setsum_diff1' ~> setsum.remove
4192 setsum_empty ~> setsum.empty
4193 setsum_infinite ~> setsum.infinite
4194 setsum_insert ~> setsum.insert
4195 setsum_inter_restrict'' ~> setsum.inter_filter
4196 setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left
4197 setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right
4198 setsum_mono_zero_left ~> setsum.mono_neutral_left
4199 setsum_mono_zero_right ~> setsum.mono_neutral_right
4200 setsum_reindex ~> setsum.reindex
4201 setsum_reindex_cong ~> setsum.reindex_cong
4202 setsum_reindex_nonzero ~> setsum.reindex_nontrivial
4203 setsum_restrict_set ~> setsum.inter_restrict
4204 setsum_Plus ~> setsum.Plus
4205 setsum_setsum_restrict ~> setsum.commute_restrict
4206 setsum_Sigma ~> setsum.Sigma
4207 setsum_subset_diff ~> setsum.subset_diff
4208 setsum_Un_disjoint ~> setsum.union_disjoint
4209 setsum_UN_disjoint ~> setsum.UNION_disjoint
4210 setsum_Un_Int ~> setsum.union_inter
4211 setsum_Union_disjoint ~> setsum.Union_disjoint
4212 setsum_UNION_zero ~> setsum.Union_comp
4213 setsum_Un_zero ~> setsum.union_inter_neutral
4214 strong_setprod_cong ~> setprod.strong_cong
4215 strong_setsum_cong ~> setsum.strong_cong
4216 setprod_1' ~> setprod.neutral
4217 setprod_1 ~> setprod.neutral_const
4218 setprod_cartesian_product ~> setprod.cartesian_product
4219 setprod_cong ~> setprod.cong
4220 setprod_delta ~> setprod.delta
4221 setprod_delta' ~> setprod.delta'
4222 setprod_empty ~> setprod.empty
4223 setprod_infinite ~> setprod.infinite
4224 setprod_insert ~> setprod.insert
4225 setprod_mono_one_cong_left ~> setprod.mono_neutral_cong_left
4226 setprod_mono_one_cong_right ~> setprod.mono_neutral_cong_right
4227 setprod_mono_one_left ~> setprod.mono_neutral_left
4228 setprod_mono_one_right ~> setprod.mono_neutral_right
4229 setprod_reindex ~> setprod.reindex
4230 setprod_reindex_cong ~> setprod.reindex_cong
4231 setprod_reindex_nonzero ~> setprod.reindex_nontrivial
4232 setprod_Sigma ~> setprod.Sigma
4233 setprod_subset_diff ~> setprod.subset_diff
4234 setprod_timesf ~> setprod.distrib
4235 setprod_Un2 ~> setprod.union_diff2
4236 setprod_Un_disjoint ~> setprod.union_disjoint
4237 setprod_UN_disjoint ~> setprod.UNION_disjoint
4238 setprod_Un_Int ~> setprod.union_inter
4239 setprod_Union_disjoint ~> setprod.Union_disjoint
4240 setprod_Un_one ~> setprod.union_inter_neutral
4242 Dropped setsum_cong2 (simple variant of setsum.cong).
4243 Dropped setsum_inter_restrict' (simple variant of setsum.inter_restrict)
4244 Dropped setsum_reindex_id, setprod_reindex_id
4245 (simple variants of setsum.reindex [symmetric], setprod.reindex [symmetric]).
4249 * Abolished slightly odd global lattice interpretation for min/max.
4251 Fact consolidations:
4252 min_max.inf_assoc ~> min.assoc
4253 min_max.inf_commute ~> min.commute
4254 min_max.inf_left_commute ~> min.left_commute
4255 min_max.inf_idem ~> min.idem
4256 min_max.inf_left_idem ~> min.left_idem
4257 min_max.inf_right_idem ~> min.right_idem
4258 min_max.sup_assoc ~> max.assoc
4259 min_max.sup_commute ~> max.commute
4260 min_max.sup_left_commute ~> max.left_commute
4261 min_max.sup_idem ~> max.idem
4262 min_max.sup_left_idem ~> max.left_idem
4263 min_max.sup_inf_distrib1 ~> max_min_distrib2
4264 min_max.sup_inf_distrib2 ~> max_min_distrib1
4265 min_max.inf_sup_distrib1 ~> min_max_distrib2
4266 min_max.inf_sup_distrib2 ~> min_max_distrib1
4267 min_max.distrib ~> min_max_distribs
4268 min_max.inf_absorb1 ~> min.absorb1
4269 min_max.inf_absorb2 ~> min.absorb2
4270 min_max.sup_absorb1 ~> max.absorb1
4271 min_max.sup_absorb2 ~> max.absorb2
4272 min_max.le_iff_inf ~> min.absorb_iff1
4273 min_max.le_iff_sup ~> max.absorb_iff2
4274 min_max.inf_le1 ~> min.cobounded1
4275 min_max.inf_le2 ~> min.cobounded2
4276 le_maxI1, min_max.sup_ge1 ~> max.cobounded1
4277 le_maxI2, min_max.sup_ge2 ~> max.cobounded2
4278 min_max.le_infI1 ~> min.coboundedI1
4279 min_max.le_infI2 ~> min.coboundedI2
4280 min_max.le_supI1 ~> max.coboundedI1
4281 min_max.le_supI2 ~> max.coboundedI2
4282 min_max.less_infI1 ~> min.strict_coboundedI1
4283 min_max.less_infI2 ~> min.strict_coboundedI2
4284 min_max.less_supI1 ~> max.strict_coboundedI1
4285 min_max.less_supI2 ~> max.strict_coboundedI2
4286 min_max.inf_mono ~> min.mono
4287 min_max.sup_mono ~> max.mono
4288 min_max.le_infI, min_max.inf_greatest ~> min.boundedI
4289 min_max.le_supI, min_max.sup_least ~> max.boundedI
4290 min_max.le_inf_iff ~> min.bounded_iff
4291 min_max.le_sup_iff ~> max.bounded_iff
4293 For min_max.inf_sup_aci, prefer (one of) min.commute, min.assoc,
4294 min.left_commute, min.left_idem, max.commute, max.assoc,
4295 max.left_commute, max.left_idem directly.
4297 For min_max.inf_sup_ord, prefer (one of) min.cobounded1,
4298 min.cobounded2, max.cobounded1m max.cobounded2 directly.
4300 For min_ac or max_ac, prefer more general collection ac_simps.
4304 * Theorem disambiguation Inf_le_Sup (on finite sets) ~>
4305 Inf_fin_le_Sup_fin. INCOMPATIBILITY.
4307 * Qualified constant names Wellfounded.acc, Wellfounded.accp.
4310 * Fact generalization and consolidation:
4311 neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1
4315 * Purely algebraic definition of even. Fact generalization and
4317 nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd
4318 even_zero_(nat|int) ~> even_zero
4322 * Abolished neg_numeral.
4323 - Canonical representation for minus one is "- 1".
4324 - Canonical representation for other negative numbers is "- (numeral _)".
4325 - When devising rule sets for number calculation, consider the
4326 following canonical cases: 0, 1, numeral _, - 1, - numeral _.
4327 - HOLogic.dest_number also recognizes numerals in non-canonical forms
4328 like "numeral One", "- numeral One", "- 0" and even "- ... - _".
4329 - Syntax for negative numerals is mere input syntax.
4333 * Reduced name variants for rules on associativity and commutativity:
4335 add_assoc ~> add.assoc
4336 add_commute ~> add.commute
4337 add_left_commute ~> add.left_commute
4338 mult_assoc ~> mult.assoc
4339 mult_commute ~> mult.commute
4340 mult_left_commute ~> mult.left_commute
4341 nat_add_assoc ~> add.assoc
4342 nat_add_commute ~> add.commute
4343 nat_add_left_commute ~> add.left_commute
4344 nat_mult_assoc ~> mult.assoc
4345 nat_mult_commute ~> mult.commute
4346 eq_assoc ~> iff_assoc
4347 eq_left_commute ~> iff_left_commute
4351 * Fact collections add_ac and mult_ac are considered old-fashioned.
4352 Prefer ac_simps instead, or specify rules
4353 (add|mult).(assoc|commute|left_commute) individually.
4355 * Elimination of fact duplicates:
4356 equals_zero_I ~> minus_unique
4357 diff_eq_0_iff_eq ~> right_minus_eq
4358 nat_infinite ~> infinite_UNIV_nat
4359 int_infinite ~> infinite_UNIV_int
4363 * Fact name consolidation:
4364 diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
4365 minus_le_self_iff ~> neg_less_eq_nonneg
4366 le_minus_self_iff ~> less_eq_neg_nonpos
4367 neg_less_nonneg ~> neg_less_pos
4368 less_minus_self_iff ~> less_neg_neg [simp]
4372 * More simplification rules on unary and binary minus:
4373 add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
4374 add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
4375 add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
4376 le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
4377 minus_add_cancel, uminus_add_conv_diff. These correspondingly have
4378 been taken away from fact collections algebra_simps and field_simps.
4381 To restore proofs, the following patterns are helpful:
4383 a) Arbitrary failing proof not involving "diff_def":
4384 Consider simplification with algebra_simps or field_simps.
4386 b) Lifting rules from addition to subtraction:
4387 Try with "using <rule for addition> of [... "- _" ...]" by simp".
4389 c) Simplification with "diff_def": just drop "diff_def".
4390 Consider simplification with algebra_simps or field_simps;
4391 or the brute way with
4392 "simp add: diff_conv_add_uminus del: add_uminus_conv_diff".
4394 * Introduce bdd_above and bdd_below in theory
4395 Conditionally_Complete_Lattices, use them instead of explicitly
4396 stating boundedness of sets.
4398 * ccpo.admissible quantifies only over non-empty chains to allow more
4399 syntax-directed proof rules; the case of the empty chain shows up as
4400 additional case in fixpoint induction proofs. INCOMPATIBILITY.
4402 * Removed and renamed theorems in Series:
4403 summable_le ~> suminf_le
4404 suminf_le ~> suminf_le_const
4405 series_pos_le ~> setsum_le_suminf
4406 series_pos_less ~> setsum_less_suminf
4407 suminf_ge_zero ~> suminf_nonneg
4408 suminf_gt_zero ~> suminf_pos
4409 suminf_gt_zero_iff ~> suminf_pos_iff
4410 summable_sumr_LIMSEQ_suminf ~> summable_LIMSEQ
4411 suminf_0_le ~> suminf_nonneg [rotate]
4412 pos_summable ~> summableI_nonneg_bounded
4413 ratio_test ~> summable_ratio_test
4415 removed series_zero, replaced by sums_finite
4417 removed auxiliary lemmas:
4419 sumr_offset, sumr_offset2, sumr_offset3, sumr_offset4, sumr_group,
4420 half, le_Suc_ex_iff, lemma_realpow_diff_sumr,
4421 real_setsum_nat_ivl_bounded, summable_le2, ratio_test_lemma2,
4422 sumr_minus_one_realpow_zerom, sumr_one_lb_realpow_zero,
4423 summable_convergent_sumr_iff, sumr_diff_mult_const
4427 * Replace (F)DERIV syntax by has_derivative:
4428 - "(f has_derivative f') (at x within s)" replaces "FDERIV f x : s : f'"
4430 - "(f has_field_derivative f') (at x within s)" replaces "DERIV f x : s : f'"
4432 - "f differentiable at x within s" replaces "_ differentiable _ in _" syntax
4434 - removed constant isDiff
4436 - "DERIV f x : f'" and "FDERIV f x : f'" syntax is only available as
4439 - "DERIV f x : s : f'" and "FDERIV f x : s : f'" syntax removed.
4441 - Renamed FDERIV_... lemmas to has_derivative_...
4443 - renamed deriv (the syntax constant used for "DERIV _ _ :> _") to DERIV
4445 - removed DERIV_intros, has_derivative_eq_intros
4447 - introduced derivative_intros and deriative_eq_intros which
4448 includes now rules for DERIV, has_derivative and
4449 has_vector_derivative.
4452 differentiable_def ~> real_differentiable_def
4453 differentiableE ~> real_differentiableE
4454 fderiv_def ~> has_derivative_at
4455 field_fderiv_def ~> field_has_derivative_at
4456 isDiff_der ~> differentiable_def
4457 deriv_fderiv ~> has_field_derivative_def
4458 deriv_def ~> DERIV_def
4462 * Include more theorems in continuous_intros. Remove the
4463 continuous_on_intros, isCont_intros collections, these facts are now
4464 in continuous_intros.
4466 * Theorems about complex numbers are now stated only using Re and Im,
4467 the Complex constructor is not used anymore. It is possible to use
4468 primcorec to defined the behaviour of a complex-valued function.
4470 Removed theorems about the Complex constructor from the simpset, they
4471 are available as the lemma collection legacy_Complex_simps. This
4474 i_complex_of_real: "ii * complex_of_real r = Complex 0 r".
4476 Instead the reverse direction is supported with
4477 Complex_eq: "Complex a b = a + \<i> * b"
4479 Moved csqrt from Fundamental_Algebra_Theorem to Complex.
4482 Re/Im ~> complex.sel
4483 complex_Re/Im_zero ~> zero_complex.sel
4484 complex_Re/Im_add ~> plus_complex.sel
4485 complex_Re/Im_minus ~> uminus_complex.sel
4486 complex_Re/Im_diff ~> minus_complex.sel
4487 complex_Re/Im_one ~> one_complex.sel
4488 complex_Re/Im_mult ~> times_complex.sel
4489 complex_Re/Im_inverse ~> inverse_complex.sel
4490 complex_Re/Im_scaleR ~> scaleR_complex.sel
4491 complex_Re/Im_i ~> ii.sel
4492 complex_Re/Im_cnj ~> cnj.sel
4493 Re/Im_cis ~> cis.sel
4495 complex_divide_def ~> divide_complex_def
4496 complex_norm_def ~> norm_complex_def
4497 cmod_def ~> norm_complex_de
4511 * Theory Lubs moved HOL image to HOL-Library. It is replaced by
4512 Conditionally_Complete_Lattices. INCOMPATIBILITY.
4514 * HOL-Library: new theory src/HOL/Library/Tree.thy.
4516 * HOL-Library: removed theory src/HOL/Library/Kleene_Algebra.thy; it
4517 is subsumed by session Kleene_Algebra in AFP.
4519 * HOL-Library / theory RBT: various constants and facts are hidden;
4520 lifting setup is unregistered. INCOMPATIBILITY.
4522 * HOL-Cardinals: new theory src/HOL/Cardinals/Ordinal_Arithmetic.thy.
4524 * HOL-Word: bit representations prefer type bool over type bit.
4528 - Abandoned fact collection "word_arith_alts", which is a duplicate
4529 of "word_arith_wis".
4530 - Dropped first (duplicated) element in fact collections
4531 "sint_word_ariths", "word_arith_alts", "uint_word_ariths",
4532 "uint_word_arith_bintrs".
4534 * HOL-Number_Theory:
4535 - consolidated the proofs of the binomial theorem
4536 - the function fib is again of type nat => nat and not overloaded
4537 - no more references to Old_Number_Theory in the HOL libraries
4542 * HOL-Multivariate_Analysis:
4543 - Type class ordered_real_vector for ordered vector spaces.
4544 - New theory Complex_Basic_Analysis defining complex derivatives,
4545 holomorphic functions, etc., ported from HOL Light's canal.ml.
4546 - Changed order of ordered_euclidean_space to be compatible with
4547 pointwise ordering on products. Therefore instance of
4548 conditionally_complete_lattice and ordered_real_vector.
4549 INCOMPATIBILITY: use box instead of greaterThanLessThan or
4550 explicit set-comprehensions with eucl_less for other (half-)open
4552 - removed dependencies on type class ordered_euclidean_space with
4553 introduction of "cbox" on euclidean_space
4556 mem_interval ~> mem_box
4557 interval_eq_empty ~> box_eq_empty
4558 interval_ne_empty ~> box_ne_empty
4559 interval_sing(1) ~> cbox_sing
4560 interval_sing(2) ~> box_sing
4561 subset_interval_imp ~> subset_box_imp
4562 subset_interval ~> subset_box
4563 open_interval ~> open_box
4564 closed_interval ~> closed_cbox
4565 interior_closed_interval ~> interior_cbox
4566 bounded_closed_interval ~> bounded_cbox
4567 compact_interval ~> compact_cbox
4568 bounded_subset_closed_interval_symmetric ~> bounded_subset_cbox_symmetric
4569 bounded_subset_closed_interval ~> bounded_subset_cbox
4570 mem_interval_componentwiseI ~> mem_box_componentwiseI
4571 convex_box ~> convex_prod
4572 rel_interior_real_interval ~> rel_interior_real_box
4573 convex_interval ~> convex_box
4574 convex_hull_eq_real_interval ~> convex_hull_eq_real_cbox
4575 frechet_derivative_within_closed_interval ~> frechet_derivative_within_cbox
4576 content_closed_interval' ~> content_cbox'
4577 elementary_subset_interval ~> elementary_subset_box
4578 diameter_closed_interval ~> diameter_cbox
4579 frontier_closed_interval ~> frontier_cbox
4580 frontier_open_interval ~> frontier_box
4581 bounded_subset_open_interval_symmetric ~> bounded_subset_box_symmetric
4582 closure_open_interval ~> closure_box
4583 open_closed_interval_convex ~> open_cbox_convex
4584 open_interval_midpoint ~> box_midpoint
4585 content_image_affinity_interval ~> content_image_affinity_cbox
4586 is_interval_interval ~> is_interval_cbox + is_interval_box + is_interval_closed_interval
4587 bounded_interval ~> bounded_closed_interval + bounded_boxes
4589 - respective theorems for intervals over the reals:
4590 content_closed_interval + content_cbox
4591 has_integral + has_integral_real
4592 fine_division_exists + fine_division_exists_real
4593 has_integral_null + has_integral_null_real
4594 tagged_division_union_interval + tagged_division_union_interval_real
4595 has_integral_const + has_integral_const_real
4596 integral_const + integral_const_real
4597 has_integral_bound + has_integral_bound_real
4598 integrable_continuous + integrable_continuous_real
4599 integrable_subinterval + integrable_subinterval_real
4600 has_integral_reflect_lemma + has_integral_reflect_lemma_real
4601 integrable_reflect + integrable_reflect_real
4602 integral_reflect + integral_reflect_real
4603 image_affinity_interval + image_affinity_cbox
4604 image_smult_interval + image_smult_cbox
4605 integrable_const + integrable_const_ivl
4606 integrable_on_subinterval + integrable_on_subcbox
4609 derivative_linear ~> has_derivative_bounded_linear
4610 derivative_is_linear ~> has_derivative_linear
4611 bounded_linear_imp_linear ~> bounded_linear.linear
4614 - Renamed positive_integral to nn_integral:
4616 . Renamed all lemmas "*positive_integral*" to *nn_integral*"
4617 positive_integral_positive ~> nn_integral_nonneg
4619 . Renamed abbreviation integral\<^sup>P to integral\<^sup>N.
4621 - replaced the Lebesgue integral on real numbers by the more general
4622 Bochner integral for functions into a real-normed vector space.
4624 integral_zero ~> integral_zero / integrable_zero
4625 integral_minus ~> integral_minus / integrable_minus
4626 integral_add ~> integral_add / integrable_add
4627 integral_diff ~> integral_diff / integrable_diff
4628 integral_setsum ~> integral_setsum / integrable_setsum
4629 integral_multc ~> integral_mult_left / integrable_mult_left
4630 integral_cmult ~> integral_mult_right / integrable_mult_right
4631 integral_triangle_inequality~> integral_norm_bound
4632 integrable_nonneg ~> integrableI_nonneg
4633 integral_positive ~> integral_nonneg_AE
4634 integrable_abs_iff ~> integrable_abs_cancel
4635 positive_integral_lim_INF ~> nn_integral_liminf
4636 lebesgue_real_affine ~> lborel_real_affine
4637 borel_integral_has_integral ~> has_integral_lebesgue_integral
4638 integral_indicator ~>
4639 integral_real_indicator / integrable_real_indicator
4640 positive_integral_fst ~> nn_integral_fst'
4641 positive_integral_fst_measurable ~> nn_integral_fst
4642 positive_integral_snd_measurable ~> nn_integral_snd
4644 integrable_fst_measurable ~>
4645 integral_fst / integrable_fst / AE_integrable_fst
4647 integrable_snd_measurable ~>
4648 integral_snd / integrable_snd / AE_integrable_snd
4650 integral_monotone_convergence ~>
4651 integral_monotone_convergence / integrable_monotone_convergence
4653 integral_monotone_convergence_at_top ~>
4654 integral_monotone_convergence_at_top /
4655 integrable_monotone_convergence_at_top
4657 has_integral_iff_positive_integral_lebesgue ~>
4658 has_integral_iff_has_bochner_integral_lebesgue_nonneg
4660 lebesgue_integral_has_integral ~>
4661 has_integral_integrable_lebesgue_nonneg
4663 positive_integral_lebesgue_has_integral ~>
4664 integral_has_integral_lebesgue_nonneg /
4665 integrable_has_integral_lebesgue_nonneg
4667 lebesgue_integral_real_affine ~>
4668 nn_integral_real_affine
4670 has_integral_iff_positive_integral_lborel ~>
4671 integral_has_integral_nonneg / integrable_has_integral_nonneg
4673 The following theorems where removed:
4675 lebesgue_integral_nonneg
4676 lebesgue_integral_uminus
4677 lebesgue_integral_cmult
4678 lebesgue_integral_multc
4679 lebesgue_integral_cmult_nonneg
4680 integral_cmul_indicator
4683 - Formalized properties about exponentially, Erlang, and normal
4684 distributed random variables.
4686 * HOL-Decision_Procs: Separate command 'approximate' for approximative
4687 computation in src/HOL/Decision_Procs/Approximation. Minor
4693 * The signature and semantics of Document.Snapshot.cumulate_markup /
4694 select_markup have been clarified. Markup is now traversed in the
4695 order of reports given by the prover: later markup is usually more
4696 specific and may override results accumulated so far. The elements
4697 guard is mandatory and checked precisely. Subtle INCOMPATIBILITY.
4699 * Substantial reworking of internal PIDE protocol communication
4700 channels. INCOMPATIBILITY.
4705 * Subtle change of semantics of Thm.eq_thm: theory stamps are not
4706 compared (according to Thm.thm_ord), but assumed to be covered by the
4707 current background theory. Thus equivalent data produced in different
4708 branches of the theory graph usually coincides (e.g. relevant for
4709 theory merge). Note that the softer Thm.eq_thm_prop is often more
4710 appropriate than Thm.eq_thm.
4712 * Proper context for basic Simplifier operations: rewrite_rule,
4713 rewrite_goals_rule, rewrite_goals_tac etc. INCOMPATIBILITY, need to
4714 pass runtime Proof.context (and ensure that the simplified entity
4715 actually belongs to it).
4717 * Proper context discipline for read_instantiate and instantiate_tac:
4718 variables that are meant to become schematic need to be given as
4719 fixed, and are generalized by the explicit context of local variables.
4720 This corresponds to Isar attributes "where" and "of" with 'for'
4721 declaration. INCOMPATIBILITY, also due to potential change of indices
4722 of schematic variables.
4724 * Moved ML_Compiler.exn_trace and other operations on exceptions to
4725 structure Runtime. Minor INCOMPATIBILITY.
4727 * Discontinued old Toplevel.debug in favour of system option
4728 "ML_exception_trace", which may be also declared within the context
4729 via "declare [[ML_exception_trace = true]]". Minor INCOMPATIBILITY.
4731 * Renamed configuration option "ML_trace" to "ML_source_trace". Minor
4734 * Configuration option "ML_print_depth" controls the pretty-printing
4735 depth of the ML compiler within the context. The old print_depth in
4736 ML is still available as default_print_depth, but rarely used. Minor
4739 * Toplevel function "use" refers to raw ML bootstrap environment,
4740 without Isar context nor antiquotations. Potential INCOMPATIBILITY.
4741 Note that 'ML_file' is the canonical command to load ML files into the
4744 * Simplified programming interface to define ML antiquotations, see
4745 structure ML_Antiquotation. Minor INCOMPATIBILITY.
4747 * ML antiquotation @{here} refers to its source position, which is
4748 occasionally useful for experimentation and diagnostic purposes.
4750 * ML antiquotation @{path} produces a Path.T value, similarly to
4751 Path.explode, but with compile-time check against the file-system and
4752 some PIDE markup. Note that unlike theory source, ML does not have a
4753 well-defined master directory, so an absolute symbolic path
4754 specification is usually required, e.g. "~~/src/HOL".
4756 * ML antiquotation @{print} inlines a function to print an arbitrary
4757 ML value, which is occasionally useful for diagnostic or demonstration
4763 * Proof General with its traditional helper scripts is now an optional
4764 Isabelle component, e.g. see ProofGeneral-4.2-2 from the Isabelle
4765 component repository http://isabelle.in.tum.de/components/. Note that
4766 the "system" manual provides general explanations about add-on
4767 components, especially those that are not bundled with the release.
4769 * The raw Isabelle process executable has been renamed from
4770 "isabelle-process" to "isabelle_process", which conforms to common
4771 shell naming conventions, and allows to define a shell function within
4772 the Isabelle environment to avoid dynamic path lookup. Rare
4773 incompatibility for old tools that do not use the ISABELLE_PROCESS
4776 * Former "isabelle tty" has been superseded by "isabelle console",
4777 with implicit build like "isabelle jedit", and without the mostly
4778 obsolete Isar TTY loop.
4780 * Simplified "isabelle display" tool. Settings variables DVI_VIEWER
4781 and PDF_VIEWER now refer to the actual programs, not shell
4782 command-lines. Discontinued option -c: invocation may be asynchronous
4783 via desktop environment, without any special precautions. Potential
4784 INCOMPATIBILITY with ambitious private settings.
4786 * Removed obsolete "isabelle unsymbolize". Note that the usual format
4787 for email communication is the Unicode rendering of Isabelle symbols,
4788 as produced by Isabelle/jEdit, for example.
4790 * Removed obsolete tool "wwwfind". Similar functionality may be
4791 integrated into Isabelle/jEdit eventually.
4793 * Improved 'display_drafts' concerning desktop integration and
4794 repeated invocation in PIDE front-end: re-use single file
4795 $ISABELLE_HOME_USER/tmp/drafts.pdf and corresponding views.
4797 * Session ROOT specifications require explicit 'document_files' for
4798 robust dependencies on LaTeX sources. Only these explicitly given
4799 files are copied to the document output directory, before document
4800 processing is started.
4802 * Windows: support for regular TeX installation (e.g. MiKTeX) instead
4803 of TeX Live from Cygwin.
4807 New in Isabelle2013-2 (December 2013)
4808 -------------------------------------
4810 *** Prover IDE -- Isabelle/Scala/jEdit ***
4812 * More robust editing of running commands with internal forks,
4813 e.g. non-terminating 'by' steps.
4815 * More relaxed Sledgehammer panel: avoid repeated application of query
4816 after edits surrounding the command location.
4818 * More status information about commands that are interrupted
4819 accidentally (via physical event or Poly/ML runtime system signal,
4820 e.g. out-of-memory).
4825 * More robust termination of external processes managed by
4826 Isabelle/ML: support cancellation of tasks within the range of
4827 milliseconds, as required for PIDE document editing with automatically
4828 tried tools (e.g. Sledgehammer).
4830 * Reactivated Isabelle/Scala kill command for external processes on
4831 Mac OS X, which was accidentally broken in Isabelle2013-1 due to a
4832 workaround for some Debian/Ubuntu Linux versions from 2013.
4836 New in Isabelle2013-1 (November 2013)
4837 -------------------------------------
4841 * Discontinued obsolete 'uses' within theory header. Note that
4842 commands like 'ML_file' work without separate declaration of file
4843 dependencies. Minor INCOMPATIBILITY.
4845 * Discontinued redundant 'use' command, which was superseded by
4846 'ML_file' in Isabelle2013. Minor INCOMPATIBILITY.
4848 * Simplified subscripts within identifiers, using plain \<^sub>
4849 instead of the second copy \<^isub> and \<^isup>. Superscripts are
4850 only for literal tokens within notation; explicit mixfix annotations
4851 for consts or fixed variables may be used as fall-back for unusual
4852 names. Obsolete \<twosuperior> has been expanded to \<^sup>2 in
4853 Isabelle/HOL. INCOMPATIBILITY, use "isabelle update_sub_sup" to
4854 standardize symbols as a starting point for further manual cleanup.
4855 The ML reference variable "legacy_isub_isup" may be set as temporary
4856 workaround, to make the prover accept a subset of the old identifier
4859 * Document antiquotations: term style "isub" has been renamed to
4860 "sub". Minor INCOMPATIBILITY.
4862 * Uniform management of "quick_and_dirty" as system option (see also
4863 "isabelle options"), configuration option within the context (see also
4864 Config.get in Isabelle/ML), and attribute in Isabelle/Isar. Minor
4865 INCOMPATIBILITY, need to use more official Isabelle means to access
4866 quick_and_dirty, instead of historical poking into mutable reference.
4868 * Renamed command 'print_configs' to 'print_options'. Minor
4871 * Proper diagnostic command 'print_state'. Old 'pr' (with its
4872 implicit change of some global references) is retained for now as
4873 control command, e.g. for ProofGeneral 3.7.x.
4875 * Discontinued 'print_drafts' command with its old-fashioned PS output
4876 and Unix command-line print spooling. Minor INCOMPATIBILITY: use
4877 'display_drafts' instead and print via the regular document viewer.
4879 * Updated and extended "isar-ref" and "implementation" manual,
4880 eliminated old "ref" manual.
4883 *** Prover IDE -- Isabelle/Scala/jEdit ***
4885 * New manual "jedit" for Isabelle/jEdit, see isabelle doc or
4886 Documentation panel.
4888 * Dockable window "Documentation" provides access to Isabelle
4891 * Dockable window "Find" provides query operations for formal entities
4892 (GUI front-end to 'find_theorems' command).
4894 * Dockable window "Sledgehammer" manages asynchronous / parallel
4895 sledgehammer runs over existing document sources, independently of
4896 normal editing and checking process.
4898 * Dockable window "Timing" provides an overview of relevant command
4899 timing information, depending on option jedit_timing_threshold. The
4900 same timing information is shown in the extended tooltip of the
4901 command keyword, when hovering the mouse over it while the CONTROL or
4902 COMMAND modifier is pressed.
4904 * Improved dockable window "Theories": Continuous checking of proof
4905 document (visible and required parts) may be controlled explicitly,
4906 using check box or shortcut "C+e ENTER". Individual theory nodes may
4907 be marked explicitly as required and checked in full, using check box
4908 or shortcut "C+e SPACE".
4910 * Improved completion mechanism, which is now managed by the
4911 Isabelle/jEdit plugin instead of SideKick. Refined table of Isabelle
4912 symbol abbreviations (see $ISABELLE_HOME/etc/symbols).
4914 * Standard jEdit keyboard shortcut C+b complete-word is remapped to
4915 isabelle.complete for explicit completion in Isabelle sources.
4916 INCOMPATIBILITY wrt. jEdit defaults, may have to invent new shortcuts
4917 to resolve conflict.
4919 * Improved support of various "minor modes" for Isabelle NEWS,
4920 options, session ROOT etc., with completion and SideKick tree view.
4922 * Strictly monotonic document update, without premature cancellation of
4923 running transactions that are still needed: avoid reset/restart of
4924 such command executions while editing.
4926 * Support for asynchronous print functions, as overlay to existing
4929 * Support for automatic tools in HOL, which try to prove or disprove
4930 toplevel theorem statements.
4932 * Action isabelle.reset-font-size resets main text area font size
4933 according to Isabelle/Scala plugin option "jedit_font_reset_size" (see
4934 also "Plugin Options / Isabelle / General"). It can be bound to some
4935 keyboard shortcut by the user (e.g. C+0 and/or C+NUMPAD0).
4937 * File specifications in jEdit (e.g. file browser) may refer to
4938 $ISABELLE_HOME and $ISABELLE_HOME_USER on all platforms. Discontinued
4939 obsolete $ISABELLE_HOME_WINDOWS variable.
4941 * Improved support for Linux look-and-feel "GTK+", see also "Utilities
4942 / Global Options / Appearance".
4944 * Improved support of native Mac OS X functionality via "MacOSX"
4945 plugin, which is now enabled by default.
4950 * Commands 'interpretation' and 'sublocale' are now target-sensitive.
4951 In particular, 'interpretation' allows for non-persistent
4952 interpretation within "context ... begin ... end" blocks offering a
4953 light-weight alternative to 'sublocale'. See "isar-ref" manual for
4956 * Improved locales diagnostic command 'print_dependencies'.
4958 * Discontinued obsolete 'axioms' command, which has been marked as
4959 legacy since Isabelle2009-2. INCOMPATIBILITY, use 'axiomatization'
4960 instead, while observing its uniform scope for polymorphism.
4962 * Discontinued empty name bindings in 'axiomatization'.
4965 * System option "proofs" has been discontinued. Instead the global
4966 state of Proofterm.proofs is persistently compiled into logic images
4967 as required, notably HOL-Proofs. Users no longer need to change
4968 Proofterm.proofs dynamically. Minor INCOMPATIBILITY.
4970 * Syntax translation functions (print_translation etc.) always depend
4971 on Proof.context. Discontinued former "(advanced)" option -- this is
4972 now the default. Minor INCOMPATIBILITY.
4974 * Former global reference trace_unify_fail is now available as
4975 configuration option "unify_trace_failure" (global context only).
4977 * SELECT_GOAL now retains the syntactic context of the overall goal
4978 state (schematic variables etc.). Potential INCOMPATIBILITY in rare
4984 * Stronger precedence of syntax for big intersection and union on
4985 sets, in accordance with corresponding lattice operations.
4988 * Notation "{p:A. P}" now allows tuple patterns as well.
4990 * Nested case expressions are now translated in a separate check phase
4991 rather than during parsing. The data for case combinators is separated
4992 from the datatype package. The declaration attribute
4993 "case_translation" can be used to register new case combinators:
4995 declare [[case_translation case_combinator constructor1 ... constructorN]]
4998 - 'code_printing' unifies 'code_const' / 'code_type' / 'code_class' /
5000 - 'code_identifier' declares name hints for arbitrary identifiers in
5001 generated code, subsuming 'code_modulename'.
5003 See the isar-ref manual for syntax diagrams, and the HOL theories for
5006 * Attibute 'code': 'code' now declares concrete and abstract code
5007 equations uniformly. Use explicit 'code equation' and 'code abstract'
5008 to distinguish both when desired.
5010 * Discontinued theories Code_Integer and Efficient_Nat by a more
5011 fine-grain stack of theories Code_Target_Int, Code_Binary_Nat,
5012 Code_Target_Nat and Code_Target_Numeral. See the tutorial on code
5013 generation for details. INCOMPATIBILITY.
5015 * Numeric types are mapped by default to target language numerals:
5016 natural (replaces former code_numeral) and integer (replaces former
5017 code_int). Conversions are available as integer_of_natural /
5018 natural_of_integer / integer_of_nat / nat_of_integer (in HOL) and
5019 Code_Numeral.integer_of_natural / Code_Numeral.natural_of_integer (in
5020 ML). INCOMPATIBILITY.
5022 * Function package: For mutually recursive functions f and g, separate
5023 cases rules f.cases and g.cases are generated instead of unusable
5024 f_g.cases which exposed internal sum types. Potential INCOMPATIBILITY,
5025 in the case that the unusable rule was used nevertheless.
5027 * Function package: For each function f, new rules f.elims are
5028 generated, which eliminate equalities of the form "f x = t".
5030 * New command 'fun_cases' derives ad-hoc elimination rules for
5031 function equations as simplified instances of f.elims, analogous to
5032 inductive_cases. See ~~/src/HOL/ex/Fundefs.thy for some examples.
5035 - parametrized correspondence relations are now supported:
5036 + parametricity theorems for the raw term can be specified in
5037 the command lift_definition, which allow us to generate stronger
5039 + setup_lifting generates stronger transfer rules if parametric
5040 correspondence relation can be generated
5041 + various new properties of the relator must be specified to support
5043 + parametricity theorem for the Quotient relation can be specified
5044 - setup_lifting generates domain rules for the Transfer package
5045 - stronger reflexivity prover of respectfulness theorems for type
5047 - ===> and --> are now local. The symbols can be introduced
5048 by interpreting the locale lifting_syntax (typically in an
5050 - Lifting/Transfer relevant parts of Library/Quotient_* are now in
5051 Main. Potential INCOMPATIBILITY
5052 - new commands for restoring and deleting Lifting/Transfer context:
5053 lifting_forget, lifting_update
5054 - the command print_quotmaps was renamed to print_quot_maps.
5058 - better support for domains in Transfer: replace Domainp T
5059 by the actual invariant in a transferred goal
5060 - transfer rules can have as assumptions other transfer rules
5061 - Experimental support for transferring from the raw level to the
5062 abstract level: Transfer.transferred attribute
5063 - Attribute version of the transfer method: untransferred attribute
5065 * Reification and reflection:
5066 - Reification is now directly available in HOL-Main in structure
5068 - Reflection now handles multiple lists with variables also.
5069 - The whole reflection stack has been decomposed into conversions.
5072 * Revised devices for recursive definitions over finite sets:
5073 - Only one fundamental fold combinator on finite set remains:
5074 Finite_Set.fold :: ('a => 'b => 'b) => 'b => 'a set => 'b
5075 This is now identity on infinite sets.
5076 - Locales ("mini packages") for fundamental definitions with
5077 Finite_Set.fold: folding, folding_idem.
5078 - Locales comm_monoid_set, semilattice_order_set and
5079 semilattice_neutr_order_set for big operators on sets.
5080 See theory Big_Operators for canonical examples.
5081 Note that foundational constants comm_monoid_set.F and
5082 semilattice_set.F correspond to former combinators fold_image
5083 and fold1 respectively. These are now gone. You may use
5084 those foundational constants as substitutes, but it is
5085 preferable to interpret the above locales accordingly.
5086 - Dropped class ab_semigroup_idem_mult (special case of lattice,
5087 no longer needed in connection with Finite_Set.fold etc.)
5089 card.union_inter ~> card_Un_Int [symmetric]
5090 card.union_disjoint ~> card_Un_disjoint
5093 * Locale hierarchy for abstract orderings and (semi)lattices.
5095 * Complete_Partial_Order.admissible is defined outside the type class
5096 ccpo, but with mandatory prefix ccpo. Admissibility theorems lose the
5097 class predicate assumption or sort constraint when possible.
5100 * Introduce type class "conditionally_complete_lattice": Like a
5101 complete lattice but does not assume the existence of the top and
5102 bottom elements. Allows to generalize some lemmas about reals and
5103 extended reals. Removed SupInf and replaced it by the instantiation
5104 of conditionally_complete_lattice for real. Renamed lemmas about
5105 conditionally-complete lattice from Sup_... to cSup_... and from
5106 Inf_... to cInf_... to avoid hidding of similar complete lattice
5109 * Introduce type class linear_continuum as combination of
5110 conditionally-complete lattices and inner dense linorders which have
5111 more than one element. INCOMPATIBILITY.
5113 * Introduced type classes order_top and order_bot. The old classes top
5114 and bot only contain the syntax without assumptions. INCOMPATIBILITY:
5115 Rename bot -> order_bot, top -> order_top
5117 * Introduce type classes "no_top" and "no_bot" for orderings without
5118 top and bottom elements.
5120 * Split dense_linorder into inner_dense_order and no_top, no_bot.
5122 * Complex_Main: Unify and move various concepts from
5123 HOL-Multivariate_Analysis to HOL-Complex_Main.
5125 - Introduce type class (lin)order_topology and
5126 linear_continuum_topology. Allows to generalize theorems about
5127 limits and order. Instances are reals and extended reals.
5129 - continuous and continuos_on from Multivariate_Analysis:
5130 "continuous" is the continuity of a function at a filter. "isCont"
5131 is now an abbrevitation: "isCont x f == continuous (at _) f".
5133 Generalized continuity lemmas from isCont to continuous on an
5136 - compact from Multivariate_Analysis. Use Bolzano's lemma to prove
5137 compactness of closed intervals on reals. Continuous functions
5138 attain infimum and supremum on compact sets. The inverse of a
5139 continuous function is continuous, when the function is continuous
5142 - connected from Multivariate_Analysis. Use it to prove the
5143 intermediate value theorem. Show connectedness of intervals on
5144 linear_continuum_topology).
5146 - first_countable_topology from Multivariate_Analysis. Is used to
5147 show equivalence of properties on the neighbourhood filter of x and
5148 on all sequences converging to x.
5150 - FDERIV: Definition of has_derivative moved to Deriv.thy. Moved
5151 theorems from Library/FDERIV.thy to Deriv.thy and base the
5152 definition of DERIV on FDERIV. Add variants of DERIV and FDERIV
5153 which are restricted to sets, i.e. to represent derivatives from
5156 - Removed the within-filter. It is replaced by the principal filter:
5158 F within X = inf F (principal X)
5160 - Introduce "at x within U" as a single constant, "at x" is now an
5161 abbreviation for "at x within UNIV"
5163 - Introduce named theorem collections tendsto_intros,
5164 continuous_intros, continuous_on_intros and FDERIV_intros. Theorems
5165 in tendsto_intros (or FDERIV_intros) are also available as
5166 tendsto_eq_intros (or FDERIV_eq_intros) where the right-hand side
5167 is replaced by a congruence rule. This allows to apply them as
5168 intro rules and then proving equivalence by the simplifier.
5170 - Restructured theories in HOL-Complex_Main:
5172 + Moved RealDef and RComplete into Real
5174 + Introduced Topological_Spaces and moved theorems about
5175 topological spaces, filters, limits and continuity to it
5177 + Renamed RealVector to Real_Vector_Spaces
5179 + Split Lim, SEQ, Series into Topological_Spaces,
5180 Real_Vector_Spaces, and Limits
5182 + Moved Ln and Log to Transcendental
5184 + Moved theorems about continuity from Deriv to Topological_Spaces
5186 - Remove various auxiliary lemmas.
5191 - Added option "spy".
5192 - Reduce incidence of "too high arity" errors.
5196 isar_shrink ~> isar_compress
5198 - Added options "isar_try0", "spy".
5199 - Better support for "isar_proofs".
5200 - MaSh has been fined-tuned and now runs as a local server.
5202 * Improved support for ad hoc overloading of constants (see also
5203 isar-ref manual and ~~/src/HOL/ex/Adhoc_Overloading_Examples.thy).
5205 * Library/Polynomial.thy:
5206 - Use lifting for primitive definitions.
5207 - Explicit conversions from and to lists of coefficients, used for
5209 - Replaced recursion operator poly_rec by fold_coeffs.
5210 - Prefer pre-existing gcd operation for gcd.
5212 poly_eq_iff ~> poly_eq_poly_eq_iff
5213 poly_ext ~> poly_eqI
5214 expand_poly_eq ~> poly_eq_iff
5217 * New Library/Simps_Case_Conv.thy: Provides commands simps_of_case and
5218 case_of_simps to convert function definitions between a list of
5219 equations with patterns on the lhs and a single equation with case
5220 expressions on the rhs. See also Ex/Simps_Case_Conv_Examples.thy.
5222 * New Library/FSet.thy: type of finite sets defined as a subtype of
5223 sets defined by Lifting/Transfer.
5225 * Discontinued theory src/HOL/Library/Eval_Witness. INCOMPATIBILITY.
5227 * Consolidation of library theories on product orders:
5229 Product_Lattice ~> Product_Order -- pointwise order on products
5230 Product_ord ~> Product_Lexorder -- lexicographic order on products
5234 * Imperative-HOL: The MREC combinator is considered legacy and no
5235 longer included by default. INCOMPATIBILITY, use partial_function
5236 instead, or import theory Legacy_Mrec as a fallback.
5238 * HOL-Algebra: Discontinued theories ~~/src/HOL/Algebra/abstract and
5239 ~~/src/HOL/Algebra/poly. Existing theories should be based on
5240 ~~/src/HOL/Library/Polynomial instead. The latter provides
5241 integration with HOL's type classes for rings. INCOMPATIBILITY.
5244 - Various improvements to BNF-based (co)datatype package, including
5245 new commands "primrec_new", "primcorec", and
5246 "datatype_new_compat", as well as documentation. See
5247 "datatypes.pdf" for details.
5248 - New "coinduction" method to avoid some boilerplate (compared to
5251 data ~> datatype_new
5252 codata ~> codatatype
5254 - Renamed many generated theorems, including
5256 map_comp' ~> map_comp
5266 * Spec_Check is a Quickcheck tool for Isabelle/ML. The ML function
5267 "check_property" allows to check specifications of the form "ALL x y
5268 z. prop x y z". See also ~~/src/Tools/Spec_Check/ with its
5269 Examples.thy in particular.
5271 * Improved printing of exception trace in Poly/ML 5.5.1, with regular
5272 tracing output in the command transaction context instead of physical
5273 stdout. See also Toplevel.debug, Toplevel.debugging and
5274 ML_Compiler.exn_trace.
5276 * ML type "theory" is now immutable, without any special treatment of
5277 drafts or linear updates (which could lead to "stale theory" errors in
5278 the past). Discontinued obsolete operations like Theory.copy,
5279 Theory.checkpoint, and the auxiliary type theory_ref. Minor
5282 * More uniform naming of goal functions for skipped proofs:
5284 Skip_Proof.prove ~> Goal.prove_sorry
5285 Skip_Proof.prove_global ~> Goal.prove_sorry_global
5287 Minor INCOMPATIBILITY.
5289 * Simplifier tactics and tools use proper Proof.context instead of
5290 historic type simpset. Old-style declarations like addsimps,
5291 addsimprocs etc. operate directly on Proof.context. Raw type simpset
5292 retains its use as snapshot of the main Simplifier context, using
5293 simpset_of and put_simpset on Proof.context. INCOMPATIBILITY -- port
5294 old tools by making them depend on (ctxt : Proof.context) instead of
5295 (ss : simpset), then turn (simpset_of ctxt) into ctxt.
5297 * Modifiers for classical wrappers (e.g. addWrapper, delWrapper)
5298 operate on Proof.context instead of claset, for uniformity with addIs,
5299 addEs, addDs etc. Note that claset_of and put_claset allow to manage
5300 clasets separately from the context.
5302 * Discontinued obsolete ML antiquotations @{claset} and @{simpset}.
5303 INCOMPATIBILITY, use @{context} instead.
5305 * Antiquotation @{theory_context A} is similar to @{theory A}, but
5306 presents the result as initial Proof.context.
5311 * Discontinued obsolete isabelle usedir, mkdir, make -- superseded by
5312 "isabelle build" in Isabelle2013. INCOMPATIBILITY.
5314 * Discontinued obsolete isabelle-process options -f and -u (former
5315 administrative aliases of option -e). Minor INCOMPATIBILITY.
5317 * Discontinued obsolete isabelle print tool, and PRINT_COMMAND
5320 * Discontinued ISABELLE_DOC_FORMAT settings variable and historic
5321 document formats: dvi.gz, ps, ps.gz -- the default document format is
5324 * Isabelle settings variable ISABELLE_BUILD_JAVA_OPTIONS allows to
5325 specify global resources of the JVM process run by isabelle build.
5327 * Toplevel executable $ISABELLE_HOME/bin/isabelle_scala_script allows
5328 to run Isabelle/Scala source files as standalone programs.
5330 * Improved "isabelle keywords" tool (for old-style ProofGeneral
5331 keyword tables): use Isabelle/Scala operations, which inspect outer
5332 syntax without requiring to build sessions first.
5334 * Sessions may be organized via 'chapter' specifications in the ROOT
5335 file, which determines a two-level hierarchy of browser info. The old
5336 tree-like organization via implicit sub-session relation (with its
5337 tendency towards erratic fluctuation of URLs) has been discontinued.
5338 The default chapter is called "Unsorted". Potential INCOMPATIBILITY
5339 for HTML presentation of theories.
5343 New in Isabelle2013 (February 2013)
5344 -----------------------------------
5348 * Theorem status about oracles and unfinished/failed future proofs is
5349 no longer printed by default, since it is incompatible with
5350 incremental / parallel checking of the persistent document model. ML
5351 function Thm.peek_status may be used to inspect a snapshot of the
5352 ongoing evaluation process. Note that in batch mode --- notably
5353 isabelle build --- the system ensures that future proofs of all
5354 accessible theorems in the theory context are finished (as before).
5356 * Configuration option show_markup controls direct inlining of markup
5357 into the printed representation of formal entities --- notably type
5358 and sort constraints. This enables Prover IDE users to retrieve that
5359 information via tooltips in the output window, for example.
5361 * Command 'ML_file' evaluates ML text from a file directly within the
5362 theory, without any predeclaration via 'uses' in the theory header.
5364 * Old command 'use' command and corresponding keyword 'uses' in the
5365 theory header are legacy features and will be discontinued soon.
5366 Tools that load their additional source files may imitate the
5367 'ML_file' implementation, such that the system can take care of
5368 dependencies properly.
5370 * Discontinued obsolete method fastsimp / tactic fast_simp_tac, which
5371 is called fastforce / fast_force_tac already since Isabelle2011-1.
5373 * Updated and extended "isar-ref" and "implementation" manual, reduced
5374 remaining material in old "ref" manual.
5376 * Improved support for auxiliary contexts that indicate block structure
5377 for specifications. Nesting of "context fixes ... context assumes ..."
5378 and "class ... context ...".
5380 * Attribute "consumes" allows a negative value as well, which is
5381 interpreted relatively to the total number of premises of the rule in
5382 the target context. This form of declaration is stable when exported
5383 from a nested 'context' with additional assumptions. It is the
5384 preferred form for definitional packages, notably cases/rules produced
5385 in HOL/inductive and HOL/function.
5387 * More informative error messages for Isar proof commands involving
5388 lazy enumerations (method applications etc.).
5390 * Refined 'help' command to retrieve outer syntax commands according
5391 to name patterns (with clickable results).
5394 *** Prover IDE -- Isabelle/Scala/jEdit ***
5396 * Parallel terminal proofs ('by') are enabled by default, likewise
5397 proofs that are built into packages like 'datatype', 'function'. This
5398 allows to "run ahead" checking the theory specifications on the
5399 surface, while the prover is still crunching on internal
5400 justifications. Unfinished / cancelled proofs are restarted as
5401 required to complete full proof checking eventually.
5403 * Improved output panel with tooltips, hyperlinks etc. based on the
5404 same Rich_Text_Area as regular Isabelle/jEdit buffers. Activation of
5405 tooltips leads to some window that supports the same recursively,
5406 which can lead to stacks of tooltips as the semantic document content
5407 is explored. ESCAPE closes the whole stack, individual windows may be
5408 closed separately, or detached to become independent jEdit dockables.
5410 * Improved support for commands that produce graph output: the text
5411 message contains a clickable area to open a new instance of the graph
5414 * More robust incremental parsing of outer syntax (partial comments,
5415 malformed symbols). Changing the balance of open/close quotes and
5416 comment delimiters works more conveniently with unfinished situations
5417 that frequently occur in user interaction.
5419 * More efficient painting and improved reactivity when editing large
5420 files. More scalable management of formal document content.
5422 * Smarter handling of tracing messages: prover process pauses after
5423 certain number of messages per command transaction, with some user
5424 dialog to stop or continue. This avoids swamping the front-end with
5425 potentially infinite message streams.
5427 * More plugin options and preferences, based on Isabelle/Scala. The
5428 jEdit plugin option panel provides access to some Isabelle/Scala
5429 options, including tuning parameters for editor reactivity and color
5432 * Dockable window "Symbols" provides some editing support for Isabelle
5435 * Dockable window "Monitor" shows ML runtime statistics. Note that
5436 continuous display of the chart slows down the system.
5438 * Improved editing support for control styles: subscript, superscript,
5439 bold, reset of style -- operating on single symbols or text
5440 selections. Cf. keyboard shortcuts C+e DOWN/UP/RIGHT/LEFT.
5442 * Actions isabelle.increase-font-size and isabelle.decrease-font-size
5443 adjust the main text area font size, and its derivatives for output,
5444 tooltips etc. Cf. keyboard shortcuts C-PLUS and C-MINUS, which often
5445 need to be adapted to local keyboard layouts.
5447 * More reactive completion popup by default: use \t (TAB) instead of
5448 \n (NEWLINE) to minimize intrusion into regular flow of editing. See
5449 also "Plugin Options / SideKick / General / Code Completion Options".
5451 * Implicit check and build dialog of the specified logic session
5452 image. For example, HOL, HOLCF, HOL-Nominal can be produced on
5453 demand, without bundling big platform-dependent heap images in the
5454 Isabelle distribution.
5456 * Uniform Java 7 platform on Linux, Mac OS X, Windows: recent updates
5457 from Oracle provide better multi-platform experience. This version is
5458 now bundled exclusively with Isabelle.
5463 * Code generation for Haskell: restrict unqualified imports from
5464 Haskell Prelude to a small set of fundamental operations.
5466 * Command 'export_code': relative file names are interpreted
5467 relatively to master directory of current theory rather than the
5468 rather arbitrary current working directory. INCOMPATIBILITY.
5470 * Discontinued obsolete attribute "COMP". Potential INCOMPATIBILITY,
5471 use regular rule composition via "OF" / "THEN", or explicit proof
5472 structure instead. Note that Isabelle/ML provides a variety of
5473 operators like COMP, INCR_COMP, COMP_INCR, which need to be applied
5474 with some care where this is really required.
5476 * Command 'typ' supports an additional variant with explicit sort
5477 constraint, to infer and check the most general type conforming to a
5478 given sort. Example (in HOL):
5480 typ "_ * _ * bool * unit" :: finite
5482 * Command 'locale_deps' visualizes all locales and their relations as
5490 - Added MaSh relevance filter based on machine-learning; see the
5491 Sledgehammer manual for details.
5492 - Polished Isar proofs generated with "isar_proofs" option.
5493 - Rationalized type encodings ("type_enc" option).
5494 - Renamed "kill_provers" subcommand to "kill_all".
5496 isar_proof ~> isar_proofs
5497 isar_shrink_factor ~> isar_shrink
5498 max_relevant ~> max_facts
5499 relevance_thresholds ~> fact_thresholds
5501 * Quickcheck: added an optimisation for equality premises. It is
5502 switched on by default, and can be switched off by setting the
5503 configuration quickcheck_optimise_equality to false.
5505 * Quotient: only one quotient can be defined by quotient_type
5509 - generation of an abstraction function equation in lift_definition
5510 - quot_del attribute
5511 - renamed no_abs_code -> no_code (INCOMPATIBILITY.)
5513 * Simproc "finite_Collect" rewrites set comprehensions into pointfree
5516 * Preprocessing of the code generator rewrites set comprehensions into
5517 pointfree expressions.
5519 * The SMT solver Z3 has now by default a restricted set of directly
5520 supported features. For the full set of features (div/mod, nonlinear
5521 arithmetic, datatypes/records) with potential proof reconstruction
5522 failures, enable the configuration option "z3_with_extensions". Minor
5525 * Simplified 'typedef' specifications: historical options for implicit
5526 set definition and alternative name have been discontinued. The
5527 former behavior of "typedef (open) t = A" is now the default, but
5528 written just "typedef t = A". INCOMPATIBILITY, need to adapt theories
5531 * Removed constant "chars"; prefer "Enum.enum" on type "char"
5532 directly. INCOMPATIBILITY.
5534 * Moved operation product, sublists and n_lists from theory Enum to
5535 List. INCOMPATIBILITY.
5537 * Theorem UN_o generalized to SUP_comp. INCOMPATIBILITY.
5539 * Class "comm_monoid_diff" formalises properties of bounded
5540 subtraction, with natural numbers and multisets as typical instances.
5542 * Added combinator "Option.these" with type "'a option set => 'a set".
5544 * Theory "Transitive_Closure": renamed lemmas
5546 reflcl_tranclp -> reflclp_tranclp
5547 rtranclp_reflcl -> rtranclp_reflclp
5551 * Theory "Rings": renamed lemmas (in class semiring)
5553 left_distrib ~> distrib_right
5554 right_distrib ~> distrib_left
5558 * Generalized the definition of limits:
5560 - Introduced the predicate filterlim (LIM x F. f x :> G) which
5561 expresses that when the input values x converge to F then the
5562 output f x converges to G.
5564 - Added filters for convergence to positive (at_top) and negative
5567 - Moved infinity in the norm (at_infinity) from
5568 Multivariate_Analysis to Complex_Main.
5570 - Removed real_tendsto_inf, it is superseded by "LIM x F. f x :>
5575 * Theory "Library/Option_ord" provides instantiation of option type to
5576 lattice type classes.
5578 * Theory "Library/Multiset": renamed
5580 constant fold_mset ~> Multiset.fold
5581 fact fold_mset_commute ~> fold_mset_comm
5585 * Renamed theory Library/List_Prefix to Library/Sublist, with related
5588 - Renamed constants (and related lemmas)
5591 strict_prefix ~> prefix
5593 - Replaced constant "postfix" by "suffixeq" with swapped argument
5594 order (i.e., "postfix xs ys" is now "suffixeq ys xs") and dropped
5595 old infix syntax "xs >>= ys"; use "suffixeq ys xs" instead.
5596 Renamed lemmas accordingly.
5598 - Added constant "list_hembeq" for homeomorphic embedding on
5599 lists. Added abbreviation "sublisteq" for special case
5600 "list_hembeq (op =)".
5602 - Theory Library/Sublist no longer provides "order" and "bot" type
5603 class instances for the prefix order (merely corresponding locale
5604 interpretations). The type class instances are now in theory
5605 Library/Prefix_Order.
5607 - The sublist relation of theory Library/Sublist_Order is now based
5608 on "Sublist.sublisteq". Renamed lemmas accordingly:
5610 le_list_append_le_same_iff ~> Sublist.sublisteq_append_le_same_iff
5611 le_list_append_mono ~> Sublist.list_hembeq_append_mono
5612 le_list_below_empty ~> Sublist.list_hembeq_Nil, Sublist.list_hembeq_Nil2
5613 le_list_Cons_EX ~> Sublist.list_hembeq_ConsD
5614 le_list_drop_Cons2 ~> Sublist.sublisteq_Cons2'
5615 le_list_drop_Cons_neq ~> Sublist.sublisteq_Cons2_neq
5616 le_list_drop_Cons ~> Sublist.sublisteq_Cons'
5617 le_list_drop_many ~> Sublist.sublisteq_drop_many
5618 le_list_filter_left ~> Sublist.sublisteq_filter_left
5619 le_list_rev_drop_many ~> Sublist.sublisteq_rev_drop_many
5620 le_list_rev_take_iff ~> Sublist.sublisteq_append
5621 le_list_same_length ~> Sublist.sublisteq_same_length
5622 le_list_take_many_iff ~> Sublist.sublisteq_append'
5623 less_eq_list.drop ~> less_eq_list_drop
5624 less_eq_list.induct ~> less_eq_list_induct
5625 not_le_list_length ~> Sublist.not_sublisteq_length
5629 * New theory Library/Countable_Set.
5631 * Theory Library/Debug and Library/Parallel provide debugging and
5632 parallel execution for code generated towards Isabelle/ML.
5634 * Theory Library/FuncSet: Extended support for Pi and extensional and
5635 introduce the extensional dependent function space "PiE". Replaced
5636 extensional_funcset by an abbreviation, and renamed lemmas from
5637 extensional_funcset to PiE as follows:
5639 extensional_empty ~> PiE_empty
5640 extensional_funcset_empty_domain ~> PiE_empty_domain
5641 extensional_funcset_empty_range ~> PiE_empty_range
5642 extensional_funcset_arb ~> PiE_arb
5643 extensional_funcset_mem ~> PiE_mem
5644 extensional_funcset_extend_domainI ~> PiE_fun_upd
5645 extensional_funcset_restrict_domain ~> fun_upd_in_PiE
5646 extensional_funcset_extend_domain_eq ~> PiE_insert_eq
5647 card_extensional_funcset ~> card_PiE
5648 finite_extensional_funcset ~> finite_PiE
5652 * Theory Library/FinFun: theory of almost everywhere constant
5653 functions (supersedes the AFP entry "Code Generation for Functions as
5656 * Theory Library/Phantom: generic phantom type to make a type
5657 parameter appear in a constant's type. This alternative to adding
5658 TYPE('a) as another parameter avoids unnecessary closures in generated
5661 * Theory Library/RBT_Impl: efficient construction of red-black trees
5662 from sorted associative lists. Merging two trees with rbt_union may
5663 return a structurally different tree than before. Potential
5666 * Theory Library/IArray: immutable arrays with code generation.
5668 * Theory Library/Finite_Lattice: theory of finite lattices.
5670 * HOL/Multivariate_Analysis: replaced
5672 "basis :: 'a::euclidean_space => nat => real"
5673 "\<Chi>\<Chi> :: (nat => real) => 'a::euclidean_space"
5675 on euclidean spaces by using the inner product "_ \<bullet> _" with
5676 vectors from the Basis set: "\<Chi>\<Chi> i. f i" is superseded by
5677 "SUM i : Basis. f i * r i".
5679 With this change the following constants are also changed or removed:
5681 DIM('a) :: nat ~> card (Basis :: 'a set) (is an abbreviation)
5682 a $$ i ~> inner a i (where i : Basis)
5684 \<pi>, \<pi>' removed
5686 Theorems about these constants where removed.
5690 component_le_norm ~> Basis_le_norm
5691 euclidean_eq ~> euclidean_eq_iff
5692 differential_zero_maxmin_component ~> differential_zero_maxmin_cart
5693 euclidean_simps ~> inner_simps
5694 independent_basis ~> independent_Basis
5695 span_basis ~> span_Basis
5696 in_span_basis ~> in_span_Basis
5697 norm_bound_component_le ~> norm_boound_Basis_le
5698 norm_bound_component_lt ~> norm_boound_Basis_lt
5699 component_le_infnorm ~> Basis_le_infnorm
5705 - Added simproc "measurable" to automatically prove measurability.
5707 - Added induction rules for sigma sets with disjoint union
5708 (sigma_sets_induct_disjoint) and for Borel-measurable functions
5709 (borel_measurable_induct).
5711 - Added the Daniell-Kolmogorov theorem (the existence the limit of a
5714 * HOL/Cardinals: Theories of ordinals and cardinals (supersedes the
5715 AFP entry "Ordinals_and_Cardinals").
5717 * HOL/BNF: New (co)datatype package based on bounded natural functors
5718 with support for mixed, nested recursion and interesting non-free
5721 * HOL/Finite_Set and Relation: added new set and relation operations
5722 expressed by Finite_Set.fold.
5724 * New theory HOL/Library/RBT_Set: implementation of sets by red-black
5725 trees for the code generator.
5727 * HOL/Library/RBT and HOL/Library/Mapping have been converted to
5729 possible INCOMPATIBILITY.
5731 * HOL/Set: renamed Set.project -> Set.filter
5735 *** Document preparation ***
5737 * Dropped legacy antiquotations "term_style" and "thm_style", since
5738 styles may be given as arguments to "term" and "thm" already.
5739 Discontinued legacy styles "prem1" .. "prem19".
5741 * Default LaTeX rendering for \<euro> is now based on eurosym package,
5742 instead of slightly exotic babel/greek.
5744 * Document variant NAME may use different LaTeX entry point
5745 document/root_NAME.tex if that file exists, instead of the common
5748 * Simplified custom document/build script, instead of old-style
5749 document/IsaMakefile. Minor INCOMPATIBILITY.
5754 * The default limit for maximum number of worker threads is now 8,
5755 instead of 4, in correspondence to capabilities of contemporary
5756 hardware and Poly/ML runtime system.
5758 * Type Seq.results and related operations support embedded error
5759 messages within lazy enumerations, and thus allow to provide
5760 informative errors in the absence of any usable results.
5762 * Renamed Position.str_of to Position.here to emphasize that this is a
5763 formal device to inline positions into message text, but not
5764 necessarily printing visible text.
5769 * Advanced support for Isabelle sessions and build management, see
5770 "system" manual for the chapter of that name, especially the "isabelle
5771 build" tool and its examples. The "isabelle mkroot" tool prepares
5772 session root directories for use with "isabelle build", similar to
5773 former "isabelle mkdir" for "isabelle usedir". Note that this affects
5774 document preparation as well. INCOMPATIBILITY, isabelle usedir /
5775 mkdir / make are rendered obsolete.
5777 * Discontinued obsolete Isabelle/build script, it is superseded by the
5778 regular isabelle build tool. For example:
5780 isabelle build -s -b HOL
5782 * Discontinued obsolete "isabelle makeall".
5784 * Discontinued obsolete IsaMakefile and ROOT.ML files from the
5785 Isabelle distribution, except for rudimentary src/HOL/IsaMakefile that
5786 provides some traditional targets that invoke "isabelle build". Note
5787 that this is inefficient! Applications of Isabelle/HOL involving
5788 "isabelle make" should be upgraded to use "isabelle build" directly.
5790 * The "isabelle options" tool prints Isabelle system options, as
5791 required for "isabelle build", for example.
5793 * The "isabelle logo" tool produces EPS and PDF format simultaneously.
5794 Minor INCOMPATIBILITY in command-line options.
5796 * The "isabelle install" tool has now a simpler command-line. Minor
5799 * The "isabelle components" tool helps to resolve add-on components
5800 that are not bundled, or referenced from a bare-bones repository
5801 version of Isabelle.
5803 * Settings variable ISABELLE_PLATFORM_FAMILY refers to the general
5804 platform family: "linux", "macos", "windows".
5806 * The ML system is configured as regular component, and no longer
5807 picked up from some surrounding directory. Potential INCOMPATIBILITY
5808 for home-made settings.
5810 * Improved ML runtime statistics (heap, threads, future tasks etc.).
5812 * Discontinued support for Poly/ML 5.2.1, which was the last version
5813 without exception positions and advanced ML compiler/toplevel
5816 * Discontinued special treatment of Proof General -- no longer guess
5817 PROOFGENERAL_HOME based on accidental file-system layout. Minor
5818 INCOMPATIBILITY: provide PROOFGENERAL_HOME and PROOFGENERAL_OPTIONS
5819 settings manually, or use a Proof General version that has been
5820 bundled as Isabelle component.
5824 New in Isabelle2012 (May 2012)
5825 ------------------------------
5829 * Prover IDE (PIDE) improvements:
5831 - more robust Sledgehammer integration (as before the sledgehammer
5832 command-line needs to be typed into the source buffer)
5833 - markup for bound variables
5834 - markup for types of term variables (displayed as tooltips)
5835 - support for user-defined Isar commands within the running session
5836 - improved support for Unicode outside original 16bit range
5837 e.g. glyph for \<A> (thanks to jEdit 4.5.1)
5839 * Forward declaration of outer syntax keywords within the theory
5840 header -- minor INCOMPATIBILITY for user-defined commands. Allow new
5841 commands to be used in the same theory where defined.
5843 * Auxiliary contexts indicate block structure for specifications with
5844 additional parameters and assumptions. Such unnamed contexts may be
5845 nested within other targets, like 'theory', 'locale', 'class',
5846 'instantiation' etc. Results from the local context are generalized
5847 accordingly and applied to the enclosing target context. Example:
5851 assumes xy: "x = y" and yz: "y = z"
5854 lemma my_trans: "x = z" using xy yz by simp
5860 The most basic application is to factor-out context elements of
5861 several fixes/assumes/shows theorem statements, e.g. see
5862 ~~/src/HOL/Isar_Examples/Group_Context.thy
5864 Any other local theory specification element works within the "context
5865 ... begin ... end" block as well.
5867 * Bundled declarations associate attributed fact expressions with a
5868 given name in the context. These may be later included in other
5869 contexts. This allows to manage context extensions casually, without
5870 the logical dependencies of locales and locale interpretation. See
5871 commands 'bundle', 'include', 'including' etc. in the isar-ref manual.
5873 * Commands 'lemmas' and 'theorems' allow local variables using 'for'
5874 declaration, and results are standardized before being stored. Thus
5875 old-style "standard" after instantiation or composition of facts
5876 becomes obsolete. Minor INCOMPATIBILITY, due to potential change of
5877 indices of schematic variables.
5879 * Rule attributes in local theory declarations (e.g. locale or class)
5880 are now statically evaluated: the resulting theorem is stored instead
5881 of the original expression. INCOMPATIBILITY in rare situations, where
5882 the historic accident of dynamic re-evaluation in interpretations
5885 * New tutorial "Programming and Proving in Isabelle/HOL"
5886 ("prog-prove"). It completely supersedes "A Tutorial Introduction to
5887 Structured Isar Proofs" ("isar-overview"), which has been removed. It
5888 also supersedes "Isabelle/HOL, A Proof Assistant for Higher-Order
5889 Logic" as the recommended beginners tutorial, but does not cover all
5890 of the material of that old tutorial.
5892 * Updated and extended reference manuals: "isar-ref",
5893 "implementation", "system"; reduced remaining material in old "ref"
5899 * Command 'definition' no longer exports the foundational "raw_def"
5900 into the user context. Minor INCOMPATIBILITY, may use the regular
5901 "def" result with attribute "abs_def" to imitate the old version.
5903 * Attribute "abs_def" turns an equation of the form "f x y == t" into
5904 "f == %x y. t", which ensures that "simp" or "unfold" steps always
5905 expand it. This also works for object-logic equality. (Formerly
5906 undocumented feature.)
5908 * Sort constraints are now propagated in simultaneous statements, just
5909 like type constraints. INCOMPATIBILITY in rare situations, where
5910 distinct sorts used to be assigned accidentally. For example:
5912 lemma "P (x::'a::foo)" and "Q (y::'a::bar)" -- "now illegal"
5914 lemma "P (x::'a)" and "Q (y::'a::bar)"
5915 -- "now uniform 'a::bar instead of default sort for first occurrence (!)"
5917 * Rule composition via attribute "OF" (or ML functions OF/MRS) is more
5918 tolerant against multiple unifiers, as long as the final result is
5919 unique. (As before, rules are composed in canonical right-to-left
5920 order to accommodate newly introduced premises.)
5922 * Renamed some inner syntax categories:
5928 Minor INCOMPATIBILITY. Note that in practice "num_const" or
5929 "num_position" etc. are mainly used instead (which also include
5930 position information via constraints).
5932 * Simplified configuration options for syntax ambiguity: see
5933 "syntax_ambiguity_warning" and "syntax_ambiguity_limit" in isar-ref
5934 manual. Minor INCOMPATIBILITY.
5936 * Discontinued configuration option "syntax_positions": atomic terms
5937 in parse trees are always annotated by position constraints.
5939 * Old code generator for SML and its commands 'code_module',
5940 'code_library', 'consts_code', 'types_code' have been discontinued.
5941 Use commands of the generic code generator instead. INCOMPATIBILITY.
5943 * Redundant attribute "code_inline" has been discontinued. Use
5944 "code_unfold" instead. INCOMPATIBILITY.
5946 * Dropped attribute "code_unfold_post" in favor of the its dual
5947 "code_abbrev", which yields a common pattern in definitions like
5949 definition [code_abbrev]: "f = t"
5953 * Obsolete 'types' command has been discontinued. Use 'type_synonym'
5954 instead. INCOMPATIBILITY.
5956 * Discontinued old "prems" fact, which used to refer to the accidental
5957 collection of foundational premises in the context (already marked as
5958 legacy since Isabelle2011).
5963 * Type 'a set is now a proper type constructor (just as before
5964 Isabelle2008). Definitions mem_def and Collect_def have disappeared.
5965 Non-trivial INCOMPATIBILITY. For developments keeping predicates and
5966 sets separate, it is often sufficient to rephrase some set S that has
5967 been accidentally used as predicates by "%x. x : S", and some
5968 predicate P that has been accidentally used as set by "{x. P x}".
5969 Corresponding proofs in a first step should be pruned from any
5970 tinkering with former theorems mem_def and Collect_def as far as
5973 For developments which deliberately mix predicates and sets, a
5974 planning step is necessary to determine what should become a predicate
5975 and what a set. It can be helpful to carry out that step in
5976 Isabelle2011-1 before jumping right into the current release.
5978 * Code generation by default implements sets as container type rather
5979 than predicates. INCOMPATIBILITY.
5981 * New type synonym 'a rel = ('a * 'a) set
5983 * The representation of numerals has changed. Datatype "num"
5984 represents strictly positive binary numerals, along with functions
5985 "numeral :: num => 'a" and "neg_numeral :: num => 'a" to represent
5986 positive and negated numeric literals, respectively. See also
5987 definitions in ~~/src/HOL/Num.thy. Potential INCOMPATIBILITY, some
5988 user theories may require adaptations as follows:
5990 - Theorems with number_ring or number_semiring constraints: These
5991 classes are gone; use comm_ring_1 or comm_semiring_1 instead.
5993 - Theories defining numeric types: Remove number, number_semiring,
5994 and number_ring instances. Defer all theorems about numerals until
5995 after classes one and semigroup_add have been instantiated.
5997 - Numeral-only simp rules: Replace each rule having a "number_of v"
5998 pattern with two copies, one for numeral and one for neg_numeral.
6000 - Theorems about subclasses of semiring_1 or ring_1: These classes
6001 automatically support numerals now, so more simp rules and
6002 simprocs may now apply within the proof.
6004 - Definitions and theorems using old constructors Pls/Min/Bit0/Bit1:
6005 Redefine using other integer operations.
6007 * Transfer: New package intended to generalize the existing
6008 "descending" method and related theorem attributes from the Quotient
6009 package. (Not all functionality is implemented yet, but future
6010 development will focus on Transfer as an eventual replacement for the
6011 corresponding parts of the Quotient package.)
6013 - transfer_rule attribute: Maintains a collection of transfer rules,
6014 which relate constants at two different types. Transfer rules may
6015 relate different type instances of the same polymorphic constant,
6016 or they may relate an operation on a raw type to a corresponding
6017 operation on an abstract type (quotient or subtype). For example:
6019 ((A ===> B) ===> list_all2 A ===> list_all2 B) map map
6020 (cr_int ===> cr_int ===> cr_int) (%(x,y) (u,v). (x+u, y+v)) plus_int
6022 - transfer method: Replaces a subgoal on abstract types with an
6023 equivalent subgoal on the corresponding raw types. Constants are
6024 replaced with corresponding ones according to the transfer rules.
6025 Goals are generalized over all free variables by default; this is
6026 necessary for variables whose types change, but can be overridden
6027 for specific variables with e.g. "transfer fixing: x y z". The
6028 variant transfer' method allows replacing a subgoal with one that
6029 is logically stronger (rather than equivalent).
6031 - relator_eq attribute: Collects identity laws for relators of
6032 various type constructors, e.g. "list_all2 (op =) = (op =)". The
6033 transfer method uses these lemmas to infer transfer rules for
6034 non-polymorphic constants on the fly.
6036 - transfer_prover method: Assists with proving a transfer rule for a
6037 new constant, provided the constant is defined in terms of other
6038 constants that already have transfer rules. It should be applied
6039 after unfolding the constant definitions.
6041 - HOL/ex/Transfer_Int_Nat.thy: Example theory demonstrating transfer
6042 from type nat to type int.
6044 * Lifting: New package intended to generalize the quotient_definition
6045 facility of the Quotient package; designed to work with Transfer.
6047 - lift_definition command: Defines operations on an abstract type in
6048 terms of a corresponding operation on a representation
6049 type. Example syntax:
6051 lift_definition dlist_insert :: "'a => 'a dlist => 'a dlist"
6054 Users must discharge a respectfulness proof obligation when each
6055 constant is defined. (For a type copy, i.e. a typedef with UNIV,
6056 the proof is discharged automatically.) The obligation is
6057 presented in a user-friendly, readable form; a respectfulness
6058 theorem in the standard format and a transfer rule are generated
6061 - Integration with code_abstype: For typedefs (e.g. subtypes
6062 corresponding to a datatype invariant, such as dlist),
6063 lift_definition generates a code certificate theorem and sets up
6064 code generation for each constant.
6066 - setup_lifting command: Sets up the Lifting package to work with a
6067 user-defined type. The user must provide either a quotient theorem
6068 or a type_definition theorem. The package configures transfer
6069 rules for equality and quantifiers on the type, and sets up the
6070 lift_definition command to work with the type.
6072 - Usage examples: See Quotient_Examples/Lift_DList.thy,
6073 Quotient_Examples/Lift_RBT.thy, Quotient_Examples/Lift_FSet.thy,
6074 Word/Word.thy and Library/Float.thy.
6078 - The 'quotient_type' command now supports a 'morphisms' option with
6079 rep and abs functions, similar to typedef.
6081 - 'quotient_type' sets up new types to work with the Lifting and
6082 Transfer packages, as with 'setup_lifting'.
6084 - The 'quotient_definition' command now requires the user to prove a
6085 respectfulness property at the point where the constant is
6086 defined, similar to lift_definition; INCOMPATIBILITY.
6088 - Renamed predicate 'Quotient' to 'Quotient3', and renamed theorems
6089 accordingly, INCOMPATIBILITY.
6091 * New diagnostic command 'find_unused_assms' to find potentially
6092 superfluous assumptions in theorems using Quickcheck.
6096 - Quickcheck returns variable assignments as counterexamples, which
6097 allows to reveal the underspecification of functions under test.
6098 For example, refuting "hd xs = x", it presents the variable
6099 assignment xs = [] and x = a1 as a counterexample, assuming that
6100 any property is false whenever "hd []" occurs in it.
6102 These counterexample are marked as potentially spurious, as
6103 Quickcheck also returns "xs = []" as a counterexample to the
6104 obvious theorem "hd xs = hd xs".
6106 After finding a potentially spurious counterexample, Quickcheck
6107 continues searching for genuine ones.
6109 By default, Quickcheck shows potentially spurious and genuine
6110 counterexamples. The option "genuine_only" sets quickcheck to only
6111 show genuine counterexamples.
6113 - The command 'quickcheck_generator' creates random and exhaustive
6114 value generators for a given type and operations.
6116 It generates values by using the operations as if they were
6117 constructors of that type.
6119 - Support for multisets.
6121 - Added "use_subtype" options.
6123 - Added "quickcheck_locale" configuration to specify how to process
6124 conjectures in a locale context.
6126 * Nitpick: Fixed infinite loop caused by the 'peephole_optim' option
6127 and affecting 'rat' and 'real'.
6130 - Integrated more tightly with SPASS, as described in the ITP 2012
6131 paper "More SPASS with Isabelle".
6132 - Made it try "smt" as a fallback if "metis" fails or times out.
6133 - Added support for the following provers: Alt-Ergo (via Why3 and
6134 TFF1), iProver, iProver-Eq.
6135 - Sped up the minimizer.
6136 - Added "lam_trans", "uncurry_aliases", and "minimize" options.
6137 - Renamed "slicing" ("no_slicing") option to "slice" ("dont_slice").
6138 - Renamed "sound" option to "strict".
6140 * Metis: Added possibility to specify lambda translations scheme as a
6141 parenthesized argument (e.g., "by (metis (lifting) ...)").
6143 * SMT: Renamed "smt_fixed" option to "smt_read_only_certificates".
6145 * Command 'try0': Renamed from 'try_methods'. INCOMPATIBILITY.
6147 * New "case_product" attribute to generate a case rule doing multiple
6148 case distinctions at the same time. E.g.
6150 list.exhaust [case_product nat.exhaust]
6152 produces a rule which can be used to perform case distinction on both
6155 * New "eventually_elim" method as a generalized variant of the
6156 eventually_elim* rules. Supports structured proofs.
6158 * Typedef with implicit set definition is considered legacy. Use
6159 "typedef (open)" form instead, which will eventually become the
6162 * Record: code generation can be switched off manually with
6164 declare [[record_coden = false]] -- "default true"
6166 * Datatype: type parameters allow explicit sort constraints.
6168 * Concrete syntax for case expressions includes constraints for source
6169 positions, and thus produces Prover IDE markup for its bindings.
6170 INCOMPATIBILITY for old-style syntax translations that augment the
6171 pattern notation; e.g. see src/HOL/HOLCF/One.thy for translations of
6174 * Clarified attribute "mono_set": pure declaration without modifying
6175 the result of the fact expression.
6177 * More default pred/set conversions on a couple of relation operations
6178 and predicates. Added powers of predicate relations. Consolidation
6179 of some relation theorems:
6181 converse_def ~> converse_unfold
6182 rel_comp_def ~> relcomp_unfold
6183 symp_def ~> (modified, use symp_def and sym_def instead)
6184 transp_def ~> transp_trans
6185 Domain_def ~> Domain_unfold
6186 Range_def ~> Domain_converse [symmetric]
6188 Generalized theorems INF_INT_eq, INF_INT_eq2, SUP_UN_eq, SUP_UN_eq2.
6190 See theory "Relation" for examples for making use of pred/set
6191 conversions by means of attributes "to_set" and "to_pred".
6195 * Renamed facts about the power operation on relations, i.e., relpow
6196 to match the constant's name:
6198 rel_pow_1 ~> relpow_1
6199 rel_pow_0_I ~> relpow_0_I
6200 rel_pow_Suc_I ~> relpow_Suc_I
6201 rel_pow_Suc_I2 ~> relpow_Suc_I2
6202 rel_pow_0_E ~> relpow_0_E
6203 rel_pow_Suc_E ~> relpow_Suc_E
6204 rel_pow_E ~> relpow_E
6205 rel_pow_Suc_D2 ~> relpow_Suc_D2
6206 rel_pow_Suc_E2 ~> relpow_Suc_E2
6207 rel_pow_Suc_D2' ~> relpow_Suc_D2'
6208 rel_pow_E2 ~> relpow_E2
6209 rel_pow_add ~> relpow_add
6210 rel_pow_commute ~> relpow
6211 rel_pow_empty ~> relpow_empty:
6212 rtrancl_imp_UN_rel_pow ~> rtrancl_imp_UN_relpow
6213 rel_pow_imp_rtrancl ~> relpow_imp_rtrancl
6214 rtrancl_is_UN_rel_pow ~> rtrancl_is_UN_relpow
6215 rtrancl_imp_rel_pow ~> rtrancl_imp_relpow
6216 rel_pow_fun_conv ~> relpow_fun_conv
6217 rel_pow_finite_bounded1 ~> relpow_finite_bounded1
6218 rel_pow_finite_bounded ~> relpow_finite_bounded
6219 rtrancl_finite_eq_rel_pow ~> rtrancl_finite_eq_relpow
6220 trancl_finite_eq_rel_pow ~> trancl_finite_eq_relpow
6221 single_valued_rel_pow ~> single_valued_relpow
6225 * Theory Relation: Consolidated constant name for relation composition
6226 and corresponding theorem names:
6228 - Renamed constant rel_comp to relcomp.
6230 - Dropped abbreviation pred_comp. Use relcompp instead.
6234 rel_compI ~> relcompI
6235 rel_compEpair ~> relcompEpair
6236 rel_compE ~> relcompE
6237 pred_comp_rel_comp_eq ~> relcompp_relcomp_eq
6238 rel_comp_empty1 ~> relcomp_empty1
6239 rel_comp_mono ~> relcomp_mono
6240 rel_comp_subset_Sigma ~> relcomp_subset_Sigma
6241 rel_comp_distrib ~> relcomp_distrib
6242 rel_comp_distrib2 ~> relcomp_distrib2
6243 rel_comp_UNION_distrib ~> relcomp_UNION_distrib
6244 rel_comp_UNION_distrib2 ~> relcomp_UNION_distrib2
6245 single_valued_rel_comp ~> single_valued_relcomp
6246 rel_comp_def ~> relcomp_unfold
6247 converse_rel_comp ~> converse_relcomp
6248 pred_compI ~> relcomppI
6249 pred_compE ~> relcomppE
6250 pred_comp_bot1 ~> relcompp_bot1
6251 pred_comp_bot2 ~> relcompp_bot2
6252 transp_pred_comp_less_eq ~> transp_relcompp_less_eq
6253 pred_comp_mono ~> relcompp_mono
6254 pred_comp_distrib ~> relcompp_distrib
6255 pred_comp_distrib2 ~> relcompp_distrib2
6256 converse_pred_comp ~> converse_relcompp
6258 finite_rel_comp ~> finite_relcomp