added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
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 this Isabelle version
8 ----------------------------
12 * Marginal comments need to be written exclusively in the new-style form
13 "\<comment> \<open>text\<close>", old ASCII variants like "-- {* ... *}" are no longer
14 supported. INCOMPATIBILITY, use the command-line tool "isabelle
15 update_comments" to update existing theory files.
17 * Old-style inner comments (* ... *) within the term language are legacy
18 and will be discontinued soon: use formal comments "\<comment> \<open>...\<close>" or "\<^cancel>\<open>...\<close>"
21 * The "op <infix-op>" syntax for infix operators has been replaced by
22 "(<infix-op>)". If <infix-op> begins or ends with a "*", there needs to
23 be a space between the "*" and the corresponding parenthesis.
25 There is an Isabelle tool "update_op" that converts theory and ML files
26 to the new syntax. Because it is based on regular expression matching,
27 the result may need a bit of manual postprocessing. Invoking "isabelle
28 update_op" converts all files in the current directory (recursively).
29 In case you want to exclude conversion of ML files (because the tool
30 frequently also converts ML's "op" syntax), use option "-m".
32 * The old 'def' command has been discontinued (legacy since
33 Isbelle2016-1). INCOMPATIBILITY, use 'define' instead -- usually with
34 object-logic equality or equivalence.
36 * Session-qualified theory names are mandatory: it is no longer possible
37 to refer to unqualified theories from the parent session.
38 INCOMPATIBILITY for old developments that have not been updated to
39 Isabelle2017 yet (using the "isabelle imports" tool).
41 * Theory header 'abbrevs' specifications need to be separated by 'and'.
44 * Only the most fundamental theory names are global, usually the entry
45 points to major logic sessions: Pure, Main, Complex_Main, HOLCF, IFOL,
46 FOL, ZF, ZFC etc. INCOMPATIBILITY, need to use qualified names for
47 formerly global "HOL-Probability.Probability" and "HOL-SPARK.SPARK".
49 * Command 'external_file' declares the formal dependency on the given
50 file name, such that the Isabelle build process knows about it, but
51 without specific Prover IDE management.
53 * Session ROOT entries no longer allow specification of 'files'. Rare
54 INCOMPATIBILITY, use command 'external_file' within a proper theory
57 * Session root directories may be specified multiple times: each
58 accessible ROOT file is processed only once. This facilitates
59 specification of $ISABELLE_HOME_USER/ROOTS or command-line options like
60 -d or -D for "isabelle build" and "isabelle jedit". Example:
62 isabelle build -D '~~/src/ZF'
64 * The command 'display_drafts' has been discontinued. INCOMPATIBILITY,
65 use action "isabelle.draft" (or "print") in Isabelle/jEdit instead.
67 * Isabelle symbol "\<hyphen>" is rendered as explicit Unicode hyphen U+2010, to
68 avoid unclear meaning of the old "soft hyphen" U+00AD. Rare
69 INCOMPATIBILITY, e.g. copy-paste of historic Isabelle HTML output.
72 *** Isabelle/jEdit Prover IDE ***
74 * System options "spell_checker_include" and "spell_checker_exclude"
75 supersede former "spell_checker_elements" to determine regions of text
76 that are subject to spell-checking. Minor INCOMPATIBILITY.
78 * PIDE markup for session ROOT files: allows to complete session names,
79 follow links to theories and document files etc.
81 * Completion supports theory header imports, using theory base name.
82 E.g. "Prob" may be completed to "HOL-Probability.Probability".
84 * The command-line tool "isabelle jedit" provides more flexible options
85 for session selection:
86 - options -P opens the parent session image of -l
87 - options -A and -B modify the meaning of -l to produce a base
88 image on the spot, based on the specified ancestor (or parent)
89 - option -F focuses on the specified logic session
90 - option -R has changed: it only opens the session ROOT entry
91 - option -S sets up the development environment to edit the
92 specified session: it abbreviates -B -F -R -l
95 isabelle jedit -d '$AFP' -S Formal_SSA -A HOL
96 isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis
98 * Named control symbols (without special Unicode rendering) are shown as
99 bold-italic keyword. This is particularly useful for the short form of
100 antiquotations with control symbol: \<^name>\<open>argument\<close>. The action
101 "isabelle.antiquoted_cartouche" turns an antiquotation with 0 or 1
102 arguments into this format.
104 * Completion provides templates for named symbols with arguments,
105 e.g. "\<comment> \<open>ARGUMENT\<close>" or "\<^emph>\<open>ARGUMENT\<close>".
107 * Bibtex database files (.bib) are semantically checked.
109 * Action "isabelle.preview" is able to present more file formats,
110 notably bibtex database files and ML files.
112 * Action "isabelle.draft" is similar to "isabelle.preview", but shows a
113 plain-text document draft.
115 * When loading text files, the Isabelle symbols encoding UTF-8-Isabelle
116 is only used if there is no conflict with existing Unicode sequences in
117 the file. Otherwise, the fallback encoding is plain UTF-8 and Isabelle
118 symbols remain in literal \<symbol> form. This avoids accidental loss of
119 Unicode content when saving the file.
121 * Update to jedit-5.5.0, the latest release.
124 *** Isabelle/VSCode Prover IDE ***
126 * HTML preview of theories and other file-formats similar to
130 *** Document preparation ***
132 * Formal comments work uniformly in outer syntax, inner syntax (term
133 language), Isabelle/ML and some other embedded languages of Isabelle.
134 See also "Document comments" in the isar-ref manual. The following forms
137 - marginal text comment: \<comment> \<open>\<dots>\<close>
138 - canceled source: \<^cancel>\<open>\<dots>\<close>
139 - raw LaTeX: \<^latex>\<open>\<dots>\<close>
141 * Outside of the inner theory body, the default presentation context is
142 theory Pure. Thus elementary antiquotations may be used in markup
143 commands (e.g. 'chapter', 'section', 'text') and formal comments.
145 * System option "document_tags" specifies a default for otherwise
146 untagged commands. This is occasionally useful to control the global
147 visibility of commands via session options (e.g. in ROOT).
149 * Document markup commands ('section', 'text' etc.) are implicitly
150 tagged as "document" and visible by default. This avoids the application
151 of option "document_tags" to these commands.
153 * Isabelle names are mangled into LaTeX macro names to allow the full
154 identifier syntax with underscore, prime, digits. This is relevant for
155 antiquotations in control symbol notation, e.g. \<^const_name> becomes
156 \isactrlconstUNDERSCOREname.
158 * Document antiquotation @{cite} now checks the given Bibtex entries
159 against the Bibtex database files -- only in batch-mode session builds.
161 * Document antiquotation @{session name} checks and prints the given
162 session name verbatim.
164 * Document preparation with skip_proofs option now preserves the content
165 more accurately: only terminal proof steps ('by' etc.) are skipped.
167 * Command-line tool "isabelle document" has been re-implemented in
168 Isabelle/Scala, with simplified arguments and explicit errors from the
169 latex and bibtex process. Minor INCOMPATIBILITY.
174 * Command 'interpret' no longer exposes resulting theorems as literal
175 facts, notably for the \<open>prop\<close> notation or the "fact" proof method. This
176 improves modularity of proofs and scalability of locale interpretation.
177 Rare INCOMPATIBILITY, need to refer to explicitly named facts instead
178 (e.g. use 'find_theorems' or 'try' to figure this out).
180 * Rewrites clauses (keyword 'rewrites') were moved into the locale
181 expression syntax, where they are part of locale instances. In
182 interpretation commands rewrites clauses now need to occur before
183 'for' and 'defines'. Minor INCOMPATIBILITY.
185 * For rewrites clauses, if activating a locale instance fails, fall
186 back to reading the clause first. This helps avoid qualification of
187 locale instances where the qualifier's sole purpose is avoiding
188 duplicate constant declarations.
193 * The inner syntax category "sort" now includes notation "_" for the
194 dummy sort: it is effectively ignored in type-inference.
199 * Abstract bit operations as part of Main: push_bit, push_take,
202 * New, more general, axiomatization of complete_distrib_lattice.
204 "sup x (Inf X) = Inf (sup x ` X)" and "inf x (Sup X) = Sup (inf x ` X)"
206 "Inf (Sup ` A) <= Sup (Inf ` {f ` A | f . (! Y \<in> A . f Y \<in> Y)})"
207 The instantiations of sets and functions as complete_distrib_lattice
208 are moved to Hilbert_Choice.thy because their proofs need the Hilbert
209 choice operator. The dual of this property is also proved in
212 * New syntax for the minimum/maximum of a function over a finite set:
213 MIN x\<in>A. B and even MIN x. B (only useful for finite types), also
216 * Clarifed theorem names:
218 Min.antimono ~> Min.subset_imp
219 Max.antimono ~> Max.subset_imp
221 Minor INCOMPATIBILITY.
223 * A new command parametric_constant for proving parametricity of
224 non-recursive definitions. For constants that are not fully parametric
225 the command will infer conditions on relations (e.g., bi_unique,
226 bi_total, or type class conditions such as "respects 0") sufficient for
227 parametricity. See ~~/src/HOL/ex/Conditional_Parametricity_Examples for
231 - The 'smt_oracle' option is now necessary when using the 'smt' method
232 with a solver other than Z3. INCOMPATIBILITY.
233 - The encoding to first-order logic is now more complete in the
234 presence of higher-order quantifiers. An 'smt_explicit_application'
235 option has been added to control this. INCOMPATIBILITY.
238 - The legacy command 'old_datatype', provided by
239 '~~/src/HOL/Library/Old_Datatype.thy', has been removed. INCOMPATIBILITY.
241 * Theory HOL-Computational_Algebra.Polynomial_Factorial does not provide
242 instances of rat, real, complex as factorial rings etc. Import
243 HOL-Computational_Algebra.Field_as_Ring explicitly in case of need.
246 * Fact mod_mult_self4 (on nat) renamed to mod_mult_self3', to avoid
247 clash with fact mod_mult_self4 (on more generic semirings).
250 * Facts sum.commute(_restrict) and prod.commute(_restrict) renamed to
251 sum.swap(_restrict) and prod.swap(_restrict), to avoid name clashes on
252 interpretation of abstract locales. INCOMPATIBILITY.
254 * Predicate pairwise_coprime abolished, use "pairwise coprime" instead.
257 * The relator rel_filter on filters has been strengthened to its
258 canonical categorical definition with better properties. INCOMPATIBILITY.
260 * Generalized linear algebra involving linear, span, dependent, dim
261 from type class real_vector to locales module and vector_space.
263 - span_inc ~> span_superset
264 - span_superset ~> span_base
265 - span_eq ~> span_eq_iff
269 * HOL-Algebra: renamed (^) to [^]
271 * Session HOL-Analysis: Moebius functions, the Riemann mapping
272 theorem, the Vitali covering theorem, change-of-variables results for
273 integration and measures.
275 * Class linordered_semiring_1 covers zero_less_one also, ruling out
276 pathologic instances. Minor INCOMPATIBILITY.
278 * Removed nat-int transfer machinery. Rare INCOMPATIBILITY.
280 * Predicate coprime is now a real definition, not a mere
281 abbreviation. INCOMPATIBILITY.
283 * Code generation: Code generation takes an explicit option
284 "case_insensitive" to accomodate case-insensitive file systems.
289 * The command-line tools "isabelle server" and "isabelle client" provide
290 access to the Isabelle Server: it supports responsive session management
291 and concurrent use of theories, based on Isabelle/PIDE infrastructure.
292 See also the "system" manual.
294 * The command-line tool "isabelle update_comments" normalizes formal
295 comments in outer syntax as follows: \<comment> \<open>text\<close> (whith a single space to
296 approximate the appearance in document output). This is more specific
297 than former "isabelle update_cartouches -c": the latter tool option has
300 * Session ROOT entry: empty 'document_files' means there is no document
301 for this session. There is no need to specify options [document = false]
304 * The command-line tool "isabelle mkroot" now always produces a document
305 outline: its options have been adapted accordingly. INCOMPATIBILITY.
307 * The command-line tool "isabelle mkroot -I" initializes a Mercurial
308 repository for the generated session files.
310 * Linux and Windows/Cygwin is for x86_64 only. Old 32bit platform
311 support has been discontinued.
313 * Mac OS X 10.10 Yosemite is now the baseline version; Mavericks is no
316 * Java runtime is for x86_64 only. Corresponding Isabelle settings have
317 been renamed to ISABELLE_TOOL_JAVA_OPTIONS and JEDIT_JAVA_OPTIONS,
318 instead of former 32/64 variants. INCOMPATIBILITY.
320 * Command-line tool "isabelle build" supports new options:
321 - option -B NAME: include session NAME and all descendants
322 - option -S: only observe changes of sources, not heap images
323 - option -f: forces a fresh build
325 * Command-line tool "isabelle build" takes "condition" options with the
326 corresponding environment values into account, when determining the
327 up-to-date status of a session.
329 * Command-line tool "isabelle imports -I" also reports actual session
330 imports. This helps to minimize the session dependency graph.
332 * Update to current Poly/ML 5.7.1 with slightly improved performance and
333 PIDE markup for identifier bindings.
335 * ISABELLE_LATEX and ISABELLE_PDFLATEX now include platform-specific
336 options for improved error reporting. Potential INCOMPATIBILITY with
337 unusual LaTeX installations, may have to adapt these settings.
339 * The bundled Poly/ML 5.7.1 now uses The GNU Multiple Precision
340 Arithmetic Library (libgmp) on all platforms, notably Mac OS X with
345 New in Isabelle2017 (October 2017)
346 ----------------------------------
350 * Experimental support for Visual Studio Code (VSCode) as alternative
351 Isabelle/PIDE front-end, see also
352 https://marketplace.visualstudio.com/items?itemName=makarius.Isabelle2017
354 VSCode is a new type of application that continues the concepts of
355 "programmer's editor" and "integrated development environment" towards
356 fully semantic editing and debugging -- in a relatively light-weight
357 manner. Thus it fits nicely on top of the Isabelle/PIDE infrastructure.
358 Technically, VSCode is based on the Electron application framework
359 (Node.js + Chromium browser + V8), which is implemented in JavaScript
360 and TypeScript, while Isabelle/VSCode mainly consists of Isabelle/Scala
361 modules around a Language Server implementation.
363 * Theory names are qualified by the session name that they belong to.
364 This affects imports, but not the theory name space prefix (which is
365 just the theory base name as before).
367 In order to import theories from other sessions, the ROOT file format
368 provides a new 'sessions' keyword. In contrast, a theory that is
369 imported in the old-fashioned manner via an explicit file-system path
370 belongs to the current session, and might cause theory name conflicts
371 later on. Theories that are imported from other sessions are excluded
372 from the current session document. The command-line tool "isabelle
373 imports" helps to update theory imports.
375 * The main theory entry points for some non-HOL sessions have changed,
376 to avoid confusion with the global name "Main" of the session HOL. This
377 leads to the follow renamings:
379 CTT/Main.thy ~> CTT/CTT.thy
380 ZF/Main.thy ~> ZF/ZF.thy
381 ZF/Main_ZF.thy ~> ZF/ZF.thy
382 ZF/Main_ZFC.thy ~> ZF/ZFC.thy
383 ZF/ZF.thy ~> ZF/ZF_Base.thy
387 * Commands 'alias' and 'type_alias' introduce aliases for constants and
388 type constructors, respectively. This allows adhoc changes to name-space
389 accesses within global or local theory contexts, e.g. within a 'bundle'.
391 * Document antiquotations @{prf} and @{full_prf} output proof terms
392 (again) in the same way as commands 'prf' and 'full_prf'.
394 * Computations generated by the code generator can be embedded directly
395 into ML, alongside with @{code} antiquotations, using the following
398 @{computation ... terms: ... datatypes: ...} :
399 ((term -> term) -> 'ml option -> 'a) -> Proof.context -> term -> 'a
400 @{computation_conv ... terms: ... datatypes: ...} :
401 (Proof.context -> 'ml -> conv) -> Proof.context -> conv
402 @{computation_check terms: ... datatypes: ...} : Proof.context -> conv
404 See src/HOL/ex/Computations.thy,
405 src/HOL/Decision_Procs/Commutative_Ring.thy and
406 src/HOL/Decision_Procs/Reflective_Field.thy for examples and the
407 tutorial on code generation.
410 *** Prover IDE -- Isabelle/Scala/jEdit ***
412 * Session-qualified theory imports allow the Prover IDE to process
413 arbitrary theory hierarchies independently of the underlying logic
414 session image (e.g. option "isabelle jedit -l"), but the directory
415 structure needs to be known in advance (e.g. option "isabelle jedit -d"
416 or a line in the file $ISABELLE_HOME_USER/ROOTS).
418 * The PIDE document model maintains file content independently of the
419 status of jEdit editor buffers. Reloading jEdit buffers no longer causes
420 changes of formal document content. Theory dependencies are always
421 resolved internally, without the need for corresponding editor buffers.
422 The system option "jedit_auto_load" has been discontinued: it is
423 effectively always enabled.
425 * The Theories dockable provides a "Purge" button, in order to restrict
426 the document model to theories that are required for open editor
429 * The Theories dockable indicates the overall status of checking of each
430 entry. When all forked tasks of a theory are finished, the border is
431 painted with thick lines; remaining errors in this situation are
432 represented by a different border color.
434 * Automatic indentation is more careful to avoid redundant spaces in
435 intermediate situations. Keywords are indented after input (via typed
436 characters or completion); see also option "jedit_indent_input".
438 * Action "isabelle.preview" opens an HTML preview of the current theory
439 document in the default web browser.
441 * Command-line invocation "isabelle jedit -R -l LOGIC" opens the ROOT
442 entry of the specified logic session in the editor, while its parent is
443 used for formal checking.
445 * The main Isabelle/jEdit plugin may be restarted manually (using the
446 jEdit Plugin Manager), as long as the "Isabelle Base" plugin remains
447 enabled at all times.
449 * Update to current jedit-5.4.0.
454 * Deleting the last code equations for a particular function using
455 [code del] results in function with no equations (runtime abort) rather
456 than an unimplemented function (generation time abort). Use explicit
457 [[code drop:]] to enforce the latter. Minor INCOMPATIBILITY.
459 * Proper concept of code declarations in code.ML:
460 - Regular code declarations act only on the global theory level, being
461 ignored with warnings if syntactically malformed.
462 - Explicitly global code declarations yield errors if syntactically
464 - Default code declarations are silently ignored if syntactically
466 Minor INCOMPATIBILITY.
468 * Clarified and standardized internal data bookkeeping of code
469 declarations: history of serials allows to track potentially
470 non-monotonous declarations appropriately. Minor INCOMPATIBILITY.
475 * The Nunchaku model finder is now part of "Main".
478 - A new option, 'smt_nat_as_int', has been added to translate 'nat' to
479 'int' and benefit from the SMT solver's theory reasoning. It is
481 - The legacy module "src/HOL/Library/Old_SMT.thy" has been removed.
482 - Several small issues have been rectified in the 'smt' command.
484 * (Co)datatype package: The 'size_gen_o_map' lemma is no longer
485 generated for datatypes with type class annotations. As a result, the
486 tactic that derives it no longer fails on nested datatypes. Slight
489 * Command and antiquotation "value" with modified default strategy:
490 terms without free variables are always evaluated using plain evaluation
491 only, with no fallback on normalization by evaluation. Minor
494 * Theories "GCD" and "Binomial" are already included in "Main" (instead
497 * Constant "surj" is a full input/output abbreviation (again).
498 Minor INCOMPATIBILITY.
500 * Dropped aliasses RangeP, DomainP for Rangep, Domainp respectively.
503 * Renamed ii to imaginary_unit in order to free up ii as a variable
504 name. The syntax \<i> remains available. INCOMPATIBILITY.
506 * Dropped abbreviations transP, antisymP, single_valuedP; use constants
507 transp, antisymp, single_valuedp instead. INCOMPATIBILITY.
509 * Constant "subseq" in Topological_Spaces has been removed -- it is
510 subsumed by "strict_mono". Some basic lemmas specific to "subseq" have
511 been renamed accordingly, e.g. "subseq_o" -> "strict_mono_o" etc.
513 * Theory List: "sublist" renamed to "nths" in analogy with "nth", and
514 "sublisteq" renamed to "subseq". Minor INCOMPATIBILITY.
516 * Theory List: new generic function "sorted_wrt".
518 * Named theorems mod_simps covers various congruence rules concerning
519 mod, replacing former zmod_simps. INCOMPATIBILITY.
521 * Swapped orientation of congruence rules mod_add_left_eq,
522 mod_add_right_eq, mod_add_eq, mod_mult_left_eq, mod_mult_right_eq,
523 mod_mult_eq, mod_minus_eq, mod_diff_left_eq, mod_diff_right_eq,
524 mod_diff_eq. INCOMPATIBILITY.
526 * Generalized some facts:
529 zminus_zmod ~> mod_minus_eq
530 zdiff_zmod_left ~> mod_diff_left_eq
531 zdiff_zmod_right ~> mod_diff_right_eq
532 zmod_eq_dvd_iff ~> mod_eq_dvd_iff
535 * Algebraic type class hierarchy of euclidean (semi)rings in HOL:
536 euclidean_(semi)ring, euclidean_(semi)ring_cancel,
537 unique_euclidean_(semi)ring; instantiation requires provision of a
540 * Theory "HOL-Number_Theory.Euclidean_Algorithm" has been reworked:
541 - Euclidean induction is available as rule eucl_induct.
542 - Constants Euclidean_Algorithm.gcd, Euclidean_Algorithm.lcm,
543 Euclidean_Algorithm.Gcd and Euclidean_Algorithm.Lcm allow
544 easy instantiation of euclidean (semi)rings as GCD (semi)rings.
545 - Coefficients obtained by extended euclidean algorithm are
546 available as "bezout_coefficients".
549 * Theory "Number_Theory.Totient" introduces basic notions about Euler's
550 totient function previously hidden as solitary example in theory
551 Residues. Definition changed so that "totient 1 = 1" in agreement with
552 the literature. Minor INCOMPATIBILITY.
554 * New styles in theory "HOL-Library.LaTeXsugar":
555 - "dummy_pats" for printing equations with "_" on the lhs;
556 - "eta_expand" for printing eta-expanded terms.
558 * Theory "HOL-Library.Permutations": theorem bij_swap_ompose_bij has
559 been renamed to bij_swap_compose_bij. INCOMPATIBILITY.
561 * New theory "HOL-Library.Going_To_Filter" providing the "f going_to F"
562 filter for describing points x such that f(x) is in the filter F.
564 * Theory "HOL-Library.Formal_Power_Series": constants X/E/L/F have been
565 renamed to fps_X/fps_exp/fps_ln/fps_hypergeo to avoid polluting the name
566 space. INCOMPATIBILITY.
568 * Theory "HOL-Library.FinFun" has been moved to AFP (again).
571 * Theory "HOL-Library.FuncSet": some old and rarely used ASCII
572 replacement syntax has been removed. INCOMPATIBILITY, standard syntax
573 with symbols should be used instead. The subsequent commands help to
574 reproduce the old forms, e.g. to simplify porting old theories:
577 "_PiE" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(3PIE _:_./ _)" 10)
578 "_Pi" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(3PI _:_./ _)" 10)
579 "_lam" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("(3%_:_./ _)" [0,0,3] 3)
581 * Theory "HOL-Library.Multiset": the simprocs on subsets operators of
582 multisets have been renamed:
584 msetless_cancel_numerals ~> msetsubset_cancel
585 msetle_cancel_numerals ~> msetsubset_eq_cancel
589 * Theory "HOL-Library.Pattern_Aliases" provides input and output syntax
590 for pattern aliases as known from Haskell, Scala and ML.
592 * Theory "HOL-Library.Uprod" formalizes the type of unordered pairs.
594 * Session HOL-Analysis: more material involving arcs, paths, covering
595 spaces, innessential maps, retracts, infinite products, simplicial
596 complexes. Baire Category theorem. Major results include the Jordan
597 Curve Theorem and the Great Picard Theorem.
599 * Session HOL-Algebra has been extended by additional lattice theory:
600 the Knaster-Tarski fixed point theorem and Galois Connections.
602 * Sessions HOL-Computational_Algebra and HOL-Number_Theory: new notions
603 of squarefreeness, n-th powers, and prime powers.
605 * Session "HOL-Computional_Algebra" covers many previously scattered
606 theories, notably Euclidean_Algorithm, Factorial_Ring,
607 Formal_Power_Series, Fraction_Field, Fundamental_Theorem_Algebra,
608 Normalized_Fraction, Polynomial_FPS, Polynomial, Primes. Minor
614 * Isabelle/Scala: the SQL module supports access to relational
615 databases, either as plain file (SQLite) or full-scale server
616 (PostgreSQL via local port or remote ssh connection).
618 * Results of "isabelle build" are recorded as SQLite database (i.e.
619 "Application File Format" in the sense of
620 https://www.sqlite.org/appfileformat.html). This allows systematic
621 access via operations from module Sessions.Store in Isabelle/Scala.
623 * System option "parallel_proofs" is 1 by default (instead of more
624 aggressive 2). This requires less heap space and avoids burning parallel
625 CPU cycles, while full subproof parallelization is enabled for repeated
626 builds (according to parallel_subproofs_threshold).
628 * System option "record_proofs" allows to change the global
629 Proofterm.proofs variable for a session. Regular values are are 0, 1, 2;
630 a negative value means the current state in the ML heap image remains
633 * Isabelle settings variable ISABELLE_SCALA_BUILD_OPTIONS has been
634 renamed to ISABELLE_SCALAC_OPTIONS. Rare INCOMPATIBILITY.
636 * Isabelle settings variables ISABELLE_WINDOWS_PLATFORM,
637 ISABELLE_WINDOWS_PLATFORM32, ISABELLE_WINDOWS_PLATFORM64 indicate the
638 native Windows platform (independently of the Cygwin installation). This
639 is analogous to ISABELLE_PLATFORM, ISABELLE_PLATFORM32,
642 * Command-line tool "isabelle build_docker" builds a Docker image from
643 the Isabelle application bundle for Linux. See also
644 https://hub.docker.com/r/makarius/isabelle
646 * Command-line tool "isabelle vscode_server" provides a Language Server
647 Protocol implementation, e.g. for the Visual Studio Code editor. It
648 serves as example for alternative PIDE front-ends.
650 * Command-line tool "isabelle imports" helps to maintain theory imports
651 wrt. session structure. Examples for the main Isabelle distribution:
653 isabelle imports -I -a
654 isabelle imports -U -a
655 isabelle imports -U -i -a
656 isabelle imports -M -a -d '~~/src/Benchmarks'
660 New in Isabelle2016-1 (December 2016)
661 -------------------------------------
665 * Splitter in proof methods "simp", "auto" and friends:
666 - The syntax "split add" has been discontinued, use plain "split",
668 - For situations with many conditional or case expressions, there is
669 an alternative splitting strategy that can be much faster. It is
670 selected by writing "split!" instead of "split". It applies safe
671 introduction and elimination rules after each split rule. As a
672 result the subgoal may be split into several subgoals.
674 * Command 'bundle' provides a local theory target to define a bundle
675 from the body of specification commands (such as 'declare',
676 'declaration', 'notation', 'lemmas', 'lemma'). For example:
684 * Command 'unbundle' is like 'include', but works within a local theory
685 context. Unlike "context includes ... begin", the effect of 'unbundle'
686 on the target context persists, until different declarations are given.
688 * Simplified outer syntax: uniform category "name" includes long
689 identifiers. Former "xname" / "nameref" / "name reference" has been
692 * Embedded content (e.g. the inner syntax of types, terms, props) may be
693 delimited uniformly via cartouches. This works better than old-fashioned
694 quotes when sub-languages are nested.
696 * Mixfix annotations support general block properties, with syntax
697 "(\<open>x=a y=b z \<dots>\<close>". Notable property names are "indent", "consistent",
698 "unbreakable", "markup". The existing notation "(DIGITS" is equivalent
699 to "(\<open>indent=DIGITS\<close>". The former notation "(00" for unbreakable blocks
700 is superseded by "(\<open>unbreabable\<close>" --- rare INCOMPATIBILITY.
702 * Proof method "blast" is more robust wrt. corner cases of Pure
703 statements without object-logic judgment.
705 * Commands 'prf' and 'full_prf' are somewhat more informative (again):
706 proof terms are reconstructed and cleaned from administrative thm nodes.
708 * Code generator: config option "code_timing" triggers measurements of
709 different phases of code generation. See src/HOL/ex/Code_Timing.thy for
712 * Code generator: implicits in Scala (stemming from type class
713 instances) are generated into companion object of corresponding type
714 class, to resolve some situations where ambiguities may occur.
716 * Solve direct: option "solve_direct_strict_warnings" gives explicit
717 warnings for lemma statements with trivial proofs.
720 *** Prover IDE -- Isabelle/Scala/jEdit ***
722 * More aggressive flushing of machine-generated input, according to
723 system option editor_generated_input_delay (in addition to existing
724 editor_input_delay for regular user edits). This may affect overall PIDE
725 reactivity and CPU usage.
727 * Syntactic indentation according to Isabelle outer syntax. Action
728 "indent-lines" (shortcut C+i) indents the current line according to
729 command keywords and some command substructure. Action
730 "isabelle.newline" (shortcut ENTER) indents the old and the new line
731 according to command keywords only; see also option
732 "jedit_indent_newline".
734 * Semantic indentation for unstructured proof scripts ('apply' etc.) via
735 number of subgoals. This requires information of ongoing document
736 processing and may thus lag behind, when the user is editing too
737 quickly; see also option "jedit_script_indent" and
738 "jedit_script_indent_limit".
740 * Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed'
741 are treated as delimiters for fold structure; 'begin' and 'end'
742 structure of theory specifications is treated as well.
744 * Command 'proof' provides information about proof outline with cases,
745 e.g. for proof methods "cases", "induct", "goal_cases".
747 * Completion templates for commands involving "begin ... end" blocks,
748 e.g. 'context', 'notepad'.
750 * Sidekick parser "isabelle-context" shows nesting of context blocks
751 according to 'begin' and 'end' structure.
753 * Highlighting of entity def/ref positions wrt. cursor.
755 * Action "isabelle.select-entity" (shortcut CS+ENTER) selects all
756 occurrences of the formal entity at the caret position. This facilitates
759 * PIDE document markup works across multiple Isar commands, e.g. the
760 results established at the end of a proof are properly identified in the
763 * Cartouche abbreviations work both for " and ` to accomodate typical
764 situations where old ASCII notation may be updated.
766 * Dockable window "Symbols" also provides access to 'abbrevs' from the
767 outer syntax of the current theory buffer. This provides clickable
768 syntax templates, including entries with empty abbrevs name (which are
769 inaccessible via keyboard completion).
771 * IDE support for the Isabelle/Pure bootstrap process, with the
772 following independent stages:
777 src/Pure/ML_Bootstrap.thy
779 The ML ROOT files act like quasi-theories in the context of theory
780 ML_Bootstrap: this allows continuous checking of all loaded ML files.
781 The theory files are presented with a modified header to import Pure
782 from the running Isabelle instance. Results from changed versions of
783 each stage are *not* propagated to the next stage, and isolated from the
784 actual Isabelle/Pure that runs the IDE itself. The sequential
785 dependencies of the above files are only observed for batch build.
787 * Isabelle/ML and Standard ML files are presented in Sidekick with the
788 tree structure of section headings: this special comment format is
789 described in "implementation" chapter 0, e.g. (*** section ***).
791 * Additional abbreviations for syntactic completion may be specified
792 within the theory header as 'abbrevs'. The theory syntax for 'keywords'
793 has been simplified accordingly: optional abbrevs need to go into the
794 new 'abbrevs' section.
796 * Global abbreviations via $ISABELLE_HOME/etc/abbrevs and
797 $ISABELLE_HOME_USER/etc/abbrevs are no longer supported. Minor
798 INCOMPATIBILITY, use 'abbrevs' within theory header instead.
800 * Action "isabelle.keymap-merge" asks the user to resolve pending
801 Isabelle keymap changes that are in conflict with the current jEdit
802 keymap; non-conflicting changes are always applied implicitly. This
803 action is automatically invoked on Isabelle/jEdit startup and thus
804 increases chances that users see new keyboard shortcuts when re-using
807 * ML and document antiquotations for file-systems paths are more uniform
810 @{path NAME} -- no file-system check
811 @{file NAME} -- check for plain file
812 @{dir NAME} -- check for directory
814 Minor INCOMPATIBILITY, former uses of @{file} and @{file_unchecked} may
818 *** Document preparation ***
820 * New symbol \<circle>, e.g. for temporal operator.
822 * New document and ML antiquotation @{locale} for locales, similar to
823 existing antiquotation @{class}.
825 * Mixfix annotations support delimiters like \<^control>\<open>cartouche\<close> --
826 this allows special forms of document output.
828 * Raw LaTeX output now works via \<^latex>\<open>...\<close> instead of raw control
829 symbol \<^raw:...>. INCOMPATIBILITY, notably for LaTeXsugar.thy and its
832 * \<^raw:...> symbols are no longer supported.
834 * Old 'header' command is no longer supported (legacy since
840 * Many specification elements support structured statements with 'if' /
841 'for' eigen-context, e.g. 'axiomatization', 'abbreviation',
842 'definition', 'inductive', 'function'.
844 * Toplevel theorem statements support eigen-context notation with 'if' /
845 'for' (in postfix), which corresponds to 'assumes' / 'fixes' in the
846 traditional long statement form (in prefix). Local premises are called
847 "that" or "assms", respectively. Empty premises are *not* bound in the
848 context: INCOMPATIBILITY.
850 * Command 'define' introduces a local (non-polymorphic) definition, with
851 optional abstraction over local parameters. The syntax resembles
852 'definition' and 'obtain'. It fits better into the Isar language than
853 old 'def', which is now a legacy feature.
855 * Command 'obtain' supports structured statements with 'if' / 'for'
858 * Command '\<proof>' is an alias for 'sorry', with different
859 typesetting. E.g. to produce proof holes in examples and documentation.
861 * The defining position of a literal fact \<open>prop\<close> is maintained more
862 carefully, and made accessible as hyperlink in the Prover IDE.
864 * Commands 'finally' and 'ultimately' used to expose the result as
865 literal fact: this accidental behaviour has been discontinued. Rare
866 INCOMPATIBILITY, use more explicit means to refer to facts in Isar.
868 * Command 'axiomatization' has become more restrictive to correspond
869 better to internal axioms as singleton facts with mandatory name. Minor
872 * Proof methods may refer to the main facts via the dynamic fact
873 "method_facts". This is particularly useful for Eisbach method
876 * Proof method "use" allows to modify the main facts of a given method
880 (use facts in \<open>simp add: ...\<close>)
882 * The old proof method "default" has been removed (legacy since
883 Isabelle2016). INCOMPATIBILITY, use "standard" instead.
888 * Pure provides basic versions of proof methods "simp" and "simp_all"
889 that only know about meta-equality (==). Potential INCOMPATIBILITY in
890 theory imports that merge Pure with e.g. Main of Isabelle/HOL: the order
891 is relevant to avoid confusion of Pure.simp vs. HOL.simp.
893 * The command 'unfolding' and proof method "unfold" include a second
894 stage where given equations are passed through the attribute "abs_def"
895 before rewriting. This ensures that definitions are fully expanded,
896 regardless of the actual parameters that are provided. Rare
897 INCOMPATIBILITY in some corner cases: use proof method (simp only:)
898 instead, or declare [[unfold_abs_def = false]] in the proof context.
900 * Type-inference improves sorts of newly introduced type variables for
901 the object-logic, using its base sort (i.e. HOL.type for Isabelle/HOL).
902 Thus terms like "f x" or "\<And>x. P x" without any further syntactic context
903 produce x::'a::type in HOL instead of x::'a::{} in Pure. Rare
904 INCOMPATIBILITY, need to provide explicit type constraints for Pure
905 types where this is really intended.
910 * New proof method "argo" using the built-in Argo solver based on SMT
911 technology. The method can be used to prove goals of quantifier-free
912 propositional logic, goals based on a combination of quantifier-free
913 propositional logic with equality, and goals based on a combination of
914 quantifier-free propositional logic with linear real arithmetic
915 including min/max/abs. See HOL/ex/Argo_Examples.thy for examples.
917 * The new "nunchaku" command integrates the Nunchaku model finder. The
918 tool is experimental. See ~~/src/HOL/Nunchaku/Nunchaku.thy for details.
920 * Metis: The problem encoding has changed very slightly. This might
921 break existing proofs. INCOMPATIBILITY.
924 - The MaSh relevance filter is now faster than before.
925 - Produce syntactically correct Vampire 4.0 problem files.
927 * (Co)datatype package:
928 - New commands for defining corecursive functions and reasoning about
929 them in "~~/src/HOL/Library/BNF_Corec.thy": 'corec', 'corecursive',
930 'friend_of_corec', and 'corecursion_upto'; and 'corec_unique' proof
931 method. See 'isabelle doc corec'.
932 - The predicator :: ('a \<Rightarrow> bool) \<Rightarrow> 'a F \<Rightarrow> bool is now a first-class
933 citizen in bounded natural functors.
934 - 'primrec' now allows nested calls through the predicator in addition
936 - 'bnf' automatically discharges reflexive proof obligations.
937 - 'bnf' outputs a slightly modified proof obligation expressing rel in
939 (not giving a specification for rel makes this one reflexive).
940 - 'bnf' outputs a new proof obligation expressing pred in terms of set
941 (not giving a specification for pred makes this one reflexive).
942 INCOMPATIBILITY: manual 'bnf' declarations may need adjustment.
944 rel_prod_apply ~> rel_prod_inject
945 pred_prod_apply ~> pred_prod_inject
947 - The "size" plugin has been made compatible again with locales.
948 - The theorems about "rel" and "set" may have a slightly different (but
952 * The 'coinductive' command produces a proper coinduction rule for
953 mutual coinductive predicates. This new rule replaces the old rule,
954 which exposed details of the internal fixpoint construction and was
955 hard to use. INCOMPATIBILITY.
957 * New abbreviations for negated existence (but not bounded existence):
959 \<nexists>x. P x \<equiv> \<not> (\<exists>x. P x)
960 \<nexists>!x. P x \<equiv> \<not> (\<exists>!x. P x)
962 * The print mode "HOL" for ASCII syntax of binders "!", "?", "?!", "@"
963 has been removed for output. It is retained for input only, until it is
964 eliminated altogether.
966 * The unique existence quantifier no longer provides 'binder' syntax,
967 but uses syntax translations (as for bounded unique existence). Thus
968 iterated quantification \<exists>!x y. P x y with its slightly confusing
969 sequential meaning \<exists>!x. \<exists>!y. P x y is no longer possible. Instead,
970 pattern abstraction admits simultaneous unique existence \<exists>!(x, y). P x y
971 (analogous to existing notation \<exists>!(x, y)\<in>A. P x y). Potential
972 INCOMPATIBILITY in rare situations.
974 * Conventional syntax "%(). t" for unit abstractions. Slight syntactic
977 * Renamed constants and corresponding theorems:
982 listprod ~> prod_list
986 * Sligthly more standardized theorem names:
987 sgn_times ~> sgn_mult
988 sgn_mult' ~> Real_Vector_Spaces.sgn_mult
989 divide_zero_left ~> div_0
990 zero_mod_left ~> mod_0
991 divide_zero ~> div_by_0
993 nonzero_mult_divide_cancel_left ~> nonzero_mult_div_cancel_left
994 div_mult_self1_is_id ~> nonzero_mult_div_cancel_left
995 nonzero_mult_divide_cancel_right ~> nonzero_mult_div_cancel_right
996 div_mult_self2_is_id ~> nonzero_mult_div_cancel_right
997 is_unit_divide_mult_cancel_left ~> is_unit_div_mult_cancel_left
998 is_unit_divide_mult_cancel_right ~> is_unit_div_mult_cancel_right
999 mod_div_equality ~> div_mult_mod_eq
1000 mod_div_equality2 ~> mult_div_mod_eq
1001 mod_div_equality3 ~> mod_div_mult_eq
1002 mod_div_equality4 ~> mod_mult_div_eq
1003 minus_div_eq_mod ~> minus_div_mult_eq_mod
1004 minus_div_eq_mod2 ~> minus_mult_div_eq_mod
1005 minus_mod_eq_div ~> minus_mod_eq_div_mult
1006 minus_mod_eq_div2 ~> minus_mod_eq_mult_div
1007 div_mod_equality' ~> minus_mod_eq_div_mult [symmetric]
1008 mod_div_equality' ~> minus_div_mult_eq_mod [symmetric]
1009 zmod_zdiv_equality ~> mult_div_mod_eq [symmetric]
1010 zmod_zdiv_equality' ~> minus_div_mult_eq_mod [symmetric]
1011 Divides.mult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
1012 mult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
1013 zmult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
1014 div_1 ~> div_by_Suc_0
1015 mod_1 ~> mod_by_Suc_0
1018 * New type class "idom_abs_sgn" specifies algebraic properties
1019 of sign and absolute value functions. Type class "sgn_if" has
1020 disappeared. Slight INCOMPATIBILITY.
1022 * Dedicated syntax LENGTH('a) for length of types.
1024 * Characters (type char) are modelled as finite algebraic type
1025 corresponding to {0..255}.
1027 - Logical representation:
1028 * 0 is instantiated to the ASCII zero character.
1029 * All other characters are represented as "Char n"
1030 with n being a raw numeral expression less than 256.
1031 * Expressions of the form "Char n" with n greater than 255
1033 - Printing and parsing:
1034 * Printable characters are printed and parsed as "CHR ''\<dots>''"
1036 * The ASCII zero character is printed and parsed as "0".
1037 * All other canonical characters are printed as "CHR 0xXX"
1038 with XX being the hexadecimal character code. "CHR n"
1039 is parsable for every numeral expression n.
1040 * Non-canonical characters have no special syntax and are
1041 printed as their logical representation.
1042 - Explicit conversions from and to the natural numbers are
1043 provided as char_of_nat, nat_of_char (as before).
1044 - The auxiliary nibble type has been discontinued.
1048 * Type class "div" with operation "mod" renamed to type class "modulo"
1049 with operation "modulo", analogously to type class "divide". This
1050 eliminates the need to qualify any of those names in the presence of
1051 infix "mod" syntax. INCOMPATIBILITY.
1053 * Statements and proofs of Knaster-Tarski fixpoint combinators lfp/gfp
1054 have been clarified. The fixpoint properties are lfp_fixpoint, its
1055 symmetric lfp_unfold (as before), and the duals for gfp. Auxiliary items
1056 for the proof (lfp_lemma2 etc.) are no longer exported, but can be
1057 easily recovered by composition with eq_refl. Minor INCOMPATIBILITY.
1059 * Constant "surj" is a mere input abbreviation, to avoid hiding an
1060 equation in term output. Minor INCOMPATIBILITY.
1062 * Command 'code_reflect' accepts empty constructor lists for datatypes,
1063 which renders those abstract effectively.
1065 * Command 'export_code' checks given constants for abstraction
1066 violations: a small guarantee that given constants specify a safe
1067 interface for the generated code.
1069 * Code generation for Scala: ambiguous implicts in class diagrams are
1070 spelt out explicitly.
1072 * Static evaluators (Code_Evaluation.static_* in Isabelle/ML) rely on
1073 explicitly provided auxiliary definitions for required type class
1074 dictionaries rather than half-working magic. INCOMPATIBILITY, see the
1075 tutorial on code generation for details.
1077 * Theory Set_Interval: substantial new theorems on indexed sums and
1080 * Locale bijection establishes convenient default simp rules such as
1081 "inv f (f a) = a" for total bijections.
1083 * Abstract locales semigroup, abel_semigroup, semilattice,
1084 semilattice_neutr, ordering, ordering_top, semilattice_order,
1085 semilattice_neutr_order, comm_monoid_set, semilattice_set,
1086 semilattice_neutr_set, semilattice_order_set,
1087 semilattice_order_neutr_set monoid_list, comm_monoid_list,
1088 comm_monoid_list_set, comm_monoid_mset, comm_monoid_fun use boldified
1089 syntax uniformly that does not clash with corresponding global syntax.
1092 * Former locale lifting_syntax is now a bundle, which is easier to
1093 include in a local context or theorem statement, e.g. "context includes
1094 lifting_syntax begin ... end". Minor INCOMPATIBILITY.
1096 * Some old / obsolete theorems have been renamed / removed, potential
1099 nat_less_cases -- removed, use linorder_cases instead
1100 inv_image_comp -- removed, use image_inv_f_f instead
1101 image_surj_f_inv_f ~> image_f_inv_f
1103 * Some theorems about groups and orders have been generalised from
1104 groups to semi-groups that are also monoids:
1107 less_add_same_cancel1
1108 less_add_same_cancel2
1111 add_less_same_cancel1
1112 add_less_same_cancel2
1114 * Some simplifications theorems about rings have been removed, since
1115 superseeded by a more general version:
1116 less_add_cancel_left_greater_zero ~> less_add_same_cancel1
1117 less_add_cancel_right_greater_zero ~> less_add_same_cancel2
1118 less_eq_add_cancel_left_greater_eq_zero ~> le_add_same_cancel1
1119 less_eq_add_cancel_right_greater_eq_zero ~> le_add_same_cancel2
1120 less_eq_add_cancel_left_less_eq_zero ~> add_le_same_cancel1
1121 less_eq_add_cancel_right_less_eq_zero ~> add_le_same_cancel2
1122 less_add_cancel_left_less_zero ~> add_less_same_cancel1
1123 less_add_cancel_right_less_zero ~> add_less_same_cancel2
1126 * Renamed split_if -> if_split and split_if_asm -> if_split_asm to
1127 resemble the f.split naming convention, INCOMPATIBILITY.
1129 * Added class topological_monoid.
1131 * The following theorems have been renamed:
1133 setsum_left_distrib ~> sum_distrib_right
1134 setsum_right_distrib ~> sum_distrib_left
1138 * Compound constants INFIMUM and SUPREMUM are mere abbreviations now.
1141 * "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional
1142 comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f `
1145 * Class semiring_Lcd merged into semiring_Gcd. INCOMPATIBILITY.
1147 * The type class ordered_comm_monoid_add is now called
1148 ordered_cancel_comm_monoid_add. A new type class ordered_comm_monoid_add
1149 is introduced as the combination of ordered_ab_semigroup_add +
1150 comm_monoid_add. INCOMPATIBILITY.
1152 * Introduced the type classes canonically_ordered_comm_monoid_add and
1155 * Introduced the type class ordered_ab_semigroup_monoid_add_imp_le. When
1156 instantiating linordered_semiring_strict and ordered_ab_group_add, an
1157 explicit instantiation of ordered_ab_semigroup_monoid_add_imp_le might
1158 be required. INCOMPATIBILITY.
1160 * Dropped various legacy fact bindings, whose replacements are often
1161 of a more general type also:
1162 lcm_left_commute_nat ~> lcm.left_commute
1163 lcm_left_commute_int ~> lcm.left_commute
1164 gcd_left_commute_nat ~> gcd.left_commute
1165 gcd_left_commute_int ~> gcd.left_commute
1166 gcd_greatest_iff_nat ~> gcd_greatest_iff
1167 gcd_greatest_iff_int ~> gcd_greatest_iff
1168 coprime_dvd_mult_nat ~> coprime_dvd_mult
1169 coprime_dvd_mult_int ~> coprime_dvd_mult
1170 zpower_numeral_even ~> power_numeral_even
1171 gcd_mult_cancel_nat ~> gcd_mult_cancel
1172 gcd_mult_cancel_int ~> gcd_mult_cancel
1173 div_gcd_coprime_nat ~> div_gcd_coprime
1174 div_gcd_coprime_int ~> div_gcd_coprime
1175 zpower_numeral_odd ~> power_numeral_odd
1176 zero_less_int_conv ~> of_nat_0_less_iff
1177 gcd_greatest_nat ~> gcd_greatest
1178 gcd_greatest_int ~> gcd_greatest
1179 coprime_mult_nat ~> coprime_mult
1180 coprime_mult_int ~> coprime_mult
1181 lcm_commute_nat ~> lcm.commute
1182 lcm_commute_int ~> lcm.commute
1183 int_less_0_conv ~> of_nat_less_0_iff
1184 gcd_commute_nat ~> gcd.commute
1185 gcd_commute_int ~> gcd.commute
1186 Gcd_insert_nat ~> Gcd_insert
1187 Gcd_insert_int ~> Gcd_insert
1188 of_int_int_eq ~> of_int_of_nat_eq
1189 lcm_least_nat ~> lcm_least
1190 lcm_least_int ~> lcm_least
1191 lcm_assoc_nat ~> lcm.assoc
1192 lcm_assoc_int ~> lcm.assoc
1193 int_le_0_conv ~> of_nat_le_0_iff
1194 int_eq_0_conv ~> of_nat_eq_0_iff
1195 Gcd_empty_nat ~> Gcd_empty
1196 Gcd_empty_int ~> Gcd_empty
1197 gcd_assoc_nat ~> gcd.assoc
1198 gcd_assoc_int ~> gcd.assoc
1199 zero_zle_int ~> of_nat_0_le_iff
1200 lcm_dvd2_nat ~> dvd_lcm2
1201 lcm_dvd2_int ~> dvd_lcm2
1202 lcm_dvd1_nat ~> dvd_lcm1
1203 lcm_dvd1_int ~> dvd_lcm1
1204 gcd_zero_nat ~> gcd_eq_0_iff
1205 gcd_zero_int ~> gcd_eq_0_iff
1206 gcd_dvd2_nat ~> gcd_dvd2
1207 gcd_dvd2_int ~> gcd_dvd2
1208 gcd_dvd1_nat ~> gcd_dvd1
1209 gcd_dvd1_int ~> gcd_dvd1
1210 int_numeral ~> of_nat_numeral
1211 lcm_ac_nat ~> ac_simps
1212 lcm_ac_int ~> ac_simps
1213 gcd_ac_nat ~> ac_simps
1214 gcd_ac_int ~> ac_simps
1215 abs_int_eq ~> abs_of_nat
1216 zless_int ~> of_nat_less_iff
1217 zdiff_int ~> of_nat_diff
1218 zadd_int ~> of_nat_add
1219 int_mult ~> of_nat_mult
1220 int_Suc ~> of_nat_Suc
1221 inj_int ~> inj_of_nat
1224 Lcm_empty_nat ~> Lcm_empty
1225 Lcm_empty_int ~> Lcm_empty
1226 Lcm_insert_nat ~> Lcm_insert
1227 Lcm_insert_int ~> Lcm_insert
1228 comp_fun_idem_gcd_nat ~> comp_fun_idem_gcd
1229 comp_fun_idem_gcd_int ~> comp_fun_idem_gcd
1230 comp_fun_idem_lcm_nat ~> comp_fun_idem_lcm
1231 comp_fun_idem_lcm_int ~> comp_fun_idem_lcm
1232 Lcm_eq_0 ~> Lcm_eq_0_I
1233 Lcm0_iff ~> Lcm_0_iff
1234 Lcm_dvd_int ~> Lcm_least
1235 divides_mult_nat ~> divides_mult
1236 divides_mult_int ~> divides_mult
1237 lcm_0_nat ~> lcm_0_right
1238 lcm_0_int ~> lcm_0_right
1239 lcm_0_left_nat ~> lcm_0_left
1240 lcm_0_left_int ~> lcm_0_left
1241 dvd_gcd_D1_nat ~> dvd_gcdD1
1242 dvd_gcd_D1_int ~> dvd_gcdD1
1243 dvd_gcd_D2_nat ~> dvd_gcdD2
1244 dvd_gcd_D2_int ~> dvd_gcdD2
1245 coprime_dvd_mult_iff_nat ~> coprime_dvd_mult_iff
1246 coprime_dvd_mult_iff_int ~> coprime_dvd_mult_iff
1247 realpow_minus_mult ~> power_minus_mult
1248 realpow_Suc_le_self ~> power_Suc_le_self
1249 dvd_Gcd, dvd_Gcd_nat, dvd_Gcd_int removed in favour of Gcd_greatest
1252 * Renamed HOL/Quotient_Examples/FSet.thy to
1253 HOL/Quotient_Examples/Quotient_FSet.thy INCOMPATIBILITY.
1255 * Session HOL-Library: theory FinFun bundles "finfun_syntax" and
1256 "no_finfun_syntax" allow to control optional syntax in local contexts;
1257 this supersedes former theory FinFun_Syntax. INCOMPATIBILITY, e.g. use
1258 "unbundle finfun_syntax" to imitate import of
1259 "~~/src/HOL/Library/FinFun_Syntax".
1261 * Session HOL-Library: theory Multiset_Permutations (executably) defines
1262 the set of permutations of a given set or multiset, i.e. the set of all
1263 lists that contain every element of the carrier (multi-)set exactly
1266 * Session HOL-Library: multiset membership is now expressed using
1267 set_mset rather than count.
1269 - Expressions "count M a > 0" and similar simplify to membership
1272 - Converting between "count M a = 0" and non-membership happens using
1273 equations count_eq_zero_iff and not_in_iff.
1275 - Rules count_inI and in_countE obtain facts of the form
1276 "count M a = n" from membership.
1278 - Rules count_in_diffI and in_diff_countE obtain facts of the form
1279 "count M a = n + count N a" from membership on difference sets.
1283 * Session HOL-Library: theory LaTeXsugar uses new-style "dummy_pats" for
1284 displaying equations in functional programming style --- variables
1285 present on the left-hand but not on the righ-hand side are replaced by
1288 * Session HOL-Library: theory Combinator_PER provides combinator to
1289 build partial equivalence relations from a predicate and an equivalence
1292 * Session HOL-Library: theory Perm provides basic facts about almost
1293 everywhere fix bijections.
1295 * Session HOL-Library: theory Normalized_Fraction allows viewing an
1296 element of a field of fractions as a normalized fraction (i.e. a pair of
1297 numerator and denominator such that the two are coprime and the
1298 denominator is normalized wrt. unit factors).
1300 * Session HOL-NSA has been renamed to HOL-Nonstandard_Analysis.
1302 * Session HOL-Multivariate_Analysis has been renamed to HOL-Analysis.
1304 * Session HOL-Analysis: measure theory has been moved here from
1305 HOL-Probability. When importing HOL-Analysis some theorems need
1306 additional name spaces prefixes due to name clashes. INCOMPATIBILITY.
1308 * Session HOL-Analysis: more complex analysis including Cauchy's
1309 inequality, Liouville theorem, open mapping theorem, maximum modulus
1310 principle, Residue theorem, Schwarz Lemma.
1312 * Session HOL-Analysis: Theory of polyhedra: faces, extreme points,
1313 polytopes, and the Krein–Milman Minkowski theorem.
1315 * Session HOL-Analysis: Numerous results ported from the HOL Light
1316 libraries: homeomorphisms, continuous function extensions, invariance of
1319 * Session HOL-Probability: the type of emeasure and nn_integral was
1320 changed from ereal to ennreal, INCOMPATIBILITY.
1322 emeasure :: 'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal
1323 nn_integral :: 'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal
1325 * Session HOL-Probability: Code generation and QuickCheck for
1326 Probability Mass Functions.
1328 * Session HOL-Probability: theory Random_Permutations contains some
1329 theory about choosing a permutation of a set uniformly at random and
1330 folding over a list in random order.
1332 * Session HOL-Probability: theory SPMF formalises discrete
1333 subprobability distributions.
1335 * Session HOL-Library: the names of multiset theorems have been
1336 normalised to distinguish which ordering the theorems are about
1338 mset_less_eqI ~> mset_subset_eqI
1339 mset_less_insertD ~> mset_subset_insertD
1340 mset_less_eq_count ~> mset_subset_eq_count
1341 mset_less_diff_self ~> mset_subset_diff_self
1342 mset_le_exists_conv ~> mset_subset_eq_exists_conv
1343 mset_le_mono_add_right_cancel ~> mset_subset_eq_mono_add_right_cancel
1344 mset_le_mono_add_left_cancel ~> mset_subset_eq_mono_add_left_cancel
1345 mset_le_mono_add ~> mset_subset_eq_mono_add
1346 mset_le_add_left ~> mset_subset_eq_add_left
1347 mset_le_add_right ~> mset_subset_eq_add_right
1348 mset_le_single ~> mset_subset_eq_single
1349 mset_le_multiset_union_diff_commute ~> mset_subset_eq_multiset_union_diff_commute
1350 diff_le_self ~> diff_subset_eq_self
1351 mset_leD ~> mset_subset_eqD
1352 mset_lessD ~> mset_subsetD
1353 mset_le_insertD ~> mset_subset_eq_insertD
1354 mset_less_of_empty ~> mset_subset_of_empty
1355 mset_less_size ~> mset_subset_size
1356 wf_less_mset_rel ~> wf_subset_mset_rel
1357 count_le_replicate_mset_le ~> count_le_replicate_mset_subset_eq
1358 mset_remdups_le ~> mset_remdups_subset_eq
1359 ms_lesseq_impl ~> subset_eq_mset_impl
1361 Some functions have been renamed:
1362 ms_lesseq_impl -> subset_eq_mset_impl
1364 * HOL-Library: multisets are now ordered with the multiset ordering
1365 #\<subseteq># ~> \<le>
1367 le_multiset ~> less_eq_multiset
1368 less_multiset ~> le_multiset
1371 * Session HOL-Library: the prefix multiset_order has been discontinued:
1372 the theorems can be directly accessed. As a consequence, the lemmas
1373 "order_multiset" and "linorder_multiset" have been discontinued, and the
1374 interpretations "multiset_linorder" and "multiset_wellorder" have been
1375 replaced by instantiations. INCOMPATIBILITY.
1377 * Session HOL-Library: some theorems about the multiset ordering have
1380 le_multiset_def ~> less_eq_multiset_def
1381 less_multiset_def ~> le_multiset_def
1382 less_eq_imp_le_multiset ~> subset_eq_imp_le_multiset
1383 mult_less_not_refl ~> mset_le_not_refl
1384 mult_less_trans ~> mset_le_trans
1385 mult_less_not_sym ~> mset_le_not_sym
1386 mult_less_asym ~> mset_le_asym
1387 mult_less_irrefl ~> mset_le_irrefl
1388 union_less_mono2{,1,2} ~> union_le_mono2{,1,2}
1390 le_multiset\<^sub>H\<^sub>O ~> less_eq_multiset\<^sub>H\<^sub>O
1391 le_multiset_total ~> less_eq_multiset_total
1392 less_multiset_right_total ~> subset_eq_imp_le_multiset
1393 le_multiset_empty_left ~> less_eq_multiset_empty_left
1394 le_multiset_empty_right ~> less_eq_multiset_empty_right
1395 less_multiset_empty_right ~> le_multiset_empty_left
1396 less_multiset_empty_left ~> le_multiset_empty_right
1397 union_less_diff_plus ~> union_le_diff_plus
1398 ex_gt_count_imp_less_multiset ~> ex_gt_count_imp_le_multiset
1399 less_multiset_plus_left_nonempty ~> le_multiset_plus_left_nonempty
1400 le_multiset_plus_right_nonempty ~> le_multiset_plus_right_nonempty
1403 * Session HOL-Library: the lemma mset_map has now the attribute [simp].
1406 * Session HOL-Library: some theorems about multisets have been removed.
1407 INCOMPATIBILITY, use the following replacements:
1409 le_multiset_plus_plus_left_iff ~> add_less_cancel_right
1410 less_multiset_plus_plus_left_iff ~> add_less_cancel_right
1411 le_multiset_plus_plus_right_iff ~> add_less_cancel_left
1412 less_multiset_plus_plus_right_iff ~> add_less_cancel_left
1413 add_eq_self_empty_iff ~> add_cancel_left_right
1414 mset_subset_add_bothsides ~> subset_mset.add_less_cancel_right
1415 mset_less_add_bothsides ~> subset_mset.add_less_cancel_right
1416 mset_le_add_bothsides ~> subset_mset.add_less_cancel_right
1417 empty_inter ~> subset_mset.inf_bot_left
1418 inter_empty ~> subset_mset.inf_bot_right
1419 empty_sup ~> subset_mset.sup_bot_left
1420 sup_empty ~> subset_mset.sup_bot_right
1421 bdd_below_multiset ~> subset_mset.bdd_above_bot
1422 subset_eq_empty ~> subset_mset.le_zero_eq
1423 le_empty ~> subset_mset.le_zero_eq
1424 mset_subset_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero
1425 mset_less_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero
1427 * Session HOL-Library: some typeclass constraints about multisets have
1428 been reduced from ordered or linordered to preorder. Multisets have the
1429 additional typeclasses order_bot, no_top,
1430 ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add,
1431 linordered_cancel_ab_semigroup_add, and
1432 ordered_ab_semigroup_monoid_add_imp_le. INCOMPATIBILITY.
1434 * Session HOL-Library: there are some new simplification rules about
1435 multisets, the multiset ordering, and the subset ordering on multisets.
1438 * Session HOL-Library: the subset ordering on multisets has now the
1439 interpretations ordered_ab_semigroup_monoid_add_imp_le and
1440 bounded_lattice_bot. INCOMPATIBILITY.
1442 * Session HOL-Library, theory Multiset: single has been removed in favor
1443 of add_mset that roughly corresponds to Set.insert. Some theorems have
1446 single_not_empty ~> add_mset_not_empty or empty_not_add_mset
1447 fold_mset_insert ~> fold_mset_add_mset
1448 image_mset_insert ~> image_mset_add_mset
1449 union_single_eq_diff
1450 multi_self_add_other_not_self
1451 diff_single_eq_union
1454 * Session HOL-Library, theory Multiset: some theorems have been changed
1455 to use add_mset instead of single:
1458 multi_self_add_other_not_self
1459 diff_single_eq_union
1460 union_single_eq_diff
1461 union_single_eq_member
1466 multiset_add_sub_el_shuffle
1467 mset_subset_eq_insertD
1469 insert_subset_eq_iff
1470 insert_union_subset_iff
1471 multi_psub_of_add_self
1480 size_eq_Suc_imp_eq_union
1481 multi_nonempty_split
1486 mset_zip_take_Cons_drop_twice
1492 le_multiset_right_total
1494 multiset_induct2_size
1498 * Session HOL-Library, theory Multiset: the definitions of some
1499 constants have changed to use add_mset instead of adding a single
1512 * Session HOL-Library, theory Multiset: due to the above changes, the
1513 attributes of some multiset theorems have been changed:
1515 insert_DiffM [] ~> [simp]
1516 insert_DiffM2 [simp] ~> []
1517 diff_add_mset_swap [simp]
1518 fold_mset_add_mset [simp]
1519 diff_diff_add [simp] (for multisets only)
1520 diff_cancel [simp] ~> []
1521 count_single [simp] ~> []
1522 set_mset_single [simp] ~> []
1523 size_multiset_single [simp] ~> []
1524 size_single [simp] ~> []
1525 image_mset_single [simp] ~> []
1526 mset_subset_eq_mono_add_right_cancel [simp] ~> []
1527 mset_subset_eq_mono_add_left_cancel [simp] ~> []
1528 fold_mset_single [simp] ~> []
1529 subset_eq_empty [simp] ~> []
1530 empty_sup [simp] ~> []
1531 sup_empty [simp] ~> []
1532 inter_empty [simp] ~> []
1533 empty_inter [simp] ~> []
1536 * Session HOL-Library, theory Multiset: the order of the variables in
1537 the second cases of multiset_induct, multiset_induct2_size,
1538 multiset_induct2 has been changed (e.g. Add A a ~> Add a A).
1541 * Session HOL-Library, theory Multiset: there is now a simplification
1542 procedure on multisets. It mimics the behavior of the procedure on
1543 natural numbers. INCOMPATIBILITY.
1545 * Session HOL-Library, theory Multiset: renamed sums and products of
1549 msetprod ~> prod_mset
1551 * Session HOL-Library, theory Multiset: the notation for intersection
1552 and union of multisets have been changed:
1554 #\<inter> ~> \<inter>#
1555 #\<union> ~> \<union>#
1559 * Session HOL-Library, theory Multiset: the lemma
1560 one_step_implies_mult_aux on multisets has been removed, use
1561 one_step_implies_mult instead. INCOMPATIBILITY.
1563 * Session HOL-Library: theory Complete_Partial_Order2 provides reasoning
1564 support for monotonicity and continuity in chain-complete partial orders
1565 and about admissibility conditions for fixpoint inductions.
1567 * Session HOL-Library: theory Library/Polynomial contains also
1568 derivation of polynomials (formerly in Library/Poly_Deriv) but not
1569 gcd/lcm on polynomials over fields. This has been moved to a separate
1570 theory Library/Polynomial_GCD_euclidean.thy, to pave way for a possible
1571 future different type class instantiation for polynomials over factorial
1572 rings. INCOMPATIBILITY.
1574 * Session HOL-Library: theory Sublist provides function "prefixes" with
1575 the following renaming
1578 prefix -> strict_prefix
1580 suffix -> strict_suffix
1582 Added theory of longest common prefixes.
1584 * Session HOL-Number_Theory: algebraic foundation for primes:
1585 Generalisation of predicate "prime" and introduction of predicates
1586 "prime_elem", "irreducible", a "prime_factorization" function, and the
1587 "factorial_ring" typeclass with instance proofs for nat, int, poly. Some
1588 theorems now have different names, most notably "prime_def" is now
1589 "prime_nat_iff". INCOMPATIBILITY.
1591 * Session Old_Number_Theory has been removed, after porting remaining
1594 * Session HOL-Types_To_Sets provides an experimental extension of
1595 Higher-Order Logic to allow translation of types to sets.
1600 * Integer.gcd and Integer.lcm use efficient operations from the Poly/ML
1601 library (notably for big integers). Subtle change of semantics:
1602 Integer.gcd and Integer.lcm both normalize the sign, results are never
1603 negative. This coincides with the definitions in HOL/GCD.thy.
1606 * Structure Rat for rational numbers is now an integral part of
1607 Isabelle/ML, with special notation @int/nat or @int for numerals (an
1608 abbreviation for antiquotation @{Pure.rat argument}) and ML pretty
1609 printing. Standard operations on type Rat.rat are provided via ad-hoc
1610 overloading of + - * / < <= > >= ~ abs. INCOMPATIBILITY, need to
1611 use + instead of +/ etc. Moreover, exception Rat.DIVZERO has been
1612 superseded by General.Div.
1614 * ML antiquotation @{path} is superseded by @{file}, which ensures that
1615 the argument is a plain file. Minor INCOMPATIBILITY.
1617 * Antiquotation @{make_string} is available during Pure bootstrap --
1618 with approximative output quality.
1620 * Low-level ML system structures (like PolyML and RunCall) are no longer
1621 exposed to Isabelle/ML user-space. Potential INCOMPATIBILITY.
1623 * The ML function "ML" provides easy access to run-time compilation.
1624 This is particularly useful for conditional compilation, without
1625 requiring separate files.
1627 * Option ML_exception_debugger controls detailed exception trace via the
1628 Poly/ML debugger. Relevant ML modules need to be compiled beforehand
1629 with ML_file_debug, or with ML_file and option ML_debugger enabled. Note
1630 debugger information requires consirable time and space: main
1631 Isabelle/HOL with full debugger support may need ML_system_64.
1633 * Local_Theory.restore has been renamed to Local_Theory.reset to
1634 emphasize its disruptive impact on the cumulative context, notably the
1635 scope of 'private' or 'qualified' names. Note that Local_Theory.reset is
1636 only appropriate when targets are managed, e.g. starting from a global
1637 theory and returning to it. Regular definitional packages should use
1638 balanced blocks of Local_Theory.open_target versus
1639 Local_Theory.close_target instead. Rare INCOMPATIBILITY.
1641 * Structure TimeLimit (originally from the SML/NJ library) has been
1642 replaced by structure Timeout, with slightly different signature.
1645 * Discontinued cd and pwd operations, which are not well-defined in a
1646 multi-threaded environment. Note that files are usually located
1647 relatively to the master directory of a theory (see also
1648 File.full_path). Potential INCOMPATIBILITY.
1650 * Binding.empty_atts supersedes Thm.empty_binding and
1651 Attrib.empty_binding. Minor INCOMPATIBILITY.
1656 * SML/NJ and old versions of Poly/ML are no longer supported.
1658 * Poly/ML heaps now follow the hierarchy of sessions, and thus require
1659 much less disk space.
1661 * The Isabelle ML process is now managed directly by Isabelle/Scala, and
1662 shell scripts merely provide optional command-line access. In
1665 . Scala module ML_Process to connect to the raw ML process,
1666 with interaction via stdin/stdout/stderr or in batch mode;
1667 . command-line tool "isabelle console" as interactive wrapper;
1668 . command-line tool "isabelle process" as batch mode wrapper.
1670 * The executable "isabelle_process" has been discontinued. Tools and
1671 prover front-ends should use ML_Process or Isabelle_Process in
1672 Isabelle/Scala. INCOMPATIBILITY.
1674 * New command-line tool "isabelle process" supports ML evaluation of
1675 literal expressions (option -e) or files (option -f) in the context of a
1676 given heap image. Errors lead to premature exit of the ML process with
1679 * The command-line tool "isabelle build" supports option -N for cyclic
1680 shuffling of NUMA CPU nodes. This may help performance tuning on Linux
1681 servers with separate CPU/memory modules.
1683 * System option "threads" (for the size of the Isabelle/ML thread farm)
1684 is also passed to the underlying ML runtime system as --gcthreads,
1685 unless there is already a default provided via ML_OPTIONS settings.
1687 * System option "checkpoint" helps to fine-tune the global heap space
1688 management of isabelle build. This is relevant for big sessions that may
1689 exhaust the small 32-bit address space of the ML process (which is used
1692 * System option "profiling" specifies the mode for global ML profiling
1693 in "isabelle build". Possible values are "time", "allocations". The
1694 command-line tool "isabelle profiling_report" helps to digest the
1695 resulting log files.
1697 * System option "ML_process_policy" specifies an optional command prefix
1698 for the underlying ML process, e.g. to control CPU affinity on
1699 multiprocessor systems. The "isabelle jedit" tool allows to override the
1700 implicit default via option -p.
1702 * Command-line tool "isabelle console" provides option -r to help to
1703 bootstrapping Isabelle/Pure interactively.
1705 * Command-line tool "isabelle yxml" has been discontinued.
1706 INCOMPATIBILITY, use operations from the modules "XML" and "YXML" in
1707 Isabelle/ML or Isabelle/Scala.
1709 * Many Isabelle tools that require a Java runtime system refer to the
1710 settings ISABELLE_TOOL_JAVA_OPTIONS32 / ISABELLE_TOOL_JAVA_OPTIONS64,
1711 depending on the underlying platform. The settings for "isabelle build"
1712 ISABELLE_BUILD_JAVA_OPTIONS32 / ISABELLE_BUILD_JAVA_OPTIONS64 have been
1713 discontinued. Potential INCOMPATIBILITY.
1715 * The Isabelle system environment always ensures that the main
1716 executables are found within the shell search $PATH: "isabelle" and
1717 "isabelle_scala_script".
1719 * Isabelle tools may consist of .scala files: the Scala compiler is
1720 invoked on the spot. The source needs to define some object that extends
1723 * File.bash_string, File.bash_path etc. represent Isabelle/ML and
1724 Isabelle/Scala strings authentically within GNU bash. This is useful to
1725 produce robust shell scripts under program control, without worrying
1726 about spaces or special characters. Note that user output works via
1727 Path.print (ML) or Path.toString (Scala). INCOMPATIBILITY, the old (and
1728 less versatile) operations File.shell_quote, File.shell_path etc. have
1731 * The isabelle_java executable allows to run a Java process within the
1732 name space of Java and Scala components that are bundled with Isabelle,
1733 but without the Isabelle settings environment.
1735 * Isabelle/Scala: the SSH module supports ssh and sftp connections, for
1736 remote command-execution and file-system access. This resembles
1737 operations from module File and Isabelle_System to some extent. Note
1738 that Path specifications need to be resolved remotely via
1739 ssh.remote_path instead of File.standard_path: the implicit process
1740 environment is different, Isabelle settings are not available remotely.
1742 * Isabelle/Scala: the Mercurial module supports repositories via the
1743 regular hg command-line interface. The repositroy clone and working
1744 directory may reside on a local or remote file-system (via ssh
1749 New in Isabelle2016 (February 2016)
1750 -----------------------------------
1754 * Eisbach is now based on Pure instead of HOL. Objects-logics may import
1755 either the theory ~~/src/HOL/Eisbach/Eisbach (for HOL etc.) or
1756 ~~/src/HOL/Eisbach/Eisbach_Old_Appl_Syntax (for FOL, ZF etc.). Note that
1757 the HOL-Eisbach session located in ~~/src/HOL/Eisbach/ contains further
1758 examples that do require HOL.
1760 * Better resource usage on all platforms (Linux, Windows, Mac OS X) for
1761 both Isabelle/ML and Isabelle/Scala. Slightly reduced heap space usage.
1763 * Former "xsymbols" syntax with Isabelle symbols is used by default,
1764 without any special print mode. Important ASCII replacement syntax
1765 remains available under print mode "ASCII", but less important syntax
1766 has been removed (see below).
1768 * Support for more arrow symbols, with rendering in LaTeX and Isabelle
1769 fonts: \<Lleftarrow> \<Rrightarrow> \<longlongleftarrow> \<longlongrightarrow> \<longlonglongleftarrow> \<longlonglongrightarrow>.
1771 * Special notation \<struct> for the first implicit 'structure' in the
1772 context has been discontinued. Rare INCOMPATIBILITY, use explicit
1773 structure name instead, notably in indexed notation with block-subscript
1774 (e.g. \<odot>\<^bsub>A\<^esub>).
1776 * The glyph for \<diamond> in the IsabelleText font now corresponds better to its
1777 counterpart \<box> as quantifier-like symbol. A small diamond is available as
1778 \<diamondop>; the old symbol \<struct> loses this rendering and any special
1781 * Syntax for formal comments "-- text" now also supports the symbolic
1782 form "\<comment> text". Command-line tool "isabelle update_cartouches -c" helps
1783 to update old sources.
1785 * Toplevel theorem statements have been simplified as follows:
1788 schematic_lemma ~> schematic_goal
1789 schematic_theorem ~> schematic_goal
1790 schematic_corollary ~> schematic_goal
1792 Command-line tool "isabelle update_theorems" updates theory sources
1795 * Toplevel theorem statement 'proposition' is another alias for
1798 * The old 'defs' command has been removed (legacy since Isabelle2014).
1799 INCOMPATIBILITY, use regular 'definition' instead. Overloaded and/or
1800 deferred definitions require a surrounding 'overloading' block.
1803 *** Prover IDE -- Isabelle/Scala/jEdit ***
1805 * IDE support for the source-level debugger of Poly/ML, to work with
1806 Isabelle/ML and official Standard ML. Option "ML_debugger" and commands
1807 'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug',
1808 'SML_file_no_debug' control compilation of sources with or without
1809 debugging information. The Debugger panel allows to set breakpoints (via
1810 context menu), step through stopped threads, evaluate local ML
1811 expressions etc. At least one Debugger view needs to be active to have
1812 any effect on the running ML program.
1814 * The State panel manages explicit proof state output, with dynamic
1815 auto-update according to cursor movement. Alternatively, the jEdit
1816 action "isabelle.update-state" (shortcut S+ENTER) triggers manual
1819 * The Output panel no longer shows proof state output by default, to
1820 avoid GUI overcrowding. INCOMPATIBILITY, use the State panel instead or
1821 enable option "editor_output_state".
1823 * The text overview column (status of errors, warnings etc.) is updated
1824 asynchronously, leading to much better editor reactivity. Moreover, the
1825 full document node content is taken into account. The width of the
1826 column is scaled according to the main text area font, for improved
1829 * The main text area no longer changes its color hue in outdated
1830 situations. The text overview column takes over the role to indicate
1831 unfinished edits in the PIDE pipeline. This avoids flashing text display
1832 due to ad-hoc updates by auxiliary GUI components, such as the State
1835 * Slightly improved scheduling for urgent print tasks (e.g. command
1836 state output, interactive queries) wrt. long-running background tasks.
1838 * Completion of symbols via prefix of \<name> or \<^name> or \name is
1839 always possible, independently of the language context. It is never
1840 implicit: a popup will show up unconditionally.
1842 * Additional abbreviations for syntactic completion may be specified in
1843 $ISABELLE_HOME/etc/abbrevs and $ISABELLE_HOME_USER/etc/abbrevs, with
1844 support for simple templates using ASCII 007 (bell) as placeholder.
1846 * Symbols \<oplus>, \<Oplus>, \<otimes>, \<Otimes>, \<odot>, \<Odot>, \<ominus>, \<oslash> no longer provide abbreviations for
1847 completion like "+o", "*o", ".o" etc. -- due to conflicts with other
1848 ASCII syntax. INCOMPATIBILITY, use plain backslash-completion or define
1849 suitable abbreviations in $ISABELLE_HOME_USER/etc/abbrevs.
1851 * Action "isabelle-emph" (with keyboard shortcut C+e LEFT) controls
1852 emphasized text style; the effect is visible in document output, not in
1855 * Action "isabelle-reset" now uses keyboard shortcut C+e BACK_SPACE,
1856 instead of former C+e LEFT.
1858 * The command-line tool "isabelle jedit" and the isabelle.Main
1859 application wrapper treat the default $USER_HOME/Scratch.thy more
1860 uniformly, and allow the dummy file argument ":" to open an empty buffer
1863 * New command-line tool "isabelle jedit_client" allows to connect to an
1864 already running Isabelle/jEdit process. This achieves the effect of
1865 single-instance applications seen on common GUI desktops.
1867 * The default look-and-feel for Linux is the traditional "Metal", which
1868 works better with GUI scaling for very high-resolution displays (e.g.
1869 4K). Moreover, it is generally more robust than "Nimbus".
1871 * Update to jedit-5.3.0, with improved GUI scaling and support of
1872 high-resolution displays (e.g. 4K).
1874 * The main Isabelle executable is managed as single-instance Desktop
1875 application uniformly on all platforms: Linux, Windows, Mac OS X.
1878 *** Document preparation ***
1880 * Commands 'paragraph' and 'subparagraph' provide additional section
1881 headings. Thus there are 6 levels of standard headings, as in HTML.
1883 * Command 'text_raw' has been clarified: input text is processed as in
1884 'text' (with antiquotations and control symbols). The key difference is
1885 the lack of the surrounding isabelle markup environment in output.
1887 * Text is structured in paragraphs and nested lists, using notation that
1888 is similar to Markdown. The control symbols for list items are as
1893 \<^descr> description
1895 * There is a new short form for antiquotations with a single argument
1896 that is a cartouche: \<^name>\<open>...\<close> is equivalent to @{name \<open>...\<close>} and
1897 \<open>...\<close> without control symbol is equivalent to @{cartouche \<open>...\<close>}.
1898 \<^name> without following cartouche is equivalent to @{name}. The
1899 standard Isabelle fonts provide glyphs to render important control
1900 symbols, e.g. "\<^verbatim>", "\<^emph>", "\<^bold>".
1902 * Antiquotations @{noindent}, @{smallskip}, @{medskip}, @{bigskip} with
1903 corresponding control symbols \<^noindent>, \<^smallskip>, \<^medskip>, \<^bigskip> specify spacing formally, using
1904 standard LaTeX macros of the same names.
1906 * Antiquotation @{cartouche} in Isabelle/Pure is the same as @{text}.
1907 Consequently, \<open>...\<close> without any decoration prints literal quasi-formal
1908 text. Command-line tool "isabelle update_cartouches -t" helps to update
1909 old sources, by approximative patching of the content of string and
1910 cartouche tokens seen in theory sources.
1912 * The @{text} antiquotation now ignores the antiquotation option
1913 "source". The given text content is output unconditionally, without any
1914 surrounding quotes etc. Subtle INCOMPATIBILITY, put quotes into the
1915 argument where they are really intended, e.g. @{text \<open>"foo"\<close>}. Initial
1916 or terminal spaces are ignored.
1918 * Antiquotations @{emph} and @{bold} output LaTeX source recursively,
1919 adding appropriate text style markup. These may be used in the short
1920 form \<^emph>\<open>...\<close> and \<^bold>\<open>...\<close>.
1922 * Document antiquotation @{footnote} outputs LaTeX source recursively,
1923 marked as \footnote{}. This may be used in the short form \<^footnote>\<open>...\<close>.
1925 * Antiquotation @{verbatim [display]} supports option "indent".
1927 * Antiquotation @{theory_text} prints uninterpreted theory source text
1928 (Isar outer syntax with command keywords etc.). This may be used in the
1929 short form \<^theory_text>\<open>...\<close>. @{theory_text [display]} supports option "indent".
1931 * Antiquotation @{doc ENTRY} provides a reference to the given
1932 documentation, with a hyperlink in the Prover IDE.
1934 * Antiquotations @{command}, @{method}, @{attribute} print checked
1935 entities of the Isar language.
1937 * HTML presentation uses the standard IsabelleText font and Unicode
1938 rendering of Isabelle symbols like Isabelle/Scala/jEdit. The former
1939 print mode "HTML" loses its special meaning.
1944 * Local goals ('have', 'show', 'hence', 'thus') allow structured rule
1945 statements like fixes/assumes/shows in theorem specifications, but the
1946 notation is postfix with keywords 'if' (or 'when') and 'for'. For
1949 have result: "C x y"
1951 for x :: 'a and y :: 'a
1954 The local assumptions are bound to the name "that". The result is
1955 exported from context of the statement as usual. The above roughly
1956 corresponds to a raw proof block like this:
1959 fix x :: 'a and y :: 'a
1960 assume that: "A x" "B y"
1961 have "C x y" <proof>
1965 The keyword 'when' may be used instead of 'if', to indicate 'presume'
1966 instead of 'assume' above.
1968 * Assumptions ('assume', 'presume') allow structured rule statements
1969 using 'if' and 'for', similar to 'have' etc. above. For example:
1971 assume result: "C x y"
1973 for x :: 'a and y :: 'a
1975 This assumes "\<And>x y::'a. A x \<Longrightarrow> B y \<Longrightarrow> C x y" and produces a general
1976 result as usual: "A ?x \<Longrightarrow> B ?y \<Longrightarrow> C ?x ?y".
1978 Vacuous quantification in assumptions is omitted, i.e. a for-context
1979 only effects propositions according to actual use of variables. For
1982 assume "A x" and "B y" for x and y
1986 assume "\<And>x. A x" and "\<And>y. B y"
1988 * The meaning of 'show' with Pure rule statements has changed: premises
1989 are treated in the sense of 'assume', instead of 'presume'. This means,
1990 a goal like "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x" can be solved completely as
1993 show "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
1997 show "C x" if "A x" "B x" for x
1999 Rare INCOMPATIBILITY, the old behaviour may be recovered as follows:
2001 show "C x" when "A x" "B x" for x
2003 * New command 'consider' states rules for generalized elimination and
2004 case splitting. This is like a toplevel statement "theorem obtains" used
2005 within a proof body; or like a multi-branch 'obtain' without activation
2006 of the local context elements yet.
2008 * Proof method "cases" allows to specify the rule as first entry of
2009 chained facts. This is particularly useful with 'consider':
2011 consider (a) A | (b) B | (c) C <proof>
2015 then show ?thesis <proof>
2018 then show ?thesis <proof>
2021 then show ?thesis <proof>
2024 * Command 'case' allows fact name and attribute specification like this:
2027 case a [attributes]: (c xs)
2029 Facts that are introduced by invoking the case context are uniformly
2030 qualified by "a"; the same name is used for the cumulative fact. The old
2031 form "case (c xs) [attributes]" is no longer supported. Rare
2032 INCOMPATIBILITY, need to adapt uses of case facts in exotic situations,
2033 and always put attributes in front.
2035 * The standard proof method of commands 'proof' and '..' is now called
2036 "standard" to make semantically clear what it is; the old name "default"
2037 is still available as legacy for some time. Documentation now explains
2038 '..' more accurately as "by standard" instead of "by rule".
2040 * Nesting of Isar goal structure has been clarified: the context after
2041 the initial backwards refinement is retained for the whole proof, within
2042 all its context sections (as indicated via 'next'). This is e.g.
2043 relevant for 'using', 'including', 'supply':
2045 have "A \<and> A" if a: A for A
2053 * Command 'obtain' binds term abbreviations (via 'is' patterns) in the
2054 proof body as well, abstracted over relevant parameters.
2056 * Improved type-inference for theorem statement 'obtains': separate
2057 parameter scope for of each clause.
2059 * Term abbreviations via 'is' patterns also work for schematic
2060 statements: result is abstracted over unknowns.
2062 * Command 'subgoal' allows to impose some structure on backward
2063 refinements, to avoid proof scripts degenerating into long of 'apply'
2064 sequences. Further explanations and examples are given in the isar-ref
2067 * Command 'supply' supports fact definitions during goal refinement
2070 * Proof method "goal_cases" turns the current subgoals into cases within
2071 the context; the conclusion is bound to variable ?case in each case. For
2074 lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
2075 and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
2078 then show ?case using \<open>A x\<close> \<open>B x\<close> sorry
2081 then show ?case using \<open>U y\<close> \<open>V z\<close> sorry
2084 lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
2085 and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
2088 then show ?case using prems sorry
2091 then show ?case using prems sorry
2094 * The undocumented feature of implicit cases goal1, goal2, goal3, etc.
2095 is marked as legacy, and will be removed eventually. The proof method
2096 "goals" achieves a similar effect within regular Isar; often it can be
2097 done more adequately by other means (e.g. 'consider').
2099 * The vacuous fact "TERM x" may be established "by fact" or as `TERM x`
2100 as well, not just "by this" or "." as before.
2102 * Method "sleep" succeeds after a real-time delay (in seconds). This is
2103 occasionally useful for demonstration and testing purposes.
2108 * Qualifiers in locale expressions default to mandatory ('!') regardless
2109 of the command. Previously, for 'locale' and 'sublocale' the default was
2110 optional ('?'). The old synatx '!' has been discontinued.
2111 INCOMPATIBILITY, remove '!' and add '?' as required.
2113 * Keyword 'rewrites' identifies rewrite morphisms in interpretation
2114 commands. Previously, the keyword was 'where'. INCOMPATIBILITY.
2116 * More gentle suppression of syntax along locale morphisms while
2117 printing terms. Previously 'abbreviation' and 'notation' declarations
2118 would be suppressed for morphisms except term identity. Now
2119 'abbreviation' is also kept for morphims that only change the involved
2120 parameters, and only 'notation' is suppressed. This can be of great help
2121 when working with complex locale hierarchies, because proof states are
2122 displayed much more succinctly. It also means that only notation needs
2123 to be redeclared if desired, as illustrated by this example:
2125 locale struct = fixes composition :: "'a => 'a => 'a" (infixl "\<cdot>" 65)
2127 definition derived (infixl "\<odot>" 65) where ...
2131 left: struct composition + right: struct composition'
2132 for composition (infix "\<cdot>" 65) and composition' (infix "\<cdot>''" 65)
2134 notation right.derived ("\<odot>''")
2137 * Command 'global_interpretation' issues interpretations into global
2138 theories, with optional rewrite definitions following keyword 'defines'.
2140 * Command 'sublocale' accepts optional rewrite definitions after keyword
2143 * Command 'permanent_interpretation' has been discontinued. Use
2144 'global_interpretation' or 'sublocale' instead. INCOMPATIBILITY.
2146 * Command 'print_definitions' prints dependencies of definitional
2147 specifications. This functionality used to be part of 'print_theory'.
2149 * Configuration option rule_insts_schematic has been discontinued
2150 (intermediate legacy feature in Isabelle2015). INCOMPATIBILITY.
2152 * Abbreviations in type classes now carry proper sort constraint. Rare
2153 INCOMPATIBILITY in situations where the previous misbehaviour has been
2156 * Refinement of user-space type system in type classes: pseudo-local
2157 operations behave more similar to abbreviations. Potential
2158 INCOMPATIBILITY in exotic situations.
2163 * The 'typedef' command has been upgraded from a partially checked
2164 "axiomatization", to a full definitional specification that takes the
2165 global collection of overloaded constant / type definitions into
2166 account. Type definitions with open dependencies on overloaded
2167 definitions need to be specified as "typedef (overloaded)". This
2168 provides extra robustness in theory construction. Rare INCOMPATIBILITY.
2170 * Qualification of various formal entities in the libraries is done more
2171 uniformly via "context begin qualified definition ... end" instead of
2172 old-style "hide_const (open) ...". Consequently, both the defined
2173 constant and its defining fact become qualified, e.g. Option.is_none and
2174 Option.is_none_def. Occasional INCOMPATIBILITY in applications.
2176 * Some old and rarely used ASCII replacement syntax has been removed.
2177 INCOMPATIBILITY, standard syntax with symbols should be used instead.
2178 The subsequent commands help to reproduce the old forms, e.g. to
2179 simplify porting old theories:
2181 notation iff (infixr "<->" 25)
2183 notation Times (infixr "<*>" 80)
2185 type_notation Map.map (infixr "~=>" 0)
2186 notation Map.map_comp (infixl "o'_m" 55)
2188 type_notation FinFun.finfun ("(_ =>f /_)" [22, 21] 21)
2190 notation FuncSet.funcset (infixr "->" 60)
2191 notation FuncSet.extensional_funcset (infixr "->\<^sub>E" 60)
2193 notation Omega_Words_Fun.conc (infixr "conc" 65)
2195 notation Preorder.equiv ("op ~~")
2196 and Preorder.equiv ("(_/ ~~ _)" [51, 51] 50)
2198 notation (in topological_space) tendsto (infixr "--->" 55)
2199 notation (in topological_space) LIMSEQ ("((_)/ ----> (_))" [60, 60] 60)
2200 notation LIM ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
2202 notation NSA.approx (infixl "@=" 50)
2203 notation NSLIMSEQ ("((_)/ ----NS> (_))" [60, 60] 60)
2204 notation NSLIM ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
2206 * The alternative notation "\<Colon>" for type and sort constraints has been
2207 removed: in LaTeX document output it looks the same as "::".
2208 INCOMPATIBILITY, use plain "::" instead.
2210 * Commands 'inductive' and 'inductive_set' work better when names for
2211 intro rules are omitted: the "cases" and "induct" rules no longer
2212 declare empty case_names, but no case_names at all. This allows to use
2213 numbered cases in proofs, without requiring method "goal_cases".
2215 * Inductive definitions ('inductive', 'coinductive', etc.) expose
2216 low-level facts of the internal construction only if the option
2217 "inductive_internals" is enabled. This refers to the internal predicate
2218 definition and its monotonicity result. Rare INCOMPATIBILITY.
2220 * Recursive function definitions ('fun', 'function', 'partial_function')
2221 expose low-level facts of the internal construction only if the option
2222 "function_internals" is enabled. Its internal inductive definition is
2223 also subject to "inductive_internals". Rare INCOMPATIBILITY.
2225 * BNF datatypes ('datatype', 'codatatype', etc.) expose low-level facts
2226 of the internal construction only if the option "bnf_internals" is
2227 enabled. This supersedes the former option "bnf_note_all". Rare
2230 * Combinator to represent case distinction on products is named
2231 "case_prod", uniformly, discontinuing any input aliasses. Very popular
2232 theorem aliasses have been retained.
2235 PairE ~> prod.exhaust
2236 Pair_eq ~> prod.inject
2237 pair_collapse ~> prod.collapse
2238 Pair_fst_snd_eq ~> prod_eq_iff
2239 split_twice ~> prod.case_distrib
2240 split_weak_cong ~> prod.case_cong_weak
2241 split_split ~> prod.split
2242 split_split_asm ~> prod.split_asm
2243 splitI ~> case_prodI
2244 splitD ~> case_prodD
2245 splitI2 ~> case_prodI2
2246 splitI2' ~> case_prodI2'
2247 splitE ~> case_prodE
2248 splitE' ~> case_prodE'
2249 split_pair ~> case_prod_Pair
2250 split_eta ~> case_prod_eta
2251 split_comp ~> case_prod_comp
2252 mem_splitI ~> mem_case_prodI
2253 mem_splitI2 ~> mem_case_prodI2
2254 mem_splitE ~> mem_case_prodE
2255 The_split ~> The_case_prod
2256 cond_split_eta ~> cond_case_prod_eta
2257 Collect_split_in_rel_leE ~> Collect_case_prod_in_rel_leE
2258 Collect_split_in_rel_leI ~> Collect_case_prod_in_rel_leI
2259 in_rel_Collect_split_eq ~> in_rel_Collect_case_prod_eq
2260 Collect_split_Grp_eqD ~> Collect_case_prod_Grp_eqD
2261 Collect_split_Grp_inD ~> Collect_case_prod_Grp_in
2262 Domain_Collect_split ~> Domain_Collect_case_prod
2263 Image_Collect_split ~> Image_Collect_case_prod
2264 Range_Collect_split ~> Range_Collect_case_prod
2265 Eps_split ~> Eps_case_prod
2266 Eps_split_eq ~> Eps_case_prod_eq
2267 split_rsp ~> case_prod_rsp
2268 curry_split ~> curry_case_prod
2269 split_curry ~> case_prod_curry
2271 Changes in structure HOLogic:
2272 split_const ~> case_prod_const
2273 mk_split ~> mk_case_prod
2274 mk_psplits ~> mk_ptupleabs
2275 strip_psplits ~> strip_ptupleabs
2279 * The coercions to type 'real' have been reorganised. The function
2280 'real' is no longer overloaded, but has type 'nat => real' and
2281 abbreviates of_nat for that type. Also 'real_of_int :: int => real'
2282 abbreviates of_int for that type. Other overloaded instances of 'real'
2283 have been replaced by 'real_of_ereal' and 'real_of_float'.
2285 Consolidated facts (among others):
2286 real_of_nat_le_iff -> of_nat_le_iff
2287 real_of_nat_numeral of_nat_numeral
2288 real_of_int_zero of_int_0
2289 real_of_nat_zero of_nat_0
2290 real_of_one of_int_1
2291 real_of_int_add of_int_add
2292 real_of_nat_add of_nat_add
2293 real_of_int_diff of_int_diff
2294 real_of_nat_diff of_nat_diff
2295 floor_subtract floor_diff_of_int
2296 real_of_int_inject of_int_eq_iff
2297 real_of_int_gt_zero_cancel_iff of_int_0_less_iff
2298 real_of_int_ge_zero_cancel_iff of_int_0_le_iff
2299 real_of_nat_ge_zero of_nat_0_le_iff
2300 real_of_int_ceiling_ge le_of_int_ceiling
2301 ceiling_less_eq ceiling_less_iff
2302 ceiling_le_eq ceiling_le_iff
2303 less_floor_eq less_floor_iff
2304 floor_less_eq floor_less_iff
2305 floor_divide_eq_div floor_divide_of_int_eq
2306 real_of_int_zero_cancel of_nat_eq_0_iff
2307 ceiling_real_of_int ceiling_of_int
2311 * Theory Map: lemma map_of_is_SomeD was a clone of map_of_SomeD and has
2312 been removed. INCOMPATIBILITY.
2314 * Quickcheck setup for finite sets.
2316 * Discontinued simp_legacy_precond. Potential INCOMPATIBILITY.
2319 - The MaSh relevance filter has been sped up.
2320 - Proof reconstruction has been improved, to minimize the incidence of
2321 cases where Sledgehammer gives a proof that does not work.
2322 - Auto Sledgehammer now minimizes and preplays the results.
2323 - Handle Vampire 4.0 proof output without raising exception.
2324 - Eliminated "MASH" environment variable. Use the "MaSh" option in
2325 Isabelle/jEdit instead. INCOMPATIBILITY.
2326 - Eliminated obsolete "blocking" option and related subcommands.
2329 - Fixed soundness bug in translation of "finite" predicate.
2330 - Fixed soundness bug in "destroy_constrs" optimization.
2331 - Fixed soundness bug in translation of "rat" type.
2332 - Removed "check_potential" and "check_genuine" options.
2333 - Eliminated obsolete "blocking" option.
2335 * (Co)datatype package:
2336 - New commands "lift_bnf" and "copy_bnf" for lifting (copying) a BNF
2337 structure on the raw type to an abstract type defined using typedef.
2338 - Always generate "case_transfer" theorem.
2339 - For mutual types, generate slightly stronger "rel_induct",
2340 "rel_coinduct", and "coinduct" theorems. INCOMPATIBILITY.
2341 - Allow discriminators and selectors with the same name as the type
2343 - Avoid various internal name clashes (e.g., 'datatype f = f').
2345 * Transfer: new methods for interactive debugging of 'transfer' and
2346 'transfer_prover': 'transfer_start', 'transfer_step', 'transfer_end',
2347 'transfer_prover_start' and 'transfer_prover_end'.
2349 * New diagnostic command print_record for displaying record definitions.
2351 * Division on integers is bootstrapped directly from division on
2352 naturals and uses generic numeral algorithm for computations. Slight
2353 INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes former
2354 simprocs binary_int_div and binary_int_mod
2356 * Tightened specification of class semiring_no_zero_divisors. Minor
2359 * Class algebraic_semidom introduces common algebraic notions of
2360 integral (semi)domains, particularly units. Although logically subsumed
2361 by fields, is is not a super class of these in order not to burden
2362 fields with notions that are trivial there.
2364 * Class normalization_semidom specifies canonical representants for
2365 equivalence classes of associated elements in an integral (semi)domain.
2366 This formalizes associated elements as well.
2368 * Abstract specification of gcd/lcm operations in classes semiring_gcd,
2369 semiring_Gcd, semiring_Lcd. Minor INCOMPATIBILITY: facts gcd_nat.commute
2370 and gcd_int.commute are subsumed by gcd.commute, as well as
2371 gcd_nat.assoc and gcd_int.assoc by gcd.assoc.
2373 * Former constants Fields.divide (_ / _) and Divides.div (_ div _) are
2374 logically unified to Rings.divide in syntactic type class Rings.divide,
2375 with infix syntax (_ div _). Infix syntax (_ / _) for field division is
2376 added later as abbreviation in class Fields.inverse. INCOMPATIBILITY,
2377 instantiations must refer to Rings.divide rather than the former
2378 separate constants, hence infix syntax (_ / _) is usually not available
2379 during instantiation.
2381 * New cancellation simprocs for boolean algebras to cancel complementary
2382 terms for sup and inf. For example, "sup x (sup y (- x))" simplifies to
2383 "top". INCOMPATIBILITY.
2385 * Class uniform_space introduces uniform spaces btw topological spaces
2386 and metric spaces. Minor INCOMPATIBILITY: open_<type>_def needs to be
2387 introduced in the form of an uniformity. Some constants are more general
2388 now, it may be necessary to add type class constraints.
2390 open_real_def \<leadsto> open_dist
2391 open_complex_def \<leadsto> open_dist
2393 * Library/Monad_Syntax: notation uses symbols \<bind> and \<then>. INCOMPATIBILITY.
2396 - Renamed multiset inclusion operators:
2404 - Added multiset inclusion operator syntax:
2409 - "'a multiset" is no longer an instance of the "order",
2410 "ordered_ab_semigroup_add_imp_le", "ordered_cancel_comm_monoid_diff",
2411 "semilattice_inf", and "semilattice_sup" type classes. The theorems
2412 previously provided by these type classes (directly or indirectly)
2413 are now available through the "subset_mset" interpretation
2414 (e.g. add_mono ~> subset_mset.add_mono).
2416 - Renamed conversions:
2418 multiset_of_set ~> mset_set
2422 mset_le_def ~> subseteq_mset_def
2423 mset_less_def ~> subset_mset_def
2424 less_eq_multiset.rep_eq ~> subseteq_mset_def
2426 - Removed lemmas generated by lift_definition:
2427 less_eq_multiset.abs_eq, less_eq_multiset.rsp,
2428 less_eq_multiset.transfer, less_eq_multiset_def
2431 * Library/Omega_Words_Fun: Infinite words modeled as functions nat \<Rightarrow> 'a.
2433 * Library/Bourbaki_Witt_Fixpoint: Added formalisation of the
2434 Bourbaki-Witt fixpoint theorem for increasing functions in
2435 chain-complete partial orders.
2437 * Library/Old_Recdef: discontinued obsolete 'defer_recdef' command.
2438 Minor INCOMPATIBILITY, use 'function' instead.
2440 * Library/Periodic_Fun: a locale that provides convenient lemmas for
2443 * Library/Formal_Power_Series: proper definition of division (with
2444 remainder) for formal power series; instances for Euclidean Ring and
2447 * HOL-Imperative_HOL: obsolete theory Legacy_Mrec has been removed.
2449 * HOL-Statespace: command 'statespace' uses mandatory qualifier for
2450 import of parent, as for general 'locale' expressions. INCOMPATIBILITY,
2451 remove '!' and add '?' as required.
2453 * HOL-Decision_Procs: The "approximation" method works with "powr"
2454 (exponentiation on real numbers) again.
2456 * HOL-Multivariate_Analysis: theory Cauchy_Integral_Thm with Contour
2457 integrals (= complex path integrals), Cauchy's integral theorem, winding
2458 numbers and Cauchy's integral formula, Liouville theorem, Fundamental
2459 Theorem of Algebra. Ported from HOL Light.
2461 * HOL-Multivariate_Analysis: topological concepts such as connected
2462 components, homotopic paths and the inside or outside of a set.
2464 * HOL-Multivariate_Analysis: radius of convergence of power series and
2465 various summability tests; Harmonic numbers and the Euler–Mascheroni
2466 constant; the Generalised Binomial Theorem; the complex and real
2467 Gamma/log-Gamma/Digamma/ Polygamma functions and their most important
2470 * HOL-Probability: The central limit theorem based on Levy's uniqueness
2471 and continuity theorems, weak convergence, and characterisitc functions.
2473 * HOL-Data_Structures: new and growing session of standard data
2479 * The following combinators for low-level profiling of the ML runtime
2480 system are available:
2482 profile_time (*CPU time*)
2483 profile_time_thread (*CPU time on this thread*)
2484 profile_allocations (*overall heap allocations*)
2486 * Antiquotation @{undefined} or \<^undefined> inlines (raise Match).
2488 * Antiquotation @{method NAME} inlines the (checked) name of the given
2491 * Pretty printing of Poly/ML compiler output in Isabelle has been
2492 improved: proper treatment of break offsets and blocks with consistent
2495 * The auxiliary module Pure/display.ML has been eliminated. Its
2496 elementary thm print operations are now in Pure/more_thm.ML and thus
2497 called Thm.pretty_thm, Thm.string_of_thm etc. INCOMPATIBILITY.
2499 * Simproc programming interfaces have been simplified:
2500 Simplifier.make_simproc and Simplifier.define_simproc supersede various
2501 forms of Simplifier.mk_simproc, Simplifier.simproc_global etc. Note that
2502 term patterns for the left-hand sides are specified with implicitly
2503 fixed variables, like top-level theorem statements. INCOMPATIBILITY.
2505 * Instantiation rules have been re-organized as follows:
2507 Thm.instantiate (*low-level instantiation with named arguments*)
2508 Thm.instantiate' (*version with positional arguments*)
2510 Drule.infer_instantiate (*instantiation with type inference*)
2511 Drule.infer_instantiate' (*version with positional arguments*)
2513 The LHS only requires variable specifications, instead of full terms.
2514 Old cterm_instantiate is superseded by infer_instantiate.
2515 INCOMPATIBILITY, need to re-adjust some ML names and types accordingly.
2517 * Old tactic shorthands atac, rtac, etac, dtac, ftac have been
2518 discontinued. INCOMPATIBILITY, use regular assume_tac, resolve_tac etc.
2519 instead (with proper context).
2521 * Thm.instantiate (and derivatives) no longer require the LHS of the
2522 instantiation to be certified: plain variables are given directly.
2524 * Subgoal.SUBPROOF and Subgoal.FOCUS combinators use anonymous
2525 quasi-bound variables (like the Simplifier), instead of accidentally
2526 named local fixes. This has the potential to improve stability of proof
2527 tools, but can also cause INCOMPATIBILITY for tools that don't observe
2528 the proof context discipline.
2530 * Isar proof methods are based on a slightly more general type
2531 context_tactic, which allows to change the proof context dynamically
2532 (e.g. to update cases) and indicate explicit Seq.Error results. Former
2533 METHOD_CASES is superseded by CONTEXT_METHOD; further combinators are
2534 provided in src/Pure/Isar/method.ML for convenience. INCOMPATIBILITY.
2539 * Command-line tool "isabelle console" enables print mode "ASCII".
2541 * Command-line tool "isabelle update_then" expands old Isar command
2547 This syntax is more orthogonal and improves readability and
2548 maintainability of proofs.
2550 * Global session timeout is multiplied by timeout_scale factor. This
2551 allows to adjust large-scale tests (e.g. AFP) to overall hardware
2554 * Property values in etc/symbols may contain spaces, if written with the
2555 replacement character "␣" (Unicode point 0x2324). For example:
2557 \<star> code: 0x0022c6 group: operator font: Deja␣Vu␣Sans␣Mono
2559 * Java runtime environment for x86_64-windows allows to use larger heap
2562 * Java runtime options are determined separately for 32bit vs. 64bit
2563 platforms as follows.
2565 - Isabelle desktop application: platform-specific files that are
2566 associated with the main app bundle
2568 - isabelle jedit: settings
2569 JEDIT_JAVA_SYSTEM_OPTIONS
2570 JEDIT_JAVA_OPTIONS32 vs. JEDIT_JAVA_OPTIONS64
2572 - isabelle build: settings
2573 ISABELLE_BUILD_JAVA_OPTIONS32 vs. ISABELLE_BUILD_JAVA_OPTIONS64
2575 * Bash shell function "jvmpath" has been renamed to "platform_path": it
2576 is relevant both for Poly/ML and JVM processes.
2578 * Poly/ML default platform architecture may be changed from 32bit to
2579 64bit via system option ML_system_64. A system restart (and rebuild) is
2580 required after change.
2582 * Poly/ML 5.6 runs natively on x86-windows and x86_64-windows, which
2583 both allow larger heap space than former x86-cygwin.
2585 * Heap images are 10-15% smaller due to less wasteful persistent theory
2586 content (using ML type theory_id instead of theory);
2590 New in Isabelle2015 (May 2015)
2591 ------------------------------
2595 * Local theory specification commands may have a 'private' or
2596 'qualified' modifier to restrict name space accesses to the local scope,
2597 as provided by some "context begin ... end" block. For example:
2602 private definition ...
2605 qualified definition ...
2613 * Command 'experiment' opens an anonymous locale context with private
2616 * Command 'notepad' requires proper nesting of begin/end and its proof
2617 structure in the body: 'oops' is no longer supported here. Minor
2618 INCOMPATIBILITY, use 'sorry' instead.
2620 * Command 'named_theorems' declares a dynamic fact within the context,
2621 together with an attribute to maintain the content incrementally. This
2622 supersedes functor Named_Thms in Isabelle/ML, but with a subtle change
2623 of semantics due to external visual order vs. internal reverse order.
2625 * 'find_theorems': search patterns which are abstractions are
2626 schematically expanded before search. Search results match the naive
2627 expectation more closely, particularly wrt. abbreviations.
2630 * Commands 'method_setup' and 'attribute_setup' now work within a local
2633 * Outer syntax commands are managed authentically within the theory
2634 context, without implicit global state. Potential for accidental
2635 INCOMPATIBILITY, make sure that required theories are really imported.
2637 * Historical command-line terminator ";" is no longer accepted (and
2638 already used differently in Isar). Minor INCOMPATIBILITY, use "isabelle
2639 update_semicolons" to remove obsolete semicolons from old theory
2642 * Structural composition of proof methods (meth1; meth2) in Isar
2643 corresponds to (tac1 THEN_ALL_NEW tac2) in ML.
2645 * The Eisbach proof method language allows to define new proof methods
2646 by combining existing ones with their usual syntax. The "match" proof
2647 method provides basic fact/term matching in addition to
2648 premise/conclusion matching through Subgoal.focus, and binds fact names
2649 from matches as well as term patterns within matches. The Isabelle
2650 documentation provides an entry "eisbach" for the Eisbach User Manual.
2651 Sources and various examples are in ~~/src/HOL/Eisbach/.
2654 *** Prover IDE -- Isabelle/Scala/jEdit ***
2656 * Improved folding mode "isabelle" based on Isar syntax. Alternatively,
2657 the "sidekick" mode may be used for document structure.
2659 * Extended bracket matching based on Isar language structure. System
2660 option jedit_structure_limit determines maximum number of lines to scan
2663 * Support for BibTeX files: context menu, context-sensitive token
2664 marker, SideKick parser.
2666 * Document antiquotation @{cite} provides formal markup, which is
2667 interpreted semi-formally based on .bib files that happen to be open in
2668 the editor (hyperlinks, completion etc.).
2670 * Less waste of vertical space via negative line spacing (see Global
2671 Options / Text Area).
2673 * Improved graphview panel with optional output of PNG or PDF, for
2674 display of 'thy_deps', 'class_deps' etc.
2676 * The commands 'thy_deps' and 'class_deps' allow optional bounds to
2677 restrict the visualized hierarchy.
2679 * Improved scheduling for asynchronous print commands (e.g. provers
2680 managed by the Sledgehammer panel) wrt. ongoing document processing.
2683 *** Document preparation ***
2685 * Document markup commands 'chapter', 'section', 'subsection',
2686 'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any
2687 context, even before the initial 'theory' command. Obsolete proof
2688 commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been
2689 discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw'
2690 instead. The old 'header' command is still retained for some time, but
2691 should be replaced by 'chapter', 'section' etc. (using "isabelle
2692 update_header"). Minor INCOMPATIBILITY.
2694 * Official support for "tt" style variants, via \isatt{...} or
2695 \begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or
2696 verbatim environment of LaTeX is no longer used. This allows @{ML} etc.
2697 as argument to other macros (such as footnotes).
2699 * Document antiquotation @{verbatim} prints ASCII text literally in "tt"
2702 * Discontinued obsolete option "document_graph": session_graph.pdf is
2703 produced unconditionally for HTML browser_info and PDF-LaTeX document.
2705 * Diagnostic commands and document markup commands within a proof do not
2706 affect the command tag for output. Thus commands like 'thm' are subject
2707 to proof document structure, and no longer "stick out" accidentally.
2708 Commands 'text' and 'txt' merely differ in the LaTeX style, not their
2709 tags. Potential INCOMPATIBILITY in exotic situations.
2711 * System option "pretty_margin" is superseded by "thy_output_margin",
2712 which is also accessible via document antiquotation option "margin".
2713 Only the margin for document output may be changed, but not the global
2714 pretty printing: that is 76 for plain console output, and adapted
2715 dynamically in GUI front-ends. Implementations of document
2716 antiquotations need to observe the margin explicitly according to
2717 Thy_Output.string_of_margin. Minor INCOMPATIBILITY.
2719 * Specification of 'document_files' in the session ROOT file is
2720 mandatory for document preparation. The legacy mode with implicit
2721 copying of the document/ directory is no longer supported. Minor
2727 * Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
2728 etc.) allow an optional context of local variables ('for' declaration):
2729 these variables become schematic in the instantiated theorem; this
2730 behaviour is analogous to 'for' in attributes "where" and "of".
2731 Configuration option rule_insts_schematic (default false) controls use
2732 of schematic variables outside the context. Minor INCOMPATIBILITY,
2733 declare rule_insts_schematic = true temporarily and update to use local
2734 variable declarations or dummy patterns instead.
2736 * Explicit instantiation via attributes "where", "of", and proof methods
2737 "rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns
2738 ("_") that stand for anonymous local variables.
2740 * Generated schematic variables in standard format of exported facts are
2741 incremented to avoid material in the proof context. Rare
2742 INCOMPATIBILITY, explicit instantiation sometimes needs to refer to
2745 * Lexical separation of signed and unsigned numerals: categories "num"
2746 and "float" are unsigned. INCOMPATIBILITY: subtle change in precedence
2747 of numeral signs, particularly in expressions involving infix syntax
2750 * Old inner token category "xnum" has been discontinued. Potential
2751 INCOMPATIBILITY for exotic syntax: may use mixfix grammar with "num"
2752 token category instead.
2757 * New (co)datatype package:
2758 - The 'datatype_new' command has been renamed 'datatype'. The old
2759 command of that name is now called 'old_datatype' and is provided
2760 by "~~/src/HOL/Library/Old_Datatype.thy". See
2761 'isabelle doc datatypes' for information on porting.
2764 disc_corec ~> corec_disc
2765 disc_corec_iff ~> corec_disc_iff
2766 disc_exclude ~> distinct_disc
2767 disc_exhaust ~> exhaust_disc
2768 disc_map_iff ~> map_disc_iff
2769 sel_corec ~> corec_sel
2770 sel_exhaust ~> exhaust_sel
2773 sel_split ~> split_sel
2774 sel_split_asm ~> split_sel_asm
2775 strong_coinduct ~> coinduct_strong
2776 weak_case_cong ~> case_cong_weak
2778 - The "no_code" option to "free_constructors", "datatype_new", and
2779 "codatatype" has been renamed "plugins del: code".
2781 - The rules "set_empty" have been removed. They are easy
2782 consequences of other set rules "by auto".
2784 - The rule "set_cases" is now registered with the "[cases set]"
2785 attribute. This can influence the behavior of the "cases" proof
2786 method when more than one case rule is applicable (e.g., an
2787 assumption is of the form "w : set ws" and the method "cases w"
2788 is invoked). The solution is to specify the case rule explicitly
2789 (e.g. "cases w rule: widget.exhaust").
2792 BNF_Comp ~> BNF_Composition
2793 BNF_FP_Base ~> BNF_Fixpoint_Base
2794 BNF_GFP ~> BNF_Greatest_Fixpoint
2795 BNF_LFP ~> BNF_Least_Fixpoint
2796 BNF_Constructions_on_Wellorders ~> BNF_Wellorder_Constructions
2797 Cardinals/Constructions_on_Wellorders ~> Cardinals/Wellorder_Constructions
2799 - Lifting and Transfer setup for basic HOL types sum and prod (also
2800 option) is now performed by the BNF package. Theories Lifting_Sum,
2801 Lifting_Product and Lifting_Option from Main became obsolete and
2802 were removed. Changed definitions of the relators rel_prod and
2803 rel_sum (using inductive).
2804 INCOMPATIBILITY: use rel_prod.simps and rel_sum.simps instead
2805 of rel_prod_def and rel_sum_def.
2806 Minor INCOMPATIBILITY: (rarely used by name) transfer theorem names
2807 changed (e.g. map_prod_transfer ~> prod.map_transfer).
2808 - Parametricity theorems for map functions, relators, set functions,
2809 constructors, case combinators, discriminators, selectors and
2810 (co)recursors are automatically proved and registered as transfer
2813 * Old datatype package:
2814 - The old 'datatype' command has been renamed 'old_datatype', and
2815 'rep_datatype' has been renamed 'old_rep_datatype'. They are
2816 provided by "~~/src/HOL/Library/Old_Datatype.thy". See
2817 'isabelle doc datatypes' for information on porting.
2820 weak_case_cong ~> case_cong_weak
2823 ~~/src/HOL/Datatype.thy ~> ~~/src/HOL/Library/Old_Datatype.thy
2827 - Fixed soundness bug related to the strict and non-strict subset
2831 - CVC4 is now included with Isabelle instead of CVC3 and run by
2833 - Z3 is now always enabled by default, now that it is fully open
2834 source. The "z3_non_commercial" option is discontinued.
2835 - Minimization is now always enabled by default.
2836 Removed sub-command:
2838 - Proof reconstruction, both one-liners and Isar, has been
2839 dramatically improved.
2840 - Improved support for CVC4 and veriT.
2842 * Old and new SMT modules:
2843 - The old 'smt' method has been renamed 'old_smt' and moved to
2844 'src/HOL/Library/Old_SMT.thy'. It is provided for compatibility,
2845 until applications have been ported to use the new 'smt' method. For
2846 the method to work, an older version of Z3 (e.g. Z3 3.2 or 4.0) must
2847 be installed, and the environment variable "OLD_Z3_SOLVER" must
2850 - The 'smt2' method has been renamed 'smt'.
2852 - New option 'smt_reconstruction_step_timeout' to limit the
2853 reconstruction time of Z3 proof steps in the new 'smt' method.
2854 - New option 'smt_statistics' to display statistics of the new 'smt'
2855 method, especially runtime statistics of Z3 proof reconstruction.
2857 * Lifting: command 'lift_definition' allows to execute lifted constants
2858 that have as a return type a datatype containing a subtype. This
2859 overcomes long-time limitations in the area of code generation and
2860 lifting, and avoids tedious workarounds.
2862 * Command and antiquotation "value" provide different evaluation slots
2863 (again), where the previous strategy (NBE after ML) serves as default.
2864 Minor INCOMPATIBILITY.
2866 * Add NO_MATCH-simproc, allows to check for syntactic non-equality.
2868 * field_simps: Use NO_MATCH-simproc for distribution rules, to avoid
2869 non-termination in case of distributing a division. With this change
2870 field_simps is in some cases slightly less powerful, if it fails try to
2871 add algebra_simps, or use divide_simps. Minor INCOMPATIBILITY.
2873 * Separate class no_zero_divisors has been given up in favour of fully
2874 algebraic semiring_no_zero_divisors. INCOMPATIBILITY.
2876 * Class linordered_semidom really requires no zero divisors.
2879 * Classes division_ring, field and linordered_field always demand
2880 "inverse 0 = 0". Given up separate classes division_ring_inverse_zero,
2881 field_inverse_zero and linordered_field_inverse_zero. INCOMPATIBILITY.
2883 * Classes cancel_ab_semigroup_add / cancel_monoid_add specify explicit
2884 additive inverse operation. INCOMPATIBILITY.
2886 * Complex powers and square roots. The functions "ln" and "powr" are now
2887 overloaded for types real and complex, and 0 powr y = 0 by definition.
2888 INCOMPATIBILITY: type constraints may be necessary.
2890 * The functions "sin" and "cos" are now defined for any type of sort
2891 "{real_normed_algebra_1,banach}" type, so in particular on "real" and
2892 "complex" uniformly. Minor INCOMPATIBILITY: type constraints may be
2895 * New library of properties of the complex transcendental functions sin,
2896 cos, tan, exp, Ln, Arctan, Arcsin, Arccos. Ported from HOL Light.
2898 * The factorial function, "fact", now has type "nat => 'a" (of a sort
2899 that admits numeric types including nat, int, real and complex.
2900 INCOMPATIBILITY: an expression such as "fact 3 = 6" may require a type
2901 constraint, and the combination "real (fact k)" is likely to be
2902 unsatisfactory. If a type conversion is still necessary, then use
2903 "of_nat (fact k)" or "real_of_nat (fact k)".
2905 * Removed functions "natfloor" and "natceiling", use "nat o floor" and
2906 "nat o ceiling" instead. A few of the lemmas have been retained and
2907 adapted: in their names "natfloor"/"natceiling" has been replaced by
2908 "nat_floor"/"nat_ceiling".
2910 * Qualified some duplicated fact names required for boostrapping the
2911 type class hierarchy:
2912 ab_add_uminus_conv_diff ~> diff_conv_add_uminus
2913 field_inverse_zero ~> inverse_zero
2914 field_divide_inverse ~> divide_inverse
2915 field_inverse ~> left_inverse
2916 Minor INCOMPATIBILITY.
2918 * Eliminated fact duplicates:
2919 mult_less_imp_less_right ~> mult_right_less_imp_less
2920 mult_less_imp_less_left ~> mult_left_less_imp_less
2921 Minor INCOMPATIBILITY.
2923 * Fact consolidation: even_less_0_iff is subsumed by
2924 double_add_less_zero_iff_single_add_less_zero (simp by default anyway).
2926 * Generalized and consolidated some theorems concerning divsibility:
2927 dvd_reduce ~> dvd_add_triv_right_iff
2928 dvd_plus_eq_right ~> dvd_add_right_iff
2929 dvd_plus_eq_left ~> dvd_add_left_iff
2930 Minor INCOMPATIBILITY.
2932 * "even" and "odd" are mere abbreviations for "2 dvd _" and "~ 2 dvd _"
2933 and part of theory Main.
2934 even_def ~> even_iff_mod_2_eq_zero
2937 * Lemma name consolidation: divide_Numeral1 ~> divide_numeral_1. Minor
2940 * Bootstrap of listsum as special case of abstract product over lists.
2942 listsum_def ~> listsum.eq_foldr
2945 * Product over lists via constant "listprod".
2947 * Theory List: renamed drop_Suc_conv_tl and nth_drop' to
2950 * New infrastructure for compiling, running, evaluating and testing
2951 generated code in target languages in HOL/Library/Code_Test. See
2952 HOL/Codegenerator_Test/Code_Test* for examples.
2955 - Introduced "replicate_mset" operation.
2956 - Introduced alternative characterizations of the multiset ordering in
2957 "Library/Multiset_Order".
2958 - Renamed multiset ordering:
2961 \<subset># ~> #\<subset>#
2962 \<subseteq># ~> #\<subseteq>#
2964 - Introduced abbreviations for ill-named multiset operations:
2965 <#, \<subset># abbreviate < (strict subset)
2966 <=#, \<le>#, \<subseteq># abbreviate <= (subset or equal)
2969 in_multiset_of ~> in_multiset_in_set
2970 Multiset.fold ~> fold_mset
2971 Multiset.filter ~> filter_mset
2973 - Removed mcard, is equal to size.
2975 image_mset.id [simp]
2976 image_mset_id [simp]
2977 elem_multiset_of_set [simp, intro]
2978 comp_fun_commute_plus_mset [simp]
2979 comp_fun_commute.fold_mset_insert [OF comp_fun_commute_plus_mset, simp]
2980 in_mset_fold_plus_iff [iff]
2981 set_of_Union_mset [simp]
2982 in_Union_mset_iff [iff]
2985 * Library/Sum_of_Squares: simplified and improved "sos" method. Always
2986 use local CSDP executable, which is much faster than the NEOS server.
2987 The "sos_cert" functionality is invoked as "sos" with additional
2988 argument. Minor INCOMPATIBILITY.
2990 * HOL-Decision_Procs: New counterexample generator quickcheck
2991 [approximation] for inequalities of transcendental functions. Uses
2992 hardware floating point arithmetic to randomly discover potential
2993 counterexamples. Counterexamples are certified with the "approximation"
2994 method. See HOL/Decision_Procs/ex/Approximation_Quickcheck_Ex.thy for
2997 * HOL-Probability: Reworked measurability prover
2998 - applies destructor rules repeatedly
2999 - removed application splitting (replaced by destructor rule)
3000 - added congruence rules to rewrite measure spaces under the sets
3003 * New proof method "rewrite" (in theory ~~/src/HOL/Library/Rewrite) for
3004 single-step rewriting with subterm selection based on patterns.
3009 * Subtle change of name space policy: undeclared entries are now
3010 considered inaccessible, instead of accessible via the fully-qualified
3011 internal name. This mainly affects Name_Space.intern (and derivatives),
3012 which may produce an unexpected Long_Name.hidden prefix. Note that
3013 contemporary applications use the strict Name_Space.check (and
3014 derivatives) instead, which is not affected by the change. Potential
3015 INCOMPATIBILITY in rare applications of Name_Space.intern.
3017 * Subtle change of error semantics of Toplevel.proof_of: regular user
3018 ERROR instead of internal Toplevel.UNDEF.
3020 * Basic combinators map, fold, fold_map, split_list, apply are available
3021 as parameterized antiquotations, e.g. @{map 4} for lists of quadruples.
3023 * Renamed "pairself" to "apply2", in accordance to @{apply 2}.
3026 * Former combinators NAMED_CRITICAL and CRITICAL for central critical
3027 sections have been discontinued, in favour of the more elementary
3028 Multithreading.synchronized and its high-level derivative
3029 Synchronized.var (which is usually sufficient in applications). Subtle
3030 INCOMPATIBILITY: synchronized access needs to be atomic and cannot be
3033 * Synchronized.value (ML) is actually synchronized (as in Scala): subtle
3034 change of semantics with minimal potential for INCOMPATIBILITY.
3036 * The main operations to certify logical entities are Thm.ctyp_of and
3037 Thm.cterm_of with a local context; old-style global theory variants are
3038 available as Thm.global_ctyp_of and Thm.global_cterm_of.
3041 * Elementary operations in module Thm are no longer pervasive.
3042 INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of,
3045 * Proper context for various elementary tactics: assume_tac,
3046 resolve_tac, eresolve_tac, dresolve_tac, forward_tac, match_tac,
3047 compose_tac, Splitter.split_tac etc. INCOMPATIBILITY.
3049 * Tactical PARALLEL_ALLGOALS is the most common way to refer to
3052 * Goal.prove_multi is superseded by the fully general Goal.prove_common,
3053 which also allows to specify a fork priority.
3055 * Antiquotation @{command_spec "COMMAND"} is superseded by
3056 @{command_keyword COMMAND} (usually without quotes and with PIDE
3057 markup). Minor INCOMPATIBILITY.
3059 * Cartouches within ML sources are turned into values of type
3060 Input.source (with formal position information).
3065 * The Isabelle tool "update_cartouches" changes theory files to use
3066 cartouches instead of old-style {* verbatim *} or `alt_string` tokens.
3068 * The Isabelle tool "build" provides new options -X, -k, -x.
3070 * Discontinued old-fashioned "codegen" tool. Code generation can always
3071 be externally triggered using an appropriate ROOT file plus a
3072 corresponding theory. Parametrization is possible using environment
3073 variables, or ML snippets in the most extreme cases. Minor
3076 * JVM system property "isabelle.threads" determines size of Scala thread
3077 pool, like Isabelle system option "threads" for ML.
3079 * JVM system property "isabelle.laf" determines the default Swing
3080 look-and-feel, via internal class name or symbolic name as in the jEdit
3081 menu Global Options / Appearance.
3083 * Support for Proof General and Isar TTY loop has been discontinued.
3084 Minor INCOMPATIBILITY, use standard PIDE infrastructure instead.
3088 New in Isabelle2014 (August 2014)
3089 ---------------------------------
3093 * Support for official Standard ML within the Isabelle context.
3094 Command 'SML_file' reads and evaluates the given Standard ML file.
3095 Toplevel bindings are stored within the theory context; the initial
3096 environment is restricted to the Standard ML implementation of
3097 Poly/ML, without the add-ons of Isabelle/ML. Commands 'SML_import'
3098 and 'SML_export' allow to exchange toplevel bindings between the two
3099 separate environments. See also ~~/src/Tools/SML/Examples.thy for
3102 * Standard tactics and proof methods such as "clarsimp", "auto" and
3103 "safe" now preserve equality hypotheses "x = expr" where x is a free
3104 variable. Locale assumptions and chained facts containing "x"
3105 continue to be useful. The new method "hypsubst_thin" and the
3106 configuration option "hypsubst_thin" (within the attribute name space)
3107 restore the previous behavior. INCOMPATIBILITY, especially where
3108 induction is done after these methods or when the names of free and
3109 bound variables clash. As first approximation, old proofs may be
3110 repaired by "using [[hypsubst_thin = true]]" in the critical spot.
3112 * More static checking of proof methods, which allows the system to
3113 form a closure over the concrete syntax. Method arguments should be
3114 processed in the original proof context as far as possible, before
3115 operating on the goal state. In any case, the standard discipline for
3116 subgoal-addressing needs to be observed: no subgoals or a subgoal
3117 number that is out of range produces an empty result sequence, not an
3118 exception. Potential INCOMPATIBILITY for non-conformant tactical
3121 * Lexical syntax (inner and outer) supports text cartouches with
3122 arbitrary nesting, and without escapes of quotes etc. The Prover IDE
3123 supports input via ` (backquote).
3125 * The outer syntax categories "text" (for formal comments and document
3126 markup commands) and "altstring" (for literal fact references) allow
3127 cartouches as well, in addition to the traditional mix of quotations.
3129 * Syntax of document antiquotation @{rail} now uses \<newline> instead
3130 of "\\", to avoid the optical illusion of escaped backslash within
3131 string token. General renovation of its syntax using text cartouches.
3132 Minor INCOMPATIBILITY.
3134 * Discontinued legacy_isub_isup, which was a temporary workaround for
3135 Isabelle/ML in Isabelle2013-1. The prover process no longer accepts
3136 old identifier syntax with \<^isub> or \<^isup>. Potential
3139 * Document antiquotation @{url} produces markup for the given URL,
3140 which results in an active hyperlink within the text.
3142 * Document antiquotation @{file_unchecked} is like @{file}, but does
3143 not check existence within the file-system.
3145 * Updated and extended manuals: codegen, datatypes, implementation,
3146 isar-ref, jedit, system.
3149 *** Prover IDE -- Isabelle/Scala/jEdit ***
3151 * Improved Document panel: simplified interaction where every single
3152 mouse click (re)opens document via desktop environment or as jEdit
3155 * Support for Navigator plugin (with toolbar buttons), with connection
3158 * Auxiliary files ('ML_file' etc.) are managed by the Prover IDE.
3159 Open text buffers take precedence over copies within the file-system.
3161 * Improved support for Isabelle/ML, with jEdit mode "isabelle-ml" for
3164 * Improved syntactic and semantic completion mechanism, with simple
3165 templates, completion language context, name-space completion,
3166 file-name completion, spell-checker completion.
3168 * Refined GUI popup for completion: more robust key/mouse event
3169 handling and propagation to enclosing text area -- avoid loosing
3170 keystrokes with slow / remote graphics displays.
3172 * Completion popup supports both ENTER and TAB (default) to select an
3173 item, depending on Isabelle options.
3175 * Refined insertion of completion items wrt. jEdit text: multiple
3176 selections, rectangular selections, rectangular selection as "tall
3179 * Integrated spell-checker for document text, comments etc. with
3180 completion popup and context-menu.
3182 * More general "Query" panel supersedes "Find" panel, with GUI access
3183 to commands 'find_theorems' and 'find_consts', as well as print
3184 operations for the context. Minor incompatibility in keyboard
3185 shortcuts etc.: replace action isabelle-find by isabelle-query.
3187 * Search field for all output panels ("Output", "Query", "Info" etc.)
3188 to highlight text via regular expression.
3190 * Option "jedit_print_mode" (see also "Plugin Options / Isabelle /
3191 General") allows to specify additional print modes for the prover
3192 process, without requiring old-fashioned command-line invocation of
3193 "isabelle jedit -m MODE".
3195 * More support for remote files (e.g. http) using standard Java
3196 networking operations instead of jEdit virtual file-systems.
3198 * Empty editors buffers that are no longer required (e.g.\ via theory
3199 imports) are automatically removed from the document model.
3201 * Improved monitor panel.
3203 * Improved Console/Scala plugin: more uniform scala.Console output,
3204 more robust treatment of threads and interrupts.
3206 * Improved management of dockable windows: clarified keyboard focus
3207 and window placement wrt. main editor view; optional menu item to
3208 "Detach" a copy where this makes sense.
3210 * New Simplifier Trace panel provides an interactive view of the
3211 simplification process, enabled by the "simp_trace_new" attribute
3217 * Low-level type-class commands 'classes', 'classrel', 'arities' have
3218 been discontinued to avoid the danger of non-trivial axiomatization
3219 that is not immediately visible. INCOMPATIBILITY, use regular
3220 'instance' command with proof. The required OFCLASS(...) theorem
3221 might be postulated via 'axiomatization' beforehand, or the proof
3222 finished trivially if the underlying class definition is made vacuous
3223 (without any assumptions). See also Isabelle/ML operations
3224 Axclass.class_axiomatization, Axclass.classrel_axiomatization,
3225 Axclass.arity_axiomatization.
3227 * Basic constants of Pure use more conventional names and are always
3228 qualified. Rare INCOMPATIBILITY, but with potentially serious
3229 consequences, notably for tools in Isabelle/ML. The following
3230 renaming needs to be applied:
3236 dummy_pattern ~> Pure.dummy_pattern
3238 Systematic porting works by using the following theory setup on a
3239 *previous* Isabelle version to introduce the new name accesses for the
3245 |> Sign.const_alias (Binding.qualify true "Pure" @{binding eq}) "=="
3246 |> Sign.const_alias (Binding.qualify true "Pure" @{binding imp}) "==>"
3247 |> Sign.const_alias (Binding.qualify true "Pure" @{binding all}) "all"
3248 |> Sign.restore_naming thy
3251 Thus ML antiquotations like @{const_name Pure.eq} may be used already.
3252 Later the application is moved to the current Isabelle version, and
3253 the auxiliary aliases are deleted.
3255 * Attributes "where" and "of" allow an optional context of local
3256 variables ('for' declaration): these variables become schematic in the
3257 instantiated theorem.
3259 * Obsolete attribute "standard" has been discontinued (legacy since
3260 Isabelle2012). Potential INCOMPATIBILITY, use explicit 'for' context
3261 where instantiations with schematic variables are intended (for
3262 declaration commands like 'lemmas' or attributes like "of"). The
3263 following temporary definition may help to port old applications:
3265 attribute_setup standard =
3266 "Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))"
3268 * More thorough check of proof context for goal statements and
3269 attributed fact expressions (concerning background theory, declared
3270 hyps). Potential INCOMPATIBILITY, tools need to observe standard
3271 context discipline. See also Assumption.add_assumes and the more
3272 primitive Thm.assume_hyps.
3274 * Inner syntax token language allows regular quoted strings "..."
3275 (only makes sense in practice, if outer syntax is delimited
3276 differently, e.g. via cartouches).
3278 * Command 'print_term_bindings' supersedes 'print_binds' for clarity,
3279 but the latter is retained some time as Proof General legacy.
3281 * Code generator preprocessor: explicit control of simp tracing on a
3282 per-constant basis. See attribute "code_preproc".
3287 * Code generator: enforce case of identifiers only for strict target
3288 language requirements. INCOMPATIBILITY.
3290 * Code generator: explicit proof contexts in many ML interfaces.
3293 * Code generator: minimize exported identifiers by default. Minor
3296 * Code generation for SML and OCaml: dropped arcane "no_signatures"
3297 option. Minor INCOMPATIBILITY.
3299 * "declare [[code abort: ...]]" replaces "code_abort ...".
3302 * "declare [[code drop: ...]]" drops all code equations associated
3303 with the given constants.
3305 * Code generations are provided for make, fields, extend and truncate
3306 operations on records.
3308 * Command and antiquotation "value" are now hardcoded against nbe and
3309 ML. Minor INCOMPATIBILITY.
3311 * Renamed command 'enriched_type' to 'functor'. INCOMPATIBILITY.
3313 * The symbol "\<newline>" may be used within char or string literals
3314 to represent (Char Nibble0 NibbleA), i.e. ASCII newline.
3316 * Qualified String.implode and String.explode. INCOMPATIBILITY.
3318 * Simplifier: Enhanced solver of preconditions of rewrite rules can
3319 now deal with conjunctions. For help with converting proofs, the old
3320 behaviour of the simplifier can be restored like this: declare/using
3321 [[simp_legacy_precond]]. This configuration option will disappear
3322 again in the future. INCOMPATIBILITY.
3324 * Simproc "finite_Collect" is no longer enabled by default, due to
3325 spurious crashes and other surprises. Potential INCOMPATIBILITY.
3327 * Moved new (co)datatype package and its dependencies from session
3328 "HOL-BNF" to "HOL". The commands 'bnf', 'wrap_free_constructors',
3329 'datatype_new', 'codatatype', 'primcorec', 'primcorecursive' are now
3330 part of theory "Main".
3333 FunDef.thy ~> Fun_Def.thy (and Fun_Def_Base.thy)
3334 Library/Wfrec.thy ~> Wfrec.thy
3335 Library/Zorn.thy ~> Zorn.thy
3336 Cardinals/Order_Relation.thy ~> Order_Relation.thy
3337 Library/Order_Union.thy ~> Cardinals/Order_Union.thy
3338 Cardinals/Cardinal_Arithmetic_Base.thy ~> BNF_Cardinal_Arithmetic.thy
3339 Cardinals/Cardinal_Order_Relation_Base.thy ~> BNF_Cardinal_Order_Relation.thy
3340 Cardinals/Constructions_on_Wellorders_Base.thy ~> BNF_Constructions_on_Wellorders.thy
3341 Cardinals/Wellorder_Embedding_Base.thy ~> BNF_Wellorder_Embedding.thy
3342 Cardinals/Wellorder_Relation_Base.thy ~> BNF_Wellorder_Relation.thy
3343 BNF/Ctr_Sugar.thy ~> Ctr_Sugar.thy
3344 BNF/Basic_BNFs.thy ~> Basic_BNFs.thy
3345 BNF/BNF_Comp.thy ~> BNF_Comp.thy
3346 BNF/BNF_Def.thy ~> BNF_Def.thy
3347 BNF/BNF_FP_Base.thy ~> BNF_FP_Base.thy
3348 BNF/BNF_GFP.thy ~> BNF_GFP.thy
3349 BNF/BNF_LFP.thy ~> BNF_LFP.thy
3350 BNF/BNF_Util.thy ~> BNF_Util.thy
3351 BNF/Coinduction.thy ~> Coinduction.thy
3352 BNF/More_BNFs.thy ~> Library/More_BNFs.thy
3353 BNF/Countable_Type.thy ~> Library/Countable_Set_Type.thy
3354 BNF/Examples/* ~> BNF_Examples/*
3357 Wellorder_Extension.thy (split from Zorn.thy)
3358 Library/Cardinal_Notations.thy
3359 Library/BNF_Axomatization.thy
3360 BNF_Examples/Misc_Primcorec.thy
3361 BNF_Examples/Stream_Processor.thy
3363 Discontinued theories:
3365 BNF/Equiv_Relations_More.thy
3369 * New (co)datatype package:
3370 - Command 'primcorec' is fully implemented.
3371 - Command 'datatype_new' generates size functions ("size_xxx" and
3372 "size") as required by 'fun'.
3373 - BNFs are integrated with the Lifting tool and new-style
3374 (co)datatypes with Transfer.
3376 datatype_new_compat ~> datatype_compat
3377 primrec_new ~> primrec
3378 wrap_free_constructors ~> free_constructors
3380 - The generated constants "xxx_case" and "xxx_rec" have been renamed
3381 "case_xxx" and "rec_xxx" (e.g., "prod_case" ~> "case_prod").
3383 - The constant "xxx_(un)fold" and related theorems are no longer
3384 generated. Use "xxx_(co)rec" or define "xxx_(un)fold" manually
3385 using "prim(co)rec".
3387 - No discriminators are generated for nullary constructors by
3388 default, eliminating the need for the odd "=:" syntax.
3390 - No discriminators or selectors are generated by default by
3391 "datatype_new", unless custom names are specified or the new
3392 "discs_sels" option is passed.
3395 * Old datatype package:
3396 - The generated theorems "xxx.cases" and "xxx.recs" have been
3397 renamed "xxx.case" and "xxx.rec" (e.g., "sum.cases" ->
3398 "sum.case"). INCOMPATIBILITY.
3399 - The generated constants "xxx_case", "xxx_rec", and "xxx_size" have
3400 been renamed "case_xxx", "rec_xxx", and "size_xxx" (e.g.,
3401 "prod_case" ~> "case_prod"). INCOMPATIBILITY.
3403 * The types "'a list" and "'a option", their set and map functions,
3404 their relators, and their selectors are now produced using the new
3405 BNF-based datatype package.
3408 Option.set ~> set_option
3409 Option.map ~> map_option
3410 option_rel ~> rel_option
3413 set_def ~> set_rec[abs_def]
3414 map_def ~> map_rec[abs_def]
3415 Option.map_def ~> map_option_case[abs_def] (with "case_option" instead of "rec_option")
3416 option.recs ~> option.rec
3417 list_all2_def ~> list_all2_iff
3418 set.simps ~> set_simps (or the slightly different "list.set")
3419 map.simps ~> list.map
3420 hd.simps ~> list.sel(1)
3421 tl.simps ~> list.sel(2-3)
3422 the.simps ~> option.sel
3426 * The following map functions and relators have been renamed:
3428 map_pair ~> map_prod
3429 prod_rel ~> rel_prod
3433 filter_rel ~> rel_filter
3434 fset_rel ~> rel_fset (in "src/HOL/Library/FSet.thy")
3435 cset_rel ~> rel_cset (in "src/HOL/Library/Countable_Set_Type.thy")
3436 vset ~> rel_vset (in "src/HOL/Library/Quotient_Set.thy")
3440 * Lifting and Transfer:
3441 - a type variable as a raw type is supported
3442 - stronger reflexivity prover
3443 - rep_eq is always generated by lift_definition
3444 - setup for Lifting/Transfer is now automated for BNFs
3445 + holds for BNFs that do not contain a dead variable
3446 + relator_eq, relator_mono, relator_distr, relator_domain,
3447 relator_eq_onp, quot_map, transfer rules for bi_unique, bi_total,
3448 right_unique, right_total, left_unique, left_total are proved
3450 + definition of a predicator is generated automatically
3451 + simplification rules for a predicator definition are proved
3452 automatically for datatypes
3453 - consolidation of the setup of Lifting/Transfer
3454 + property that a relator preservers reflexivity is not needed any
3456 Minor INCOMPATIBILITY.
3457 + left_total and left_unique rules are now transfer rules
3458 (reflexivity_rule attribute not needed anymore)
3460 + Domainp does not have to be a separate assumption in
3461 relator_domain theorems (=> more natural statement)
3463 - registration of code equations is more robust
3464 Potential INCOMPATIBILITY.
3465 - respectfulness proof obligation is preprocessed to a more readable
3467 Potential INCOMPATIBILITY.
3468 - eq_onp is always unfolded in respectfulness proof obligation
3469 Potential INCOMPATIBILITY.
3470 - unregister lifting setup for Code_Numeral.integer and
3471 Code_Numeral.natural
3472 Potential INCOMPATIBILITY.
3473 - Lifting.invariant -> eq_onp
3476 * New internal SAT solver "cdclite" that produces models and proof
3477 traces. This solver replaces the internal SAT solvers "enumerate" and
3478 "dpll". Applications that explicitly used one of these two SAT
3479 solvers should use "cdclite" instead. In addition, "cdclite" is now
3480 the default SAT solver for the "sat" and "satx" proof methods and
3481 corresponding tactics; the old default can be restored using "declare
3482 [[sat_solver = zchaff_with_proofs]]". Minor INCOMPATIBILITY.
3484 * SMT module: A new version of the SMT module, temporarily called
3485 "SMT2", uses SMT-LIB 2 and supports recent versions of Z3 (e.g.,
3486 4.3). The new proof method is called "smt2". CVC3 and CVC4 are also
3487 supported as oracles. Yices is no longer supported, because no version
3488 of the solver can handle both SMT-LIB 2 and quantifiers.
3490 * Activation of Z3 now works via "z3_non_commercial" system option
3491 (without requiring restart), instead of former settings variable
3492 "Z3_NON_COMMERCIAL". The option can be edited in Isabelle/jEdit menu
3493 Plugin Options / Isabelle / General.
3496 - Z3 can now produce Isar proofs.
3498 . New SML-based learning algorithms eliminate the dependency on
3499 Python and increase performance and reliability.
3500 . MaSh and MeSh are now used by default together with the
3501 traditional MePo (Meng-Paulson) relevance filter. To disable
3502 MaSh, set the "MaSh" system option in Isabelle/jEdit Plugin
3503 Options / Isabelle / General to "none".
3507 isar_compress ~> compress
3512 * Removed solvers remote_cvc3 and remote_z3. Use cvc3 and z3 instead.
3515 - Fixed soundness bug whereby mutually recursive datatypes could
3516 take infinite values.
3517 - Fixed soundness bug with low-level number functions such as
3518 "Abs_Integ" and "Rep_Integ".
3519 - Removed "std" option.
3520 - Renamed "show_datatypes" to "show_types" and "hide_datatypes" to
3523 * Metis: Removed legacy proof method 'metisFT'. Use 'metis
3524 (full_types)' instead. INCOMPATIBILITY.
3526 * Try0: Added 'algebra' and 'meson' to the set of proof methods.
3528 * Adjustion of INF and SUP operations:
3529 - Elongated constants INFI and SUPR to INFIMUM and SUPREMUM.
3530 - Consolidated theorem names containing INFI and SUPR: have INF and
3531 SUP instead uniformly.
3532 - More aggressive normalization of expressions involving INF and Inf
3534 - INF_image and SUP_image do not unfold composition.
3535 - Dropped facts INF_comp, SUP_comp.
3536 - Default congruence rules strong_INF_cong and strong_SUP_cong, with
3537 simplifier implication in premises. Generalize and replace former
3542 * SUP and INF generalized to conditionally_complete_lattice.
3544 * Swapped orientation of facts image_comp and vimage_comp:
3546 image_compose ~> image_comp [symmetric]
3547 image_comp ~> image_comp [symmetric]
3548 vimage_compose ~> vimage_comp [symmetric]
3549 vimage_comp ~> vimage_comp [symmetric]
3553 * Theory reorganization: split of Big_Operators.thy into
3554 Groups_Big.thy and Lattices_Big.thy.
3556 * Consolidated some facts about big group operators:
3558 setsum_0' ~> setsum.neutral
3559 setsum_0 ~> setsum.neutral_const
3560 setsum_addf ~> setsum.distrib
3561 setsum_cartesian_product ~> setsum.cartesian_product
3562 setsum_cases ~> setsum.If_cases
3563 setsum_commute ~> setsum.commute
3564 setsum_cong ~> setsum.cong
3565 setsum_delta ~> setsum.delta
3566 setsum_delta' ~> setsum.delta'
3567 setsum_diff1' ~> setsum.remove
3568 setsum_empty ~> setsum.empty
3569 setsum_infinite ~> setsum.infinite
3570 setsum_insert ~> setsum.insert
3571 setsum_inter_restrict'' ~> setsum.inter_filter
3572 setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left
3573 setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right
3574 setsum_mono_zero_left ~> setsum.mono_neutral_left
3575 setsum_mono_zero_right ~> setsum.mono_neutral_right
3576 setsum_reindex ~> setsum.reindex
3577 setsum_reindex_cong ~> setsum.reindex_cong
3578 setsum_reindex_nonzero ~> setsum.reindex_nontrivial
3579 setsum_restrict_set ~> setsum.inter_restrict
3580 setsum_Plus ~> setsum.Plus
3581 setsum_setsum_restrict ~> setsum.commute_restrict
3582 setsum_Sigma ~> setsum.Sigma
3583 setsum_subset_diff ~> setsum.subset_diff
3584 setsum_Un_disjoint ~> setsum.union_disjoint
3585 setsum_UN_disjoint ~> setsum.UNION_disjoint
3586 setsum_Un_Int ~> setsum.union_inter
3587 setsum_Union_disjoint ~> setsum.Union_disjoint
3588 setsum_UNION_zero ~> setsum.Union_comp
3589 setsum_Un_zero ~> setsum.union_inter_neutral
3590 strong_setprod_cong ~> setprod.strong_cong
3591 strong_setsum_cong ~> setsum.strong_cong
3592 setprod_1' ~> setprod.neutral
3593 setprod_1 ~> setprod.neutral_const
3594 setprod_cartesian_product ~> setprod.cartesian_product
3595 setprod_cong ~> setprod.cong
3596 setprod_delta ~> setprod.delta
3597 setprod_delta' ~> setprod.delta'
3598 setprod_empty ~> setprod.empty
3599 setprod_infinite ~> setprod.infinite
3600 setprod_insert ~> setprod.insert
3601 setprod_mono_one_cong_left ~> setprod.mono_neutral_cong_left
3602 setprod_mono_one_cong_right ~> setprod.mono_neutral_cong_right
3603 setprod_mono_one_left ~> setprod.mono_neutral_left
3604 setprod_mono_one_right ~> setprod.mono_neutral_right
3605 setprod_reindex ~> setprod.reindex
3606 setprod_reindex_cong ~> setprod.reindex_cong
3607 setprod_reindex_nonzero ~> setprod.reindex_nontrivial
3608 setprod_Sigma ~> setprod.Sigma
3609 setprod_subset_diff ~> setprod.subset_diff
3610 setprod_timesf ~> setprod.distrib
3611 setprod_Un2 ~> setprod.union_diff2
3612 setprod_Un_disjoint ~> setprod.union_disjoint
3613 setprod_UN_disjoint ~> setprod.UNION_disjoint
3614 setprod_Un_Int ~> setprod.union_inter
3615 setprod_Union_disjoint ~> setprod.Union_disjoint
3616 setprod_Un_one ~> setprod.union_inter_neutral
3618 Dropped setsum_cong2 (simple variant of setsum.cong).
3619 Dropped setsum_inter_restrict' (simple variant of setsum.inter_restrict)
3620 Dropped setsum_reindex_id, setprod_reindex_id
3621 (simple variants of setsum.reindex [symmetric], setprod.reindex [symmetric]).
3625 * Abolished slightly odd global lattice interpretation for min/max.
3627 Fact consolidations:
3628 min_max.inf_assoc ~> min.assoc
3629 min_max.inf_commute ~> min.commute
3630 min_max.inf_left_commute ~> min.left_commute
3631 min_max.inf_idem ~> min.idem
3632 min_max.inf_left_idem ~> min.left_idem
3633 min_max.inf_right_idem ~> min.right_idem
3634 min_max.sup_assoc ~> max.assoc
3635 min_max.sup_commute ~> max.commute
3636 min_max.sup_left_commute ~> max.left_commute
3637 min_max.sup_idem ~> max.idem
3638 min_max.sup_left_idem ~> max.left_idem
3639 min_max.sup_inf_distrib1 ~> max_min_distrib2
3640 min_max.sup_inf_distrib2 ~> max_min_distrib1
3641 min_max.inf_sup_distrib1 ~> min_max_distrib2
3642 min_max.inf_sup_distrib2 ~> min_max_distrib1
3643 min_max.distrib ~> min_max_distribs
3644 min_max.inf_absorb1 ~> min.absorb1
3645 min_max.inf_absorb2 ~> min.absorb2
3646 min_max.sup_absorb1 ~> max.absorb1
3647 min_max.sup_absorb2 ~> max.absorb2
3648 min_max.le_iff_inf ~> min.absorb_iff1
3649 min_max.le_iff_sup ~> max.absorb_iff2
3650 min_max.inf_le1 ~> min.cobounded1
3651 min_max.inf_le2 ~> min.cobounded2
3652 le_maxI1, min_max.sup_ge1 ~> max.cobounded1
3653 le_maxI2, min_max.sup_ge2 ~> max.cobounded2
3654 min_max.le_infI1 ~> min.coboundedI1
3655 min_max.le_infI2 ~> min.coboundedI2
3656 min_max.le_supI1 ~> max.coboundedI1
3657 min_max.le_supI2 ~> max.coboundedI2
3658 min_max.less_infI1 ~> min.strict_coboundedI1
3659 min_max.less_infI2 ~> min.strict_coboundedI2
3660 min_max.less_supI1 ~> max.strict_coboundedI1
3661 min_max.less_supI2 ~> max.strict_coboundedI2
3662 min_max.inf_mono ~> min.mono
3663 min_max.sup_mono ~> max.mono
3664 min_max.le_infI, min_max.inf_greatest ~> min.boundedI
3665 min_max.le_supI, min_max.sup_least ~> max.boundedI
3666 min_max.le_inf_iff ~> min.bounded_iff
3667 min_max.le_sup_iff ~> max.bounded_iff
3669 For min_max.inf_sup_aci, prefer (one of) min.commute, min.assoc,
3670 min.left_commute, min.left_idem, max.commute, max.assoc,
3671 max.left_commute, max.left_idem directly.
3673 For min_max.inf_sup_ord, prefer (one of) min.cobounded1,
3674 min.cobounded2, max.cobounded1m max.cobounded2 directly.
3676 For min_ac or max_ac, prefer more general collection ac_simps.
3680 * Theorem disambiguation Inf_le_Sup (on finite sets) ~>
3681 Inf_fin_le_Sup_fin. INCOMPATIBILITY.
3683 * Qualified constant names Wellfounded.acc, Wellfounded.accp.
3686 * Fact generalization and consolidation:
3687 neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1
3691 * Purely algebraic definition of even. Fact generalization and
3693 nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd
3694 even_zero_(nat|int) ~> even_zero
3698 * Abolished neg_numeral.
3699 - Canonical representation for minus one is "- 1".
3700 - Canonical representation for other negative numbers is "- (numeral _)".
3701 - When devising rule sets for number calculation, consider the
3702 following canonical cases: 0, 1, numeral _, - 1, - numeral _.
3703 - HOLogic.dest_number also recognizes numerals in non-canonical forms
3704 like "numeral One", "- numeral One", "- 0" and even "- ... - _".
3705 - Syntax for negative numerals is mere input syntax.
3709 * Reduced name variants for rules on associativity and commutativity:
3711 add_assoc ~> add.assoc
3712 add_commute ~> add.commute
3713 add_left_commute ~> add.left_commute
3714 mult_assoc ~> mult.assoc
3715 mult_commute ~> mult.commute
3716 mult_left_commute ~> mult.left_commute
3717 nat_add_assoc ~> add.assoc
3718 nat_add_commute ~> add.commute
3719 nat_add_left_commute ~> add.left_commute
3720 nat_mult_assoc ~> mult.assoc
3721 nat_mult_commute ~> mult.commute
3722 eq_assoc ~> iff_assoc
3723 eq_left_commute ~> iff_left_commute
3727 * Fact collections add_ac and mult_ac are considered old-fashioned.
3728 Prefer ac_simps instead, or specify rules
3729 (add|mult).(assoc|commute|left_commute) individually.
3731 * Elimination of fact duplicates:
3732 equals_zero_I ~> minus_unique
3733 diff_eq_0_iff_eq ~> right_minus_eq
3734 nat_infinite ~> infinite_UNIV_nat
3735 int_infinite ~> infinite_UNIV_int
3739 * Fact name consolidation:
3740 diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
3741 minus_le_self_iff ~> neg_less_eq_nonneg
3742 le_minus_self_iff ~> less_eq_neg_nonpos
3743 neg_less_nonneg ~> neg_less_pos
3744 less_minus_self_iff ~> less_neg_neg [simp]
3748 * More simplification rules on unary and binary minus:
3749 add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
3750 add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
3751 add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
3752 le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
3753 minus_add_cancel, uminus_add_conv_diff. These correspondingly have
3754 been taken away from fact collections algebra_simps and field_simps.
3757 To restore proofs, the following patterns are helpful:
3759 a) Arbitrary failing proof not involving "diff_def":
3760 Consider simplification with algebra_simps or field_simps.
3762 b) Lifting rules from addition to subtraction:
3763 Try with "using <rule for addition> of [... "- _" ...]" by simp".
3765 c) Simplification with "diff_def": just drop "diff_def".
3766 Consider simplification with algebra_simps or field_simps;
3767 or the brute way with
3768 "simp add: diff_conv_add_uminus del: add_uminus_conv_diff".
3770 * Introduce bdd_above and bdd_below in theory
3771 Conditionally_Complete_Lattices, use them instead of explicitly
3772 stating boundedness of sets.
3774 * ccpo.admissible quantifies only over non-empty chains to allow more
3775 syntax-directed proof rules; the case of the empty chain shows up as
3776 additional case in fixpoint induction proofs. INCOMPATIBILITY.
3778 * Removed and renamed theorems in Series:
3779 summable_le ~> suminf_le
3780 suminf_le ~> suminf_le_const
3781 series_pos_le ~> setsum_le_suminf
3782 series_pos_less ~> setsum_less_suminf
3783 suminf_ge_zero ~> suminf_nonneg
3784 suminf_gt_zero ~> suminf_pos
3785 suminf_gt_zero_iff ~> suminf_pos_iff
3786 summable_sumr_LIMSEQ_suminf ~> summable_LIMSEQ
3787 suminf_0_le ~> suminf_nonneg [rotate]
3788 pos_summable ~> summableI_nonneg_bounded
3789 ratio_test ~> summable_ratio_test
3791 removed series_zero, replaced by sums_finite
3793 removed auxiliary lemmas:
3795 sumr_offset, sumr_offset2, sumr_offset3, sumr_offset4, sumr_group,
3796 half, le_Suc_ex_iff, lemma_realpow_diff_sumr,
3797 real_setsum_nat_ivl_bounded, summable_le2, ratio_test_lemma2,
3798 sumr_minus_one_realpow_zerom, sumr_one_lb_realpow_zero,
3799 summable_convergent_sumr_iff, sumr_diff_mult_const
3803 * Replace (F)DERIV syntax by has_derivative:
3804 - "(f has_derivative f') (at x within s)" replaces "FDERIV f x : s : f'"
3806 - "(f has_field_derivative f') (at x within s)" replaces "DERIV f x : s : f'"
3808 - "f differentiable at x within s" replaces "_ differentiable _ in _" syntax
3810 - removed constant isDiff
3812 - "DERIV f x : f'" and "FDERIV f x : f'" syntax is only available as
3815 - "DERIV f x : s : f'" and "FDERIV f x : s : f'" syntax removed.
3817 - Renamed FDERIV_... lemmas to has_derivative_...
3819 - renamed deriv (the syntax constant used for "DERIV _ _ :> _") to DERIV
3821 - removed DERIV_intros, has_derivative_eq_intros
3823 - introduced derivative_intros and deriative_eq_intros which
3824 includes now rules for DERIV, has_derivative and
3825 has_vector_derivative.
3828 differentiable_def ~> real_differentiable_def
3829 differentiableE ~> real_differentiableE
3830 fderiv_def ~> has_derivative_at
3831 field_fderiv_def ~> field_has_derivative_at
3832 isDiff_der ~> differentiable_def
3833 deriv_fderiv ~> has_field_derivative_def
3834 deriv_def ~> DERIV_def
3838 * Include more theorems in continuous_intros. Remove the
3839 continuous_on_intros, isCont_intros collections, these facts are now
3840 in continuous_intros.
3842 * Theorems about complex numbers are now stated only using Re and Im,
3843 the Complex constructor is not used anymore. It is possible to use
3844 primcorec to defined the behaviour of a complex-valued function.
3846 Removed theorems about the Complex constructor from the simpset, they
3847 are available as the lemma collection legacy_Complex_simps. This
3850 i_complex_of_real: "ii * complex_of_real r = Complex 0 r".
3852 Instead the reverse direction is supported with
3853 Complex_eq: "Complex a b = a + \<i> * b"
3855 Moved csqrt from Fundamental_Algebra_Theorem to Complex.
3858 Re/Im ~> complex.sel
3859 complex_Re/Im_zero ~> zero_complex.sel
3860 complex_Re/Im_add ~> plus_complex.sel
3861 complex_Re/Im_minus ~> uminus_complex.sel
3862 complex_Re/Im_diff ~> minus_complex.sel
3863 complex_Re/Im_one ~> one_complex.sel
3864 complex_Re/Im_mult ~> times_complex.sel
3865 complex_Re/Im_inverse ~> inverse_complex.sel
3866 complex_Re/Im_scaleR ~> scaleR_complex.sel
3867 complex_Re/Im_i ~> ii.sel
3868 complex_Re/Im_cnj ~> cnj.sel
3869 Re/Im_cis ~> cis.sel
3871 complex_divide_def ~> divide_complex_def
3872 complex_norm_def ~> norm_complex_def
3873 cmod_def ~> norm_complex_de
3887 * Theory Lubs moved HOL image to HOL-Library. It is replaced by
3888 Conditionally_Complete_Lattices. INCOMPATIBILITY.
3890 * HOL-Library: new theory src/HOL/Library/Tree.thy.
3892 * HOL-Library: removed theory src/HOL/Library/Kleene_Algebra.thy; it
3893 is subsumed by session Kleene_Algebra in AFP.
3895 * HOL-Library / theory RBT: various constants and facts are hidden;
3896 lifting setup is unregistered. INCOMPATIBILITY.
3898 * HOL-Cardinals: new theory src/HOL/Cardinals/Ordinal_Arithmetic.thy.
3900 * HOL-Word: bit representations prefer type bool over type bit.
3904 - Abandoned fact collection "word_arith_alts", which is a duplicate
3905 of "word_arith_wis".
3906 - Dropped first (duplicated) element in fact collections
3907 "sint_word_ariths", "word_arith_alts", "uint_word_ariths",
3908 "uint_word_arith_bintrs".
3910 * HOL-Number_Theory:
3911 - consolidated the proofs of the binomial theorem
3912 - the function fib is again of type nat => nat and not overloaded
3913 - no more references to Old_Number_Theory in the HOL libraries
3918 * HOL-Multivariate_Analysis:
3919 - Type class ordered_real_vector for ordered vector spaces.
3920 - New theory Complex_Basic_Analysis defining complex derivatives,
3921 holomorphic functions, etc., ported from HOL Light's canal.ml.
3922 - Changed order of ordered_euclidean_space to be compatible with
3923 pointwise ordering on products. Therefore instance of
3924 conditionally_complete_lattice and ordered_real_vector.
3925 INCOMPATIBILITY: use box instead of greaterThanLessThan or
3926 explicit set-comprehensions with eucl_less for other (half-)open
3928 - removed dependencies on type class ordered_euclidean_space with
3929 introduction of "cbox" on euclidean_space
3932 mem_interval ~> mem_box
3933 interval_eq_empty ~> box_eq_empty
3934 interval_ne_empty ~> box_ne_empty
3935 interval_sing(1) ~> cbox_sing
3936 interval_sing(2) ~> box_sing
3937 subset_interval_imp ~> subset_box_imp
3938 subset_interval ~> subset_box
3939 open_interval ~> open_box
3940 closed_interval ~> closed_cbox
3941 interior_closed_interval ~> interior_cbox
3942 bounded_closed_interval ~> bounded_cbox
3943 compact_interval ~> compact_cbox
3944 bounded_subset_closed_interval_symmetric ~> bounded_subset_cbox_symmetric
3945 bounded_subset_closed_interval ~> bounded_subset_cbox
3946 mem_interval_componentwiseI ~> mem_box_componentwiseI
3947 convex_box ~> convex_prod
3948 rel_interior_real_interval ~> rel_interior_real_box
3949 convex_interval ~> convex_box
3950 convex_hull_eq_real_interval ~> convex_hull_eq_real_cbox
3951 frechet_derivative_within_closed_interval ~> frechet_derivative_within_cbox
3952 content_closed_interval' ~> content_cbox'
3953 elementary_subset_interval ~> elementary_subset_box
3954 diameter_closed_interval ~> diameter_cbox
3955 frontier_closed_interval ~> frontier_cbox
3956 frontier_open_interval ~> frontier_box
3957 bounded_subset_open_interval_symmetric ~> bounded_subset_box_symmetric
3958 closure_open_interval ~> closure_box
3959 open_closed_interval_convex ~> open_cbox_convex
3960 open_interval_midpoint ~> box_midpoint
3961 content_image_affinity_interval ~> content_image_affinity_cbox
3962 is_interval_interval ~> is_interval_cbox + is_interval_box + is_interval_closed_interval
3963 bounded_interval ~> bounded_closed_interval + bounded_boxes
3965 - respective theorems for intervals over the reals:
3966 content_closed_interval + content_cbox
3967 has_integral + has_integral_real
3968 fine_division_exists + fine_division_exists_real
3969 has_integral_null + has_integral_null_real
3970 tagged_division_union_interval + tagged_division_union_interval_real
3971 has_integral_const + has_integral_const_real
3972 integral_const + integral_const_real
3973 has_integral_bound + has_integral_bound_real
3974 integrable_continuous + integrable_continuous_real
3975 integrable_subinterval + integrable_subinterval_real
3976 has_integral_reflect_lemma + has_integral_reflect_lemma_real
3977 integrable_reflect + integrable_reflect_real
3978 integral_reflect + integral_reflect_real
3979 image_affinity_interval + image_affinity_cbox
3980 image_smult_interval + image_smult_cbox
3981 integrable_const + integrable_const_ivl
3982 integrable_on_subinterval + integrable_on_subcbox
3985 derivative_linear ~> has_derivative_bounded_linear
3986 derivative_is_linear ~> has_derivative_linear
3987 bounded_linear_imp_linear ~> bounded_linear.linear
3990 - Renamed positive_integral to nn_integral:
3992 . Renamed all lemmas "*positive_integral*" to *nn_integral*"
3993 positive_integral_positive ~> nn_integral_nonneg
3995 . Renamed abbreviation integral\<^sup>P to integral\<^sup>N.
3997 - replaced the Lebesgue integral on real numbers by the more general
3998 Bochner integral for functions into a real-normed vector space.
4000 integral_zero ~> integral_zero / integrable_zero
4001 integral_minus ~> integral_minus / integrable_minus
4002 integral_add ~> integral_add / integrable_add
4003 integral_diff ~> integral_diff / integrable_diff
4004 integral_setsum ~> integral_setsum / integrable_setsum
4005 integral_multc ~> integral_mult_left / integrable_mult_left
4006 integral_cmult ~> integral_mult_right / integrable_mult_right
4007 integral_triangle_inequality~> integral_norm_bound
4008 integrable_nonneg ~> integrableI_nonneg
4009 integral_positive ~> integral_nonneg_AE
4010 integrable_abs_iff ~> integrable_abs_cancel
4011 positive_integral_lim_INF ~> nn_integral_liminf
4012 lebesgue_real_affine ~> lborel_real_affine
4013 borel_integral_has_integral ~> has_integral_lebesgue_integral
4014 integral_indicator ~>
4015 integral_real_indicator / integrable_real_indicator
4016 positive_integral_fst ~> nn_integral_fst'
4017 positive_integral_fst_measurable ~> nn_integral_fst
4018 positive_integral_snd_measurable ~> nn_integral_snd
4020 integrable_fst_measurable ~>
4021 integral_fst / integrable_fst / AE_integrable_fst
4023 integrable_snd_measurable ~>
4024 integral_snd / integrable_snd / AE_integrable_snd
4026 integral_monotone_convergence ~>
4027 integral_monotone_convergence / integrable_monotone_convergence
4029 integral_monotone_convergence_at_top ~>
4030 integral_monotone_convergence_at_top /
4031 integrable_monotone_convergence_at_top
4033 has_integral_iff_positive_integral_lebesgue ~>
4034 has_integral_iff_has_bochner_integral_lebesgue_nonneg
4036 lebesgue_integral_has_integral ~>
4037 has_integral_integrable_lebesgue_nonneg
4039 positive_integral_lebesgue_has_integral ~>
4040 integral_has_integral_lebesgue_nonneg /
4041 integrable_has_integral_lebesgue_nonneg
4043 lebesgue_integral_real_affine ~>
4044 nn_integral_real_affine
4046 has_integral_iff_positive_integral_lborel ~>
4047 integral_has_integral_nonneg / integrable_has_integral_nonneg
4049 The following theorems where removed:
4051 lebesgue_integral_nonneg
4052 lebesgue_integral_uminus
4053 lebesgue_integral_cmult
4054 lebesgue_integral_multc
4055 lebesgue_integral_cmult_nonneg
4056 integral_cmul_indicator
4059 - Formalized properties about exponentially, Erlang, and normal
4060 distributed random variables.
4062 * HOL-Decision_Procs: Separate command 'approximate' for approximative
4063 computation in src/HOL/Decision_Procs/Approximation. Minor
4069 * The signature and semantics of Document.Snapshot.cumulate_markup /
4070 select_markup have been clarified. Markup is now traversed in the
4071 order of reports given by the prover: later markup is usually more
4072 specific and may override results accumulated so far. The elements
4073 guard is mandatory and checked precisely. Subtle INCOMPATIBILITY.
4075 * Substantial reworking of internal PIDE protocol communication
4076 channels. INCOMPATIBILITY.
4081 * Subtle change of semantics of Thm.eq_thm: theory stamps are not
4082 compared (according to Thm.thm_ord), but assumed to be covered by the
4083 current background theory. Thus equivalent data produced in different
4084 branches of the theory graph usually coincides (e.g. relevant for
4085 theory merge). Note that the softer Thm.eq_thm_prop is often more
4086 appropriate than Thm.eq_thm.
4088 * Proper context for basic Simplifier operations: rewrite_rule,
4089 rewrite_goals_rule, rewrite_goals_tac etc. INCOMPATIBILITY, need to
4090 pass runtime Proof.context (and ensure that the simplified entity
4091 actually belongs to it).
4093 * Proper context discipline for read_instantiate and instantiate_tac:
4094 variables that are meant to become schematic need to be given as
4095 fixed, and are generalized by the explicit context of local variables.
4096 This corresponds to Isar attributes "where" and "of" with 'for'
4097 declaration. INCOMPATIBILITY, also due to potential change of indices
4098 of schematic variables.
4100 * Moved ML_Compiler.exn_trace and other operations on exceptions to
4101 structure Runtime. Minor INCOMPATIBILITY.
4103 * Discontinued old Toplevel.debug in favour of system option
4104 "ML_exception_trace", which may be also declared within the context
4105 via "declare [[ML_exception_trace = true]]". Minor INCOMPATIBILITY.
4107 * Renamed configuration option "ML_trace" to "ML_source_trace". Minor
4110 * Configuration option "ML_print_depth" controls the pretty-printing
4111 depth of the ML compiler within the context. The old print_depth in
4112 ML is still available as default_print_depth, but rarely used. Minor
4115 * Toplevel function "use" refers to raw ML bootstrap environment,
4116 without Isar context nor antiquotations. Potential INCOMPATIBILITY.
4117 Note that 'ML_file' is the canonical command to load ML files into the
4120 * Simplified programming interface to define ML antiquotations, see
4121 structure ML_Antiquotation. Minor INCOMPATIBILITY.
4123 * ML antiquotation @{here} refers to its source position, which is
4124 occasionally useful for experimentation and diagnostic purposes.
4126 * ML antiquotation @{path} produces a Path.T value, similarly to
4127 Path.explode, but with compile-time check against the file-system and
4128 some PIDE markup. Note that unlike theory source, ML does not have a
4129 well-defined master directory, so an absolute symbolic path
4130 specification is usually required, e.g. "~~/src/HOL".
4132 * ML antiquotation @{print} inlines a function to print an arbitrary
4133 ML value, which is occasionally useful for diagnostic or demonstration
4139 * Proof General with its traditional helper scripts is now an optional
4140 Isabelle component, e.g. see ProofGeneral-4.2-2 from the Isabelle
4141 component repository http://isabelle.in.tum.de/components/. Note that
4142 the "system" manual provides general explanations about add-on
4143 components, especially those that are not bundled with the release.
4145 * The raw Isabelle process executable has been renamed from
4146 "isabelle-process" to "isabelle_process", which conforms to common
4147 shell naming conventions, and allows to define a shell function within
4148 the Isabelle environment to avoid dynamic path lookup. Rare
4149 incompatibility for old tools that do not use the ISABELLE_PROCESS
4152 * Former "isabelle tty" has been superseded by "isabelle console",
4153 with implicit build like "isabelle jedit", and without the mostly
4154 obsolete Isar TTY loop.
4156 * Simplified "isabelle display" tool. Settings variables DVI_VIEWER
4157 and PDF_VIEWER now refer to the actual programs, not shell
4158 command-lines. Discontinued option -c: invocation may be asynchronous
4159 via desktop environment, without any special precautions. Potential
4160 INCOMPATIBILITY with ambitious private settings.
4162 * Removed obsolete "isabelle unsymbolize". Note that the usual format
4163 for email communication is the Unicode rendering of Isabelle symbols,
4164 as produced by Isabelle/jEdit, for example.
4166 * Removed obsolete tool "wwwfind". Similar functionality may be
4167 integrated into Isabelle/jEdit eventually.
4169 * Improved 'display_drafts' concerning desktop integration and
4170 repeated invocation in PIDE front-end: re-use single file
4171 $ISABELLE_HOME_USER/tmp/drafts.pdf and corresponding views.
4173 * Session ROOT specifications require explicit 'document_files' for
4174 robust dependencies on LaTeX sources. Only these explicitly given
4175 files are copied to the document output directory, before document
4176 processing is started.
4178 * Windows: support for regular TeX installation (e.g. MiKTeX) instead
4179 of TeX Live from Cygwin.
4183 New in Isabelle2013-2 (December 2013)
4184 -------------------------------------
4186 *** Prover IDE -- Isabelle/Scala/jEdit ***
4188 * More robust editing of running commands with internal forks,
4189 e.g. non-terminating 'by' steps.
4191 * More relaxed Sledgehammer panel: avoid repeated application of query
4192 after edits surrounding the command location.
4194 * More status information about commands that are interrupted
4195 accidentally (via physical event or Poly/ML runtime system signal,
4196 e.g. out-of-memory).
4201 * More robust termination of external processes managed by
4202 Isabelle/ML: support cancellation of tasks within the range of
4203 milliseconds, as required for PIDE document editing with automatically
4204 tried tools (e.g. Sledgehammer).
4206 * Reactivated Isabelle/Scala kill command for external processes on
4207 Mac OS X, which was accidentally broken in Isabelle2013-1 due to a
4208 workaround for some Debian/Ubuntu Linux versions from 2013.
4212 New in Isabelle2013-1 (November 2013)
4213 -------------------------------------
4217 * Discontinued obsolete 'uses' within theory header. Note that
4218 commands like 'ML_file' work without separate declaration of file
4219 dependencies. Minor INCOMPATIBILITY.
4221 * Discontinued redundant 'use' command, which was superseded by
4222 'ML_file' in Isabelle2013. Minor INCOMPATIBILITY.
4224 * Simplified subscripts within identifiers, using plain \<^sub>
4225 instead of the second copy \<^isub> and \<^isup>. Superscripts are
4226 only for literal tokens within notation; explicit mixfix annotations
4227 for consts or fixed variables may be used as fall-back for unusual
4228 names. Obsolete \<twosuperior> has been expanded to \<^sup>2 in
4229 Isabelle/HOL. INCOMPATIBILITY, use "isabelle update_sub_sup" to
4230 standardize symbols as a starting point for further manual cleanup.
4231 The ML reference variable "legacy_isub_isup" may be set as temporary
4232 workaround, to make the prover accept a subset of the old identifier
4235 * Document antiquotations: term style "isub" has been renamed to
4236 "sub". Minor INCOMPATIBILITY.
4238 * Uniform management of "quick_and_dirty" as system option (see also
4239 "isabelle options"), configuration option within the context (see also
4240 Config.get in Isabelle/ML), and attribute in Isabelle/Isar. Minor
4241 INCOMPATIBILITY, need to use more official Isabelle means to access
4242 quick_and_dirty, instead of historical poking into mutable reference.
4244 * Renamed command 'print_configs' to 'print_options'. Minor
4247 * Proper diagnostic command 'print_state'. Old 'pr' (with its
4248 implicit change of some global references) is retained for now as
4249 control command, e.g. for ProofGeneral 3.7.x.
4251 * Discontinued 'print_drafts' command with its old-fashioned PS output
4252 and Unix command-line print spooling. Minor INCOMPATIBILITY: use
4253 'display_drafts' instead and print via the regular document viewer.
4255 * Updated and extended "isar-ref" and "implementation" manual,
4256 eliminated old "ref" manual.
4259 *** Prover IDE -- Isabelle/Scala/jEdit ***
4261 * New manual "jedit" for Isabelle/jEdit, see isabelle doc or
4262 Documentation panel.
4264 * Dockable window "Documentation" provides access to Isabelle
4267 * Dockable window "Find" provides query operations for formal entities
4268 (GUI front-end to 'find_theorems' command).
4270 * Dockable window "Sledgehammer" manages asynchronous / parallel
4271 sledgehammer runs over existing document sources, independently of
4272 normal editing and checking process.
4274 * Dockable window "Timing" provides an overview of relevant command
4275 timing information, depending on option jedit_timing_threshold. The
4276 same timing information is shown in the extended tooltip of the
4277 command keyword, when hovering the mouse over it while the CONTROL or
4278 COMMAND modifier is pressed.
4280 * Improved dockable window "Theories": Continuous checking of proof
4281 document (visible and required parts) may be controlled explicitly,
4282 using check box or shortcut "C+e ENTER". Individual theory nodes may
4283 be marked explicitly as required and checked in full, using check box
4284 or shortcut "C+e SPACE".
4286 * Improved completion mechanism, which is now managed by the
4287 Isabelle/jEdit plugin instead of SideKick. Refined table of Isabelle
4288 symbol abbreviations (see $ISABELLE_HOME/etc/symbols).
4290 * Standard jEdit keyboard shortcut C+b complete-word is remapped to
4291 isabelle.complete for explicit completion in Isabelle sources.
4292 INCOMPATIBILITY wrt. jEdit defaults, may have to invent new shortcuts
4293 to resolve conflict.
4295 * Improved support of various "minor modes" for Isabelle NEWS,
4296 options, session ROOT etc., with completion and SideKick tree view.
4298 * Strictly monotonic document update, without premature cancellation of
4299 running transactions that are still needed: avoid reset/restart of
4300 such command executions while editing.
4302 * Support for asynchronous print functions, as overlay to existing
4305 * Support for automatic tools in HOL, which try to prove or disprove
4306 toplevel theorem statements.
4308 * Action isabelle.reset-font-size resets main text area font size
4309 according to Isabelle/Scala plugin option "jedit_font_reset_size" (see
4310 also "Plugin Options / Isabelle / General"). It can be bound to some
4311 keyboard shortcut by the user (e.g. C+0 and/or C+NUMPAD0).
4313 * File specifications in jEdit (e.g. file browser) may refer to
4314 $ISABELLE_HOME and $ISABELLE_HOME_USER on all platforms. Discontinued
4315 obsolete $ISABELLE_HOME_WINDOWS variable.
4317 * Improved support for Linux look-and-feel "GTK+", see also "Utilities
4318 / Global Options / Appearance".
4320 * Improved support of native Mac OS X functionality via "MacOSX"
4321 plugin, which is now enabled by default.
4326 * Commands 'interpretation' and 'sublocale' are now target-sensitive.
4327 In particular, 'interpretation' allows for non-persistent
4328 interpretation within "context ... begin ... end" blocks offering a
4329 light-weight alternative to 'sublocale'. See "isar-ref" manual for
4332 * Improved locales diagnostic command 'print_dependencies'.
4334 * Discontinued obsolete 'axioms' command, which has been marked as
4335 legacy since Isabelle2009-2. INCOMPATIBILITY, use 'axiomatization'
4336 instead, while observing its uniform scope for polymorphism.
4338 * Discontinued empty name bindings in 'axiomatization'.
4341 * System option "proofs" has been discontinued. Instead the global
4342 state of Proofterm.proofs is persistently compiled into logic images
4343 as required, notably HOL-Proofs. Users no longer need to change
4344 Proofterm.proofs dynamically. Minor INCOMPATIBILITY.
4346 * Syntax translation functions (print_translation etc.) always depend
4347 on Proof.context. Discontinued former "(advanced)" option -- this is
4348 now the default. Minor INCOMPATIBILITY.
4350 * Former global reference trace_unify_fail is now available as
4351 configuration option "unify_trace_failure" (global context only).
4353 * SELECT_GOAL now retains the syntactic context of the overall goal
4354 state (schematic variables etc.). Potential INCOMPATIBILITY in rare
4360 * Stronger precedence of syntax for big intersection and union on
4361 sets, in accordance with corresponding lattice operations.
4364 * Notation "{p:A. P}" now allows tuple patterns as well.
4366 * Nested case expressions are now translated in a separate check phase
4367 rather than during parsing. The data for case combinators is separated
4368 from the datatype package. The declaration attribute
4369 "case_translation" can be used to register new case combinators:
4371 declare [[case_translation case_combinator constructor1 ... constructorN]]
4374 - 'code_printing' unifies 'code_const' / 'code_type' / 'code_class' /
4376 - 'code_identifier' declares name hints for arbitrary identifiers in
4377 generated code, subsuming 'code_modulename'.
4379 See the isar-ref manual for syntax diagrams, and the HOL theories for
4382 * Attibute 'code': 'code' now declares concrete and abstract code
4383 equations uniformly. Use explicit 'code equation' and 'code abstract'
4384 to distinguish both when desired.
4386 * Discontinued theories Code_Integer and Efficient_Nat by a more
4387 fine-grain stack of theories Code_Target_Int, Code_Binary_Nat,
4388 Code_Target_Nat and Code_Target_Numeral. See the tutorial on code
4389 generation for details. INCOMPATIBILITY.
4391 * Numeric types are mapped by default to target language numerals:
4392 natural (replaces former code_numeral) and integer (replaces former
4393 code_int). Conversions are available as integer_of_natural /
4394 natural_of_integer / integer_of_nat / nat_of_integer (in HOL) and
4395 Code_Numeral.integer_of_natural / Code_Numeral.natural_of_integer (in
4396 ML). INCOMPATIBILITY.
4398 * Function package: For mutually recursive functions f and g, separate
4399 cases rules f.cases and g.cases are generated instead of unusable
4400 f_g.cases which exposed internal sum types. Potential INCOMPATIBILITY,
4401 in the case that the unusable rule was used nevertheless.
4403 * Function package: For each function f, new rules f.elims are
4404 generated, which eliminate equalities of the form "f x = t".
4406 * New command 'fun_cases' derives ad-hoc elimination rules for
4407 function equations as simplified instances of f.elims, analogous to
4408 inductive_cases. See ~~/src/HOL/ex/Fundefs.thy for some examples.
4411 - parametrized correspondence relations are now supported:
4412 + parametricity theorems for the raw term can be specified in
4413 the command lift_definition, which allow us to generate stronger
4415 + setup_lifting generates stronger transfer rules if parametric
4416 correspondence relation can be generated
4417 + various new properties of the relator must be specified to support
4419 + parametricity theorem for the Quotient relation can be specified
4420 - setup_lifting generates domain rules for the Transfer package
4421 - stronger reflexivity prover of respectfulness theorems for type
4423 - ===> and --> are now local. The symbols can be introduced
4424 by interpreting the locale lifting_syntax (typically in an
4426 - Lifting/Transfer relevant parts of Library/Quotient_* are now in
4427 Main. Potential INCOMPATIBILITY
4428 - new commands for restoring and deleting Lifting/Transfer context:
4429 lifting_forget, lifting_update
4430 - the command print_quotmaps was renamed to print_quot_maps.
4434 - better support for domains in Transfer: replace Domainp T
4435 by the actual invariant in a transferred goal
4436 - transfer rules can have as assumptions other transfer rules
4437 - Experimental support for transferring from the raw level to the
4438 abstract level: Transfer.transferred attribute
4439 - Attribute version of the transfer method: untransferred attribute
4441 * Reification and reflection:
4442 - Reification is now directly available in HOL-Main in structure
4444 - Reflection now handles multiple lists with variables also.
4445 - The whole reflection stack has been decomposed into conversions.
4448 * Revised devices for recursive definitions over finite sets:
4449 - Only one fundamental fold combinator on finite set remains:
4450 Finite_Set.fold :: ('a => 'b => 'b) => 'b => 'a set => 'b
4451 This is now identity on infinite sets.
4452 - Locales ("mini packages") for fundamental definitions with
4453 Finite_Set.fold: folding, folding_idem.
4454 - Locales comm_monoid_set, semilattice_order_set and
4455 semilattice_neutr_order_set for big operators on sets.
4456 See theory Big_Operators for canonical examples.
4457 Note that foundational constants comm_monoid_set.F and
4458 semilattice_set.F correspond to former combinators fold_image
4459 and fold1 respectively. These are now gone. You may use
4460 those foundational constants as substitutes, but it is
4461 preferable to interpret the above locales accordingly.
4462 - Dropped class ab_semigroup_idem_mult (special case of lattice,
4463 no longer needed in connection with Finite_Set.fold etc.)
4465 card.union_inter ~> card_Un_Int [symmetric]
4466 card.union_disjoint ~> card_Un_disjoint
4469 * Locale hierarchy for abstract orderings and (semi)lattices.
4471 * Complete_Partial_Order.admissible is defined outside the type class
4472 ccpo, but with mandatory prefix ccpo. Admissibility theorems lose the
4473 class predicate assumption or sort constraint when possible.
4476 * Introduce type class "conditionally_complete_lattice": Like a
4477 complete lattice but does not assume the existence of the top and
4478 bottom elements. Allows to generalize some lemmas about reals and
4479 extended reals. Removed SupInf and replaced it by the instantiation
4480 of conditionally_complete_lattice for real. Renamed lemmas about
4481 conditionally-complete lattice from Sup_... to cSup_... and from
4482 Inf_... to cInf_... to avoid hidding of similar complete lattice
4485 * Introduce type class linear_continuum as combination of
4486 conditionally-complete lattices and inner dense linorders which have
4487 more than one element. INCOMPATIBILITY.
4489 * Introduced type classes order_top and order_bot. The old classes top
4490 and bot only contain the syntax without assumptions. INCOMPATIBILITY:
4491 Rename bot -> order_bot, top -> order_top
4493 * Introduce type classes "no_top" and "no_bot" for orderings without
4494 top and bottom elements.
4496 * Split dense_linorder into inner_dense_order and no_top, no_bot.
4498 * Complex_Main: Unify and move various concepts from
4499 HOL-Multivariate_Analysis to HOL-Complex_Main.
4501 - Introduce type class (lin)order_topology and
4502 linear_continuum_topology. Allows to generalize theorems about
4503 limits and order. Instances are reals and extended reals.
4505 - continuous and continuos_on from Multivariate_Analysis:
4506 "continuous" is the continuity of a function at a filter. "isCont"
4507 is now an abbrevitation: "isCont x f == continuous (at _) f".
4509 Generalized continuity lemmas from isCont to continuous on an
4512 - compact from Multivariate_Analysis. Use Bolzano's lemma to prove
4513 compactness of closed intervals on reals. Continuous functions
4514 attain infimum and supremum on compact sets. The inverse of a
4515 continuous function is continuous, when the function is continuous
4518 - connected from Multivariate_Analysis. Use it to prove the
4519 intermediate value theorem. Show connectedness of intervals on
4520 linear_continuum_topology).
4522 - first_countable_topology from Multivariate_Analysis. Is used to
4523 show equivalence of properties on the neighbourhood filter of x and
4524 on all sequences converging to x.
4526 - FDERIV: Definition of has_derivative moved to Deriv.thy. Moved
4527 theorems from Library/FDERIV.thy to Deriv.thy and base the
4528 definition of DERIV on FDERIV. Add variants of DERIV and FDERIV
4529 which are restricted to sets, i.e. to represent derivatives from
4532 - Removed the within-filter. It is replaced by the principal filter:
4534 F within X = inf F (principal X)
4536 - Introduce "at x within U" as a single constant, "at x" is now an
4537 abbreviation for "at x within UNIV"
4539 - Introduce named theorem collections tendsto_intros,
4540 continuous_intros, continuous_on_intros and FDERIV_intros. Theorems
4541 in tendsto_intros (or FDERIV_intros) are also available as
4542 tendsto_eq_intros (or FDERIV_eq_intros) where the right-hand side
4543 is replaced by a congruence rule. This allows to apply them as
4544 intro rules and then proving equivalence by the simplifier.
4546 - Restructured theories in HOL-Complex_Main:
4548 + Moved RealDef and RComplete into Real
4550 + Introduced Topological_Spaces and moved theorems about
4551 topological spaces, filters, limits and continuity to it
4553 + Renamed RealVector to Real_Vector_Spaces
4555 + Split Lim, SEQ, Series into Topological_Spaces,
4556 Real_Vector_Spaces, and Limits
4558 + Moved Ln and Log to Transcendental
4560 + Moved theorems about continuity from Deriv to Topological_Spaces
4562 - Remove various auxiliary lemmas.
4567 - Added option "spy".
4568 - Reduce incidence of "too high arity" errors.
4572 isar_shrink ~> isar_compress
4574 - Added options "isar_try0", "spy".
4575 - Better support for "isar_proofs".
4576 - MaSh has been fined-tuned and now runs as a local server.
4578 * Improved support for ad hoc overloading of constants (see also
4579 isar-ref manual and ~~/src/HOL/ex/Adhoc_Overloading_Examples.thy).
4581 * Library/Polynomial.thy:
4582 - Use lifting for primitive definitions.
4583 - Explicit conversions from and to lists of coefficients, used for
4585 - Replaced recursion operator poly_rec by fold_coeffs.
4586 - Prefer pre-existing gcd operation for gcd.
4588 poly_eq_iff ~> poly_eq_poly_eq_iff
4589 poly_ext ~> poly_eqI
4590 expand_poly_eq ~> poly_eq_iff
4593 * New Library/Simps_Case_Conv.thy: Provides commands simps_of_case and
4594 case_of_simps to convert function definitions between a list of
4595 equations with patterns on the lhs and a single equation with case
4596 expressions on the rhs. See also Ex/Simps_Case_Conv_Examples.thy.
4598 * New Library/FSet.thy: type of finite sets defined as a subtype of
4599 sets defined by Lifting/Transfer.
4601 * Discontinued theory src/HOL/Library/Eval_Witness. INCOMPATIBILITY.
4603 * Consolidation of library theories on product orders:
4605 Product_Lattice ~> Product_Order -- pointwise order on products
4606 Product_ord ~> Product_Lexorder -- lexicographic order on products
4610 * Imperative-HOL: The MREC combinator is considered legacy and no
4611 longer included by default. INCOMPATIBILITY, use partial_function
4612 instead, or import theory Legacy_Mrec as a fallback.
4614 * HOL-Algebra: Discontinued theories ~~/src/HOL/Algebra/abstract and
4615 ~~/src/HOL/Algebra/poly. Existing theories should be based on
4616 ~~/src/HOL/Library/Polynomial instead. The latter provides
4617 integration with HOL's type classes for rings. INCOMPATIBILITY.
4620 - Various improvements to BNF-based (co)datatype package, including
4621 new commands "primrec_new", "primcorec", and
4622 "datatype_new_compat", as well as documentation. See
4623 "datatypes.pdf" for details.
4624 - New "coinduction" method to avoid some boilerplate (compared to
4627 data ~> datatype_new
4628 codata ~> codatatype
4630 - Renamed many generated theorems, including
4632 map_comp' ~> map_comp
4642 * Spec_Check is a Quickcheck tool for Isabelle/ML. The ML function
4643 "check_property" allows to check specifications of the form "ALL x y
4644 z. prop x y z". See also ~~/src/Tools/Spec_Check/ with its
4645 Examples.thy in particular.
4647 * Improved printing of exception trace in Poly/ML 5.5.1, with regular
4648 tracing output in the command transaction context instead of physical
4649 stdout. See also Toplevel.debug, Toplevel.debugging and
4650 ML_Compiler.exn_trace.
4652 * ML type "theory" is now immutable, without any special treatment of
4653 drafts or linear updates (which could lead to "stale theory" errors in
4654 the past). Discontinued obsolete operations like Theory.copy,
4655 Theory.checkpoint, and the auxiliary type theory_ref. Minor
4658 * More uniform naming of goal functions for skipped proofs:
4660 Skip_Proof.prove ~> Goal.prove_sorry
4661 Skip_Proof.prove_global ~> Goal.prove_sorry_global
4663 Minor INCOMPATIBILITY.
4665 * Simplifier tactics and tools use proper Proof.context instead of
4666 historic type simpset. Old-style declarations like addsimps,
4667 addsimprocs etc. operate directly on Proof.context. Raw type simpset
4668 retains its use as snapshot of the main Simplifier context, using
4669 simpset_of and put_simpset on Proof.context. INCOMPATIBILITY -- port
4670 old tools by making them depend on (ctxt : Proof.context) instead of
4671 (ss : simpset), then turn (simpset_of ctxt) into ctxt.
4673 * Modifiers for classical wrappers (e.g. addWrapper, delWrapper)
4674 operate on Proof.context instead of claset, for uniformity with addIs,
4675 addEs, addDs etc. Note that claset_of and put_claset allow to manage
4676 clasets separately from the context.
4678 * Discontinued obsolete ML antiquotations @{claset} and @{simpset}.
4679 INCOMPATIBILITY, use @{context} instead.
4681 * Antiquotation @{theory_context A} is similar to @{theory A}, but
4682 presents the result as initial Proof.context.
4687 * Discontinued obsolete isabelle usedir, mkdir, make -- superseded by
4688 "isabelle build" in Isabelle2013. INCOMPATIBILITY.
4690 * Discontinued obsolete isabelle-process options -f and -u (former
4691 administrative aliases of option -e). Minor INCOMPATIBILITY.
4693 * Discontinued obsolete isabelle print tool, and PRINT_COMMAND
4696 * Discontinued ISABELLE_DOC_FORMAT settings variable and historic
4697 document formats: dvi.gz, ps, ps.gz -- the default document format is
4700 * Isabelle settings variable ISABELLE_BUILD_JAVA_OPTIONS allows to
4701 specify global resources of the JVM process run by isabelle build.
4703 * Toplevel executable $ISABELLE_HOME/bin/isabelle_scala_script allows
4704 to run Isabelle/Scala source files as standalone programs.
4706 * Improved "isabelle keywords" tool (for old-style ProofGeneral
4707 keyword tables): use Isabelle/Scala operations, which inspect outer
4708 syntax without requiring to build sessions first.
4710 * Sessions may be organized via 'chapter' specifications in the ROOT
4711 file, which determines a two-level hierarchy of browser info. The old
4712 tree-like organization via implicit sub-session relation (with its
4713 tendency towards erratic fluctuation of URLs) has been discontinued.
4714 The default chapter is called "Unsorted". Potential INCOMPATIBILITY
4715 for HTML presentation of theories.
4719 New in Isabelle2013 (February 2013)
4720 -----------------------------------
4724 * Theorem status about oracles and unfinished/failed future proofs is
4725 no longer printed by default, since it is incompatible with
4726 incremental / parallel checking of the persistent document model. ML
4727 function Thm.peek_status may be used to inspect a snapshot of the
4728 ongoing evaluation process. Note that in batch mode --- notably
4729 isabelle build --- the system ensures that future proofs of all
4730 accessible theorems in the theory context are finished (as before).
4732 * Configuration option show_markup controls direct inlining of markup
4733 into the printed representation of formal entities --- notably type
4734 and sort constraints. This enables Prover IDE users to retrieve that
4735 information via tooltips in the output window, for example.
4737 * Command 'ML_file' evaluates ML text from a file directly within the
4738 theory, without any predeclaration via 'uses' in the theory header.
4740 * Old command 'use' command and corresponding keyword 'uses' in the
4741 theory header are legacy features and will be discontinued soon.
4742 Tools that load their additional source files may imitate the
4743 'ML_file' implementation, such that the system can take care of
4744 dependencies properly.
4746 * Discontinued obsolete method fastsimp / tactic fast_simp_tac, which
4747 is called fastforce / fast_force_tac already since Isabelle2011-1.
4749 * Updated and extended "isar-ref" and "implementation" manual, reduced
4750 remaining material in old "ref" manual.
4752 * Improved support for auxiliary contexts that indicate block structure
4753 for specifications. Nesting of "context fixes ... context assumes ..."
4754 and "class ... context ...".
4756 * Attribute "consumes" allows a negative value as well, which is
4757 interpreted relatively to the total number of premises of the rule in
4758 the target context. This form of declaration is stable when exported
4759 from a nested 'context' with additional assumptions. It is the
4760 preferred form for definitional packages, notably cases/rules produced
4761 in HOL/inductive and HOL/function.
4763 * More informative error messages for Isar proof commands involving
4764 lazy enumerations (method applications etc.).
4766 * Refined 'help' command to retrieve outer syntax commands according
4767 to name patterns (with clickable results).
4770 *** Prover IDE -- Isabelle/Scala/jEdit ***
4772 * Parallel terminal proofs ('by') are enabled by default, likewise
4773 proofs that are built into packages like 'datatype', 'function'. This
4774 allows to "run ahead" checking the theory specifications on the
4775 surface, while the prover is still crunching on internal
4776 justifications. Unfinished / cancelled proofs are restarted as
4777 required to complete full proof checking eventually.
4779 * Improved output panel with tooltips, hyperlinks etc. based on the
4780 same Rich_Text_Area as regular Isabelle/jEdit buffers. Activation of
4781 tooltips leads to some window that supports the same recursively,
4782 which can lead to stacks of tooltips as the semantic document content
4783 is explored. ESCAPE closes the whole stack, individual windows may be
4784 closed separately, or detached to become independent jEdit dockables.
4786 * Improved support for commands that produce graph output: the text
4787 message contains a clickable area to open a new instance of the graph
4790 * More robust incremental parsing of outer syntax (partial comments,
4791 malformed symbols). Changing the balance of open/close quotes and
4792 comment delimiters works more conveniently with unfinished situations
4793 that frequently occur in user interaction.
4795 * More efficient painting and improved reactivity when editing large
4796 files. More scalable management of formal document content.
4798 * Smarter handling of tracing messages: prover process pauses after
4799 certain number of messages per command transaction, with some user
4800 dialog to stop or continue. This avoids swamping the front-end with
4801 potentially infinite message streams.
4803 * More plugin options and preferences, based on Isabelle/Scala. The
4804 jEdit plugin option panel provides access to some Isabelle/Scala
4805 options, including tuning parameters for editor reactivity and color
4808 * Dockable window "Symbols" provides some editing support for Isabelle
4811 * Dockable window "Monitor" shows ML runtime statistics. Note that
4812 continuous display of the chart slows down the system.
4814 * Improved editing support for control styles: subscript, superscript,
4815 bold, reset of style -- operating on single symbols or text
4816 selections. Cf. keyboard shortcuts C+e DOWN/UP/RIGHT/LEFT.
4818 * Actions isabelle.increase-font-size and isabelle.decrease-font-size
4819 adjust the main text area font size, and its derivatives for output,
4820 tooltips etc. Cf. keyboard shortcuts C-PLUS and C-MINUS, which often
4821 need to be adapted to local keyboard layouts.
4823 * More reactive completion popup by default: use \t (TAB) instead of
4824 \n (NEWLINE) to minimize intrusion into regular flow of editing. See
4825 also "Plugin Options / SideKick / General / Code Completion Options".
4827 * Implicit check and build dialog of the specified logic session
4828 image. For example, HOL, HOLCF, HOL-Nominal can be produced on
4829 demand, without bundling big platform-dependent heap images in the
4830 Isabelle distribution.
4832 * Uniform Java 7 platform on Linux, Mac OS X, Windows: recent updates
4833 from Oracle provide better multi-platform experience. This version is
4834 now bundled exclusively with Isabelle.
4839 * Code generation for Haskell: restrict unqualified imports from
4840 Haskell Prelude to a small set of fundamental operations.
4842 * Command 'export_code': relative file names are interpreted
4843 relatively to master directory of current theory rather than the
4844 rather arbitrary current working directory. INCOMPATIBILITY.
4846 * Discontinued obsolete attribute "COMP". Potential INCOMPATIBILITY,
4847 use regular rule composition via "OF" / "THEN", or explicit proof
4848 structure instead. Note that Isabelle/ML provides a variety of
4849 operators like COMP, INCR_COMP, COMP_INCR, which need to be applied
4850 with some care where this is really required.
4852 * Command 'typ' supports an additional variant with explicit sort
4853 constraint, to infer and check the most general type conforming to a
4854 given sort. Example (in HOL):
4856 typ "_ * _ * bool * unit" :: finite
4858 * Command 'locale_deps' visualizes all locales and their relations as
4866 - Added MaSh relevance filter based on machine-learning; see the
4867 Sledgehammer manual for details.
4868 - Polished Isar proofs generated with "isar_proofs" option.
4869 - Rationalized type encodings ("type_enc" option).
4870 - Renamed "kill_provers" subcommand to "kill_all".
4872 isar_proof ~> isar_proofs
4873 isar_shrink_factor ~> isar_shrink
4874 max_relevant ~> max_facts
4875 relevance_thresholds ~> fact_thresholds
4877 * Quickcheck: added an optimisation for equality premises. It is
4878 switched on by default, and can be switched off by setting the
4879 configuration quickcheck_optimise_equality to false.
4881 * Quotient: only one quotient can be defined by quotient_type
4885 - generation of an abstraction function equation in lift_definition
4886 - quot_del attribute
4887 - renamed no_abs_code -> no_code (INCOMPATIBILITY.)
4889 * Simproc "finite_Collect" rewrites set comprehensions into pointfree
4892 * Preprocessing of the code generator rewrites set comprehensions into
4893 pointfree expressions.
4895 * The SMT solver Z3 has now by default a restricted set of directly
4896 supported features. For the full set of features (div/mod, nonlinear
4897 arithmetic, datatypes/records) with potential proof reconstruction
4898 failures, enable the configuration option "z3_with_extensions". Minor
4901 * Simplified 'typedef' specifications: historical options for implicit
4902 set definition and alternative name have been discontinued. The
4903 former behavior of "typedef (open) t = A" is now the default, but
4904 written just "typedef t = A". INCOMPATIBILITY, need to adapt theories
4907 * Removed constant "chars"; prefer "Enum.enum" on type "char"
4908 directly. INCOMPATIBILITY.
4910 * Moved operation product, sublists and n_lists from theory Enum to
4911 List. INCOMPATIBILITY.
4913 * Theorem UN_o generalized to SUP_comp. INCOMPATIBILITY.
4915 * Class "comm_monoid_diff" formalises properties of bounded
4916 subtraction, with natural numbers and multisets as typical instances.
4918 * Added combinator "Option.these" with type "'a option set => 'a set".
4920 * Theory "Transitive_Closure": renamed lemmas
4922 reflcl_tranclp -> reflclp_tranclp
4923 rtranclp_reflcl -> rtranclp_reflclp
4927 * Theory "Rings": renamed lemmas (in class semiring)
4929 left_distrib ~> distrib_right
4930 right_distrib ~> distrib_left
4934 * Generalized the definition of limits:
4936 - Introduced the predicate filterlim (LIM x F. f x :> G) which
4937 expresses that when the input values x converge to F then the
4938 output f x converges to G.
4940 - Added filters for convergence to positive (at_top) and negative
4943 - Moved infinity in the norm (at_infinity) from
4944 Multivariate_Analysis to Complex_Main.
4946 - Removed real_tendsto_inf, it is superseded by "LIM x F. f x :>
4951 * Theory "Library/Option_ord" provides instantiation of option type to
4952 lattice type classes.
4954 * Theory "Library/Multiset": renamed
4956 constant fold_mset ~> Multiset.fold
4957 fact fold_mset_commute ~> fold_mset_comm
4961 * Renamed theory Library/List_Prefix to Library/Sublist, with related
4964 - Renamed constants (and related lemmas)
4967 strict_prefix ~> prefix
4969 - Replaced constant "postfix" by "suffixeq" with swapped argument
4970 order (i.e., "postfix xs ys" is now "suffixeq ys xs") and dropped
4971 old infix syntax "xs >>= ys"; use "suffixeq ys xs" instead.
4972 Renamed lemmas accordingly.
4974 - Added constant "list_hembeq" for homeomorphic embedding on
4975 lists. Added abbreviation "sublisteq" for special case
4976 "list_hembeq (op =)".
4978 - Theory Library/Sublist no longer provides "order" and "bot" type
4979 class instances for the prefix order (merely corresponding locale
4980 interpretations). The type class instances are now in theory
4981 Library/Prefix_Order.
4983 - The sublist relation of theory Library/Sublist_Order is now based
4984 on "Sublist.sublisteq". Renamed lemmas accordingly:
4986 le_list_append_le_same_iff ~> Sublist.sublisteq_append_le_same_iff
4987 le_list_append_mono ~> Sublist.list_hembeq_append_mono
4988 le_list_below_empty ~> Sublist.list_hembeq_Nil, Sublist.list_hembeq_Nil2
4989 le_list_Cons_EX ~> Sublist.list_hembeq_ConsD
4990 le_list_drop_Cons2 ~> Sublist.sublisteq_Cons2'
4991 le_list_drop_Cons_neq ~> Sublist.sublisteq_Cons2_neq
4992 le_list_drop_Cons ~> Sublist.sublisteq_Cons'
4993 le_list_drop_many ~> Sublist.sublisteq_drop_many
4994 le_list_filter_left ~> Sublist.sublisteq_filter_left
4995 le_list_rev_drop_many ~> Sublist.sublisteq_rev_drop_many
4996 le_list_rev_take_iff ~> Sublist.sublisteq_append
4997 le_list_same_length ~> Sublist.sublisteq_same_length
4998 le_list_take_many_iff ~> Sublist.sublisteq_append'
4999 less_eq_list.drop ~> less_eq_list_drop
5000 less_eq_list.induct ~> less_eq_list_induct
5001 not_le_list_length ~> Sublist.not_sublisteq_length
5005 * New theory Library/Countable_Set.
5007 * Theory Library/Debug and Library/Parallel provide debugging and
5008 parallel execution for code generated towards Isabelle/ML.
5010 * Theory Library/FuncSet: Extended support for Pi and extensional and
5011 introduce the extensional dependent function space "PiE". Replaced
5012 extensional_funcset by an abbreviation, and renamed lemmas from
5013 extensional_funcset to PiE as follows:
5015 extensional_empty ~> PiE_empty
5016 extensional_funcset_empty_domain ~> PiE_empty_domain
5017 extensional_funcset_empty_range ~> PiE_empty_range
5018 extensional_funcset_arb ~> PiE_arb
5019 extensional_funcset_mem ~> PiE_mem
5020 extensional_funcset_extend_domainI ~> PiE_fun_upd
5021 extensional_funcset_restrict_domain ~> fun_upd_in_PiE
5022 extensional_funcset_extend_domain_eq ~> PiE_insert_eq
5023 card_extensional_funcset ~> card_PiE
5024 finite_extensional_funcset ~> finite_PiE
5028 * Theory Library/FinFun: theory of almost everywhere constant
5029 functions (supersedes the AFP entry "Code Generation for Functions as
5032 * Theory Library/Phantom: generic phantom type to make a type
5033 parameter appear in a constant's type. This alternative to adding
5034 TYPE('a) as another parameter avoids unnecessary closures in generated
5037 * Theory Library/RBT_Impl: efficient construction of red-black trees
5038 from sorted associative lists. Merging two trees with rbt_union may
5039 return a structurally different tree than before. Potential
5042 * Theory Library/IArray: immutable arrays with code generation.
5044 * Theory Library/Finite_Lattice: theory of finite lattices.
5046 * HOL/Multivariate_Analysis: replaced
5048 "basis :: 'a::euclidean_space => nat => real"
5049 "\<Chi>\<Chi> :: (nat => real) => 'a::euclidean_space"
5051 on euclidean spaces by using the inner product "_ \<bullet> _" with
5052 vectors from the Basis set: "\<Chi>\<Chi> i. f i" is superseded by
5053 "SUM i : Basis. f i * r i".
5055 With this change the following constants are also changed or removed:
5057 DIM('a) :: nat ~> card (Basis :: 'a set) (is an abbreviation)
5058 a $$ i ~> inner a i (where i : Basis)
5060 \<pi>, \<pi>' removed
5062 Theorems about these constants where removed.
5066 component_le_norm ~> Basis_le_norm
5067 euclidean_eq ~> euclidean_eq_iff
5068 differential_zero_maxmin_component ~> differential_zero_maxmin_cart
5069 euclidean_simps ~> inner_simps
5070 independent_basis ~> independent_Basis
5071 span_basis ~> span_Basis
5072 in_span_basis ~> in_span_Basis
5073 norm_bound_component_le ~> norm_boound_Basis_le
5074 norm_bound_component_lt ~> norm_boound_Basis_lt
5075 component_le_infnorm ~> Basis_le_infnorm
5081 - Added simproc "measurable" to automatically prove measurability.
5083 - Added induction rules for sigma sets with disjoint union
5084 (sigma_sets_induct_disjoint) and for Borel-measurable functions
5085 (borel_measurable_induct).
5087 - Added the Daniell-Kolmogorov theorem (the existence the limit of a
5090 * HOL/Cardinals: Theories of ordinals and cardinals (supersedes the
5091 AFP entry "Ordinals_and_Cardinals").
5093 * HOL/BNF: New (co)datatype package based on bounded natural functors
5094 with support for mixed, nested recursion and interesting non-free
5097 * HOL/Finite_Set and Relation: added new set and relation operations
5098 expressed by Finite_Set.fold.
5100 * New theory HOL/Library/RBT_Set: implementation of sets by red-black
5101 trees for the code generator.
5103 * HOL/Library/RBT and HOL/Library/Mapping have been converted to
5105 possible INCOMPATIBILITY.
5107 * HOL/Set: renamed Set.project -> Set.filter
5111 *** Document preparation ***
5113 * Dropped legacy antiquotations "term_style" and "thm_style", since
5114 styles may be given as arguments to "term" and "thm" already.
5115 Discontinued legacy styles "prem1" .. "prem19".
5117 * Default LaTeX rendering for \<euro> is now based on eurosym package,
5118 instead of slightly exotic babel/greek.
5120 * Document variant NAME may use different LaTeX entry point
5121 document/root_NAME.tex if that file exists, instead of the common
5124 * Simplified custom document/build script, instead of old-style
5125 document/IsaMakefile. Minor INCOMPATIBILITY.
5130 * The default limit for maximum number of worker threads is now 8,
5131 instead of 4, in correspondence to capabilities of contemporary
5132 hardware and Poly/ML runtime system.
5134 * Type Seq.results and related operations support embedded error
5135 messages within lazy enumerations, and thus allow to provide
5136 informative errors in the absence of any usable results.
5138 * Renamed Position.str_of to Position.here to emphasize that this is a
5139 formal device to inline positions into message text, but not
5140 necessarily printing visible text.
5145 * Advanced support for Isabelle sessions and build management, see
5146 "system" manual for the chapter of that name, especially the "isabelle
5147 build" tool and its examples. The "isabelle mkroot" tool prepares
5148 session root directories for use with "isabelle build", similar to
5149 former "isabelle mkdir" for "isabelle usedir". Note that this affects
5150 document preparation as well. INCOMPATIBILITY, isabelle usedir /
5151 mkdir / make are rendered obsolete.
5153 * Discontinued obsolete Isabelle/build script, it is superseded by the
5154 regular isabelle build tool. For example:
5156 isabelle build -s -b HOL
5158 * Discontinued obsolete "isabelle makeall".
5160 * Discontinued obsolete IsaMakefile and ROOT.ML files from the
5161 Isabelle distribution, except for rudimentary src/HOL/IsaMakefile that
5162 provides some traditional targets that invoke "isabelle build". Note
5163 that this is inefficient! Applications of Isabelle/HOL involving
5164 "isabelle make" should be upgraded to use "isabelle build" directly.
5166 * The "isabelle options" tool prints Isabelle system options, as
5167 required for "isabelle build", for example.
5169 * The "isabelle logo" tool produces EPS and PDF format simultaneously.
5170 Minor INCOMPATIBILITY in command-line options.
5172 * The "isabelle install" tool has now a simpler command-line. Minor
5175 * The "isabelle components" tool helps to resolve add-on components
5176 that are not bundled, or referenced from a bare-bones repository
5177 version of Isabelle.
5179 * Settings variable ISABELLE_PLATFORM_FAMILY refers to the general
5180 platform family: "linux", "macos", "windows".
5182 * The ML system is configured as regular component, and no longer
5183 picked up from some surrounding directory. Potential INCOMPATIBILITY
5184 for home-made settings.
5186 * Improved ML runtime statistics (heap, threads, future tasks etc.).
5188 * Discontinued support for Poly/ML 5.2.1, which was the last version
5189 without exception positions and advanced ML compiler/toplevel
5192 * Discontinued special treatment of Proof General -- no longer guess
5193 PROOFGENERAL_HOME based on accidental file-system layout. Minor
5194 INCOMPATIBILITY: provide PROOFGENERAL_HOME and PROOFGENERAL_OPTIONS
5195 settings manually, or use a Proof General version that has been
5196 bundled as Isabelle component.
5200 New in Isabelle2012 (May 2012)
5201 ------------------------------
5205 * Prover IDE (PIDE) improvements:
5207 - more robust Sledgehammer integration (as before the sledgehammer
5208 command-line needs to be typed into the source buffer)
5209 - markup for bound variables
5210 - markup for types of term variables (displayed as tooltips)
5211 - support for user-defined Isar commands within the running session
5212 - improved support for Unicode outside original 16bit range
5213 e.g. glyph for \<A> (thanks to jEdit 4.5.1)
5215 * Forward declaration of outer syntax keywords within the theory
5216 header -- minor INCOMPATIBILITY for user-defined commands. Allow new
5217 commands to be used in the same theory where defined.
5219 * Auxiliary contexts indicate block structure for specifications with
5220 additional parameters and assumptions. Such unnamed contexts may be
5221 nested within other targets, like 'theory', 'locale', 'class',
5222 'instantiation' etc. Results from the local context are generalized
5223 accordingly and applied to the enclosing target context. Example:
5227 assumes xy: "x = y" and yz: "y = z"
5230 lemma my_trans: "x = z" using xy yz by simp
5236 The most basic application is to factor-out context elements of
5237 several fixes/assumes/shows theorem statements, e.g. see
5238 ~~/src/HOL/Isar_Examples/Group_Context.thy
5240 Any other local theory specification element works within the "context
5241 ... begin ... end" block as well.
5243 * Bundled declarations associate attributed fact expressions with a
5244 given name in the context. These may be later included in other
5245 contexts. This allows to manage context extensions casually, without
5246 the logical dependencies of locales and locale interpretation. See
5247 commands 'bundle', 'include', 'including' etc. in the isar-ref manual.
5249 * Commands 'lemmas' and 'theorems' allow local variables using 'for'
5250 declaration, and results are standardized before being stored. Thus
5251 old-style "standard" after instantiation or composition of facts
5252 becomes obsolete. Minor INCOMPATIBILITY, due to potential change of
5253 indices of schematic variables.
5255 * Rule attributes in local theory declarations (e.g. locale or class)
5256 are now statically evaluated: the resulting theorem is stored instead
5257 of the original expression. INCOMPATIBILITY in rare situations, where
5258 the historic accident of dynamic re-evaluation in interpretations
5261 * New tutorial "Programming and Proving in Isabelle/HOL"
5262 ("prog-prove"). It completely supersedes "A Tutorial Introduction to
5263 Structured Isar Proofs" ("isar-overview"), which has been removed. It
5264 also supersedes "Isabelle/HOL, A Proof Assistant for Higher-Order
5265 Logic" as the recommended beginners tutorial, but does not cover all
5266 of the material of that old tutorial.
5268 * Updated and extended reference manuals: "isar-ref",
5269 "implementation", "system"; reduced remaining material in old "ref"
5275 * Command 'definition' no longer exports the foundational "raw_def"
5276 into the user context. Minor INCOMPATIBILITY, may use the regular
5277 "def" result with attribute "abs_def" to imitate the old version.
5279 * Attribute "abs_def" turns an equation of the form "f x y == t" into
5280 "f == %x y. t", which ensures that "simp" or "unfold" steps always
5281 expand it. This also works for object-logic equality. (Formerly
5282 undocumented feature.)
5284 * Sort constraints are now propagated in simultaneous statements, just
5285 like type constraints. INCOMPATIBILITY in rare situations, where
5286 distinct sorts used to be assigned accidentally. For example:
5288 lemma "P (x::'a::foo)" and "Q (y::'a::bar)" -- "now illegal"
5290 lemma "P (x::'a)" and "Q (y::'a::bar)"
5291 -- "now uniform 'a::bar instead of default sort for first occurrence (!)"
5293 * Rule composition via attribute "OF" (or ML functions OF/MRS) is more
5294 tolerant against multiple unifiers, as long as the final result is
5295 unique. (As before, rules are composed in canonical right-to-left
5296 order to accommodate newly introduced premises.)
5298 * Renamed some inner syntax categories:
5304 Minor INCOMPATIBILITY. Note that in practice "num_const" or
5305 "num_position" etc. are mainly used instead (which also include
5306 position information via constraints).
5308 * Simplified configuration options for syntax ambiguity: see
5309 "syntax_ambiguity_warning" and "syntax_ambiguity_limit" in isar-ref
5310 manual. Minor INCOMPATIBILITY.
5312 * Discontinued configuration option "syntax_positions": atomic terms
5313 in parse trees are always annotated by position constraints.
5315 * Old code generator for SML and its commands 'code_module',
5316 'code_library', 'consts_code', 'types_code' have been discontinued.
5317 Use commands of the generic code generator instead. INCOMPATIBILITY.
5319 * Redundant attribute "code_inline" has been discontinued. Use
5320 "code_unfold" instead. INCOMPATIBILITY.
5322 * Dropped attribute "code_unfold_post" in favor of the its dual
5323 "code_abbrev", which yields a common pattern in definitions like
5325 definition [code_abbrev]: "f = t"
5329 * Obsolete 'types' command has been discontinued. Use 'type_synonym'
5330 instead. INCOMPATIBILITY.
5332 * Discontinued old "prems" fact, which used to refer to the accidental
5333 collection of foundational premises in the context (already marked as
5334 legacy since Isabelle2011).
5339 * Type 'a set is now a proper type constructor (just as before
5340 Isabelle2008). Definitions mem_def and Collect_def have disappeared.
5341 Non-trivial INCOMPATIBILITY. For developments keeping predicates and
5342 sets separate, it is often sufficient to rephrase some set S that has
5343 been accidentally used as predicates by "%x. x : S", and some
5344 predicate P that has been accidentally used as set by "{x. P x}".
5345 Corresponding proofs in a first step should be pruned from any
5346 tinkering with former theorems mem_def and Collect_def as far as
5349 For developments which deliberately mix predicates and sets, a
5350 planning step is necessary to determine what should become a predicate
5351 and what a set. It can be helpful to carry out that step in
5352 Isabelle2011-1 before jumping right into the current release.
5354 * Code generation by default implements sets as container type rather
5355 than predicates. INCOMPATIBILITY.
5357 * New type synonym 'a rel = ('a * 'a) set
5359 * The representation of numerals has changed. Datatype "num"
5360 represents strictly positive binary numerals, along with functions
5361 "numeral :: num => 'a" and "neg_numeral :: num => 'a" to represent
5362 positive and negated numeric literals, respectively. See also
5363 definitions in ~~/src/HOL/Num.thy. Potential INCOMPATIBILITY, some
5364 user theories may require adaptations as follows:
5366 - Theorems with number_ring or number_semiring constraints: These
5367 classes are gone; use comm_ring_1 or comm_semiring_1 instead.
5369 - Theories defining numeric types: Remove number, number_semiring,
5370 and number_ring instances. Defer all theorems about numerals until
5371 after classes one and semigroup_add have been instantiated.
5373 - Numeral-only simp rules: Replace each rule having a "number_of v"
5374 pattern with two copies, one for numeral and one for neg_numeral.
5376 - Theorems about subclasses of semiring_1 or ring_1: These classes
5377 automatically support numerals now, so more simp rules and
5378 simprocs may now apply within the proof.
5380 - Definitions and theorems using old constructors Pls/Min/Bit0/Bit1:
5381 Redefine using other integer operations.
5383 * Transfer: New package intended to generalize the existing
5384 "descending" method and related theorem attributes from the Quotient
5385 package. (Not all functionality is implemented yet, but future
5386 development will focus on Transfer as an eventual replacement for the
5387 corresponding parts of the Quotient package.)
5389 - transfer_rule attribute: Maintains a collection of transfer rules,
5390 which relate constants at two different types. Transfer rules may
5391 relate different type instances of the same polymorphic constant,
5392 or they may relate an operation on a raw type to a corresponding
5393 operation on an abstract type (quotient or subtype). For example:
5395 ((A ===> B) ===> list_all2 A ===> list_all2 B) map map
5396 (cr_int ===> cr_int ===> cr_int) (%(x,y) (u,v). (x+u, y+v)) plus_int
5398 - transfer method: Replaces a subgoal on abstract types with an
5399 equivalent subgoal on the corresponding raw types. Constants are
5400 replaced with corresponding ones according to the transfer rules.
5401 Goals are generalized over all free variables by default; this is
5402 necessary for variables whose types change, but can be overridden
5403 for specific variables with e.g. "transfer fixing: x y z". The
5404 variant transfer' method allows replacing a subgoal with one that
5405 is logically stronger (rather than equivalent).
5407 - relator_eq attribute: Collects identity laws for relators of
5408 various type constructors, e.g. "list_all2 (op =) = (op =)". The
5409 transfer method uses these lemmas to infer transfer rules for
5410 non-polymorphic constants on the fly.
5412 - transfer_prover method: Assists with proving a transfer rule for a
5413 new constant, provided the constant is defined in terms of other
5414 constants that already have transfer rules. It should be applied
5415 after unfolding the constant definitions.
5417 - HOL/ex/Transfer_Int_Nat.thy: Example theory demonstrating transfer
5418 from type nat to type int.
5420 * Lifting: New package intended to generalize the quotient_definition
5421 facility of the Quotient package; designed to work with Transfer.
5423 - lift_definition command: Defines operations on an abstract type in
5424 terms of a corresponding operation on a representation
5425 type. Example syntax:
5427 lift_definition dlist_insert :: "'a => 'a dlist => 'a dlist"
5430 Users must discharge a respectfulness proof obligation when each
5431 constant is defined. (For a type copy, i.e. a typedef with UNIV,
5432 the proof is discharged automatically.) The obligation is
5433 presented in a user-friendly, readable form; a respectfulness
5434 theorem in the standard format and a transfer rule are generated
5437 - Integration with code_abstype: For typedefs (e.g. subtypes
5438 corresponding to a datatype invariant, such as dlist),
5439 lift_definition generates a code certificate theorem and sets up
5440 code generation for each constant.
5442 - setup_lifting command: Sets up the Lifting package to work with a
5443 user-defined type. The user must provide either a quotient theorem
5444 or a type_definition theorem. The package configures transfer
5445 rules for equality and quantifiers on the type, and sets up the
5446 lift_definition command to work with the type.
5448 - Usage examples: See Quotient_Examples/Lift_DList.thy,
5449 Quotient_Examples/Lift_RBT.thy, Quotient_Examples/Lift_FSet.thy,
5450 Word/Word.thy and Library/Float.thy.
5454 - The 'quotient_type' command now supports a 'morphisms' option with
5455 rep and abs functions, similar to typedef.
5457 - 'quotient_type' sets up new types to work with the Lifting and
5458 Transfer packages, as with 'setup_lifting'.
5460 - The 'quotient_definition' command now requires the user to prove a
5461 respectfulness property at the point where the constant is
5462 defined, similar to lift_definition; INCOMPATIBILITY.
5464 - Renamed predicate 'Quotient' to 'Quotient3', and renamed theorems
5465 accordingly, INCOMPATIBILITY.
5467 * New diagnostic command 'find_unused_assms' to find potentially
5468 superfluous assumptions in theorems using Quickcheck.
5472 - Quickcheck returns variable assignments as counterexamples, which
5473 allows to reveal the underspecification of functions under test.
5474 For example, refuting "hd xs = x", it presents the variable
5475 assignment xs = [] and x = a1 as a counterexample, assuming that
5476 any property is false whenever "hd []" occurs in it.
5478 These counterexample are marked as potentially spurious, as
5479 Quickcheck also returns "xs = []" as a counterexample to the
5480 obvious theorem "hd xs = hd xs".
5482 After finding a potentially spurious counterexample, Quickcheck
5483 continues searching for genuine ones.
5485 By default, Quickcheck shows potentially spurious and genuine
5486 counterexamples. The option "genuine_only" sets quickcheck to only
5487 show genuine counterexamples.
5489 - The command 'quickcheck_generator' creates random and exhaustive
5490 value generators for a given type and operations.
5492 It generates values by using the operations as if they were
5493 constructors of that type.
5495 - Support for multisets.
5497 - Added "use_subtype" options.
5499 - Added "quickcheck_locale" configuration to specify how to process
5500 conjectures in a locale context.
5502 * Nitpick: Fixed infinite loop caused by the 'peephole_optim' option
5503 and affecting 'rat' and 'real'.
5506 - Integrated more tightly with SPASS, as described in the ITP 2012
5507 paper "More SPASS with Isabelle".
5508 - Made it try "smt" as a fallback if "metis" fails or times out.
5509 - Added support for the following provers: Alt-Ergo (via Why3 and
5510 TFF1), iProver, iProver-Eq.
5511 - Sped up the minimizer.
5512 - Added "lam_trans", "uncurry_aliases", and "minimize" options.
5513 - Renamed "slicing" ("no_slicing") option to "slice" ("dont_slice").
5514 - Renamed "sound" option to "strict".
5516 * Metis: Added possibility to specify lambda translations scheme as a
5517 parenthesized argument (e.g., "by (metis (lifting) ...)").
5519 * SMT: Renamed "smt_fixed" option to "smt_read_only_certificates".
5521 * Command 'try0': Renamed from 'try_methods'. INCOMPATIBILITY.
5523 * New "case_product" attribute to generate a case rule doing multiple
5524 case distinctions at the same time. E.g.
5526 list.exhaust [case_product nat.exhaust]
5528 produces a rule which can be used to perform case distinction on both
5531 * New "eventually_elim" method as a generalized variant of the
5532 eventually_elim* rules. Supports structured proofs.
5534 * Typedef with implicit set definition is considered legacy. Use
5535 "typedef (open)" form instead, which will eventually become the
5538 * Record: code generation can be switched off manually with
5540 declare [[record_coden = false]] -- "default true"
5542 * Datatype: type parameters allow explicit sort constraints.
5544 * Concrete syntax for case expressions includes constraints for source
5545 positions, and thus produces Prover IDE markup for its bindings.
5546 INCOMPATIBILITY for old-style syntax translations that augment the
5547 pattern notation; e.g. see src/HOL/HOLCF/One.thy for translations of
5550 * Clarified attribute "mono_set": pure declaration without modifying
5551 the result of the fact expression.
5553 * More default pred/set conversions on a couple of relation operations
5554 and predicates. Added powers of predicate relations. Consolidation
5555 of some relation theorems:
5557 converse_def ~> converse_unfold
5558 rel_comp_def ~> relcomp_unfold
5559 symp_def ~> (modified, use symp_def and sym_def instead)
5560 transp_def ~> transp_trans
5561 Domain_def ~> Domain_unfold
5562 Range_def ~> Domain_converse [symmetric]
5564 Generalized theorems INF_INT_eq, INF_INT_eq2, SUP_UN_eq, SUP_UN_eq2.
5566 See theory "Relation" for examples for making use of pred/set
5567 conversions by means of attributes "to_set" and "to_pred".
5571 * Renamed facts about the power operation on relations, i.e., relpow
5572 to match the constant's name:
5574 rel_pow_1 ~> relpow_1
5575 rel_pow_0_I ~> relpow_0_I
5576 rel_pow_Suc_I ~> relpow_Suc_I
5577 rel_pow_Suc_I2 ~> relpow_Suc_I2
5578 rel_pow_0_E ~> relpow_0_E
5579 rel_pow_Suc_E ~> relpow_Suc_E
5580 rel_pow_E ~> relpow_E
5581 rel_pow_Suc_D2 ~> relpow_Suc_D2
5582 rel_pow_Suc_E2 ~> relpow_Suc_E2
5583 rel_pow_Suc_D2' ~> relpow_Suc_D2'
5584 rel_pow_E2 ~> relpow_E2
5585 rel_pow_add ~> relpow_add
5586 rel_pow_commute ~> relpow
5587 rel_pow_empty ~> relpow_empty:
5588 rtrancl_imp_UN_rel_pow ~> rtrancl_imp_UN_relpow
5589 rel_pow_imp_rtrancl ~> relpow_imp_rtrancl
5590 rtrancl_is_UN_rel_pow ~> rtrancl_is_UN_relpow
5591 rtrancl_imp_rel_pow ~> rtrancl_imp_relpow
5592 rel_pow_fun_conv ~> relpow_fun_conv
5593 rel_pow_finite_bounded1 ~> relpow_finite_bounded1
5594 rel_pow_finite_bounded ~> relpow_finite_bounded
5595 rtrancl_finite_eq_rel_pow ~> rtrancl_finite_eq_relpow
5596 trancl_finite_eq_rel_pow ~> trancl_finite_eq_relpow
5597 single_valued_rel_pow ~> single_valued_relpow
5601 * Theory Relation: Consolidated constant name for relation composition
5602 and corresponding theorem names:
5604 - Renamed constant rel_comp to relcomp.
5606 - Dropped abbreviation pred_comp. Use relcompp instead.
5610 rel_compI ~> relcompI
5611 rel_compEpair ~> relcompEpair
5612 rel_compE ~> relcompE
5613 pred_comp_rel_comp_eq ~> relcompp_relcomp_eq
5614 rel_comp_empty1 ~> relcomp_empty1
5615 rel_comp_mono ~> relcomp_mono
5616 rel_comp_subset_Sigma ~> relcomp_subset_Sigma
5617 rel_comp_distrib ~> relcomp_distrib
5618 rel_comp_distrib2 ~> relcomp_distrib2
5619 rel_comp_UNION_distrib ~> relcomp_UNION_distrib
5620 rel_comp_UNION_distrib2 ~> relcomp_UNION_distrib2
5621 single_valued_rel_comp ~> single_valued_relcomp
5622 rel_comp_def ~> relcomp_unfold
5623 converse_rel_comp ~> converse_relcomp
5624 pred_compI ~> relcomppI
5625 pred_compE ~> relcomppE
5626 pred_comp_bot1 ~> relcompp_bot1
5627 pred_comp_bot2 ~> relcompp_bot2
5628 transp_pred_comp_less_eq ~> transp_relcompp_less_eq
5629 pred_comp_mono ~> relcompp_mono
5630 pred_comp_distrib ~> relcompp_distrib
5631 pred_comp_distrib2 ~> relcompp_distrib2
5632 converse_pred_comp ~> converse_relcompp
5634 finite_rel_comp ~> finite_relcomp
5636 set_rel_comp ~> set_relcomp
5640 * Theory Divides: Discontinued redundant theorems about div and mod.
5641 INCOMPATIBILITY, use the corresponding generic theorems instead.
5643 DIVISION_BY_ZERO ~> div_by_0, mod_by_0
5644 zdiv_self ~> div_self
5645 zmod_self ~> mod_self
5648 zdiv_zmod_equality ~> div_mod_equality2
5649 zdiv_zmod_equality2 ~> div_mod_equality
5650 zmod_zdiv_trivial ~> mod_div_trivial
5651 zdiv_zminus_zminus ~> div_minus_minus
5652 zmod_zminus_zminus ~> mod_minus_minus
5653 zdiv_zminus2 ~> div_minus_right
5654 zmod_zminus2 ~> mod_minus_right
5655 zdiv_minus1_right ~> div_minus1_right
5656 zmod_minus1_right ~> mod_minus1_right
5657 zdvd_mult_div_cancel ~> dvd_mult_div_cancel
5658 zmod_zmult1_eq ~> mod_mult_right_eq
5659 zpower_zmod ~> power_mod
5660 zdvd_zmod ~> dvd_mod
5661 zdvd_zmod_imp_zdvd ~> dvd_mod_imp_dvd
5662 mod_mult_distrib ~> mult_mod_left
5663 mod_mult_distrib2 ~> mult_mod_right
5665 * Removed redundant theorems nat_mult_2 and nat_mult_2_right; use
5666 generic mult_2 and mult_2_right instead. INCOMPATIBILITY.
5668 * Finite_Set.fold now qualified. INCOMPATIBILITY.
5670 * Consolidated theorem names concerning fold combinators:
5672 inf_INFI_fold_inf ~> inf_INF_fold_inf
5673 sup_SUPR_fold_sup ~> sup_SUP_fold_sup
5674 INFI_fold_inf ~> INF_fold_inf
5675 SUPR_fold_sup ~> SUP_fold_sup
5676 union_set ~> union_set_fold
5677 minus_set ~> minus_set_fold
5678 INFI_set_fold ~> INF_set_fold
5679 SUPR_set_fold ~> SUP_set_fold
5680 INF_code ~> INF_set_foldr
5681 SUP_code ~> SUP_set_foldr
5682 foldr.simps ~> foldr.simps (in point-free formulation)
5683 foldr_fold_rev ~> foldr_conv_fold
5684 foldl_fold ~> foldl_conv_fold
5685 foldr_foldr ~> foldr_conv_foldl
5686 foldl_foldr ~> foldl_conv_foldr
5687 fold_set_remdups ~> fold_set_fold_remdups
5688 fold_set ~> fold_set_fold
5689 fold1_set ~> fold1_set_fold
5693 * Dropped rarely useful theorems concerning fold combinators:
5694 foldl_apply, foldl_fun_comm, foldl_rev, fold_weak_invariant,
5695 rev_foldl_cons, fold_set_remdups, fold_set, fold_set1,
5696 concat_conv_foldl, foldl_weak_invariant, foldl_invariant,
5697 foldr_invariant, foldl_absorb0, foldl_foldr1_lemma, foldl_foldr1,
5698 listsum_conv_fold, listsum_foldl, sort_foldl_insort, foldl_assoc,
5699 foldr_conv_foldl, start_le_sum, elem_le_sum, sum_eq_0_conv.
5700 INCOMPATIBILITY. For the common phrases "%xs. List.foldr plus xs 0"
5701 and "List.foldl plus 0", prefer "List.listsum". Otherwise it can be
5702 useful to boil down "List.foldr" and "List.foldl" to "List.fold" by
5703 unfolding "foldr_conv_fold" and "foldl_conv_fold".
5705 * Dropped lemmas minus_set_foldr, union_set_foldr, union_coset_foldr,
5706 inter_coset_foldr, Inf_fin_set_foldr, Sup_fin_set_foldr,
5707 Min_fin_set_foldr, Max_fin_set_foldr, Inf_set_foldr, Sup_set_foldr,
5708 INF_set_foldr, SUP_set_foldr. INCOMPATIBILITY. Prefer corresponding
5709 lemmas over fold rather than foldr, or make use of lemmas
5710 fold_conv_foldr and fold_rev.
5712 * Congruence rules Option.map_cong and Option.bind_cong for recursion
5713 through option types.
5715 * "Transitive_Closure.ntrancl": bounded transitive closure on
5718 * Constant "Set.not_member" now qualified. INCOMPATIBILITY.
5720 * Theory Int: Discontinued many legacy theorems specific to type int.
5721 INCOMPATIBILITY, use the corresponding generic theorems instead.
5723 zminus_zminus ~> minus_minus
5724 zminus_0 ~> minus_zero
5725 zminus_zadd_distrib ~> minus_add_distrib
5726 zadd_commute ~> add_commute
5727 zadd_assoc ~> add_assoc
5728 zadd_left_commute ~> add_left_commute
5731 zadd_0 ~> add_0_left
5732 zadd_0_right ~> add_0_right
5733 zadd_zminus_inverse2 ~> left_minus
5734 zmult_zminus ~> mult_minus_left
5735 zmult_commute ~> mult_commute
5736 zmult_assoc ~> mult_assoc
5737 zadd_zmult_distrib ~> left_distrib
5738 zadd_zmult_distrib2 ~> right_distrib
5739 zdiff_zmult_distrib ~> left_diff_distrib
5740 zdiff_zmult_distrib2 ~> right_diff_distrib
5741 zmult_1 ~> mult_1_left
5742 zmult_1_right ~> mult_1_right
5743 zle_refl ~> order_refl
5744 zle_trans ~> order_trans
5745 zle_antisym ~> order_antisym
5746 zle_linear ~> linorder_linear
5747 zless_linear ~> linorder_less_linear
5748 zadd_left_mono ~> add_left_mono
5749 zadd_strict_right_mono ~> add_strict_right_mono
5750 zadd_zless_mono ~> add_less_le_mono
5751 int_0_less_1 ~> zero_less_one
5752 int_0_neq_1 ~> zero_neq_one
5754 zpower_zadd_distrib ~> power_add
5755 zero_less_zpower_abs_iff ~> zero_less_power_abs_iff
5756 zero_le_zpower_abs ~> zero_le_power_abs
5758 * Theory Deriv: Renamed
5760 DERIV_nonneg_imp_nonincreasing ~> DERIV_nonneg_imp_nondecreasing
5762 * Theory Library/Multiset: Improved code generation of multisets.
5764 * Theory HOL/Library/Set_Algebras: Addition and multiplication on sets
5765 are expressed via type classes again. The special syntax
5766 \<oplus>/\<otimes> has been replaced by plain +/*. Removed constant
5767 setsum_set, which is now subsumed by Big_Operators.setsum.
5770 * Theory HOL/Library/Diagonalize has been removed. INCOMPATIBILITY,
5771 use theory HOL/Library/Nat_Bijection instead.
5773 * Theory HOL/Library/RBT_Impl: Backing implementation of red-black
5774 trees is now inside a type class context. Names of affected
5775 operations and lemmas have been prefixed by rbt_. INCOMPATIBILITY for
5776 theories working directly with raw red-black trees, adapt the names as
5780 bulkload -> rbt_bulkload
5781 del_from_left -> rbt_del_from_left
5782 del_from_right -> rbt_del_from_right
5784 delete -> rbt_delete
5786 insert -> rbt_insert
5787 insertw -> rbt_insert_with
5788 insert_with_key -> rbt_insert_with_key
5789 map_entry -> rbt_map_entry
5790 lookup -> rbt_lookup
5791 sorted -> rbt_sorted
5792 tree_greater -> rbt_greater
5793 tree_less -> rbt_less
5794 tree_less_symbol -> rbt_less_symbol
5796 union_with -> rbt_union_with
5797 union_with_key -> rbt_union_with_key
5800 balance_left_sorted -> balance_left_rbt_sorted
5801 balance_left_tree_greater -> balance_left_rbt_greater
5802 balance_left_tree_less -> balance_left_rbt_less
5803 balance_right_sorted -> balance_right_rbt_sorted
5804 balance_right_tree_greater -> balance_right_rbt_greater
5805 balance_right_tree_less -> balance_right_rbt_less
5806 balance_sorted -> balance_rbt_sorted
5807 balance_tree_greater -> balance_rbt_greater
5808 balance_tree_less -> balance_rbt_less
5809 bulkload_is_rbt -> rbt_bulkload_is_rbt
5810 combine_sorted -> combine_rbt_sorted
5811 combine_tree_greater -> combine_rbt_greater
5812 combine_tree_less -> combine_rbt_less
5813 delete_in_tree -> rbt_delete_in_tree
5814 delete_is_rbt -> rbt_delete_is_rbt
5815 del_from_left_tree_greater -> rbt_del_from_left_rbt_greater
5816 del_from_left_tree_less -> rbt_del_from_left_rbt_less
5817 del_from_right_tree_greater -> rbt_del_from_right_rbt_greater
5818 del_from_right_tree_less -> rbt_del_from_right_rbt_less
5819 del_in_tree -> rbt_del_in_tree
5820 del_inv1_inv2 -> rbt_del_inv1_inv2
5821 del_sorted -> rbt_del_rbt_sorted
5822 del_tree_greater -> rbt_del_rbt_greater
5823 del_tree_less -> rbt_del_rbt_less
5824 dom_lookup_Branch -> dom_rbt_lookup_Branch
5825 entries_lookup -> entries_rbt_lookup
5826 finite_dom_lookup -> finite_dom_rbt_lookup
5827 insert_sorted -> rbt_insert_rbt_sorted
5828 insertw_is_rbt -> rbt_insertw_is_rbt
5829 insertwk_is_rbt -> rbt_insertwk_is_rbt
5830 insertwk_sorted -> rbt_insertwk_rbt_sorted
5831 insertw_sorted -> rbt_insertw_rbt_sorted
5832 ins_sorted -> ins_rbt_sorted
5833 ins_tree_greater -> ins_rbt_greater
5834 ins_tree_less -> ins_rbt_less
5835 is_rbt_sorted -> is_rbt_rbt_sorted
5836 lookup_balance -> rbt_lookup_balance
5837 lookup_bulkload -> rbt_lookup_rbt_bulkload
5838 lookup_delete -> rbt_lookup_rbt_delete
5839 lookup_Empty -> rbt_lookup_Empty
5840 lookup_from_in_tree -> rbt_lookup_from_in_tree
5841 lookup_in_tree -> rbt_lookup_in_tree
5842 lookup_ins -> rbt_lookup_ins
5843 lookup_insert -> rbt_lookup_rbt_insert
5844 lookup_insertw -> rbt_lookup_rbt_insertw
5845 lookup_insertwk -> rbt_lookup_rbt_insertwk
5846 lookup_keys -> rbt_lookup_keys
5847 lookup_map -> rbt_lookup_map
5848 lookup_map_entry -> rbt_lookup_rbt_map_entry
5849 lookup_tree_greater -> rbt_lookup_rbt_greater
5850 lookup_tree_less -> rbt_lookup_rbt_less
5851 lookup_union -> rbt_lookup_rbt_union
5852 map_entry_color_of -> rbt_map_entry_color_of
5853 map_entry_inv1 -> rbt_map_entry_inv1
5854 map_entry_inv2 -> rbt_map_entry_inv2
5855 map_entry_is_rbt -> rbt_map_entry_is_rbt
5856 map_entry_sorted -> rbt_map_entry_rbt_sorted
5857 map_entry_tree_greater -> rbt_map_entry_rbt_greater
5858 map_entry_tree_less -> rbt_map_entry_rbt_less
5859 map_tree_greater -> map_rbt_greater
5860 map_tree_less -> map_rbt_less
5861 map_sorted -> map_rbt_sorted
5862 paint_sorted -> paint_rbt_sorted
5863 paint_lookup -> paint_rbt_lookup
5864 paint_tree_greater -> paint_rbt_greater
5865 paint_tree_less -> paint_rbt_less
5866 sorted_entries -> rbt_sorted_entries
5867 tree_greater_eq_trans -> rbt_greater_eq_trans
5868 tree_greater_nit -> rbt_greater_nit
5869 tree_greater_prop -> rbt_greater_prop
5870 tree_greater_simps -> rbt_greater_simps
5871 tree_greater_trans -> rbt_greater_trans
5872 tree_less_eq_trans -> rbt_less_eq_trans
5873 tree_less_nit -> rbt_less_nit
5874 tree_less_prop -> rbt_less_prop
5875 tree_less_simps -> rbt_less_simps
5876 tree_less_trans -> rbt_less_trans
5877 tree_ord_props -> rbt_ord_props
5878 union_Branch -> rbt_union_Branch
5879 union_is_rbt -> rbt_union_is_rbt
5880 unionw_is_rbt -> rbt_unionw_is_rbt
5881 unionwk_is_rbt -> rbt_unionwk_is_rbt
5882 unionwk_sorted -> rbt_unionwk_rbt_sorted
5884 * Theory HOL/Library/Float: Floating point numbers are now defined as
5885 a subset of the real numbers. All operations are defined using the
5886 lifing-framework and proofs use the transfer method. INCOMPATIBILITY.
5893 round_down -> float_round_down
5894 round_up -> float_round_up
5898 ceiling_fl, lb_mult, lb_mod, ub_mult, ub_mod
5901 abs_float_def -> Float.compute_float_abs
5902 bitlen_ge0 -> bitlen_nonneg
5903 bitlen.simps -> Float.compute_bitlen
5904 float_components -> Float_mantissa_exponent
5905 float_divl.simps -> Float.compute_float_divl
5906 float_divr.simps -> Float.compute_float_divr
5907 float_eq_odd -> mult_powr_eq_mult_powr_iff
5908 float_power -> real_of_float_power
5909 lapprox_posrat_def -> Float.compute_lapprox_posrat
5910 lapprox_rat.simps -> Float.compute_lapprox_rat
5911 le_float_def' -> Float.compute_float_le
5912 le_float_def -> less_eq_float.rep_eq
5913 less_float_def' -> Float.compute_float_less
5914 less_float_def -> less_float.rep_eq
5915 normfloat_def -> Float.compute_normfloat
5916 normfloat_imp_odd_or_zero -> mantissa_not_dvd and mantissa_noteq_0
5917 normfloat -> normfloat_def
5918 normfloat_unique -> use normfloat_def
5919 number_of_float_Float -> Float.compute_float_numeral, Float.compute_float_neg_numeral
5920 one_float_def -> Float.compute_float_one
5921 plus_float_def -> Float.compute_float_plus
5922 rapprox_posrat_def -> Float.compute_rapprox_posrat
5923 rapprox_rat.simps -> Float.compute_rapprox_rat
5924 real_of_float_0 -> zero_float.rep_eq
5925 real_of_float_1 -> one_float.rep_eq
5926 real_of_float_abs -> abs_float.rep_eq
5927 real_of_float_add -> plus_float.rep_eq
5928 real_of_float_minus -> uminus_float.rep_eq
5929 real_of_float_mult -> times_float.rep_eq
5930 real_of_float_simp -> Float.rep_eq
5931 real_of_float_sub -> minus_float.rep_eq
5932 round_down.simps -> Float.compute_float_round_down
5933 round_up.simps -> Float.compute_float_round_up
5934 times_float_def -> Float.compute_float_times
5935 uminus_float_def -> Float.compute_float_uminus
5936 zero_float_def -> Float.compute_float_zero
5938 Lemmas not necessary anymore, use the transfer method:
5939 bitlen_B0, bitlen_B1, bitlen_ge1, bitlen_Min, bitlen_Pls, float_divl,
5940 float_divr, float_le_simp, float_less1_mantissa_bound,
5941 float_less_simp, float_less_zero, float_le_zero,
5942 float_pos_less1_e_neg, float_pos_m_pos, float_split, float_split2,
5943 floor_pos_exp, lapprox_posrat, lapprox_posrat_bottom, lapprox_rat,
5944 lapprox_rat_bottom, normalized_float, rapprox_posrat,
5945 rapprox_posrat_le1, rapprox_rat, real_of_float_ge0_exp,
5946 real_of_float_neg_exp, real_of_float_nge0_exp, round_down floor_fl,
5947 round_up, zero_le_float, zero_less_float
5949 * New theory HOL/Library/DAList provides an abstract type for
5950 association lists with distinct keys.
5952 * Session HOL/IMP: Added new theory of abstract interpretation of
5955 * Session HOL-Import: Re-implementation from scratch is faster,
5956 simpler, and more scalable. Requires a proof bundle, which is
5957 available as an external component. Discontinued old (and mostly
5958 dead) Importer for HOL4 and HOL Light. INCOMPATIBILITY.
5960 * Session HOL-Word: Discontinued many redundant theorems specific to
5961 type 'a word. INCOMPATIBILITY, use the corresponding generic theorems
5964 word_sub_alt ~> word_sub_wi
5965 word_add_alt ~> word_add_def
5966 word_mult_alt ~> word_mult_def
5967 word_minus_alt ~> word_minus_def
5968 word_0_alt ~> word_0_wi
5969 word_1_alt ~> word_1_wi
5970 word_add_0 ~> add_0_left
5971 word_add_0_right ~> add_0_right
5972 word_mult_1 ~> mult_1_left
5973 word_mult_1_right ~> mult_1_right
5974 word_add_commute ~> add_commute
5975 word_add_assoc ~> add_assoc
5976 word_add_left_commute ~> add_left_commute
5977 word_mult_commute ~> mult_commute
5978 word_mult_assoc ~> mult_assoc
5979 word_mult_left_commute ~> mult_left_commute
5980 word_left_distrib ~> left_distrib
5981 word_right_distrib ~> right_distrib
5982 word_left_minus ~> left_minus
5983 word_diff_0_right ~> diff_0_right
5984 word_diff_self ~> diff_self
5985 word_sub_def ~> diff_minus
5986 word_diff_minus ~> diff_minus
5987 word_add_ac ~> add_ac
5988 word_mult_ac ~> mult_ac
5989 word_plus_ac0 ~> add_0_left add_0_right add_ac
5990 word_times_ac1 ~> mult_1_left mult_1_right mult_ac
5991 word_order_trans ~> order_trans
5992 word_order_refl ~> order_refl
5993 word_order_antisym ~> order_antisym
5994 word_order_linear ~> linorder_linear
5995 lenw1_zero_neq_one ~> zero_neq_one
5996 word_number_of_eq ~> number_of_eq
5997 word_of_int_add_hom ~> wi_hom_add
5998 word_of_int_sub_hom ~> wi_hom_sub
5999 word_of_int_mult_hom ~> wi_hom_mult
6000 word_of_int_minus_hom ~> wi_hom_neg
6001 word_of_int_succ_hom ~> wi_hom_succ
6002 word_of_int_pred_hom ~> wi_hom_pred
6003 word_of_int_0_hom ~> word_0_wi
6004 word_of_int_1_hom ~> word_1_wi
6006 * Session HOL-Word: New proof method "word_bitwise" for splitting
6007 machine word equalities and inequalities into logical circuits,
6008 defined in HOL/Word/WordBitwise.thy. Supports addition, subtraction,
6009 multiplication, shifting by constants, bitwise operators and numeric
6010 constants. Requires fixed-length word types, not 'a word. Solves
6011 many standard word identities outright and converts more into first
6012 order problems amenable to blast or similar. See also examples in
6013 HOL/Word/Examples/WordExamples.thy.
6015 * Session HOL-Probability: Introduced the type "'a measure" to
6016 represent measures, this replaces the records 'a algebra and 'a
6017 measure_space. The locales based on subset_class now have two
6018 locale-parameters the space \<Omega> and the set of measurable sets M.
6019 The product of probability spaces uses now the same constant as the
6020 finite product of sigma-finite measure spaces "PiM :: ('i => 'a)
6021 measure". Most constants are defined now outside of locales and gain
6022 an additional parameter, like null_sets, almost_eventually or \<mu>'.
6023 Measure space constructions for distributions and densities now got
6024 their own constants distr and density. Instead of using locales to
6025 describe measure spaces with a finite space, the measure count_space
6026 and point_measure is introduced. INCOMPATIBILITY.
6030 finite_measure.\<mu>' -> measure
6031 product_algebra_generator -> prod_algebra
6032 product_prob_space.emb -> prod_emb
6033 product_prob_space.infprod_algebra -> PiM
6036 completeable_measure_space
6037 finite_measure_space
6039 finite_product_finite_prob_space
6040 finite_product_sigma_algebra
6041 finite_sigma_algebra
6043 pair_finite_prob_space
6044 pair_finite_sigma_algebra
6047 product_sigma_algebra
6051 distribution -> use distr measure, or distributed predicate
6053 joint_distribution -> use distr measure, or distributed predicate
6054 pair_measure_generator
6055 product_prob_space.infprod_algebra -> use PiM
6058 Replacement theorems:
6059 finite_additivity_sufficient -> ring_of_sets.countably_additiveI_finite
6060 finite_measure.empty_measure -> measure_empty
6061 finite_measure.finite_continuity_from_above -> finite_measure.finite_Lim_measure_decseq
6062 finite_measure.finite_continuity_from_below -> finite_measure.finite_Lim_measure_incseq
6063 finite_measure.finite_measure_countably_subadditive -> finite_measure.finite_measure_subadditive_countably
6064 finite_measure.finite_measure_eq -> finite_measure.emeasure_eq_measure
6065 finite_measure.finite_measure -> finite_measure.emeasure_finite
6066 finite_measure.finite_measure_finite_singleton -> finite_measure.finite_measure_eq_setsum_singleton
6067 finite_measure.positive_measure' -> measure_nonneg
6068 finite_measure.real_measure -> finite_measure.emeasure_real
6069 finite_product_prob_space.finite_measure_times -> finite_product_prob_space.finite_measure_PiM_emb
6070 finite_product_sigma_algebra.in_P -> sets_PiM_I_finite
6071 finite_product_sigma_algebra.P_empty -> space_PiM_empty, sets_PiM_empty
6072 information_space.conditional_entropy_eq -> information_space.conditional_entropy_simple_distributed
6073 information_space.conditional_entropy_positive -> information_space.conditional_entropy_nonneg_simple
6074 information_space.conditional_mutual_information_eq_mutual_information -> information_space.conditional_mutual_information_eq_mutual_information_simple
6075 information_space.conditional_mutual_information_generic_positive -> information_space.conditional_mutual_information_nonneg_simple
6076 information_space.conditional_mutual_information_positive -> information_space.conditional_mutual_information_nonneg_simple
6077 information_space.entropy_commute -> information_space.entropy_commute_simple
6078 information_space.entropy_eq -> information_space.entropy_simple_distributed
6079 information_space.entropy_generic_eq -> information_space.entropy_simple_distributed
6080 information_space.entropy_positive -> information_space.entropy_nonneg_simple
6081 information_space.entropy_uniform_max -> information_space.entropy_uniform
6082 information_space.KL_eq_0_imp -> information_space.KL_eq_0_iff_eq
6083 information_space.KL_eq_0 -> information_space.KL_same_eq_0
6084 information_space.KL_ge_0 -> information_space.KL_nonneg
6085 information_space.mutual_information_eq -> information_space.mutual_information_simple_distributed
6086 information_space.mutual_information_positive -> information_space.mutual_information_nonneg_simple
6087 Int_stable_cuboids -> Int_stable_atLeastAtMost
6088 Int_stable_product_algebra_generator -> positive_integral
6089 measure_preserving -> equality "distr M N f = N" "f : measurable M N"
6090 measure_space.additive -> emeasure_additive
6091 measure_space.AE_iff_null_set -> AE_iff_null
6092 measure_space.almost_everywhere_def -> eventually_ae_filter
6093 measure_space.almost_everywhere_vimage -> AE_distrD
6094 measure_space.continuity_from_above -> INF_emeasure_decseq
6095 measure_space.continuity_from_above_Lim -> Lim_emeasure_decseq
6096 measure_space.continuity_from_below_Lim -> Lim_emeasure_incseq
6097 measure_space.continuity_from_below -> SUP_emeasure_incseq
6098 measure_space_density -> emeasure_density
6099 measure_space.density_is_absolutely_continuous -> absolutely_continuousI_density
6100 measure_space.integrable_vimage -> integrable_distr
6101 measure_space.integral_translated_density -> integral_density
6102 measure_space.integral_vimage -> integral_distr
6103 measure_space.measure_additive -> plus_emeasure
6104 measure_space.measure_compl -> emeasure_compl
6105 measure_space.measure_countable_increasing -> emeasure_countable_increasing
6106 measure_space.measure_countably_subadditive -> emeasure_subadditive_countably
6107 measure_space.measure_decseq -> decseq_emeasure
6108 measure_space.measure_Diff -> emeasure_Diff
6109 measure_space.measure_Diff_null_set -> emeasure_Diff_null_set
6110 measure_space.measure_eq_0 -> emeasure_eq_0
6111 measure_space.measure_finitely_subadditive -> emeasure_subadditive_finite
6112 measure_space.measure_finite_singleton -> emeasure_eq_setsum_singleton
6113 measure_space.measure_incseq -> incseq_emeasure
6114 measure_space.measure_insert -> emeasure_insert
6115 measure_space.measure_mono -> emeasure_mono
6116 measure_space.measure_not_negative -> emeasure_not_MInf
6117 measure_space.measure_preserving_Int_stable -> measure_eqI_generator_eq
6118 measure_space.measure_setsum -> setsum_emeasure
6119 measure_space.measure_setsum_split -> setsum_emeasure_cover
6120 measure_space.measure_space_vimage -> emeasure_distr
6121 measure_space.measure_subadditive_finite -> emeasure_subadditive_finite
6122 measure_space.measure_subadditive -> subadditive
6123 measure_space.measure_top -> emeasure_space
6124 measure_space.measure_UN_eq_0 -> emeasure_UN_eq_0
6125 measure_space.measure_Un_null_set -> emeasure_Un_null_set
6126 measure_space.positive_integral_translated_density -> positive_integral_density
6127 measure_space.positive_integral_vimage -> positive_integral_distr
6128 measure_space.real_continuity_from_above -> Lim_measure_decseq
6129 measure_space.real_continuity_from_below -> Lim_measure_incseq
6130 measure_space.real_measure_countably_subadditive -> measure_subadditive_countably
6131 measure_space.real_measure_Diff -> measure_Diff
6132 measure_space.real_measure_finite_Union -> measure_finite_Union
6133 measure_space.real_measure_setsum_singleton -> measure_eq_setsum_singleton
6134 measure_space.real_measure_subadditive -> measure_subadditive
6135 measure_space.real_measure_Union -> measure_Union
6136 measure_space.real_measure_UNION -> measure_UNION
6137 measure_space.simple_function_vimage -> simple_function_comp
6138 measure_space.simple_integral_vimage -> simple_integral_distr
6139 measure_space.simple_integral_vimage -> simple_integral_distr
6140 measure_unique_Int_stable -> measure_eqI_generator_eq
6141 measure_unique_Int_stable_vimage -> measure_eqI_generator_eq
6142 pair_sigma_algebra.measurable_cut_fst -> sets_Pair1
6143 pair_sigma_algebra.measurable_cut_snd -> sets_Pair2
6144 pair_sigma_algebra.measurable_pair_image_fst -> measurable_Pair1
6145 pair_sigma_algebra.measurable_pair_image_snd -> measurable_Pair2
6146 pair_sigma_algebra.measurable_product_swap -> measurable_pair_swap_iff
6147 pair_sigma_algebra.pair_sigma_algebra_measurable -> measurable_pair_swap
6148 pair_sigma_algebra.pair_sigma_algebra_swap_measurable -> measurable_pair_swap'
6149 pair_sigma_algebra.sets_swap -> sets_pair_swap
6150 pair_sigma_finite.measure_cut_measurable_fst -> pair_sigma_finite.measurable_emeasure_Pair1
6151 pair_sigma_finite.measure_cut_measurable_snd -> pair_sigma_finite.measurable_emeasure_Pair2
6152 pair_sigma_finite.measure_preserving_swap -> pair_sigma_finite.distr_pair_swap
6153 pair_sigma_finite.pair_measure_alt2 -> pair_sigma_finite.emeasure_pair_measure_alt2
6154 pair_sigma_finite.pair_measure_alt -> pair_sigma_finite.emeasure_pair_measure_alt
6155 pair_sigma_finite.pair_measure_times -> pair_sigma_finite.emeasure_pair_measure_Times
6156 prob_space.indep_distribution_eq_measure -> prob_space.indep_vars_iff_distr_eq_PiM
6157 prob_space.indep_var_distributionD -> prob_space.indep_var_distribution_eq
6158 prob_space.measure_space_1 -> prob_space.emeasure_space_1
6159 prob_space.prob_space_vimage -> prob_space_distr
6160 prob_space.random_variable_restrict -> measurable_restrict
6161 prob_space_unique_Int_stable -> measure_eqI_prob_space
6162 product_algebraE -> prod_algebraE_all
6163 product_algebra_generator_der -> prod_algebra_eq_finite
6164 product_algebra_generator_into_space -> prod_algebra_sets_into_space
6165 product_algebraI -> sets_PiM_I_finite
6166 product_measure_exists -> product_sigma_finite.sigma_finite
6167 product_prob_space.finite_index_eq_finite_product -> product_prob_space.sets_PiM_generator
6168 product_prob_space.finite_measure_infprod_emb_Pi -> product_prob_space.measure_PiM_emb
6169 product_prob_space.infprod_spec -> product_prob_space.emeasure_PiM_emb_not_empty
6170 product_prob_space.measurable_component -> measurable_component_singleton
6171 product_prob_space.measurable_emb -> measurable_prod_emb
6172 product_prob_space.measurable_into_infprod_algebra -> measurable_PiM_single
6173 product_prob_space.measurable_singleton_infprod -> measurable_component_singleton
6174 product_prob_space.measure_emb -> emeasure_prod_emb
6175 product_prob_space.measure_preserving_restrict -> product_prob_space.distr_restrict
6176 product_sigma_algebra.product_algebra_into_space -> space_closed
6177 product_sigma_finite.measure_fold -> product_sigma_finite.distr_merge
6178 product_sigma_finite.measure_preserving_component_singelton -> product_sigma_finite.distr_singleton
6179 product_sigma_finite.measure_preserving_merge -> product_sigma_finite.distr_merge
6180 sequence_space.measure_infprod -> sequence_space.measure_PiM_countable
6181 sets_product_algebra -> sets_PiM
6182 sigma_algebra.measurable_sigma -> measurable_measure_of
6183 sigma_finite_measure.disjoint_sigma_finite -> sigma_finite_disjoint
6184 sigma_finite_measure.RN_deriv_vimage -> sigma_finite_measure.RN_deriv_distr
6185 sigma_product_algebra_sigma_eq -> sigma_prod_algebra_sigma_eq
6186 space_product_algebra -> space_PiM
6188 * Session HOL-TPTP: support to parse and import TPTP problems (all
6189 languages) into Isabelle/HOL.
6194 * New "case_product" attribute (see HOL).
6199 * Greater support for structured proofs involving induction or case
6202 * Much greater use of mathematical symbols.
6204 * Removal of many ML theorem bindings. INCOMPATIBILITY.
6209 * Antiquotation @{keyword "name"} produces a parser for outer syntax
6210 from a minor keyword introduced via theory header declaration.
6212 * Antiquotation @{command_spec "name"} produces the
6213 Outer_Syntax.command_spec from a major keyword introduced via theory
6214 header declaration; it can be passed to Outer_Syntax.command etc.
6216 * Local_Theory.define no longer hard-wires default theorem name
6217 "foo_def", but retains the binding as given. If that is Binding.empty
6218 / Attrib.empty_binding, the result is not registered as user-level
6219 fact. The Local_Theory.define_internal variant allows to specify a
6220 non-empty name (used for the foundation in the background theory),
6221 while omitting the fact binding in the user-context. Potential
6222 INCOMPATIBILITY for derived definitional packages: need to specify
6223 naming policy for primitive definitions more explicitly.
6225 * Renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in
6226 conformance with similar operations in structure Term and Logic.
6228 * Antiquotation @{attributes [...]} embeds attribute source
6229 representation into the ML text, which is particularly useful with
6230 declarations like Local_Theory.note.
6232 * Structure Proof_Context follows standard naming scheme. Old
6233 ProofContext has been discontinued. INCOMPATIBILITY.
6235 * Refined Local_Theory.declaration {syntax, pervasive}, with subtle
6236 change of semantics: update is applied to auxiliary local theory
6239 * Modernized some old-style infix operations:
6241 addeqcongs ~> Simplifier.add_eqcong
6242 deleqcongs ~> Simplifier.del_eqcong
6243 addcongs ~> Simplifier.add_cong
6244 delcongs ~> Simplifier.del_cong
6245 setmksimps ~> Simplifier.set_mksimps
6246 setmkcong ~> Simplifier.set_mkcong
6247 setmksym ~> Simplifier.set_mksym
6248 setmkeqTrue ~> Simplifier.set_mkeqTrue
6249 settermless ~> Simplifier.set_termless
6250 setsubgoaler ~> Simplifier.set_subgoaler
6251 addsplits ~> Splitter.add_split
6252 delsplits ~> Splitter.del_split