NEWS
author wenzelm
Sat, 15 Oct 2005 00:08:14 +0200
changeset 17865 5b0c3dcfbad2
parent 17809 195045659c06
child 17869 585c1f08499e
permissions -rw-r--r--
* antiquotations ML_type, ML_struct;
* Isar 'guess' element;
     1 Isabelle NEWS -- history user-relevant changes
     2 ==============================================
     3 
     4 New in this Isabelle release
     5 ----------------------------
     6 
     7 *** General ***
     8 
     9 * Command 'find_theorems': support "*" wildcard in "name:" criterion.
    10 
    11 
    12 *** Document preparation ***
    13 
    14 * Added antiquotations @{ML_type text} and @{ML_struct} which check
    15 the given source text as ML type/structure, printing verbatim.
    16 
    17 
    18 *** Pure ***
    19 
    20 * Isar: improper proof element 'guess' is like 'obtain', but derives
    21 the obtained context from the course of reasoning!  For example:
    22 
    23   assume "EX x y. A x & B y"   -- "any previous fact"
    24   then guess x and y by clarify
    25 
    26 This technique is potentially adventurous, depending on the facts and
    27 proof tools being involved here.
    28 
    29 * Input syntax now supports dummy variable binding "%_. b", where the
    30 body does not mention the bound variable.  Note that dummy patterns
    31 implicitly depend on their context of bounds, which makes "{_. _}"
    32 match any set comprehension as expected.  Potential INCOMPATIBILITY --
    33 parse translations need to cope with syntactic constant "_idtdummy" in
    34 the binding position.
    35 
    36 * Removed obsolete syntactic constant "_K" and its associated parse
    37 translation.  INCOMPATIBILITY -- use dummy abstraction instead, for
    38 example "A -> B" => "Pi A (%_. B)".
    39 
    40 
    41 *** HOL ***
    42 
    43 * In the context of the assumption "~(s = t)" the Simplifier rewrites
    44 "t = s" to False (by simproc "neq_simproc").  For backward
    45 compatibility this can be disabled by ML "reset use_neq_simproc".
    46 
    47 * Tactics 'sat' and 'satx' reimplemented, several improvements: goals
    48 no longer need to be stated as "<prems> ==> False", equivalences (i.e.
    49 "=" on type bool) are handled, variable names of the form "lit_<n>"
    50 are no longer reserved, significant speedup.
    51 
    52 
    53 
    54 New in Isabelle2005 (October 2005)
    55 ----------------------------------
    56 
    57 *** General ***
    58 
    59 * Theory headers: the new header syntax for Isar theories is
    60 
    61   theory <name>
    62   imports <theory1> ... <theoryN>
    63   uses <file1> ... <fileM>
    64   begin
    65 
    66 where the 'uses' part is optional.  The previous syntax
    67 
    68   theory <name> = <theory1> + ... + <theoryN>:
    69 
    70 will disappear in the next release.  Use isatool fixheaders to convert
    71 existing theory files.  Note that there is no change in ancient
    72 non-Isar theories now, but these will disappear soon.
    73 
    74 * Theory loader: parent theories can now also be referred to via
    75 relative and absolute paths.
    76 
    77 * Command 'find_theorems' searches for a list of criteria instead of a
    78 list of constants. Known criteria are: intro, elim, dest, name:string,
    79 simp:term, and any term. Criteria can be preceded by '-' to select
    80 theorems that do not match. Intro, elim, dest select theorems that
    81 match the current goal, name:s selects theorems whose fully qualified
    82 name contain s, and simp:term selects all simplification rules whose
    83 lhs match term.  Any other term is interpreted as pattern and selects
    84 all theorems matching the pattern. Available in ProofGeneral under
    85 'ProofGeneral -> Find Theorems' or C-c C-f.  Example:
    86 
    87   C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL."
    88 
    89 prints the last 100 theorems matching the pattern "(_::nat) + _ + _",
    90 matching the current goal as introduction rule and not having "HOL."
    91 in their name (i.e. not being defined in theory HOL).
    92 
    93 * Command 'thms_containing' has been discontinued in favour of
    94 'find_theorems'; INCOMPATIBILITY.
    95 
    96 * Communication with Proof General is now 8bit clean, which means that
    97 Unicode text in UTF-8 encoding may be used within theory texts (both
    98 formal and informal parts).  Cf. option -U of the Isabelle Proof
    99 General interface.  Here are some simple examples (cf. src/HOL/ex):
   100 
   101   http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html
   102   http://isabelle.in.tum.de/library/HOL/ex/Chinese.html
   103 
   104 * Improved efficiency of the Simplifier and, to a lesser degree, the
   105 Classical Reasoner.  Typical big applications run around 2 times
   106 faster.
   107 
   108 
   109 *** Document preparation ***
   110 
   111 * Commands 'display_drafts' and 'print_drafts' perform simple output
   112 of raw sources.  Only those symbols that do not require additional
   113 LaTeX packages (depending on comments in isabellesym.sty) are
   114 displayed properly, everything else is left verbatim.  isatool display
   115 and isatool print are used as front ends (these are subject to the
   116 DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively).
   117 
   118 * Command tags control specific markup of certain regions of text,
   119 notably folding and hiding.  Predefined tags include "theory" (for
   120 theory begin and end), "proof" for proof commands, and "ML" for
   121 commands involving ML code; the additional tags "visible" and
   122 "invisible" are unused by default.  Users may give explicit tag
   123 specifications in the text, e.g. ''by %invisible (auto)''.  The
   124 interpretation of tags is determined by the LaTeX job during document
   125 preparation: see option -V of isatool usedir, or options -n and -t of
   126 isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag,
   127 \isadroptag.
   128 
   129 Several document versions may be produced at the same time via isatool
   130 usedir (the generated index.html will link all of them).  Typical
   131 specifications include ''-V document=theory,proof,ML'' to present
   132 theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold
   133 proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit
   134 these parts without any formal replacement text.  The Isabelle site
   135 default settings produce ''document'' and ''outline'' versions as
   136 specified above.
   137 
   138 * Several new antiquotations:
   139 
   140   @{term_type term} prints a term with its type annotated;
   141 
   142   @{typeof term} prints the type of a term;
   143 
   144   @{const const} is the same as @{term const}, but checks that the
   145   argument is a known logical constant;
   146 
   147   @{term_style style term} and @{thm_style style thm} print a term or
   148   theorem applying a "style" to it
   149 
   150   @{ML text}
   151 
   152 Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of
   153 definitions, equations, inequations etc., 'concl' printing only the
   154 conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19'
   155 to print the specified premise.  TermStyle.add_style provides an ML
   156 interface for introducing further styles.  See also the "LaTeX Sugar"
   157 document practical applications.  The ML antiquotation prints
   158 type-checked ML expressions verbatim.
   159 
   160 * Markup commands 'chapter', 'section', 'subsection', 'subsubsection',
   161 and 'text' support optional locale specification '(in loc)', which
   162 specifies the default context for interpreting antiquotations.  For
   163 example: 'text (in lattice) {* @{thm inf_assoc}*}'.
   164 
   165 * Option 'locale=NAME' of antiquotations specifies an alternative
   166 context interpreting the subsequent argument.  For example: @{thm
   167 [locale=lattice] inf_assoc}.
   168 
   169 * Proper output of proof terms (@{prf ...} and @{full_prf ...}) within
   170 a proof context.
   171 
   172 * Proper output of antiquotations for theory commands involving a
   173 proof context (such as 'locale' or 'theorem (in loc) ...').
   174 
   175 * Delimiters of outer tokens (string etc.) now produce separate LaTeX
   176 macros (\isachardoublequoteopen, isachardoublequoteclose etc.).
   177 
   178 * isatool usedir: new option -C (default true) controls whether option
   179 -D should include a copy of the original document directory; -C false
   180 prevents unwanted effects such as copying of administrative CVS data.
   181 
   182 
   183 *** Pure ***
   184 
   185 * Considerably improved version of 'constdefs' command.  Now performs
   186 automatic type-inference of declared constants; additional support for
   187 local structure declarations (cf. locales and HOL records), see also
   188 isar-ref manual.  Potential INCOMPATIBILITY: need to observe strictly
   189 sequential dependencies of definitions within a single 'constdefs'
   190 section; moreover, the declared name needs to be an identifier.  If
   191 all fails, consider to fall back on 'consts' and 'defs' separately.
   192 
   193 * Improved indexed syntax and implicit structures.  First of all,
   194 indexed syntax provides a notational device for subscripted
   195 application, using the new syntax \<^bsub>term\<^esub> for arbitrary
   196 expressions.  Secondly, in a local context with structure
   197 declarations, number indexes \<^sub>n or the empty index (default
   198 number 1) refer to a certain fixed variable implicitly; option
   199 show_structs controls printing of implicit structures.  Typical
   200 applications of these concepts involve record types and locales.
   201 
   202 * New command 'no_syntax' removes grammar declarations (and
   203 translations) resulting from the given syntax specification, which is
   204 interpreted in the same manner as for the 'syntax' command.
   205 
   206 * 'Advanced' translation functions (parse_translation etc.) may depend
   207 on the signature of the theory context being presently used for
   208 parsing/printing, see also isar-ref manual.
   209 
   210 * Improved 'oracle' command provides a type-safe interface to turn an
   211 ML expression of type theory -> T -> term into a primitive rule of
   212 type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle
   213 is already included here); see also FOL/ex/IffExample.thy;
   214 INCOMPATIBILITY.
   215 
   216 * axclass: name space prefix for class "c" is now "c_class" (was "c"
   217 before); "cI" is no longer bound, use "c.intro" instead.
   218 INCOMPATIBILITY.  This change avoids clashes of fact bindings for
   219 axclasses vs. locales.
   220 
   221 * Improved internal renaming of symbolic identifiers -- attach primes
   222 instead of base 26 numbers.
   223 
   224 * New flag show_question_marks controls printing of leading question
   225 marks in schematic variable names.
   226 
   227 * In schematic variable names, *any* symbol following \<^isub> or
   228 \<^isup> is now treated as part of the base name.  For example, the
   229 following works without printing of awkward ".0" indexes:
   230 
   231   lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1"
   232     by simp
   233 
   234 * Inner syntax includes (*(*nested*) comments*).
   235 
   236 * Pretty printer now supports unbreakable blocks, specified in mixfix
   237 annotations as "(00...)".
   238 
   239 * Clear separation of logical types and nonterminals, where the latter
   240 may only occur in 'syntax' specifications or type abbreviations.
   241 Before that distinction was only partially implemented via type class
   242 "logic" vs. "{}".  Potential INCOMPATIBILITY in rare cases of improper
   243 use of 'types'/'consts' instead of 'nonterminals'/'syntax'.  Some very
   244 exotic syntax specifications may require further adaption
   245 (e.g. Cube/Cube.thy).
   246 
   247 * Removed obsolete type class "logic", use the top sort {} instead.
   248 Note that non-logical types should be declared as 'nonterminals'
   249 rather than 'types'.  INCOMPATIBILITY for new object-logic
   250 specifications.
   251 
   252 * Attributes 'induct' and 'cases': type or set names may now be
   253 locally fixed variables as well.
   254 
   255 * Simplifier: can now control the depth to which conditional rewriting
   256 is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth
   257 Limit.
   258 
   259 * Simplifier: simplification procedures may now take the current
   260 simpset into account (cf. Simplifier.simproc(_i) / mk_simproc
   261 interface), which is very useful for calling the Simplifier
   262 recursively.  Minor INCOMPATIBILITY: the 'prems' argument of simprocs
   263 is gone -- use prems_of_ss on the simpset instead.  Moreover, the
   264 low-level mk_simproc no longer applies Logic.varify internally, to
   265 allow for use in a context of fixed variables.
   266 
   267 * thin_tac now works even if the assumption being deleted contains !!
   268 or ==>.  More generally, erule now works even if the major premise of
   269 the elimination rule contains !! or ==>.
   270 
   271 * Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY.
   272 
   273 * Reorganized bootstrapping of the Pure theories; CPure is now derived
   274 from Pure, which contains all common declarations already.  Both
   275 theories are defined via plain Isabelle/Isar .thy files.
   276 INCOMPATIBILITY: elements of CPure (such as the CPure.intro /
   277 CPure.elim / CPure.dest attributes) now appear in the Pure name space;
   278 use isatool fixcpure to adapt your theory and ML sources.
   279 
   280 * New syntax 'name(i-j, i-, i, ...)' for referring to specific
   281 selections of theorems in named facts via index ranges.
   282 
   283 * 'print_theorems': in theory mode, really print the difference
   284 wrt. the last state (works for interactive theory development only),
   285 in proof mode print all local facts (cf. 'print_facts');
   286 
   287 * 'hide': option '(open)' hides only base names.
   288 
   289 * More efficient treatment of intermediate checkpoints in interactive
   290 theory development.
   291 
   292 * Code generator is now invoked via code_module (incremental code
   293 generation) and code_library (modular code generation, ML structures
   294 for each theory).  INCOMPATIBILITY: new keywords 'file' and 'contains'
   295 must be quoted when used as identifiers.
   296 
   297 * New 'value' command for reading, evaluating and printing terms using
   298 the code generator.  INCOMPATIBILITY: command keyword 'value' must be
   299 quoted when used as identifier.
   300 
   301 
   302 *** Locales ***
   303 
   304 * New commands for the interpretation of locale expressions in
   305 theories (1), locales (2) and proof contexts (3).  These generate
   306 proof obligations from the expression specification.  After the
   307 obligations have been discharged, theorems of the expression are added
   308 to the theory, target locale or proof context.  The synopsis of the
   309 commands is a follows:
   310 
   311   (1) interpretation expr inst
   312   (2) interpretation target < expr
   313   (3) interpret expr inst
   314 
   315 Interpretation in theories and proof contexts require a parameter
   316 instantiation of terms from the current context.  This is applied to
   317 specifications and theorems of the interpreted expression.
   318 Interpretation in locales only permits parameter renaming through the
   319 locale expression.  Interpretation is smart in that interpretations
   320 that are active already do not occur in proof obligations, neither are
   321 instantiated theorems stored in duplicate.  Use 'print_interps' to
   322 inspect active interpretations of a particular locale.  For details,
   323 see the Isar Reference manual.  Examples can be found in
   324 HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy.
   325 
   326 INCOMPATIBILITY: former 'instantiate' has been withdrawn, use
   327 'interpret' instead.
   328 
   329 * New context element 'constrains' for adding type constraints to
   330 parameters.
   331 
   332 * Context expressions: renaming of parameters with syntax
   333 redeclaration.
   334 
   335 * Locale declaration: 'includes' disallowed.
   336 
   337 * Proper static binding of attribute syntax -- i.e. types / terms /
   338 facts mentioned as arguments are always those of the locale definition
   339 context, independently of the context of later invocations.  Moreover,
   340 locale operations (renaming and type / term instantiation) are applied
   341 to attribute arguments as expected.
   342 
   343 INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of
   344 actual attributes; rare situations may require Attrib.attribute to
   345 embed those attributes into Attrib.src that lack concrete syntax.
   346 Attribute implementations need to cooperate properly with the static
   347 binding mechanism.  Basic parsers Args.XXX_typ/term/prop and
   348 Attrib.XXX_thm etc. already do the right thing without further
   349 intervention.  Only unusual applications -- such as "where" or "of"
   350 (cf. src/Pure/Isar/attrib.ML), which process arguments depending both
   351 on the context and the facts involved -- may have to assign parsed
   352 values to argument tokens explicitly.
   353 
   354 * Changed parameter management in theorem generation for long goal
   355 statements with 'includes'.  INCOMPATIBILITY: produces a different
   356 theorem statement in rare situations.
   357 
   358 * Locale inspection command 'print_locale' omits notes elements.  Use
   359 'print_locale!' to have them included in the output.
   360 
   361 
   362 *** Provers ***
   363 
   364 * Provers/hypsubst.ML: improved version of the subst method, for
   365 single-step rewriting: it now works in bound variable contexts. New is
   366 'subst (asm)', for rewriting an assumption.  INCOMPATIBILITY: may
   367 rewrite a different subterm than the original subst method, which is
   368 still available as 'simplesubst'.
   369 
   370 * Provers/quasi.ML: new transitivity reasoners for transitivity only
   371 and quasi orders.
   372 
   373 * Provers/trancl.ML: new transitivity reasoner for transitive and
   374 reflexive-transitive closure of relations.
   375 
   376 * Provers/blast.ML: new reference depth_limit to make blast's depth
   377 limit (previously hard-coded with a value of 20) user-definable.
   378 
   379 * Provers/simplifier.ML has been moved to Pure, where Simplifier.setup
   380 is peformed already.  Object-logics merely need to finish their
   381 initial simpset configuration as before.  INCOMPATIBILITY.
   382 
   383 
   384 *** HOL ***
   385 
   386 * Symbolic syntax of Hilbert Choice Operator is now as follows:
   387 
   388   syntax (epsilon)
   389     "_Eps" :: "[pttrn, bool] => 'a"    ("(3\<some>_./ _)" [0, 10] 10)
   390 
   391 The symbol \<some> is displayed as the alternative epsilon of LaTeX
   392 and x-symbol; use option '-m epsilon' to get it actually printed.
   393 Moreover, the mathematically important symbolic identifier \<epsilon>
   394 becomes available as variable, constant etc.  INCOMPATIBILITY,
   395 
   396 * "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x".
   397 Similarly for all quantifiers: "ALL x > y" etc.  The x-symbol for >=
   398 is \<ge>. New transitivity rules have been added to HOL/Orderings.thy to
   399 support corresponding Isar calculations.
   400 
   401 * "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\<in>"
   402 instead of ":".
   403 
   404 * theory SetInterval: changed the syntax for open intervals:
   405 
   406   Old       New
   407   {..n(}    {..<n}
   408   {)n..}    {n<..}
   409   {m..n(}   {m..<n}
   410   {)m..n}   {m<..n}
   411   {)m..n(}  {m<..<n}
   412 
   413 The old syntax is still supported but will disappear in the next
   414 release.  For conversion use the following Emacs search and replace
   415 patterns (these are not perfect but work quite well):
   416 
   417   {)\([^\.]*\)\.\.  ->  {\1<\.\.}
   418   \.\.\([^(}]*\)(}  ->  \.\.<\1}
   419 
   420 * Theory Commutative_Ring (in Library): method comm_ring for proving
   421 equalities in commutative rings; method 'algebra' provides a generic
   422 interface.
   423 
   424 * Theory Finite_Set: changed the syntax for 'setsum', summation over
   425 finite sets: "setsum (%x. e) A", which used to be "\<Sum>x:A. e", is
   426 now either "SUM x:A. e" or "\<Sum>x \<in> A. e". The bound variable can
   427 be a tuple pattern.
   428 
   429 Some new syntax forms are available:
   430 
   431   "\<Sum>x | P. e"      for     "setsum (%x. e) {x. P}"
   432   "\<Sum>x = a..b. e"   for     "setsum (%x. e) {a..b}"
   433   "\<Sum>x = a..<b. e"  for     "setsum (%x. e) {a..<b}"
   434   "\<Sum>x < k. e"      for     "setsum (%x. e) {..<k}"
   435 
   436 The latter form "\<Sum>x < k. e" used to be based on a separate
   437 function "Summation", which has been discontinued.
   438 
   439 * theory Finite_Set: in structured induction proofs, the insert case
   440 is now 'case (insert x F)' instead of the old counterintuitive 'case
   441 (insert F x)'.
   442 
   443 * The 'refute' command has been extended to support a much larger
   444 fragment of HOL, including axiomatic type classes, constdefs and
   445 typedefs, inductive datatypes and recursion.
   446 
   447 * New tactics 'sat' and 'satx' to prove propositional tautologies.
   448 Requires zChaff with proof generation to be installed.  See
   449 HOL/ex/SAT_Examples.thy for examples.
   450 
   451 * Datatype induction via method 'induct' now preserves the name of the
   452 induction variable. For example, when proving P(xs::'a list) by
   453 induction on xs, the induction step is now P(xs) ==> P(a#xs) rather
   454 than P(list) ==> P(a#list) as previously.  Potential INCOMPATIBILITY
   455 in unstructured proof scripts.
   456 
   457 * Reworked implementation of records.  Improved scalability for
   458 records with many fields, avoiding performance problems for type
   459 inference. Records are no longer composed of nested field types, but
   460 of nested extension types. Therefore the record type only grows linear
   461 in the number of extensions and not in the number of fields.  The
   462 top-level (users) view on records is preserved.  Potential
   463 INCOMPATIBILITY only in strange cases, where the theory depends on the
   464 old record representation. The type generated for a record is called
   465 <record_name>_ext_type.
   466 
   467 Flag record_quick_and_dirty_sensitive can be enabled to skip the
   468 proofs triggered by a record definition or a simproc (if
   469 quick_and_dirty is enabled).  Definitions of large records can take
   470 quite long.
   471 
   472 New simproc record_upd_simproc for simplification of multiple record
   473 updates enabled by default.  Moreover, trivial updates are also
   474 removed: r(|x := x r|) = r.  INCOMPATIBILITY: old proofs break
   475 occasionally, since simplification is more powerful by default.
   476 
   477 * typedef: proper support for polymorphic sets, which contain extra
   478 type-variables in the term.
   479 
   480 * Simplifier: automatically reasons about transitivity chains
   481 involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics
   482 provided by Provers/trancl.ML as additional solvers.  INCOMPATIBILITY:
   483 old proofs break occasionally as simplification may now solve more
   484 goals than previously.
   485 
   486 * Simplifier: converts x <= y into x = y if assumption y <= x is
   487 present.  Works for all partial orders (class "order"), in particular
   488 numbers and sets.  For linear orders (e.g. numbers) it treats ~ x < y
   489 just like y <= x.
   490 
   491 * Simplifier: new simproc for "let x = a in f x".  If a is a free or
   492 bound variable or a constant then the let is unfolded.  Otherwise
   493 first a is simplified to b, and then f b is simplified to g. If
   494 possible we abstract b from g arriving at "let x = b in h x",
   495 otherwise we unfold the let and arrive at g.  The simproc can be
   496 enabled/disabled by the reference use_let_simproc.  Potential
   497 INCOMPATIBILITY since simplification is more powerful by default.
   498 
   499 * Classical reasoning: the meson method now accepts theorems as arguments.
   500 
   501 * Prover support: pre-release of the Isabelle-ATP linkup, which runs background
   502 jobs to provide advice on the provability of subgoals.
   503 
   504 * Theory OrderedGroup and Ring_and_Field: various additions and
   505 improvements to faciliate calculations involving equalities and
   506 inequalities.
   507 
   508 The following theorems have been eliminated or modified
   509 (INCOMPATIBILITY):
   510 
   511   abs_eq             now named abs_of_nonneg
   512   abs_of_ge_0        now named abs_of_nonneg
   513   abs_minus_eq       now named abs_of_nonpos
   514   imp_abs_id         now named abs_of_nonneg
   515   imp_abs_neg_id     now named abs_of_nonpos
   516   mult_pos           now named mult_pos_pos
   517   mult_pos_le        now named mult_nonneg_nonneg
   518   mult_pos_neg_le    now named mult_nonneg_nonpos
   519   mult_pos_neg2_le   now named mult_nonneg_nonpos2
   520   mult_neg           now named mult_neg_neg
   521   mult_neg_le        now named mult_nonpos_nonpos
   522 
   523 * Theory Parity: added rules for simplifying exponents.
   524 
   525 * Theory List:
   526 
   527 The following theorems have been eliminated or modified
   528 (INCOMPATIBILITY):
   529 
   530   list_all_Nil       now named list_all.simps(1)
   531   list_all_Cons      now named list_all.simps(2)
   532   list_all_conv      now named list_all_iff
   533   set_mem_eq         now named mem_iff
   534 
   535 * Theories SetsAndFunctions and BigO (see HOL/Library) support
   536 asymptotic "big O" calculations.  See the notes in BigO.thy.
   537 
   538 
   539 *** HOL-Complex ***
   540 
   541 * Theory RealDef: better support for embedding natural numbers and
   542 integers in the reals.
   543 
   544 The following theorems have been eliminated or modified
   545 (INCOMPATIBILITY):
   546 
   547   exp_ge_add_one_self  now requires no hypotheses
   548   real_of_int_add      reversed direction of equality (use [symmetric])
   549   real_of_int_minus    reversed direction of equality (use [symmetric])
   550   real_of_int_diff     reversed direction of equality (use [symmetric])
   551   real_of_int_mult     reversed direction of equality (use [symmetric])
   552 
   553 * Theory RComplete: expanded support for floor and ceiling functions.
   554 
   555 * Theory Ln is new, with properties of the natural logarithm
   556 
   557 * Hyperreal: There is a new type constructor "star" for making
   558 nonstandard types.  The old type names are now type synonyms:
   559 
   560   hypreal = real star
   561   hypnat = nat star
   562   hcomplex = complex star
   563 
   564 * Hyperreal: Many groups of similarly-defined constants have been
   565 replaced by polymorphic versions (INCOMPATIBILITY):
   566 
   567   star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex
   568 
   569   starset      <-- starsetNat, starsetC
   570   *s*          <-- *sNat*, *sc*
   571   starset_n    <-- starsetNat_n, starsetC_n
   572   *sn*         <-- *sNatn*, *scn*
   573   InternalSets <-- InternalNatSets, InternalCSets
   574 
   575   starfun      <-- starfun{Nat,Nat2,C,RC,CR}
   576   *f*          <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR*
   577   starfun_n    <-- starfun{Nat,Nat2,C,RC,CR}_n
   578   *fn*         <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn*
   579   InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs
   580 
   581 * Hyperreal: Many type-specific theorems have been removed in favor of
   582 theorems specific to various axiomatic type classes (INCOMPATIBILITY):
   583 
   584   add_commute <-- {hypreal,hypnat,hcomplex}_add_commute
   585   add_assoc   <-- {hypreal,hypnat,hcomplex}_add_assocs
   586   OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left
   587   OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right
   588   right_minus <-- hypreal_add_minus
   589   left_minus <-- {hypreal,hcomplex}_add_minus_left
   590   mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute
   591   mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc
   592   mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left
   593   mult_1_right <-- hcomplex_mult_one_right
   594   mult_zero_left <-- hcomplex_mult_zero_left
   595   left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib
   596   right_distrib <-- hypnat_add_mult_distrib2
   597   zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one
   598   right_inverse <-- hypreal_mult_inverse
   599   left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left
   600   order_refl <-- {hypreal,hypnat}_le_refl
   601   order_trans <-- {hypreal,hypnat}_le_trans
   602   order_antisym <-- {hypreal,hypnat}_le_anti_sym
   603   order_less_le <-- {hypreal,hypnat}_less_le
   604   linorder_linear <-- {hypreal,hypnat}_le_linear
   605   add_left_mono <-- {hypreal,hypnat}_add_left_mono
   606   mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2
   607   add_nonneg_nonneg <-- hypreal_le_add_order
   608 
   609 * Hyperreal: Separate theorems having to do with type-specific
   610 versions of constants have been merged into theorems that apply to the
   611 new polymorphic constants (INCOMPATIBILITY):
   612 
   613   STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set
   614   STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set
   615   STAR_Un <-- {STAR,NatStar,STARC}_Un
   616   STAR_Int <-- {STAR,NatStar,STARC}_Int
   617   STAR_Compl <-- {STAR,NatStar,STARC}_Compl
   618   STAR_subset <-- {STAR,NatStar,STARC}_subset
   619   STAR_mem <-- {STAR,NatStar,STARC}_mem
   620   STAR_mem_Compl <-- {STAR,STARC}_mem_Compl
   621   STAR_diff <-- {STAR,STARC}_diff
   622   STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real,
   623     STARC_hcomplex_of_complex}_image_subset
   624   starset_n_Un <-- starset{Nat,C}_n_Un
   625   starset_n_Int <-- starset{Nat,C}_n_Int
   626   starset_n_Compl <-- starset{Nat,C}_n_Compl
   627   starset_n_diff <-- starset{Nat,C}_n_diff
   628   InternalSets_Un <-- Internal{Nat,C}Sets_Un
   629   InternalSets_Int <-- Internal{Nat,C}Sets_Int
   630   InternalSets_Compl <-- Internal{Nat,C}Sets_Compl
   631   InternalSets_diff <-- Internal{Nat,C}Sets_diff
   632   InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff
   633   InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n
   634   starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq
   635   starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C}
   636   starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR}
   637   starfun <-- starfun{Nat,Nat2,C,RC,CR}
   638   starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult
   639   starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add
   640   starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus
   641   starfun_diff <-- starfun{C,RC,CR}_diff
   642   starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o
   643   starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2
   644   starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun
   645   starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse
   646   starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq
   647   starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff
   648   starfun_Id <-- starfunC_Id
   649   starfun_approx <-- starfun{Nat,CR}_approx
   650   starfun_capprox <-- starfun{C,RC}_capprox
   651   starfun_abs <-- starfunNat_rabs
   652   starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel
   653   starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2
   654   starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox
   655   starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox
   656   starfun_add_capprox <-- starfun{C,RC}_add_capprox
   657   starfun_add_approx <-- starfunCR_add_approx
   658   starfun_inverse_inverse <-- starfunC_inverse_inverse
   659   starfun_divide <-- starfun{C,CR,RC}_divide
   660   starfun_n <-- starfun{Nat,C}_n
   661   starfun_n_mult <-- starfun{Nat,C}_n_mult
   662   starfun_n_add <-- starfun{Nat,C}_n_add
   663   starfun_n_add_minus <-- starfunNat_n_add_minus
   664   starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun
   665   starfun_n_minus <-- starfun{Nat,C}_n_minus
   666   starfun_n_eq <-- starfun{Nat,C}_n_eq
   667 
   668   star_n_add <-- {hypreal,hypnat,hcomplex}_add
   669   star_n_minus <-- {hypreal,hcomplex}_minus
   670   star_n_diff <-- {hypreal,hcomplex}_diff
   671   star_n_mult <-- {hypreal,hcomplex}_mult
   672   star_n_inverse <-- {hypreal,hcomplex}_inverse
   673   star_n_le <-- {hypreal,hypnat}_le
   674   star_n_less <-- {hypreal,hypnat}_less
   675   star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num
   676   star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num
   677   star_n_abs <-- hypreal_hrabs
   678   star_n_divide <-- hcomplex_divide
   679 
   680   star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add
   681   star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus
   682   star_of_diff <-- hypreal_of_real_diff
   683   star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult
   684   star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one
   685   star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero
   686   star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff
   687   star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff
   688   star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff
   689   star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse
   690   star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide
   691   star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat
   692   star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int
   693   star_of_number_of <-- {hypreal,hcomplex}_number_of
   694   star_of_number_less <-- number_of_less_hypreal_of_real_iff
   695   star_of_number_le <-- number_of_le_hypreal_of_real_iff
   696   star_of_eq_number <-- hypreal_of_real_eq_number_of_iff
   697   star_of_less_number <-- hypreal_of_real_less_number_of_iff
   698   star_of_le_number <-- hypreal_of_real_le_number_of_iff
   699   star_of_power <-- hypreal_of_real_power
   700   star_of_eq_0 <-- hcomplex_of_complex_zero_iff
   701 
   702 * Hyperreal: new method "transfer" that implements the transfer
   703 principle of nonstandard analysis. With a subgoal that mentions
   704 nonstandard types like "'a star", the command "apply transfer"
   705 replaces it with an equivalent one that mentions only standard types.
   706 To be successful, all free variables must have standard types; non-
   707 standard variables must have explicit universal quantifiers.
   708 
   709 * Hyperreal: A theory of Taylor series.
   710 
   711 
   712 *** HOLCF ***
   713 
   714 * Discontinued special version of 'constdefs' (which used to support
   715 continuous functions) in favor of the general Pure one with full
   716 type-inference.
   717 
   718 * New simplification procedure for solving continuity conditions; it
   719 is much faster on terms with many nested lambda abstractions (cubic
   720 instead of exponential time).
   721 
   722 * New syntax for domain package: selector names are now optional.
   723 Parentheses should be omitted unless argument is lazy, for example:
   724 
   725   domain 'a stream = cons "'a" (lazy "'a stream")
   726 
   727 * New command 'fixrec' for defining recursive functions with pattern
   728 matching; defining multiple functions with mutual recursion is also
   729 supported.  Patterns may include the constants cpair, spair, up, sinl,
   730 sinr, or any data constructor defined by the domain package. The given
   731 equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for
   732 syntax and examples.
   733 
   734 * New commands 'cpodef' and 'pcpodef' for defining predicate subtypes
   735 of cpo and pcpo types. Syntax is exactly like the 'typedef' command,
   736 but the proof obligation additionally includes an admissibility
   737 requirement. The packages generate instances of class cpo or pcpo,
   738 with continuity and strictness theorems for Rep and Abs.
   739 
   740 * HOLCF: Many theorems have been renamed according to a more standard naming
   741 scheme (INCOMPATIBILITY):
   742 
   743   foo_inject:  "foo$x = foo$y ==> x = y"
   744   foo_eq:      "(foo$x = foo$y) = (x = y)"
   745   foo_less:    "(foo$x << foo$y) = (x << y)"
   746   foo_strict:  "foo$UU = UU"
   747   foo_defined: "... ==> foo$x ~= UU"
   748   foo_defined_iff: "(foo$x = UU) = (x = UU)"
   749 
   750 
   751 *** ZF ***
   752 
   753 * ZF/ex: theories Group and Ring provide examples in abstract algebra,
   754 including the First Isomorphism Theorem (on quotienting by the kernel
   755 of a homomorphism).
   756 
   757 * ZF/Simplifier: install second copy of type solver that actually
   758 makes use of TC rules declared to Isar proof contexts (or locales);
   759 the old version is still required for ML proof scripts.
   760 
   761 
   762 *** Cube ***
   763 
   764 * Converted to Isar theory format; use locales instead of axiomatic
   765 theories.
   766 
   767 
   768 *** ML ***
   769 
   770 * Pure/library.ML no longer defines its own option datatype, but uses
   771 that of the SML basis, which has constructors NONE and SOME instead of
   772 None and Some, as well as exception Option.Option instead of OPTION.
   773 The functions the, if_none, is_some, is_none have been adapted
   774 accordingly, while Option.map replaces apsome.
   775 
   776 * Pure/library.ML: the exception LIST has been given up in favour of
   777 the standard exceptions Empty and Subscript, as well as
   778 Library.UnequalLengths.  Function like Library.hd and Library.tl are
   779 superceded by the standard hd and tl functions etc.
   780 
   781 A number of basic list functions are no longer exported to the ML
   782 toplevel, as they are variants of predefined functions.  The following
   783 suggests how one can translate existing code:
   784 
   785     rev_append xs ys = List.revAppend (xs, ys)
   786     nth_elem (i, xs) = List.nth (xs, i)
   787     last_elem xs = List.last xs
   788     flat xss = List.concat xss
   789     seq fs = List.app fs
   790     partition P xs = List.partition P xs
   791     mapfilter f xs = List.mapPartial f xs
   792 
   793 * Pure/library.ML: several combinators for linear functional
   794 transformations, notably reverse application and composition:
   795 
   796   x |> f                f #> g
   797   (x, y) |-> f          f #-> g
   798 
   799 * Pure/library.ML: introduced/changed precedence of infix operators:
   800 
   801   infix 1 |> |-> ||> ||>> |>> |>>> #> #->;
   802   infix 2 ?;
   803   infix 3 o oo ooo oooo;
   804   infix 4 ~~ upto downto;
   805 
   806 Maybe INCOMPATIBILITY when any of those is used in conjunction with other
   807 infix operators.
   808 
   809 * Pure/library.ML: natural list combinators fold, fold_rev, and
   810 fold_map support linear functional transformations and nesting.  For
   811 example:
   812 
   813   fold f [x1, ..., xN] y =
   814     y |> f x1 |> ... |> f xN
   815 
   816   (fold o fold) f [xs1, ..., xsN] y =
   817     y |> fold f xs1 |> ... |> fold f xsN
   818 
   819   fold f [x1, ..., xN] =
   820     f x1 #> ... #> f xN
   821 
   822   (fold o fold) f [xs1, ..., xsN] =
   823     fold f xs1 #> ... #> fold f xsN
   824 
   825 * Pure/library.ML: the following selectors on type 'a option are
   826 available:
   827 
   828   the:               'a option -> 'a  (*partial*)
   829   these:             'a option -> 'a  where 'a = 'b list
   830   the_default: 'a -> 'a option -> 'a
   831   the_list:          'a option -> 'a list
   832 
   833 * Pure/General: structure AList (cf. Pure/General/alist.ML) provides
   834 basic operations for association lists, following natural argument
   835 order; moreover the explicit equality predicate passed here avoids
   836 potentially expensive polymorphic runtime equality checks.
   837 The old functions may be expressed as follows:
   838 
   839   assoc = uncurry (AList.lookup (op =))
   840   assocs = these oo AList.lookup (op =)
   841   overwrite = uncurry (AList.update (op =)) o swap
   842 
   843 * Pure/General: structure AList (cf. Pure/General/alist.ML) provides
   844 
   845   val make: ('a -> 'b) -> 'a list -> ('a * 'b) list
   846   val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list
   847 
   848 replacing make_keylist and keyfilter (occassionally used)
   849 Naive rewrites:
   850 
   851   make_keylist = AList.make
   852   keyfilter = AList.find (op =)
   853 
   854 * eq_fst and eq_snd now take explicit equality parameter, thus
   855   avoiding eqtypes. Naive rewrites:
   856 
   857     eq_fst = eq_fst (op =)
   858     eq_snd = eq_snd (op =)
   859 
   860 * Removed deprecated apl and apr (rarely used).
   861   Naive rewrites:
   862 
   863     apl (n, op) =>>= curry op n
   864     apr (op, m) =>>= fn n => op (n, m)
   865 
   866 * Pure/General: structure OrdList (cf. Pure/General/ord_list.ML)
   867 provides a reasonably efficient light-weight implementation of sets as
   868 lists.
   869 
   870 * Pure/General: generic tables (cf. Pure/General/table.ML) provide a
   871 few new operations; existing lookup and update are now curried to
   872 follow natural argument order (for use with fold etc.);
   873 INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort.
   874 
   875 * Pure/General: output via the Isabelle channels of
   876 writeln/warning/error etc. is now passed through Output.output, with a
   877 hook for arbitrary transformations depending on the print_mode
   878 (cf. Output.add_mode -- the first active mode that provides a output
   879 function wins).  Already formatted output may be embedded into further
   880 text via Output.raw; the result of Pretty.string_of/str_of and derived
   881 functions (string_of_term/cterm/thm etc.) is already marked raw to
   882 accommodate easy composition of diagnostic messages etc.  Programmers
   883 rarely need to care about Output.output or Output.raw at all, with
   884 some notable exceptions: Output.output is required when bypassing the
   885 standard channels (writeln etc.), or in token translations to produce
   886 properly formatted results; Output.raw is required when capturing
   887 already output material that will eventually be presented to the user
   888 a second time.  For the default print mode, both Output.output and
   889 Output.raw have no effect.
   890 
   891 * Pure/General: Output.time_accumulator NAME creates an operator ('a
   892 -> 'b) -> 'a -> 'b to measure runtime and count invocations; the
   893 cumulative results are displayed at the end of a batch session.
   894 
   895 * Pure/General: File.sysify_path and File.quote_sysify path have been
   896 replaced by File.platform_path and File.shell_path (with appropriate
   897 hooks).  This provides a clean interface for unusual systems where the
   898 internal and external process view of file names are different.
   899 
   900 * Pure: more efficient orders for basic syntactic entities: added
   901 fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord
   902 and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is
   903 NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast
   904 orders now -- potential INCOMPATIBILITY for code that depends on a
   905 particular order for Symtab.keys, Symtab.dest, etc. (consider using
   906 Library.sort_strings on result).
   907 
   908 * Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types,
   909 fold_types traverse types/terms from left to right, observing natural
   910 argument order.  Supercedes previous foldl_XXX versions, add_frees,
   911 add_vars etc. have been adapted as well: INCOMPATIBILITY.
   912 
   913 * Pure: name spaces have been refined, with significant changes of the
   914 internal interfaces -- INCOMPATIBILITY.  Renamed cond_extern(_table)
   915 to extern(_table).  The plain name entry path is superceded by a
   916 general 'naming' context, which also includes the 'policy' to produce
   917 a fully qualified name and external accesses of a fully qualified
   918 name; NameSpace.extend is superceded by context dependent
   919 Sign.declare_name.  Several theory and proof context operations modify
   920 the naming context.  Especially note Theory.restore_naming and
   921 ProofContext.restore_naming to get back to a sane state; note that
   922 Theory.add_path is no longer sufficient to recover from
   923 Theory.absolute_path in particular.
   924 
   925 * Pure: new flags short_names (default false) and unique_names
   926 (default true) for controlling output of qualified names.  If
   927 short_names is set, names are printed unqualified.  If unique_names is
   928 reset, the name prefix is reduced to the minimum required to achieve
   929 the original result when interning again, even if there is an overlap
   930 with earlier declarations.
   931 
   932 * Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is
   933 now 'extend', and 'merge' gets an additional Pretty.pp argument
   934 (useful for printing error messages).  INCOMPATIBILITY.
   935 
   936 * Pure: major reorganization of the theory context.  Type Sign.sg and
   937 Theory.theory are now identified, referring to the universal
   938 Context.theory (see Pure/context.ML).  Actual signature and theory
   939 content is managed as theory data.  The old code and interfaces were
   940 spread over many files and structures; the new arrangement introduces
   941 considerable INCOMPATIBILITY to gain more clarity:
   942 
   943   Context -- theory management operations (name, identity, inclusion,
   944     parents, ancestors, merge, etc.), plus generic theory data;
   945 
   946   Sign -- logical signature and syntax operations (declaring consts,
   947     types, etc.), plus certify/read for common entities;
   948 
   949   Theory -- logical theory operations (stating axioms, definitions,
   950     oracles), plus a copy of logical signature operations (consts,
   951     types, etc.); also a few basic management operations (Theory.copy,
   952     Theory.merge, etc.)
   953 
   954 The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm
   955 etc.) as well as the sign field in Thm.rep_thm etc. have been retained
   956 for convenience -- they merely return the theory.
   957 
   958 * Pure: type Type.tsig is superceded by theory in most interfaces.
   959 
   960 * Pure: the Isar proof context type is already defined early in Pure
   961 as Context.proof (note that ProofContext.context and Proof.context are
   962 aliases, where the latter is the preferred name).  This enables other
   963 Isabelle components to refer to that type even before Isar is present.
   964 
   965 * Pure/sign/theory: discontinued named name spaces (i.e. classK,
   966 typeK, constK, axiomK, oracleK), but provide explicit operations for
   967 any of these kinds.  For example, Sign.intern typeK is now
   968 Sign.intern_type, Theory.hide_space Sign.typeK is now
   969 Theory.hide_types.  Also note that former
   970 Theory.hide_classes/types/consts are now
   971 Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions
   972 internalize their arguments!  INCOMPATIBILITY.
   973 
   974 * Pure: get_thm interface (of PureThy and ProofContext) expects
   975 datatype thmref (with constructors Name and NameSelection) instead of
   976 plain string -- INCOMPATIBILITY;
   977 
   978 * Pure: cases produced by proof methods specify options, where NONE
   979 means to remove case bindings -- INCOMPATIBILITY in
   980 (RAW_)METHOD_CASES.
   981 
   982 * Pure: the following operations retrieve axioms or theorems from a
   983 theory node or theory hierarchy, respectively:
   984 
   985   Theory.axioms_of: theory -> (string * term) list
   986   Theory.all_axioms_of: theory -> (string * term) list
   987   PureThy.thms_of: theory -> (string * thm) list
   988   PureThy.all_thms_of: theory -> (string * thm) list
   989 
   990 * Pure: print_tac now outputs the goal through the trace channel.
   991 
   992 * Isar toplevel: improved diagnostics, mostly for Poly/ML only.
   993 Reference Toplevel.debug (default false) controls detailed printing
   994 and tracing of low-level exceptions; Toplevel.profiling (default 0)
   995 controls execution profiling -- set to 1 for time and 2 for space
   996 (both increase the runtime).
   997 
   998 * Isar session: The initial use of ROOT.ML is now always timed,
   999 i.e. the log will show the actual process times, in contrast to the
  1000 elapsed wall-clock time that the outer shell wrapper produces.
  1001 
  1002 * Simplifier: improved handling of bound variables (nameless
  1003 representation, avoid allocating new strings).  Simprocs that invoke
  1004 the Simplifier recursively should use Simplifier.inherit_bounds to
  1005 avoid local name clashes.  Failure to do so produces warnings
  1006 "Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds
  1007 for further details.
  1008 
  1009 * ML functions legacy_bindings and use_legacy_bindings produce ML fact
  1010 bindings for all theorems stored within a given theory; this may help
  1011 in porting non-Isar theories to Isar ones, while keeping ML proof
  1012 scripts for the time being.
  1013 
  1014 * ML operator HTML.with_charset specifies the charset begin used for
  1015 generated HTML files.  For example:
  1016 
  1017   HTML.with_charset "utf-8" use_thy "Hebrew";
  1018   HTML.with_charset "utf-8" use_thy "Chinese";
  1019 
  1020 
  1021 *** System ***
  1022 
  1023 * Allow symlinks to all proper Isabelle executables (Isabelle,
  1024 isabelle, isatool etc.).
  1025 
  1026 * ISABELLE_DOC_FORMAT setting specifies preferred document format (for
  1027 isatool doc, isatool mkdir, display_drafts etc.).
  1028 
  1029 * isatool usedir: option -f allows specification of the ML file to be
  1030 used by Isabelle; default is ROOT.ML.
  1031 
  1032 * New isatool version outputs the version identifier of the Isabelle
  1033 distribution being used.
  1034 
  1035 * HOL: new isatool dimacs2hol converts files in DIMACS CNF format
  1036 (containing Boolean satisfiability problems) into Isabelle/HOL
  1037 theories.
  1038 
  1039 
  1040 
  1041 New in Isabelle2004 (April 2004)
  1042 --------------------------------
  1043 
  1044 *** General ***
  1045 
  1046 * Provers/order.ML:  new efficient reasoner for partial and linear orders.
  1047   Replaces linorder.ML.
  1048 
  1049 * Pure: Greek letters (except small lambda, \<lambda>), as well as Gothic
  1050   (\<aa>...\<zz>\<AA>...\<ZZ>), calligraphic (\<A>...\<Z>), and Euler
  1051   (\<a>...\<z>), are now considered normal letters, and can therefore
  1052   be used anywhere where an ASCII letter (a...zA...Z) has until
  1053   now. COMPATIBILITY: This obviously changes the parsing of some
  1054   terms, especially where a symbol has been used as a binder, say
  1055   '\<Pi>x. ...', which is now a type error since \<Pi>x will be parsed
  1056   as an identifier.  Fix it by inserting a space around former
  1057   symbols.  Call 'isatool fixgreek' to try to fix parsing errors in
  1058   existing theory and ML files.
  1059 
  1060 * Pure: Macintosh and Windows line-breaks are now allowed in theory files.
  1061 
  1062 * Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now
  1063   allowed in identifiers. Similar to Greek letters \<^isub> is now considered
  1064   a normal (but invisible) letter. For multiple letter subscripts repeat
  1065   \<^isub> like this: x\<^isub>1\<^isub>2.
  1066 
  1067 * Pure: There are now sub-/superscripts that can span more than one
  1068   character. Text between \<^bsub> and \<^esub> is set in subscript in
  1069   ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in
  1070   superscript. The new control characters are not identifier parts.
  1071 
  1072 * Pure: Control-symbols of the form \<^raw:...> will literally print the
  1073   content of "..." to the latex file instead of \isacntrl... . The "..."
  1074   may consist of any printable characters excluding the end bracket >.
  1075 
  1076 * Pure: Using new Isar command "finalconsts" (or the ML functions
  1077   Theory.add_finals or Theory.add_finals_i) it is now possible to
  1078   declare constants "final", which prevents their being given a definition
  1079   later.  It is useful for constants whose behaviour is fixed axiomatically
  1080   rather than definitionally, such as the meta-logic connectives.
  1081 
  1082 * Pure: 'instance' now handles general arities with general sorts
  1083   (i.e. intersections of classes),
  1084 
  1085 * Presentation: generated HTML now uses a CSS style sheet to make layout
  1086   (somewhat) independent of content. It is copied from lib/html/isabelle.css.
  1087   It can be changed to alter the colors/layout of generated pages.
  1088 
  1089 
  1090 *** Isar ***
  1091 
  1092 * Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac,
  1093   cut_tac, subgoal_tac and thin_tac:
  1094   - Now understand static (Isar) contexts.  As a consequence, users of Isar
  1095     locales are no longer forced to write Isar proof scripts.
  1096     For details see Isar Reference Manual, paragraph 4.3.2: Further tactic
  1097     emulations.
  1098   - INCOMPATIBILITY: names of variables to be instantiated may no
  1099     longer be enclosed in quotes.  Instead, precede variable name with `?'.
  1100     This is consistent with the instantiation attribute "where".
  1101 
  1102 * Attributes "where" and "of":
  1103   - Now take type variables of instantiated theorem into account when reading
  1104     the instantiation string.  This fixes a bug that caused instantiated
  1105     theorems to have too special types in some circumstances.
  1106   - "where" permits explicit instantiations of type variables.
  1107 
  1108 * Calculation commands "moreover" and "also" no longer interfere with
  1109   current facts ("this"), admitting arbitrary combinations with "then"
  1110   and derived forms.
  1111 
  1112 * Locales:
  1113   - Goal statements involving the context element "includes" no longer
  1114     generate theorems with internal delta predicates (those ending on
  1115     "_axioms") in the premise.
  1116     Resolve particular premise with <locale>.intro to obtain old form.
  1117   - Fixed bug in type inference ("unify_frozen") that prevented mix of target
  1118     specification and "includes" elements in goal statement.
  1119   - Rule sets <locale>.intro and <locale>.axioms no longer declared as
  1120     [intro?] and [elim?] (respectively) by default.
  1121   - Experimental command for instantiation of locales in proof contexts:
  1122         instantiate <label>[<attrs>]: <loc>
  1123     Instantiates locale <loc> and adds all its theorems to the current context
  1124     taking into account their attributes.  Label and attrs are optional
  1125     modifiers, like in theorem declarations.  If present, names of
  1126     instantiated theorems are qualified with <label>, and the attributes
  1127     <attrs> are applied after any attributes these theorems might have already.
  1128       If the locale has assumptions, a chained fact of the form
  1129     "<loc> t1 ... tn" is expected from which instantiations of the parameters
  1130     are derived.  The command does not support old-style locales declared
  1131     with "locale (open)".
  1132       A few (very simple) examples can be found in FOL/ex/LocaleInst.thy.
  1133 
  1134 * HOL: Tactic emulation methods induct_tac and case_tac understand static
  1135   (Isar) contexts.
  1136 
  1137 
  1138 *** HOL ***
  1139 
  1140 * Proof import: new image HOL4 contains the imported library from
  1141   the HOL4 system with about 2500 theorems. It is imported by
  1142   replaying proof terms produced by HOL4 in Isabelle. The HOL4 image
  1143   can be used like any other Isabelle image.  See
  1144   HOL/Import/HOL/README for more information.
  1145 
  1146 * Simplifier:
  1147   - Much improved handling of linear and partial orders.
  1148     Reasoners for linear and partial orders are set up for type classes
  1149     "linorder" and "order" respectively, and are added to the default simpset
  1150     as solvers.  This means that the simplifier can build transitivity chains
  1151     to solve goals from the assumptions.
  1152   - INCOMPATIBILITY: old proofs break occasionally.  Typically, applications
  1153     of blast or auto after simplification become unnecessary because the goal
  1154     is solved by simplification already.
  1155 
  1156 * Numerics: new theory Ring_and_Field contains over 250 basic numerical laws,
  1157     all proved in axiomatic type classes for semirings, rings and fields.
  1158 
  1159 * Numerics:
  1160   - Numeric types (nat, int, and in HOL-Complex rat, real, complex, etc.) are
  1161     now formalized using the Ring_and_Field theory mentioned above.
  1162   - INCOMPATIBILITY: simplification and arithmetic behaves somewhat differently
  1163     than before, because now they are set up once in a generic manner.
  1164   - INCOMPATIBILITY: many type-specific arithmetic laws have gone.
  1165     Look for the general versions in Ring_and_Field (and Power if they concern
  1166     exponentiation).
  1167 
  1168 * Type "rat" of the rational numbers is now available in HOL-Complex.
  1169 
  1170 * Records:
  1171   - Record types are now by default printed with their type abbreviation
  1172     instead of the list of all field types. This can be configured via
  1173     the reference "print_record_type_abbr".
  1174   - Simproc "record_upd_simproc" for simplification of multiple updates added
  1175     (not enabled by default).
  1176   - Simproc "record_ex_sel_eq_simproc" to simplify EX x. sel r = x resp.
  1177     EX x. x = sel r to True (not enabled by default).
  1178   - Tactic "record_split_simp_tac" to split and simplify records added.
  1179 
  1180 * 'specification' command added, allowing for definition by
  1181   specification.  There is also an 'ax_specification' command that
  1182   introduces the new constants axiomatically.
  1183 
  1184 * arith(_tac) is now able to generate counterexamples for reals as well.
  1185 
  1186 * HOL-Algebra: new locale "ring" for non-commutative rings.
  1187 
  1188 * HOL-ex: InductiveInvariant_examples illustrates advanced recursive function
  1189   definitions, thanks to Sava Krsti\'{c} and John Matthews.
  1190 
  1191 * HOL-Matrix: a first theory for matrices in HOL with an application of
  1192   matrix theory to linear programming.
  1193 
  1194 * Unions and Intersections:
  1195   The latex output syntax of UN and INT has been changed
  1196   from "\Union x \in A. B" to "\Union_{x \in A} B"
  1197   i.e. the index formulae has become a subscript.
  1198   Similarly for "\Union x. B", and for \Inter instead of \Union.
  1199 
  1200 * Unions and Intersections over Intervals:
  1201   There is new short syntax "UN i<=n. A" for "UN i:{0..n}. A". There is
  1202   also an x-symbol version with subscripts "\<Union>\<^bsub>i <= n\<^esub>. A"
  1203   like in normal math, and corresponding versions for < and for intersection.
  1204 
  1205 * HOL/List: Ordering "lexico" is renamed "lenlex" and the standard
  1206   lexicographic dictonary ordering has been added as "lexord".
  1207 
  1208 * ML: the legacy theory structures Int and List have been removed. They had
  1209   conflicted with ML Basis Library structures having the same names.
  1210 
  1211 * 'refute' command added to search for (finite) countermodels.  Only works
  1212   for a fragment of HOL.  The installation of an external SAT solver is
  1213   highly recommended.  See "HOL/Refute.thy" for details.
  1214 
  1215 * 'quickcheck' command: Allows to find counterexamples by evaluating
  1216   formulae under an assignment of free variables to random values.
  1217   In contrast to 'refute', it can deal with inductive datatypes,
  1218   but cannot handle quantifiers. See "HOL/ex/Quickcheck_Examples.thy"
  1219   for examples.
  1220 
  1221 
  1222 *** HOLCF ***
  1223 
  1224 * Streams now come with concatenation and are part of the HOLCF image
  1225 
  1226 
  1227 
  1228 New in Isabelle2003 (May 2003)
  1229 ------------------------------
  1230 
  1231 *** General ***
  1232 
  1233 * Provers/simplifier:
  1234 
  1235   - Completely reimplemented method simp (ML: Asm_full_simp_tac):
  1236     Assumptions are now subject to complete mutual simplification,
  1237     not just from left to right. The simplifier now preserves
  1238     the order of assumptions.
  1239 
  1240     Potential INCOMPATIBILITY:
  1241 
  1242     -- simp sometimes diverges where the old version did
  1243        not, e.g. invoking simp on the goal
  1244 
  1245         [| P (f x); y = x; f x = f y |] ==> Q
  1246 
  1247        now gives rise to the infinite reduction sequence
  1248 
  1249         P(f x) --(f x = f y)--> P(f y) --(y = x)--> P(f x) --(f x = f y)--> ...
  1250 
  1251        Using "simp (asm_lr)" (ML: Asm_lr_simp_tac) instead often solves this
  1252        kind of problem.
  1253 
  1254     -- Tactics combining classical reasoner and simplification (such as auto)
  1255        are also affected by this change, because many of them rely on
  1256        simp. They may sometimes diverge as well or yield a different numbers
  1257        of subgoals. Try to use e.g. force, fastsimp, or safe instead of auto
  1258        in case of problems. Sometimes subsequent calls to the classical
  1259        reasoner will fail because a preceeding call to the simplifier too
  1260        eagerly simplified the goal, e.g. deleted redundant premises.
  1261 
  1262   - The simplifier trace now shows the names of the applied rewrite rules
  1263 
  1264   - You can limit the number of recursive invocations of the simplifier
  1265     during conditional rewriting (where the simplifie tries to solve the
  1266     conditions before applying the rewrite rule):
  1267     ML "simp_depth_limit := n"
  1268     where n is an integer. Thus you can force termination where previously
  1269     the simplifier would diverge.
  1270 
  1271   - Accepts free variables as head terms in congruence rules.  Useful in Isar.
  1272 
  1273   - No longer aborts on failed congruence proof.  Instead, the
  1274     congruence is ignored.
  1275 
  1276 * Pure: New generic framework for extracting programs from constructive
  1277   proofs. See HOL/Extraction.thy for an example instantiation, as well
  1278   as HOL/Extraction for some case studies.
  1279 
  1280 * Pure: The main goal of the proof state is no longer shown by default, only
  1281 the subgoals. This behaviour is controlled by a new flag.
  1282    PG menu: Isabelle/Isar -> Settings -> Show Main Goal
  1283 (ML: Proof.show_main_goal).
  1284 
  1285 * Pure: You can find all matching introduction rules for subgoal 1, i.e. all
  1286 rules whose conclusion matches subgoal 1:
  1287       PG menu: Isabelle/Isar -> Show me -> matching rules
  1288 The rules are ordered by how closely they match the subgoal.
  1289 In particular, rules that solve a subgoal outright are displayed first
  1290 (or rather last, the way they are printed).
  1291 (ML: ProofGeneral.print_intros())
  1292 
  1293 * Pure: New flag trace_unify_fail causes unification to print
  1294 diagnostic information (PG: in trace buffer) when it fails. This is
  1295 useful for figuring out why single step proofs like rule, erule or
  1296 assumption failed.
  1297 
  1298 * Pure: Locale specifications now produce predicate definitions
  1299 according to the body of text (covering assumptions modulo local
  1300 definitions); predicate "loc_axioms" covers newly introduced text,
  1301 while "loc" is cumulative wrt. all included locale expressions; the
  1302 latter view is presented only on export into the global theory
  1303 context; potential INCOMPATIBILITY, use "(open)" option to fall back
  1304 on the old view without predicates;
  1305 
  1306 * Pure: predefined locales "var" and "struct" are useful for sharing
  1307 parameters (as in CASL, for example); just specify something like
  1308 ``var x + var y + struct M'' as import;
  1309 
  1310 * Pure: improved thms_containing: proper indexing of facts instead of
  1311 raw theorems; check validity of results wrt. current name space;
  1312 include local facts of proof configuration (also covers active
  1313 locales), cover fixed variables in index; may use "_" in term
  1314 specification; an optional limit for the number of printed facts may
  1315 be given (the default is 40);
  1316 
  1317 * Pure: disallow duplicate fact bindings within new-style theory files
  1318 (batch-mode only);
  1319 
  1320 * Provers: improved induct method: assumptions introduced by case
  1321 "foo" are split into "foo.hyps" (from the rule) and "foo.prems" (from
  1322 the goal statement); "foo" still refers to all facts collectively;
  1323 
  1324 * Provers: the function blast.overloaded has been removed: all constants
  1325 are regarded as potentially overloaded, which improves robustness in exchange
  1326 for slight decrease in efficiency;
  1327 
  1328 * Provers/linorder: New generic prover for transitivity reasoning over
  1329 linear orders.  Note: this prover is not efficient!
  1330 
  1331 * Isar: preview of problems to finish 'show' now produce an error
  1332 rather than just a warning (in interactive mode);
  1333 
  1334 
  1335 *** HOL ***
  1336 
  1337 * arith(_tac)
  1338 
  1339  - Produces a counter example if it cannot prove a goal.
  1340    Note that the counter example may be spurious if the goal is not a formula
  1341    of quantifier-free linear arithmetic.
  1342    In ProofGeneral the counter example appears in the trace buffer.
  1343 
  1344  - Knows about div k and mod k where k is a numeral of type nat or int.
  1345 
  1346  - Calls full Presburger arithmetic (by Amine Chaieb) if quantifier-free
  1347    linear arithmetic fails. This takes account of quantifiers and divisibility.
  1348    Presburger arithmetic can also be called explicitly via presburger(_tac).
  1349 
  1350 * simp's arithmetic capabilities have been enhanced a bit: it now
  1351 takes ~= in premises into account (by performing a case split);
  1352 
  1353 * simp reduces "m*(n div m) + n mod m" to n, even if the two summands
  1354 are distributed over a sum of terms;
  1355 
  1356 * New tactic "trans_tac" and method "trans" instantiate
  1357 Provers/linorder.ML for axclasses "order" and "linorder" (predicates
  1358 "<=", "<" and "=").
  1359 
  1360 * function INCOMPATIBILITIES: Pi-sets have been redefined and moved from main
  1361 HOL to Library/FuncSet; constant "Fun.op o" is now called "Fun.comp";
  1362 
  1363 * 'typedef' command has new option "open" to suppress the set
  1364 definition;
  1365 
  1366 * functions Min and Max on finite sets have been introduced (theory
  1367 Finite_Set);
  1368 
  1369 * attribute [symmetric] now works for relations as well; it turns
  1370 (x,y) : R^-1 into (y,x) : R, and vice versa;
  1371 
  1372 * induct over a !!-quantified statement (say !!x1..xn):
  1373   each "case" automatically performs "fix x1 .. xn" with exactly those names.
  1374 
  1375 * Map: `empty' is no longer a constant but a syntactic abbreviation for
  1376 %x. None. Warning: empty_def now refers to the previously hidden definition
  1377 of the empty set.
  1378 
  1379 * Algebra: formalization of classical algebra.  Intended as base for
  1380 any algebraic development in Isabelle.  Currently covers group theory
  1381 (up to Sylow's theorem) and ring theory (Universal Property of
  1382 Univariate Polynomials).  Contributions welcome;
  1383 
  1384 * GroupTheory: deleted, since its material has been moved to Algebra;
  1385 
  1386 * Complex: new directory of the complex numbers with numeric constants,
  1387 nonstandard complex numbers, and some complex analysis, standard and
  1388 nonstandard (Jacques Fleuriot);
  1389 
  1390 * HOL-Complex: new image for analysis, replacing HOL-Real and HOL-Hyperreal;
  1391 
  1392 * Hyperreal: introduced Gauge integration and hyperreal logarithms (Jacques
  1393 Fleuriot);
  1394 
  1395 * Real/HahnBanach: updated and adapted to locales;
  1396 
  1397 * NumberTheory: added Gauss's law of quadratic reciprocity (by Avigad,
  1398 Gray and Kramer);
  1399 
  1400 * UNITY: added the Meier-Sanders theory of progress sets;
  1401 
  1402 * MicroJava: bytecode verifier and lightweight bytecode verifier
  1403 as abstract algorithms, instantiated to the JVM;
  1404 
  1405 * Bali: Java source language formalization. Type system, operational
  1406 semantics, axiomatic semantics. Supported language features:
  1407 classes, interfaces, objects,virtual methods, static methods,
  1408 static/instance fields, arrays, access modifiers, definite
  1409 assignment, exceptions.
  1410 
  1411 
  1412 *** ZF ***
  1413 
  1414 * ZF/Constructible: consistency proof for AC (Gdel's constructible
  1415 universe, etc.);
  1416 
  1417 * Main ZF: virtually all theories converted to new-style format;
  1418 
  1419 
  1420 *** ML ***
  1421 
  1422 * Pure: Tactic.prove provides sane interface for internal proofs;
  1423 omits the infamous "standard" operation, so this is more appropriate
  1424 than prove_goalw_cterm in many situations (e.g. in simprocs);
  1425 
  1426 * Pure: improved error reporting of simprocs;
  1427 
  1428 * Provers: Simplifier.simproc(_i) provides sane interface for setting
  1429 up simprocs;
  1430 
  1431 
  1432 *** Document preparation ***
  1433 
  1434 * uses \par instead of \\ for line breaks in theory text. This may
  1435 shift some page breaks in large documents. To get the old behaviour
  1436 use \renewcommand{\isanewline}{\mbox{}\\\mbox{}} in root.tex.
  1437 
  1438 * minimized dependencies of isabelle.sty and isabellesym.sty on
  1439 other packages
  1440 
  1441 * \<euro> now needs package babel/greek instead of marvosym (which
  1442 broke \Rightarrow)
  1443 
  1444 * normal size for \<zero>...\<nine> (uses \mathbf instead of
  1445 textcomp package)
  1446 
  1447 
  1448 
  1449 New in Isabelle2002 (March 2002)
  1450 --------------------------------
  1451 
  1452 *** Document preparation ***
  1453 
  1454 * greatly simplified document preparation setup, including more
  1455 graceful interpretation of isatool usedir -i/-d/-D options, and more
  1456 instructive isatool mkdir; users should basically be able to get
  1457 started with "isatool mkdir HOL Test && isatool make"; alternatively,
  1458 users may run a separate document processing stage manually like this:
  1459 "isatool usedir -D output HOL Test && isatool document Test/output";
  1460 
  1461 * theory dependency graph may now be incorporated into documents;
  1462 isatool usedir -g true will produce session_graph.eps/.pdf for use
  1463 with \includegraphics of LaTeX;
  1464 
  1465 * proper spacing of consecutive markup elements, especially text
  1466 blocks after section headings;
  1467 
  1468 * support bold style (for single symbols only), input syntax is like
  1469 this: "\<^bold>\<alpha>" or "\<^bold>A";
  1470 
  1471 * \<bullet> is now output as bold \cdot by default, which looks much
  1472 better in printed text;
  1473 
  1474 * added default LaTeX bindings for \<tturnstile> and \<TTurnstile>;
  1475 note that these symbols are currently unavailable in Proof General /
  1476 X-Symbol; new symbols \<zero>, \<one>, ..., \<nine>, and \<euro>;
  1477 
  1478 * isatool latex no longer depends on changed TEXINPUTS, instead
  1479 isatool document copies the Isabelle style files to the target
  1480 location;
  1481 
  1482 
  1483 *** Isar ***
  1484 
  1485 * Pure/Provers: improved proof by cases and induction;
  1486   - 'case' command admits impromptu naming of parameters (such as
  1487     "case (Suc n)");
  1488   - 'induct' method divinates rule instantiation from the inductive
  1489     claim; no longer requires excessive ?P bindings for proper
  1490     instantiation of cases;
  1491   - 'induct' method properly enumerates all possibilities of set/type
  1492     rules; as a consequence facts may be also passed through *type*
  1493     rules without further ado;
  1494   - 'induct' method now derives symbolic cases from the *rulified*
  1495     rule (before it used to rulify cases stemming from the internal
  1496     atomized version); this means that the context of a non-atomic
  1497     statement becomes is included in the hypothesis, avoiding the
  1498     slightly cumbersome show "PROP ?case" form;
  1499   - 'induct' may now use elim-style induction rules without chaining
  1500     facts, using ``missing'' premises from the goal state; this allows
  1501     rules stemming from inductive sets to be applied in unstructured
  1502     scripts, while still benefitting from proper handling of non-atomic
  1503     statements; NB: major inductive premises need to be put first, all
  1504     the rest of the goal is passed through the induction;
  1505   - 'induct' proper support for mutual induction involving non-atomic
  1506     rule statements (uses the new concept of simultaneous goals, see
  1507     below);
  1508   - append all possible rule selections, but only use the first
  1509     success (no backtracking);
  1510   - removed obsolete "(simplified)" and "(stripped)" options of methods;
  1511   - undeclared rule case names default to numbers 1, 2, 3, ...;
  1512   - added 'print_induct_rules' (covered by help item in recent Proof
  1513     General versions);
  1514   - moved induct/cases attributes to Pure, methods to Provers;
  1515   - generic method setup instantiated for FOL and HOL;
  1516 
  1517 * Pure: support multiple simultaneous goal statements, for example
  1518 "have a: A and b: B" (same for 'theorem' etc.); being a pure
  1519 meta-level mechanism, this acts as if several individual goals had
  1520 been stated separately; in particular common proof methods need to be
  1521 repeated in order to cover all claims; note that a single elimination
  1522 step is *not* sufficient to establish the two conjunctions, so this
  1523 fails:
  1524 
  1525   assume "A & B" then have A and B ..   (*".." fails*)
  1526 
  1527 better use "obtain" in situations as above; alternative refer to
  1528 multi-step methods like 'auto', 'simp_all', 'blast+' etc.;
  1529 
  1530 * Pure: proper integration with ``locales''; unlike the original
  1531 version by Florian Kammller, Isar locales package high-level proof
  1532 contexts rather than raw logical ones (e.g. we admit to include
  1533 attributes everywhere); operations on locales include merge and
  1534 rename; support for implicit arguments (``structures''); simultaneous
  1535 type-inference over imports and text; see also HOL/ex/Locales.thy for
  1536 some examples;
  1537 
  1538 * Pure: the following commands have been ``localized'', supporting a
  1539 target locale specification "(in name)": 'lemma', 'theorem',
  1540 'corollary', 'lemmas', 'theorems', 'declare'; the results will be
  1541 stored both within the locale and at the theory level (exported and
  1542 qualified by the locale name);
  1543 
  1544 * Pure: theory goals may now be specified in ``long'' form, with
  1545 ad-hoc contexts consisting of arbitrary locale elements. for example
  1546 ``lemma foo: fixes x assumes "A x" shows "B x"'' (local syntax and
  1547 definitions may be given, too); the result is a meta-level rule with
  1548 the context elements being discharged in the obvious way;
  1549 
  1550 * Pure: new proof command 'using' allows to augment currently used
  1551 facts after a goal statement ('using' is syntactically analogous to
  1552 'apply', but acts on the goal's facts only); this allows chained facts
  1553 to be separated into parts given before and after a claim, as in
  1554 ``from a and b have C using d and e <proof>'';
  1555 
  1556 * Pure: renamed "antecedent" case to "rule_context";
  1557 
  1558 * Pure: new 'judgment' command records explicit information about the
  1559 object-logic embedding (used by several tools internally); no longer
  1560 use hard-wired "Trueprop";
  1561 
  1562 * Pure: added 'corollary' command;
  1563 
  1564 * Pure: fixed 'token_translation' command;
  1565 
  1566 * Pure: removed obsolete 'exported' attribute;
  1567 
  1568 * Pure: dummy pattern "_" in is/let is now automatically lifted over
  1569 bound variables: "ALL x. P x --> Q x" (is "ALL x. _ --> ?C x")
  1570 supersedes more cumbersome ... (is "ALL x. _ x --> ?C x");
  1571 
  1572 * Pure: method 'atomize' presents local goal premises as object-level
  1573 statements (atomic meta-level propositions); setup controlled via
  1574 rewrite rules declarations of 'atomize' attribute; example
  1575 application: 'induct' method with proper rule statements in improper
  1576 proof *scripts*;
  1577 
  1578 * Pure: emulation of instantiation tactics (rule_tac, cut_tac, etc.)
  1579 now consider the syntactic context of assumptions, giving a better
  1580 chance to get type-inference of the arguments right (this is
  1581 especially important for locales);
  1582 
  1583 * Pure: "sorry" no longer requires quick_and_dirty in interactive
  1584 mode;
  1585 
  1586 * Pure/obtain: the formal conclusion "thesis", being marked as
  1587 ``internal'', may no longer be reference directly in the text;
  1588 potential INCOMPATIBILITY, may need to use "?thesis" in rare
  1589 situations;
  1590 
  1591 * Pure: generic 'sym' attribute which declares a rule both as pure
  1592 'elim?' and for the 'symmetric' operation;
  1593 
  1594 * Pure: marginal comments ``--'' may now occur just anywhere in the
  1595 text; the fixed correlation with particular command syntax has been
  1596 discontinued;
  1597 
  1598 * Pure: new method 'rules' is particularly well-suited for proof
  1599 search in intuitionistic logic; a bit slower than 'blast' or 'fast',
  1600 but often produces more compact proof terms with less detours;
  1601 
  1602 * Pure/Provers/classical: simplified integration with pure rule
  1603 attributes and methods; the classical "intro?/elim?/dest?"
  1604 declarations coincide with the pure ones; the "rule" method no longer
  1605 includes classically swapped intros; "intro" and "elim" methods no
  1606 longer pick rules from the context; also got rid of ML declarations
  1607 AddXIs/AddXEs/AddXDs; all of this has some potential for
  1608 INCOMPATIBILITY;
  1609 
  1610 * Provers/classical: attribute 'swapped' produces classical inversions
  1611 of introduction rules;
  1612 
  1613 * Provers/simplifier: 'simplified' attribute may refer to explicit
  1614 rules instead of full simplifier context; 'iff' attribute handles
  1615 conditional rules;
  1616 
  1617 * HOL: 'typedef' now allows alternative names for Rep/Abs morphisms;
  1618 
  1619 * HOL: 'recdef' now fails on unfinished automated proofs, use
  1620 "(permissive)" option to recover old behavior;
  1621 
  1622 * HOL: 'inductive' no longer features separate (collective) attributes
  1623 for 'intros' (was found too confusing);
  1624 
  1625 * HOL: properly declared induction rules less_induct and
  1626 wf_induct_rule;
  1627 
  1628 
  1629 *** HOL ***
  1630 
  1631 * HOL: moved over to sane numeral syntax; the new policy is as
  1632 follows:
  1633 
  1634   - 0 and 1 are polymorphic constants, which are defined on any
  1635   numeric type (nat, int, real etc.);
  1636 
  1637   - 2, 3, 4, ... and -1, -2, -3, ... are polymorphic numerals, based
  1638   binary representation internally;
  1639 
  1640   - type nat has special constructor Suc, and generally prefers Suc 0
  1641   over 1::nat and Suc (Suc 0) over 2::nat;
  1642 
  1643 This change may cause significant problems of INCOMPATIBILITY; here
  1644 are some hints on converting existing sources:
  1645 
  1646   - due to the new "num" token, "-0" and "-1" etc. are now atomic
  1647   entities, so expressions involving "-" (unary or binary minus) need
  1648   to be spaced properly;
  1649 
  1650   - existing occurrences of "1" may need to be constraint "1::nat" or
  1651   even replaced by Suc 0; similar for old "2";
  1652 
  1653   - replace "#nnn" by "nnn", and "#-nnn" by "-nnn";
  1654 
  1655   - remove all special provisions on numerals in proofs;
  1656 
  1657 * HOL: simp rules nat_number expand numerals on nat to Suc/0
  1658 representation (depends on bin_arith_simps in the default context);
  1659 
  1660 * HOL: symbolic syntax for x^2 (numeral 2);
  1661 
  1662 * HOL: the class of all HOL types is now called "type" rather than
  1663 "term"; INCOMPATIBILITY, need to adapt references to this type class
  1664 in axclass/classes, instance/arities, and (usually rare) occurrences
  1665 in typings (of consts etc.); internally the class is called
  1666 "HOL.type", ML programs should refer to HOLogic.typeS;
  1667 
  1668 * HOL/record package improvements:
  1669   - new derived operations "fields" to build a partial record section,
  1670     "extend" to promote a fixed record to a record scheme, and
  1671     "truncate" for the reverse; cf. theorems "xxx.defs", which are *not*
  1672     declared as simp by default;
  1673   - shared operations ("more", "fields", etc.) now need to be always
  1674     qualified) --- potential INCOMPATIBILITY;
  1675   - removed "make_scheme" operations (use "make" with "extend") --
  1676     INCOMPATIBILITY;
  1677   - removed "more" class (simply use "term") -- INCOMPATIBILITY;
  1678   - provides cases/induct rules for use with corresponding Isar
  1679     methods (for concrete records, record schemes, concrete more
  1680     parts, and schematic more parts -- in that order);
  1681   - internal definitions directly based on a light-weight abstract
  1682     theory of product types over typedef rather than datatype;
  1683 
  1684 * HOL: generic code generator for generating executable ML code from
  1685 specifications; specific support for HOL constructs such as inductive
  1686 datatypes and sets, as well as recursive functions; can be invoked
  1687 via 'generate_code' theory section;
  1688 
  1689 * HOL: canonical cases/induct rules for n-tuples (n = 3..7);
  1690 
  1691 * HOL: consolidated and renamed several theories.  In particular:
  1692         Ord.thy has been absorbed into HOL.thy
  1693         String.thy has been absorbed into List.thy
  1694 
  1695 * HOL: concrete setsum syntax "\<Sum>i:A. b" == "setsum (%i. b) A"
  1696 (beware of argument permutation!);
  1697 
  1698 * HOL: linorder_less_split superseded by linorder_cases;
  1699 
  1700 * HOL/List: "nodups" renamed to "distinct";
  1701 
  1702 * HOL: added "The" definite description operator; move Hilbert's "Eps"
  1703 to peripheral theory "Hilbert_Choice"; some INCOMPATIBILITIES:
  1704   - Ex_def has changed, now need to use some_eq_ex
  1705 
  1706 * HOL: made split_all_tac safe; EXISTING PROOFS MAY FAIL OR LOOP, so
  1707 in this (rare) case use:
  1708 
  1709   delSWrapper "split_all_tac"
  1710   addSbefore ("unsafe_split_all_tac", unsafe_split_all_tac)
  1711 
  1712 * HOL: added safe wrapper "split_conv_tac" to claset; EXISTING PROOFS
  1713 MAY FAIL;
  1714 
  1715 * HOL: introduced f^n = f o ... o f; warning: due to the limits of
  1716 Isabelle's type classes, ^ on functions and relations has too general
  1717 a domain, namely ('a * 'b) set and 'a => 'b; this means that it may be
  1718 necessary to attach explicit type constraints;
  1719 
  1720 * HOL/Relation: the prefix name of the infix "O" has been changed from
  1721 "comp" to "rel_comp"; INCOMPATIBILITY: a few theorems have been
  1722 renamed accordingly (eg "compI" -> "rel_compI").
  1723 
  1724 * HOL: syntax translations now work properly with numerals and records
  1725 expressions;
  1726 
  1727 * HOL: bounded abstraction now uses syntax "%" / "\<lambda>" instead
  1728 of "lam" -- INCOMPATIBILITY;
  1729 
  1730 * HOL: got rid of some global declarations (potential INCOMPATIBILITY
  1731 for ML tools): const "()" renamed "Product_Type.Unity", type "unit"
  1732 renamed "Product_Type.unit";
  1733 
  1734 * HOL: renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
  1735 
  1736 * HOL: removed obsolete theorem "optionE" (use "option.exhaust", or
  1737 the "cases" method);
  1738 
  1739 * HOL/GroupTheory: group theory examples including Sylow's theorem (by
  1740 Florian Kammller);
  1741 
  1742 * HOL/IMP: updated and converted to new-style theory format; several
  1743 parts turned into readable document, with proper Isar proof texts and
  1744 some explanations (by Gerwin Klein);
  1745 
  1746 * HOL-Real: added Complex_Numbers (by Gertrud Bauer);
  1747 
  1748 * HOL-Hyperreal is now a logic image;
  1749 
  1750 
  1751 *** HOLCF ***
  1752 
  1753 * Isar: consts/constdefs supports mixfix syntax for continuous
  1754 operations;
  1755 
  1756 * Isar: domain package adapted to new-style theory format, e.g. see
  1757 HOLCF/ex/Dnat.thy;
  1758 
  1759 * theory Lift: proper use of rep_datatype lift instead of ML hacks --
  1760 potential INCOMPATIBILITY; now use plain induct_tac instead of former
  1761 lift.induct_tac, always use UU instead of Undef;
  1762 
  1763 * HOLCF/IMP: updated and converted to new-style theory;
  1764 
  1765 
  1766 *** ZF ***
  1767 
  1768 * Isar: proper integration of logic-specific tools and packages,
  1769 including theory commands '(co)inductive', '(co)datatype',
  1770 'rep_datatype', 'inductive_cases', as well as methods 'ind_cases',
  1771 'induct_tac', 'case_tac', and 'typecheck' (with attribute 'TC');
  1772 
  1773 * theory Main no longer includes AC; for the Axiom of Choice, base
  1774 your theory on Main_ZFC;
  1775 
  1776 * the integer library now covers quotients and remainders, with many
  1777 laws relating division to addition, multiplication, etc.;
  1778 
  1779 * ZF/UNITY: Chandy and Misra's UNITY is now available in ZF, giving a
  1780 typeless version of the formalism;
  1781 
  1782 * ZF/AC, Coind, IMP, Resid: updated and converted to new-style theory
  1783 format;
  1784 
  1785 * ZF/Induct: new directory for examples of inductive definitions,
  1786 including theory Multiset for multiset orderings; converted to
  1787 new-style theory format;
  1788 
  1789 * ZF: many new theorems about lists, ordinals, etc.;
  1790 
  1791 
  1792 *** General ***
  1793 
  1794 * Pure/kernel: meta-level proof terms (by Stefan Berghofer); reference
  1795 variable proof controls level of detail: 0 = no proofs (only oracle
  1796 dependencies), 1 = lemma dependencies, 2 = compact proof terms; see
  1797 also ref manual for further ML interfaces;
  1798 
  1799 * Pure/axclass: removed obsolete ML interface
  1800 goal_subclass/goal_arity;
  1801 
  1802 * Pure/syntax: new token syntax "num" for plain numerals (without "#"
  1803 of "xnum"); potential INCOMPATIBILITY, since -0, -1 etc. are now
  1804 separate tokens, so expressions involving minus need to be spaced
  1805 properly;
  1806 
  1807 * Pure/syntax: support non-oriented infixes, using keyword "infix"
  1808 rather than "infixl" or "infixr";
  1809 
  1810 * Pure/syntax: concrete syntax for dummy type variables admits genuine
  1811 sort constraint specifications in type inference; e.g. "x::_::foo"
  1812 ensures that the type of "x" is of sort "foo" (but not necessarily a
  1813 type variable);
  1814 
  1815 * Pure/syntax: print modes "type_brackets" and "no_type_brackets"
  1816 control output of nested => (types); the default behavior is
  1817 "type_brackets";
  1818 
  1819 * Pure/syntax: builtin parse translation for "_constify" turns valued
  1820 tokens into AST constants;
  1821 
  1822 * Pure/syntax: prefer later declarations of translations and print
  1823 translation functions; potential INCOMPATIBILITY: need to reverse
  1824 multiple declarations for same syntax element constant;
  1825 
  1826 * Pure/show_hyps reset by default (in accordance to existing Isar
  1827 practice);
  1828 
  1829 * Provers/classical: renamed addaltern to addafter, addSaltern to
  1830 addSafter;
  1831 
  1832 * Provers/clasimp: ``iff'' declarations now handle conditional rules
  1833 as well;
  1834 
  1835 * system: tested support for MacOS X; should be able to get Isabelle +
  1836 Proof General to work in a plain Terminal after installing Poly/ML
  1837 (e.g. from the Isabelle distribution area) and GNU bash alone
  1838 (e.g. from http://www.apple.com); full X11, XEmacs and X-Symbol
  1839 support requires further installations, e.g. from
  1840 http://fink.sourceforge.net/);
  1841 
  1842 * system: support Poly/ML 4.1.1 (able to manage larger heaps);
  1843 
  1844 * system: reduced base memory usage by Poly/ML (approx. 20 MB instead
  1845 of 40 MB), cf. ML_OPTIONS;
  1846 
  1847 * system: Proof General keywords specification is now part of the
  1848 Isabelle distribution (see etc/isar-keywords.el);
  1849 
  1850 * system: support for persistent Proof General sessions (refrain from
  1851 outdating all loaded theories on startup); user may create writable
  1852 logic images like this: ``isabelle -q HOL Test'';
  1853 
  1854 * system: smart selection of Isabelle process versus Isabelle
  1855 interface, accommodates case-insensitive file systems (e.g. HFS+); may
  1856 run both "isabelle" and "Isabelle" even if file names are badly
  1857 damaged (executable inspects the case of the first letter of its own
  1858 name); added separate "isabelle-process" and "isabelle-interface";
  1859 
  1860 * system: refrain from any attempt at filtering input streams; no
  1861 longer support ``8bit'' encoding of old isabelle font, instead proper
  1862 iso-latin characters may now be used; the related isatools
  1863 "symbolinput" and "nonascii" have disappeared as well;
  1864 
  1865 * system: removed old "xterm" interface (the print modes "xterm" and
  1866 "xterm_color" are still available for direct use in a suitable
  1867 terminal);
  1868 
  1869 
  1870 
  1871 New in Isabelle99-2 (February 2001)
  1872 -----------------------------------
  1873 
  1874 *** Overview of INCOMPATIBILITIES ***
  1875 
  1876 * HOL: please note that theories in the Library and elsewhere often use the
  1877 new-style (Isar) format; to refer to their theorems in an ML script you must
  1878 bind them to ML identifers by e.g.      val thm_name = thm "thm_name";
  1879 
  1880 * HOL: inductive package no longer splits induction rule aggressively,
  1881 but only as far as specified by the introductions given; the old
  1882 format may be recovered via ML function complete_split_rule or attribute
  1883 'split_rule (complete)';
  1884 
  1885 * HOL: induct renamed to lfp_induct, lfp_Tarski to lfp_unfold,
  1886 gfp_Tarski to gfp_unfold;
  1887 
  1888 * HOL: contrapos, contrapos2 renamed to contrapos_nn, contrapos_pp;
  1889 
  1890 * HOL: infix "dvd" now has priority 50 rather than 70 (because it is a
  1891 relation); infix "^^" has been renamed "``"; infix "``" has been
  1892 renamed "`"; "univalent" has been renamed "single_valued";
  1893 
  1894 * HOL/Real: "rinv" and "hrinv" replaced by overloaded "inverse"
  1895 operation;
  1896 
  1897 * HOLCF: infix "`" has been renamed "$"; the symbol syntax is \<cdot>;
  1898 
  1899 * Isar: 'obtain' no longer declares "that" fact as simp/intro;
  1900 
  1901 * Isar/HOL: method 'induct' now handles non-atomic goals; as a
  1902 consequence, it is no longer monotonic wrt. the local goal context
  1903 (which is now passed through the inductive cases);
  1904 
  1905 * Document preparation: renamed standard symbols \<ll> to \<lless> and
  1906 \<gg> to \<ggreater>;
  1907 
  1908 
  1909 *** Document preparation ***
  1910 
  1911 * \isabellestyle{NAME} selects version of Isabelle output (currently
  1912 available: are "it" for near math-mode best-style output, "sl" for
  1913 slanted text style, and "tt" for plain type-writer; if no
  1914 \isabellestyle command is given, output is according to slanted
  1915 type-writer);
  1916 
  1917 * support sub/super scripts (for single symbols only), input syntax is
  1918 like this: "A\<^sup>*" or "A\<^sup>\<star>";
  1919 
  1920 * some more standard symbols; see Appendix A of the system manual for
  1921 the complete list of symbols defined in isabellesym.sty;
  1922 
  1923 * improved isabelle style files; more abstract symbol implementation
  1924 (should now use \isamath{...} and \isatext{...} in custom symbol
  1925 definitions);
  1926 
  1927 * antiquotation @{goals} and @{subgoals} for output of *dynamic* goals
  1928 state; Note that presentation of goal states does not conform to
  1929 actual human-readable proof documents.  Please do not include goal
  1930 states into document output unless you really know what you are doing!
  1931 
  1932 * proper indentation of antiquoted output with proportional LaTeX
  1933 fonts;
  1934 
  1935 * no_document ML operator temporarily disables LaTeX document
  1936 generation;
  1937 
  1938 * isatool unsymbolize tunes sources for plain ASCII communication;
  1939 
  1940 
  1941 *** Isar ***
  1942 
  1943 * Pure: Isar now suffers initial goal statements to contain unbound
  1944 schematic variables (this does not conform to actual readable proof
  1945 documents, due to unpredictable outcome and non-compositional proof
  1946 checking); users who know what they are doing may use schematic goals
  1947 for Prolog-style synthesis of proven results;
  1948 
  1949 * Pure: assumption method (an implicit finishing) now handles actual
  1950 rules as well;
  1951 
  1952 * Pure: improved 'obtain' --- moved to Pure, insert "that" into
  1953 initial goal, declare "that" only as Pure intro (only for single
  1954 steps); the "that" rule assumption may now be involved in implicit
  1955 finishing, thus ".." becomes a feasible for trivial obtains;
  1956 
  1957 * Pure: default proof step now includes 'intro_classes'; thus trivial
  1958 instance proofs may be performed by "..";
  1959 
  1960 * Pure: ?thesis / ?this / "..." now work for pure meta-level
  1961 statements as well;
  1962 
  1963 * Pure: more robust selection of calculational rules;
  1964 
  1965 * Pure: the builtin notion of 'finished' goal now includes the ==-refl
  1966 rule (as well as the assumption rule);
  1967 
  1968 * Pure: 'thm_deps' command visualizes dependencies of theorems and
  1969 lemmas, using the graph browser tool;
  1970 
  1971 * Pure: predict failure of "show" in interactive mode;
  1972 
  1973 * Pure: 'thms_containing' now takes actual terms as arguments;
  1974 
  1975 * HOL: improved method 'induct' --- now handles non-atomic goals
  1976 (potential INCOMPATIBILITY); tuned error handling;
  1977 
  1978 * HOL: cases and induct rules now provide explicit hints about the
  1979 number of facts to be consumed (0 for "type" and 1 for "set" rules);
  1980 any remaining facts are inserted into the goal verbatim;
  1981 
  1982 * HOL: local contexts (aka cases) may now contain term bindings as
  1983 well; the 'cases' and 'induct' methods new provide a ?case binding for
  1984 the result to be shown in each case;
  1985 
  1986 * HOL: added 'recdef_tc' command;
  1987 
  1988 * isatool convert assists in eliminating legacy ML scripts;
  1989 
  1990 
  1991 *** HOL ***
  1992 
  1993 * HOL/Library: a collection of generic theories to be used together
  1994 with main HOL; the theory loader path already includes this directory
  1995 by default; the following existing theories have been moved here:
  1996 HOL/Induct/Multiset, HOL/Induct/Acc (as Accessible_Part), HOL/While
  1997 (as While_Combinator), HOL/Lex/Prefix (as List_Prefix);
  1998 
  1999 * HOL/Unix: "Some aspects of Unix file-system security", a typical
  2000 modelling and verification task performed in Isabelle/HOL +
  2001 Isabelle/Isar + Isabelle document preparation (by Markus Wenzel).
  2002 
  2003 * HOL/Algebra: special summation operator SUM no longer exists, it has
  2004 been replaced by setsum; infix 'assoc' now has priority 50 (like
  2005 'dvd'); axiom 'one_not_zero' has been moved from axclass 'ring' to
  2006 'domain', this makes the theory consistent with mathematical
  2007 literature;
  2008 
  2009 * HOL basics: added overloaded operations "inverse" and "divide"
  2010 (infix "/"), syntax for generic "abs" operation, generic summation
  2011 operator \<Sum>;
  2012 
  2013 * HOL/typedef: simplified package, provide more useful rules (see also
  2014 HOL/subset.thy);
  2015 
  2016 * HOL/datatype: induction rule for arbitrarily branching datatypes is
  2017 now expressed as a proper nested rule (old-style tactic scripts may
  2018 require atomize_strip_tac to cope with non-atomic premises);
  2019 
  2020 * HOL: renamed theory "Prod" to "Product_Type", renamed "split" rule
  2021 to "split_conv" (old name still available for compatibility);
  2022 
  2023 * HOL: improved concrete syntax for strings (e.g. allows translation
  2024 rules with string literals);
  2025 
  2026 * HOL-Real-Hyperreal: this extends HOL-Real with the hyperreals
  2027  and Fleuriot's mechanization of analysis, including the transcendental
  2028  functions for the reals;
  2029 
  2030 * HOL/Real, HOL/Hyperreal: improved arithmetic simplification;
  2031 
  2032 
  2033 *** CTT ***
  2034 
  2035 * CTT: x-symbol support for Pi, Sigma, -->, : (membership); note that
  2036 "lam" is displayed as TWO lambda-symbols
  2037 
  2038 * CTT: theory Main now available, containing everything (that is, Bool
  2039 and Arith);
  2040 
  2041 
  2042 *** General ***
  2043 
  2044 * Pure: the Simplifier has been implemented properly as a derived rule
  2045 outside of the actual kernel (at last!); the overall performance
  2046 penalty in practical applications is about 50%, while reliability of
  2047 the Isabelle inference kernel has been greatly improved;
  2048 
  2049 * print modes "brackets" and "no_brackets" control output of nested =>
  2050 (types) and ==> (props); the default behaviour is "brackets";
  2051 
  2052 * Provers: fast_tac (and friends) now handle actual object-logic rules
  2053 as assumptions as well;
  2054 
  2055 * system: support Poly/ML 4.0;
  2056 
  2057 * system: isatool install handles KDE version 1 or 2;
  2058 
  2059 
  2060 
  2061 New in Isabelle99-1 (October 2000)
  2062 ----------------------------------
  2063 
  2064 *** Overview of INCOMPATIBILITIES ***
  2065 
  2066 * HOL: simplification of natural numbers is much changed; to partly
  2067 recover the old behaviour (e.g. to prevent n+n rewriting to #2*n)
  2068 issue the following ML commands:
  2069 
  2070   Delsimprocs Nat_Numeral_Simprocs.cancel_numerals;
  2071   Delsimprocs [Nat_Numeral_Simprocs.combine_numerals];
  2072 
  2073 * HOL: simplification no longer dives into case-expressions; this is
  2074 controlled by "t.weak_case_cong" for each datatype t;
  2075 
  2076 * HOL: nat_less_induct renamed to less_induct;
  2077 
  2078 * HOL: systematic renaming of the SOME (Eps) rules, may use isatool
  2079 fixsome to patch .thy and .ML sources automatically;
  2080 
  2081   select_equality  -> some_equality
  2082   select_eq_Ex     -> some_eq_ex
  2083   selectI2EX       -> someI2_ex
  2084   selectI2         -> someI2
  2085   selectI          -> someI
  2086   select1_equality -> some1_equality
  2087   Eps_sym_eq       -> some_sym_eq_trivial
  2088   Eps_eq           -> some_eq_trivial
  2089 
  2090 * HOL: exhaust_tac on datatypes superceded by new generic case_tac;
  2091 
  2092 * HOL: removed obsolete theorem binding expand_if (refer to split_if
  2093 instead);
  2094 
  2095 * HOL: the recursion equations generated by 'recdef' are now called
  2096 f.simps instead of f.rules;
  2097 
  2098 * HOL: qed_spec_mp now also handles bounded ALL as well;
  2099 
  2100 * HOL: 0 is now overloaded, so the type constraint ":: nat" may
  2101 sometimes be needed;
  2102 
  2103 * HOL: the constant for "f``x" is now "image" rather than "op ``";
  2104 
  2105 * HOL: the constant for "f-``x" is now "vimage" rather than "op -``";
  2106 
  2107 * HOL: the disjoint sum is now "<+>" instead of "Plus"; the cartesian
  2108 product is now "<*>" instead of "Times"; the lexicographic product is
  2109 now "<*lex*>" instead of "**";
  2110 
  2111 * HOL: theory Sexp is now in HOL/Induct examples (it used to be part
  2112 of main HOL, but was unused); better use HOL's datatype package;
  2113 
  2114 * HOL: removed "symbols" syntax for constant "override" of theory Map;
  2115 the old syntax may be recovered as follows:
  2116 
  2117   syntax (symbols)
  2118     override  :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)"
  2119       (infixl "\\<oplus>" 100)
  2120 
  2121 * HOL/Real: "rabs" replaced by overloaded "abs" function;
  2122 
  2123 * HOL/ML: even fewer consts are declared as global (see theories Ord,
  2124 Lfp, Gfp, WF); this only affects ML packages that refer to const names
  2125 internally;
  2126 
  2127 * HOL and ZF: syntax for quotienting wrt an equivalence relation
  2128 changed from A/r to A//r;
  2129 
  2130 * ZF: new treatment of arithmetic (nat & int) may break some old
  2131 proofs;
  2132 
  2133 * Isar: renamed some attributes (RS -> THEN, simplify -> simplified,
  2134 rulify -> rule_format, elimify -> elim_format, ...);
  2135 
  2136 * Isar/Provers: intro/elim/dest attributes changed; renamed
  2137 intro/intro!/intro!! flags to intro!/intro/intro? (in most cases, one
  2138 should have to change intro!! to intro? only); replaced "delrule" by
  2139 "rule del";
  2140 
  2141 * Isar/HOL: renamed "intrs" to "intros" in inductive definitions;
  2142 
  2143 * Provers: strengthened force_tac by using new first_best_tac;
  2144 
  2145 * LaTeX document preparation: several changes of isabelle.sty (see
  2146 lib/texinputs);
  2147 
  2148 
  2149 *** Document preparation ***
  2150 
  2151 * formal comments (text blocks etc.) in new-style theories may now
  2152 contain antiquotations of thm/prop/term/typ/text to be presented
  2153 according to latex print mode; concrete syntax is like this:
  2154 @{term[show_types] "f(x) = a + x"};
  2155 
  2156 * isatool mkdir provides easy setup of Isabelle session directories,
  2157 including proper document sources;
  2158 
  2159 * generated LaTeX sources are now deleted after successful run
  2160 (isatool document -c); may retain a copy somewhere else via -D option
  2161 of isatool usedir;
  2162 
  2163 * isatool usedir -D now lets isatool latex -o sty update the Isabelle
  2164 style files, achieving self-contained LaTeX sources and simplifying
  2165 LaTeX debugging;
  2166 
  2167 * old-style theories now produce (crude) LaTeX output as well;
  2168 
  2169 * browser info session directories are now self-contained (may be put
  2170 on WWW server seperately); improved graphs of nested sessions; removed
  2171 graph for 'all sessions';
  2172 
  2173 * several improvements in isabelle style files; \isabellestyle{it}
  2174 produces fake math mode output; \isamarkupheader is now \section by
  2175 default; see lib/texinputs/isabelle.sty etc.;
  2176 
  2177 
  2178 *** Isar ***
  2179 
  2180 * Isar/Pure: local results and corresponding term bindings are now
  2181 subject to Hindley-Milner polymorphism (similar to ML); this
  2182 accommodates incremental type-inference very nicely;
  2183 
  2184 * Isar/Pure: new derived language element 'obtain' supports
  2185 generalized existence reasoning;
  2186 
  2187 * Isar/Pure: new calculational elements 'moreover' and 'ultimately'
  2188 support accumulation of results, without applying any rules yet;
  2189 useful to collect intermediate results without explicit name
  2190 references, and for use with transitivity rules with more than 2
  2191 premises;
  2192 
  2193 * Isar/Pure: scalable support for case-analysis type proofs: new
  2194 'case' language element refers to local contexts symbolically, as
  2195 produced by certain proof methods; internally, case names are attached
  2196 to theorems as "tags";
  2197 
  2198 * Isar/Pure: theory command 'hide' removes declarations from
  2199 class/type/const name spaces;
  2200 
  2201 * Isar/Pure: theory command 'defs' supports option "(overloaded)" to
  2202 indicate potential overloading;
  2203 
  2204 * Isar/Pure: changed syntax of local blocks from {{ }} to { };
  2205 
  2206 * Isar/Pure: syntax of sorts made 'inner', i.e. have to write
  2207 "{a,b,c}" instead of {a,b,c};
  2208 
  2209 * Isar/Pure now provides its own version of intro/elim/dest
  2210 attributes; useful for building new logics, but beware of confusion
  2211 with the version in Provers/classical;
  2212 
  2213 * Isar/Pure: the local context of (non-atomic) goals is provided via
  2214 case name 'antecedent';
  2215 
  2216 * Isar/Pure: removed obsolete 'transfer' attribute (transfer of thms
  2217 to the current context is now done automatically);
  2218 
  2219 * Isar/Pure: theory command 'method_setup' provides a simple interface
  2220 for definining proof methods in ML;
  2221 
  2222 * Isar/Provers: intro/elim/dest attributes changed; renamed
  2223 intro/intro!/intro!! flags to intro!/intro/intro? (INCOMPATIBILITY, in
  2224 most cases, one should have to change intro!! to intro? only);
  2225 replaced "delrule" by "rule del";
  2226 
  2227 * Isar/Provers: new 'hypsubst' method, plain 'subst' method and
  2228 'symmetric' attribute (the latter supercedes [RS sym]);
  2229 
  2230 * Isar/Provers: splitter support (via 'split' attribute and 'simp'
  2231 method modifier); 'simp' method: 'only:' modifier removes loopers as
  2232 well (including splits);
  2233 
  2234 * Isar/Provers: Simplifier and Classical methods now support all kind
  2235 of modifiers used in the past, including 'cong', 'iff', etc.
  2236 
  2237 * Isar/Provers: added 'fastsimp' and 'clarsimp' methods (combination
  2238 of Simplifier and Classical reasoner);
  2239 
  2240 * Isar/HOL: new proof method 'cases' and improved version of 'induct'
  2241 now support named cases; major packages (inductive, datatype, primrec,
  2242 recdef) support case names and properly name parameters;
  2243 
  2244 * Isar/HOL: new transitivity rules for substitution in inequalities --
  2245 monotonicity conditions are extracted to be proven at end of
  2246 calculations;
  2247 
  2248 * Isar/HOL: removed 'case_split' thm binding, should use 'cases' proof
  2249 method anyway;
  2250 
  2251 * Isar/HOL: removed old expand_if = split_if; theorems if_splits =
  2252 split_if split_if_asm; datatype package provides theorems foo.splits =
  2253 foo.split foo.split_asm for each datatype;
  2254 
  2255 * Isar/HOL: tuned inductive package, rename "intrs" to "intros"
  2256 (potential INCOMPATIBILITY), emulation of mk_cases feature for proof
  2257 scripts: new 'inductive_cases' command and 'ind_cases' method; (Note:
  2258 use "(cases (simplified))" method in proper proof texts);
  2259 
  2260 * Isar/HOL: added global 'arith_split' attribute for 'arith' method;
  2261 
  2262 * Isar: names of theorems etc. may be natural numbers as well;
  2263 
  2264 * Isar: 'pr' command: optional arguments for goals_limit and
  2265 ProofContext.prems_limit; no longer prints theory contexts, but only
  2266 proof states;
  2267 
  2268 * Isar: diagnostic commands 'pr', 'thm', 'prop', 'term', 'typ' admit
  2269 additional print modes to be specified; e.g. "pr(latex)" will print
  2270 proof state according to the Isabelle LaTeX style;
  2271 
  2272 * Isar: improved support for emulating tactic scripts, including proof
  2273 methods 'rule_tac' etc., 'cut_tac', 'thin_tac', 'subgoal_tac',
  2274 'rename_tac', 'rotate_tac', 'tactic', and 'case_tac' / 'induct_tac'
  2275 (for HOL datatypes);
  2276 
  2277 * Isar: simplified (more robust) goal selection of proof methods: 1st
  2278 goal, all goals, or explicit goal specifier (tactic emulation); thus
  2279 'proof method scripts' have to be in depth-first order;
  2280 
  2281 * Isar: tuned 'let' syntax: replaced 'as' keyword by 'and';
  2282 
  2283 * Isar: removed 'help' command, which hasn't been too helpful anyway;
  2284 should instead use individual commands for printing items
  2285 (print_commands, print_methods etc.);
  2286 
  2287 * Isar: added 'nothing' --- the empty list of theorems;
  2288 
  2289 
  2290 *** HOL ***
  2291 
  2292 * HOL/MicroJava: formalization of a fragment of Java, together with a
  2293 corresponding virtual machine and a specification of its bytecode
  2294 verifier and a lightweight bytecode verifier, including proofs of
  2295 type-safety; by Gerwin Klein, Tobias Nipkow, David von Oheimb, and
  2296 Cornelia Pusch (see also the homepage of project Bali at
  2297 http://isabelle.in.tum.de/Bali/);
  2298 
  2299 * HOL/Algebra: new theory of rings and univariate polynomials, by
  2300 Clemens Ballarin;
  2301 
  2302 * HOL/NumberTheory: fundamental Theorem of Arithmetic, Chinese
  2303 Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, by Thomas M
  2304 Rasmussen;
  2305 
  2306 * HOL/Lattice: fundamental concepts of lattice theory and order
  2307 structures, including duals, properties of bounds versus algebraic
  2308 laws, lattice operations versus set-theoretic ones, the Knaster-Tarski
  2309 Theorem for complete lattices etc.; may also serve as a demonstration
  2310 for abstract algebraic reasoning using axiomatic type classes, and
  2311 mathematics-style proof in Isabelle/Isar; by Markus Wenzel;
  2312 
  2313 * HOL/Prolog: a (bare-bones) implementation of Lambda-Prolog, by David
  2314 von Oheimb;
  2315 
  2316 * HOL/IMPP: extension of IMP with local variables and mutually
  2317 recursive procedures, by David von Oheimb;
  2318 
  2319 * HOL/Lambda: converted into new-style theory and document;
  2320 
  2321 * HOL/ex/Multiquote: example of multiple nested quotations and
  2322 anti-quotations -- basically a generalized version of de-Bruijn
  2323 representation; very useful in avoiding lifting of operations;
  2324 
  2325 * HOL/record: added general record equality rule to simpset; fixed
  2326 select-update simplification procedure to handle extended records as
  2327 well; admit "r" as field name;
  2328 
  2329 * HOL: 0 is now overloaded over the new sort "zero", allowing its use with
  2330 other numeric types and also as the identity of groups, rings, etc.;
  2331 
  2332 * HOL: new axclass plus_ac0 for addition with the AC-laws and 0 as identity.
  2333 Types nat and int belong to this axclass;
  2334 
  2335 * HOL: greatly improved simplification involving numerals of type nat, int, real:
  2336    (i + #8 + j) = Suc k simplifies to  #7 + (i + j) = k
  2337    i*j + k + j*#3*i     simplifies to  #4*(i*j) + k
  2338   two terms #m*u and #n*u are replaced by #(m+n)*u
  2339     (where #m, #n and u can implicitly be 1; this is simproc combine_numerals)
  2340   and the term/formula #m*u+x ~~ #n*u+y simplifies simplifies to #(m-n)+x ~~ y
  2341     or x ~~ #(n-m)+y, where ~~ is one of = < <= or - (simproc cancel_numerals);
  2342 
  2343 * HOL: meson_tac is available (previously in ex/meson.ML); it is a
  2344 powerful prover for predicate logic but knows nothing of clasets; see
  2345 ex/mesontest.ML and ex/mesontest2.ML for example applications;
  2346 
  2347 * HOL: new version of "case_tac" subsumes both boolean case split and
  2348 "exhaust_tac" on datatypes; INCOMPATIBILITY: exhaust_tac no longer
  2349 exists, may define val exhaust_tac = case_tac for ad-hoc portability;
  2350 
  2351 * HOL: simplification no longer dives into case-expressions: only the
  2352 selector expression is simplified, but not the remaining arms; to
  2353 enable full simplification of case-expressions for datatype t, you may
  2354 remove t.weak_case_cong from the simpset, either globally (Delcongs
  2355 [thm"t.weak_case_cong"];) or locally (delcongs [...]).
  2356 
  2357 * HOL/recdef: the recursion equations generated by 'recdef' for
  2358 function 'f' are now called f.simps instead of f.rules; if all
  2359 termination conditions are proved automatically, these simplification
  2360 rules are added to the simpset, as in primrec; rules may be named
  2361 individually as well, resulting in a separate list of theorems for
  2362 each equation;
  2363 
  2364 * HOL/While is a new theory that provides a while-combinator. It
  2365 permits the definition of tail-recursive functions without the
  2366 provision of a termination measure. The latter is necessary once the
  2367 invariant proof rule for while is applied.
  2368 
  2369 * HOL: new (overloaded) notation for the set of elements below/above
  2370 some element: {..u}, {..u(}, {l..}, {)l..}. See theory SetInterval.
  2371 
  2372 * HOL: theorems impI, allI, ballI bound as "strip";
  2373 
  2374 * HOL: new tactic induct_thm_tac: thm -> string -> int -> tactic
  2375 induct_tac th "x1 ... xn" expects th to have a conclusion of the form
  2376 P v1 ... vn and abbreviates res_inst_tac [("v1","x1"),...,("vn","xn")] th;
  2377 
  2378 * HOL/Real: "rabs" replaced by overloaded "abs" function;
  2379 
  2380 * HOL: theory Sexp now in HOL/Induct examples (it used to be part of
  2381 main HOL, but was unused);
  2382 
  2383 * HOL: fewer consts declared as global (e.g. have to refer to
  2384 "Lfp.lfp" instead of "lfp" internally; affects ML packages only);
  2385 
  2386 * HOL: tuned AST representation of nested pairs, avoiding bogus output
  2387 in case of overlap with user translations (e.g. judgements over
  2388 tuples); (note that the underlying logical represenation is still
  2389 bogus);
  2390 
  2391 
  2392 *** ZF ***
  2393 
  2394 * ZF: simplification automatically cancels common terms in arithmetic
  2395 expressions over nat and int;
  2396 
  2397 * ZF: new treatment of nat to minimize type-checking: all operators
  2398 coerce their operands to a natural number using the function natify,
  2399 making the algebraic laws unconditional;
  2400 
  2401 * ZF: as above, for int: operators coerce their operands to an integer
  2402 using the function intify;
  2403 
  2404 * ZF: the integer library now contains many of the usual laws for the
  2405 orderings, including $<=, and monotonicity laws for $+ and $*;
  2406 
  2407 * ZF: new example ZF/ex/NatSum to demonstrate integer arithmetic
  2408 simplification;
  2409 
  2410 * FOL and ZF: AddIffs now available, giving theorems of the form P<->Q
  2411 to the simplifier and classical reasoner simultaneously;
  2412 
  2413 
  2414 *** General ***
  2415 
  2416 * Provers: blast_tac now handles actual object-logic rules as
  2417 assumptions; note that auto_tac uses blast_tac internally as well;
  2418 
  2419 * Provers: new functions rulify/rulify_no_asm: thm -> thm for turning
  2420 outer -->/All/Ball into ==>/!!; qed_spec_mp now uses rulify_no_asm;
  2421 
  2422 * Provers: delrules now handles destruct rules as well (no longer need
  2423 explicit make_elim);
  2424 
  2425 * Provers: Blast_tac now warns of and ignores "weak elimination rules" e.g.
  2426   [| inj ?f;          ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
  2427 use instead the strong form,
  2428   [| inj ?f; ~ ?W ==> ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
  2429 in HOL, FOL and ZF the function cla_make_elim will create such rules
  2430 from destruct-rules;
  2431 
  2432 * Provers: Simplifier.easy_setup provides a fast path to basic
  2433 Simplifier setup for new object-logics;
  2434 
  2435 * Pure: AST translation rules no longer require constant head on LHS;
  2436 
  2437 * Pure: improved name spaces: ambiguous output is qualified; support
  2438 for hiding of names;
  2439 
  2440 * system: smart setup of canonical ML_HOME, ISABELLE_INTERFACE, and
  2441 XSYMBOL_HOME; no longer need to do manual configuration in most
  2442 situations;
  2443 
  2444 * system: compression of ML heaps images may now be controlled via -c
  2445 option of isabelle and isatool usedir (currently only observed by
  2446 Poly/ML);
  2447 
  2448 * system: isatool installfonts may handle X-Symbol fonts as well (very
  2449 useful for remote X11);
  2450 
  2451 * system: provide TAGS file for Isabelle sources;
  2452 
  2453 * ML: infix 'OF' is a version of 'MRS' with more appropriate argument
  2454 order;
  2455 
  2456 * ML: renamed flags Syntax.trace_norm_ast to Syntax.trace_ast; global
  2457 timing flag supersedes proof_timing and Toplevel.trace;
  2458 
  2459 * ML: new combinators |>> and |>>> for incremental transformations
  2460 with secondary results (e.g. certain theory extensions):
  2461 
  2462 * ML: PureThy.add_defs gets additional argument to indicate potential
  2463 overloading (usually false);
  2464 
  2465 * ML: PureThy.add_thms/add_axioms/add_defs now return theorems as
  2466 results;
  2467 
  2468 
  2469 
  2470 New in Isabelle99 (October 1999)
  2471 --------------------------------
  2472 
  2473 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  2474 
  2475 * HOL: The THEN and ELSE parts of conditional expressions (if P then x else y)
  2476 are no longer simplified.  (This allows the simplifier to unfold recursive
  2477 functional programs.)  To restore the old behaviour, declare
  2478 
  2479     Delcongs [if_weak_cong];
  2480 
  2481 * HOL: Removed the obsolete syntax "Compl A"; use -A for set
  2482 complement;
  2483 
  2484 * HOL: the predicate "inj" is now defined by translation to "inj_on";
  2485 
  2486 * HOL/datatype: mutual_induct_tac no longer exists --
  2487   use induct_tac "x_1 ... x_n" instead of mutual_induct_tac ["x_1", ..., "x_n"]
  2488 
  2489 * HOL/typedef: fixed type inference for representing set; type
  2490 arguments now have to occur explicitly on the rhs as type constraints;
  2491 
  2492 * ZF: The con_defs part of an inductive definition may no longer refer
  2493 to constants declared in the same theory;
  2494 
  2495 * HOL, ZF: the function mk_cases, generated by the inductive
  2496 definition package, has lost an argument.  To simplify its result, it
  2497 uses the default simpset instead of a supplied list of theorems.
  2498 
  2499 * HOL/List: the constructors of type list are now Nil and Cons;
  2500 
  2501 * Simplifier: the type of the infix ML functions
  2502         setSSolver addSSolver setSolver addSolver
  2503 is now  simpset * solver -> simpset  where `solver' is a new abstract type
  2504 for packaging solvers. A solver is created via
  2505         mk_solver: string -> (thm list -> int -> tactic) -> solver
  2506 where the string argument is only a comment.
  2507 
  2508 
  2509 *** Proof tools ***
  2510 
  2511 * Provers/Arith/fast_lin_arith.ML contains a functor for creating a
  2512 decision procedure for linear arithmetic. Currently it is used for
  2513 types `nat', `int', and `real' in HOL (see below); it can, should and
  2514 will be instantiated for other types and logics as well.
  2515 
  2516 * The simplifier now accepts rewrite rules with flexible heads, eg
  2517      hom ?f ==> ?f(?x+?y) = ?f ?x + ?f ?y
  2518   They are applied like any rule with a non-pattern lhs, i.e. by first-order
  2519   matching.
  2520 
  2521 
  2522 *** General ***
  2523 
  2524 * New Isabelle/Isar subsystem provides an alternative to traditional
  2525 tactical theorem proving; together with the ProofGeneral/isar user
  2526 interface it offers an interactive environment for developing human
  2527 readable proof documents (Isar == Intelligible semi-automated
  2528 reasoning); for further information see isatool doc isar-ref,
  2529 src/HOL/Isar_examples and http://isabelle.in.tum.de/Isar/
  2530 
  2531 * improved and simplified presentation of theories: better HTML markup
  2532 (including colors), graph views in several sizes; isatool usedir now
  2533 provides a proper interface for user theories (via -P option); actual
  2534 document preparation based on (PDF)LaTeX is available as well (for
  2535 new-style theories only); see isatool doc system for more information;
  2536 
  2537 * native support for Proof General, both for classic Isabelle and
  2538 Isabelle/Isar;
  2539 
  2540 * ML function thm_deps visualizes dependencies of theorems and lemmas,
  2541 using the graph browser tool;
  2542 
  2543 * Isabelle manuals now also available as PDF;
  2544 
  2545 * theory loader rewritten from scratch (may not be fully
  2546 bug-compatible); old loadpath variable has been replaced by show_path,
  2547 add_path, del_path, reset_path functions; new operations such as
  2548 update_thy, touch_thy, remove_thy, use/update_thy_only (see also
  2549 isatool doc ref);
  2550 
  2551 * improved isatool install: option -k creates KDE application icon,
  2552 option -p DIR installs standalone binaries;
  2553 
  2554 * added ML_PLATFORM setting (useful for cross-platform installations);
  2555 more robust handling of platform specific ML images for SML/NJ;
  2556 
  2557 * the settings environment is now statically scoped, i.e. it is never
  2558 created again in sub-processes invoked from isabelle, isatool, or
  2559 Isabelle;
  2560 
  2561 * path element specification '~~' refers to '$ISABELLE_HOME';
  2562 
  2563 * in locales, the "assumes" and "defines" parts may be omitted if
  2564 empty;
  2565 
  2566 * new print_mode "xsymbols" for extended symbol support (e.g. genuine
  2567 long arrows);
  2568 
  2569 * new print_mode "HTML";
  2570 
  2571 * new flag show_tags controls display of tags of theorems (which are
  2572 basically just comments that may be attached by some tools);
  2573 
  2574 * Isamode 2.6 requires patch to accomodate change of Isabelle font
  2575 mode and goal output format:
  2576 
  2577 diff -r Isamode-2.6/elisp/isa-load.el Isamode/elisp/isa-load.el
  2578 244c244
  2579 <       (list (isa-getenv "ISABELLE") "-msymbols" logic-name)
  2580 ---
  2581 >       (list (isa-getenv "ISABELLE") "-misabelle_font" "-msymbols" logic-name)
  2582 diff -r Isabelle-2.6/elisp/isa-proofstate.el Isamode/elisp/isa-proofstate.el
  2583 181c181
  2584 < (defconst proofstate-proofstart-regexp "^Level [0-9]+$"
  2585 ---
  2586 > (defconst proofstate-proofstart-regexp "^Level [0-9]+"
  2587 
  2588 * function bind_thms stores lists of theorems (cf. bind_thm);
  2589 
  2590 * new shorthand tactics ftac, eatac, datac, fatac;
  2591 
  2592 * qed (and friends) now accept "" as result name; in that case the
  2593 theorem is not stored, but proper checks and presentation of the
  2594 result still apply;
  2595 
  2596 * theorem database now also indexes constants "Trueprop", "all",
  2597 "==>", "=="; thus thms_containing, findI etc. may retrieve more rules;
  2598 
  2599 
  2600 *** HOL ***
  2601 
  2602 ** HOL arithmetic **
  2603 
  2604 * There are now decision procedures for linear arithmetic over nat and
  2605 int:
  2606 
  2607 1. arith_tac copes with arbitrary formulae involving `=', `<', `<=',
  2608 `+', `-', `Suc', `min', `max' and numerical constants; other subterms
  2609 are treated as atomic; subformulae not involving type `nat' or `int'
  2610 are ignored; quantified subformulae are ignored unless they are
  2611 positive universal or negative existential. The tactic has to be
  2612 invoked by hand and can be a little bit slow. In particular, the
  2613 running time is exponential in the number of occurrences of `min' and
  2614 `max', and `-' on `nat'.
  2615 
  2616 2. fast_arith_tac is a cut-down version of arith_tac: it only takes
  2617 (negated) (in)equalities among the premises and the conclusion into
  2618 account (i.e. no compound formulae) and does not know about `min' and
  2619 `max', and `-' on `nat'. It is fast and is used automatically by the
  2620 simplifier.
  2621 
  2622 NB: At the moment, these decision procedures do not cope with mixed
  2623 nat/int formulae where the two parts interact, such as `m < n ==>
  2624 int(m) < int(n)'.
  2625 
  2626 * HOL/Numeral provides a generic theory of numerals (encoded
  2627 efficiently as bit strings); setup for types nat/int/real is in place;
  2628 INCOMPATIBILITY: since numeral syntax is now polymorphic, rather than
  2629 int, existing theories and proof scripts may require a few additional
  2630 type constraints;
  2631 
  2632 * integer division and remainder can now be performed on constant
  2633 arguments;
  2634 
  2635 * many properties of integer multiplication, division and remainder
  2636 are now available;
  2637 
  2638 * An interface to the Stanford Validity Checker (SVC) is available through the
  2639 tactic svc_tac.  Propositional tautologies and theorems of linear arithmetic
  2640 are proved automatically.  SVC must be installed separately, and its results
  2641 must be TAKEN ON TRUST (Isabelle does not check the proofs, but tags any
  2642 invocation of the underlying oracle).  For SVC see
  2643   http://verify.stanford.edu/SVC
  2644 
  2645 * IsaMakefile: the HOL-Real target now builds an actual image;
  2646 
  2647 
  2648 ** HOL misc **
  2649 
  2650 * HOL/Real/HahnBanach: the Hahn-Banach theorem for real vector spaces
  2651 (in Isabelle/Isar) -- by Gertrud Bauer;
  2652 
  2653 * HOL/BCV: generic model of bytecode verification, i.e. data-flow
  2654 analysis for assembly languages with subtypes;
  2655 
  2656 * HOL/TLA (Lamport's Temporal Logic of Actions): major reorganization
  2657 -- avoids syntactic ambiguities and treats state, transition, and
  2658 temporal levels more uniformly; introduces INCOMPATIBILITIES due to
  2659 changed syntax and (many) tactics;
  2660 
  2661 * HOL/inductive: Now also handles more general introduction rules such
  2662   as "ALL y. (y, x) : r --> y : acc r ==> x : acc r"; monotonicity
  2663   theorems are now maintained within the theory (maintained via the
  2664   "mono" attribute);
  2665 
  2666 * HOL/datatype: Now also handles arbitrarily branching datatypes
  2667   (using function types) such as
  2668 
  2669   datatype 'a tree = Atom 'a | Branch "nat => 'a tree"
  2670 
  2671 * HOL/record: record_simproc (part of the default simpset) takes care
  2672 of selectors applied to updated records; record_split_tac is no longer
  2673 part of the default claset; update_defs may now be removed from the
  2674 simpset in many cases; COMPATIBILITY: old behavior achieved by
  2675 
  2676   claset_ref () := claset() addSWrapper record_split_wrapper;
  2677   Delsimprocs [record_simproc]
  2678 
  2679 * HOL/typedef: fixed type inference for representing set; type
  2680 arguments now have to occur explicitly on the rhs as type constraints;
  2681 
  2682 * HOL/recdef (TFL): 'congs' syntax now expects comma separated list of theorem
  2683 names rather than an ML expression;
  2684 
  2685 * HOL/defer_recdef (TFL): like recdef but the well-founded relation can be
  2686 supplied later.  Program schemes can be defined, such as
  2687     "While B C s = (if B s then While B C (C s) else s)"
  2688 where the well-founded relation can be chosen after B and C have been given.
  2689 
  2690 * HOL/List: the constructors of type list are now Nil and Cons;
  2691 INCOMPATIBILITY: while [] and infix # syntax is still there, of
  2692 course, ML tools referring to List.list.op # etc. have to be adapted;
  2693 
  2694 * HOL_quantifiers flag superseded by "HOL" print mode, which is
  2695 disabled by default; run isabelle with option -m HOL to get back to
  2696 the original Gordon/HOL-style output;
  2697 
  2698 * HOL/Ord.thy: new bounded quantifier syntax (input only): ALL x<y. P,
  2699 ALL x<=y. P, EX x<y. P, EX x<=y. P;
  2700 
  2701 * HOL basic syntax simplified (more orthogonal): all variants of
  2702 All/Ex now support plain / symbolic / HOL notation; plain syntax for
  2703 Eps operator is provided as well: "SOME x. P[x]";
  2704 
  2705 * HOL/Sum.thy: sum_case has been moved to HOL/Datatype;
  2706 
  2707 * HOL/Univ.thy: infix syntax <*>, <+>, <**>, <+> eliminated and made
  2708 thus available for user theories;
  2709 
  2710 * HOLCF/IOA/Sequents: renamed 'Cons' to 'Consq' to avoid clash with
  2711 HOL/List; hardly an INCOMPATIBILITY since '>>' syntax is used all the
  2712 time;
  2713 
  2714 * HOL: new tactic smp_tac: int -> int -> tactic, which applies spec
  2715 several times and then mp;
  2716 
  2717 
  2718 *** LK ***
  2719 
  2720 * the notation <<...>> is now available as a notation for sequences of
  2721 formulas;
  2722 
  2723 * the simplifier is now installed
  2724 
  2725 * the axiom system has been generalized (thanks to Soren Heilmann)
  2726 
  2727 * the classical reasoner now has a default rule database
  2728 
  2729 
  2730 *** ZF ***
  2731 
  2732 * new primrec section allows primitive recursive functions to be given
  2733 directly (as in HOL) over datatypes and the natural numbers;
  2734 
  2735 * new tactics induct_tac and exhaust_tac for induction (or case
  2736 analysis) over datatypes and the natural numbers;
  2737 
  2738 * the datatype declaration of type T now defines the recursor T_rec;
  2739 
  2740 * simplification automatically does freeness reasoning for datatype
  2741 constructors;
  2742 
  2743 * automatic type-inference, with AddTCs command to insert new
  2744 type-checking rules;
  2745 
  2746 * datatype introduction rules are now added as Safe Introduction rules
  2747 to the claset;
  2748 
  2749 * the syntax "if P then x else y" is now available in addition to
  2750 if(P,x,y);
  2751 
  2752 
  2753 *** Internal programming interfaces ***
  2754 
  2755 * tuned simplifier trace output; new flag debug_simp;
  2756 
  2757 * structures Vartab / Termtab (instances of TableFun) offer efficient
  2758 tables indexed by indexname_ord / term_ord (compatible with aconv);
  2759 
  2760 * AxClass.axclass_tac lost the theory argument;
  2761 
  2762 * tuned current_goals_markers semantics: begin / end goal avoids
  2763 printing empty lines;
  2764 
  2765 * removed prs and prs_fn hook, which was broken because it did not
  2766 include \n in its semantics, forcing writeln to add one
  2767 uncoditionally; replaced prs_fn by writeln_fn; consider std_output:
  2768 string -> unit if you really want to output text without newline;
  2769 
  2770 * Symbol.output subject to print mode; INCOMPATIBILITY: defaults to
  2771 plain output, interface builders may have to enable 'isabelle_font'
  2772 mode to get Isabelle font glyphs as before;
  2773 
  2774 * refined token_translation interface; INCOMPATIBILITY: output length
  2775 now of type real instead of int;
  2776 
  2777 * theory loader actions may be traced via new ThyInfo.add_hook
  2778 interface (see src/Pure/Thy/thy_info.ML); example application: keep
  2779 your own database of information attached to *whole* theories -- as
  2780 opposed to intra-theory data slots offered via TheoryDataFun;
  2781 
  2782 * proper handling of dangling sort hypotheses (at last!);
  2783 Thm.strip_shyps and Drule.strip_shyps_warning take care of removing
  2784 extra sort hypotheses that can be witnessed from the type signature;
  2785 the force_strip_shyps flag is gone, any remaining shyps are simply
  2786 left in the theorem (with a warning issued by strip_shyps_warning);
  2787 
  2788 
  2789 
  2790 New in Isabelle98-1 (October 1998)
  2791 ----------------------------------
  2792 
  2793 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  2794 
  2795 * several changes of automated proof tools;
  2796 
  2797 * HOL: major changes to the inductive and datatype packages, including
  2798 some minor incompatibilities of theory syntax;
  2799 
  2800 * HOL: renamed r^-1 to 'converse' from 'inverse'; 'inj_onto' is now
  2801 called `inj_on';
  2802 
  2803 * HOL: removed duplicate thms in Arith:
  2804   less_imp_add_less  should be replaced by  trans_less_add1
  2805   le_imp_add_le      should be replaced by  trans_le_add1
  2806 
  2807 * HOL: unary minus is now overloaded (new type constraints may be
  2808 required);
  2809 
  2810 * HOL and ZF: unary minus for integers is now #- instead of #~.  In
  2811 ZF, expressions such as n#-1 must be changed to n#- 1, since #-1 is
  2812 now taken as an integer constant.
  2813 
  2814 * Pure: ML function 'theory_of' renamed to 'theory';
  2815 
  2816 
  2817 *** Proof tools ***
  2818 
  2819 * Simplifier:
  2820   1. Asm_full_simp_tac is now more aggressive.
  2821      1. It will sometimes reorient premises if that increases their power to
  2822         simplify.
  2823      2. It does no longer proceed strictly from left to right but may also
  2824         rotate premises to achieve further simplification.
  2825      For compatibility reasons there is now Asm_lr_simp_tac which is like the
  2826      old Asm_full_simp_tac in that it does not rotate premises.
  2827   2. The simplifier now knows a little bit about nat-arithmetic.
  2828 
  2829 * Classical reasoner: wrapper mechanism for the classical reasoner now
  2830 allows for selected deletion of wrappers, by introduction of names for
  2831 wrapper functionals.  This implies that addbefore, addSbefore,
  2832 addaltern, and addSaltern now take a pair (name, tactic) as argument,
  2833 and that adding two tactics with the same name overwrites the first
  2834 one (emitting a warning).
  2835   type wrapper = (int -> tactic) -> (int -> tactic)
  2836   setWrapper, setSWrapper, compWrapper and compSWrapper are replaced by
  2837   addWrapper, addSWrapper: claset * (string * wrapper) -> claset
  2838   delWrapper, delSWrapper: claset *  string            -> claset
  2839   getWrapper is renamed to appWrappers, getSWrapper to appSWrappers;
  2840 
  2841 * Classical reasoner: addbefore/addSbefore now have APPEND/ORELSE
  2842 semantics; addbefore now affects only the unsafe part of step_tac
  2843 etc.; this affects addss/auto_tac/force_tac, so EXISTING PROOFS MAY
  2844 FAIL, but proofs should be fixable easily, e.g. by replacing Auto_tac
  2845 by Force_tac;
  2846 
  2847 * Classical reasoner: setwrapper to setWrapper and compwrapper to
  2848 compWrapper; added safe wrapper (and access functions for it);
  2849 
  2850 * HOL/split_all_tac is now much faster and fails if there is nothing
  2851 to split.  Some EXISTING PROOFS MAY REQUIRE ADAPTION because the order
  2852 and the names of the automatically generated variables have changed.
  2853 split_all_tac has moved within claset() from unsafe wrappers to safe
  2854 wrappers, which means that !!-bound variables are split much more
  2855 aggressively, and safe_tac and clarify_tac now split such variables.
  2856 If this splitting is not appropriate, use delSWrapper "split_all_tac".
  2857 Note: the same holds for record_split_tac, which does the job of
  2858 split_all_tac for record fields.
  2859 
  2860 * HOL/Simplifier: Rewrite rules for case distinctions can now be added
  2861 permanently to the default simpset using Addsplits just like
  2862 Addsimps. They can be removed via Delsplits just like
  2863 Delsimps. Lower-case versions are also available.
  2864 
  2865 * HOL/Simplifier: The rule split_if is now part of the default
  2866 simpset. This means that the simplifier will eliminate all occurrences
  2867 of if-then-else in the conclusion of a goal. To prevent this, you can
  2868 either remove split_if completely from the default simpset by
  2869 `Delsplits [split_if]' or remove it in a specific call of the
  2870 simplifier using `... delsplits [split_if]'.  You can also add/delete
  2871 other case splitting rules to/from the default simpset: every datatype
  2872 generates suitable rules `split_t_case' and `split_t_case_asm' (where
  2873 t is the name of the datatype).
  2874 
  2875 * Classical reasoner / Simplifier combination: new force_tac (and
  2876 derivatives Force_tac, force) combines rewriting and classical
  2877 reasoning (and whatever other tools) similarly to auto_tac, but is
  2878 aimed to solve the given subgoal completely.
  2879 
  2880 
  2881 *** General ***
  2882 
  2883 * new top-level commands `Goal' and `Goalw' that improve upon `goal'
  2884 and `goalw': the theory is no longer needed as an explicit argument -
  2885 the current theory context is used; assumptions are no longer returned
  2886 at the ML-level unless one of them starts with ==> or !!; it is
  2887 recommended to convert to these new commands using isatool fixgoal
  2888 (backup your sources first!);
  2889 
  2890 * new top-level commands 'thm' and 'thms' for retrieving theorems from
  2891 the current theory context, and 'theory' to lookup stored theories;
  2892 
  2893 * new theory section 'locale' for declaring constants, assumptions and
  2894 definitions that have local scope;
  2895 
  2896 * new theory section 'nonterminals' for purely syntactic types;
  2897 
  2898 * new theory section 'setup' for generic ML setup functions
  2899 (e.g. package initialization);
  2900 
  2901 * the distribution now includes Isabelle icons: see
  2902 lib/logo/isabelle-{small,tiny}.xpm;
  2903 
  2904 * isatool install - install binaries with absolute references to
  2905 ISABELLE_HOME/bin;
  2906 
  2907 * isatool logo -- create instances of the Isabelle logo (as EPS);
  2908 
  2909 * print mode 'emacs' reserved for Isamode;
  2910 
  2911 * support multiple print (ast) translations per constant name;
  2912 
  2913 * theorems involving oracles are now printed with a suffixed [!];
  2914 
  2915 
  2916 *** HOL ***
  2917 
  2918 * there is now a tutorial on Isabelle/HOL (do 'isatool doc tutorial');
  2919 
  2920 * HOL/inductive package reorganized and improved: now supports mutual
  2921 definitions such as
  2922 
  2923   inductive EVEN ODD
  2924     intrs
  2925       null "0 : EVEN"
  2926       oddI "n : EVEN ==> Suc n : ODD"
  2927       evenI "n : ODD ==> Suc n : EVEN"
  2928 
  2929 new theorem list "elims" contains an elimination rule for each of the
  2930 recursive sets; inductive definitions now handle disjunctive premises
  2931 correctly (also ZF);
  2932 
  2933 INCOMPATIBILITIES: requires Inductive as an ancestor; component
  2934 "mutual_induct" no longer exists - the induction rule is always
  2935 contained in "induct";
  2936 
  2937 
  2938 * HOL/datatype package re-implemented and greatly improved: now
  2939 supports mutually recursive datatypes such as
  2940 
  2941   datatype
  2942     'a aexp = IF_THEN_ELSE ('a bexp) ('a aexp) ('a aexp)
  2943             | SUM ('a aexp) ('a aexp)
  2944             | DIFF ('a aexp) ('a aexp)
  2945             | NUM 'a
  2946   and
  2947     'a bexp = LESS ('a aexp) ('a aexp)
  2948             | AND ('a bexp) ('a bexp)
  2949             | OR ('a bexp) ('a bexp)
  2950 
  2951 as well as indirectly recursive datatypes such as
  2952 
  2953   datatype
  2954     ('a, 'b) term = Var 'a
  2955                   | App 'b ((('a, 'b) term) list)
  2956 
  2957 The new tactic  mutual_induct_tac [<var_1>, ..., <var_n>] i  performs
  2958 induction on mutually / indirectly recursive datatypes.
  2959 
  2960 Primrec equations are now stored in theory and can be accessed via
  2961 <function_name>.simps.
  2962 
  2963 INCOMPATIBILITIES:
  2964 
  2965   - Theories using datatypes must now have theory Datatype as an
  2966     ancestor.
  2967   - The specific <typename>.induct_tac no longer exists - use the
  2968     generic induct_tac instead.
  2969   - natE has been renamed to nat.exhaust - use exhaust_tac
  2970     instead of res_inst_tac ... natE. Note that the variable
  2971     names in nat.exhaust differ from the names in natE, this
  2972     may cause some "fragile" proofs to fail.
  2973   - The theorems split_<typename>_case and split_<typename>_case_asm
  2974     have been renamed to <typename>.split and <typename>.split_asm.
  2975   - Since default sorts of type variables are now handled correctly,
  2976     some datatype definitions may have to be annotated with explicit
  2977     sort constraints.
  2978   - Primrec definitions no longer require function name and type
  2979     of recursive argument.
  2980 
  2981 Consider using isatool fixdatatype to adapt your theories and proof
  2982 scripts to the new package (backup your sources first!).
  2983 
  2984 
  2985 * HOL/record package: considerably improved implementation; now
  2986 includes concrete syntax for record types, terms, updates; theorems
  2987 for surjective pairing and splitting !!-bound record variables; proof
  2988 support is as follows:
  2989 
  2990   1) standard conversions (selectors or updates applied to record
  2991 constructor terms) are part of the standard simpset;
  2992 
  2993   2) inject equations of the form ((x, y) = (x', y')) == x=x' & y=y' are
  2994 made part of standard simpset and claset via addIffs;
  2995 
  2996   3) a tactic for record field splitting (record_split_tac) is part of
  2997 the standard claset (addSWrapper);
  2998 
  2999 To get a better idea about these rules you may retrieve them via
  3000 something like 'thms "foo.simps"' or 'thms "foo.iffs"', where "foo" is
  3001 the name of your record type.
  3002 
  3003 The split tactic 3) conceptually simplifies by the following rule:
  3004 
  3005   "(!!x. PROP ?P x) == (!!a b. PROP ?P (a, b))"
  3006 
  3007 Thus any record variable that is bound by meta-all will automatically
  3008 blow up into some record constructor term, consequently the
  3009 simplifications of 1), 2) apply.  Thus force_tac, auto_tac etc. shall
  3010 solve record problems automatically.
  3011 
  3012 
  3013 * reorganized the main HOL image: HOL/Integ and String loaded by
  3014 default; theory Main includes everything;
  3015 
  3016 * automatic simplification of integer sums and comparisons, using cancellation;
  3017 
  3018 * added option_map_eq_Some and not_Some_eq to the default simpset and claset;
  3019 
  3020 * added disj_not1 = "(~P | Q) = (P --> Q)" to the default simpset;
  3021 
  3022 * many new identities for unions, intersections, set difference, etc.;
  3023 
  3024 * expand_if, expand_split, expand_sum_case and expand_nat_case are now
  3025 called split_if, split_split, split_sum_case and split_nat_case (to go
  3026 with add/delsplits);
  3027 
  3028 * HOL/Prod introduces simplification procedure unit_eq_proc rewriting
  3029 (?x::unit) = (); this is made part of the default simpset, which COULD
  3030 MAKE EXISTING PROOFS FAIL under rare circumstances (consider
  3031 'Delsimprocs [unit_eq_proc];' as last resort); also note that
  3032 unit_abs_eta_conv is added in order to counter the effect of
  3033 unit_eq_proc on (%u::unit. f u), replacing it by f rather than by
  3034 %u.f();
  3035 
  3036 * HOL/Fun INCOMPATIBILITY: `inj_onto' is now called `inj_on' (which
  3037 makes more sense);
  3038 
  3039 * HOL/Set INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
  3040   It and 'sym RS equals0D' are now in the default  claset, giving automatic
  3041   disjointness reasoning but breaking a few old proofs.
  3042 
  3043 * HOL/Relation INCOMPATIBILITY: renamed the relational operator r^-1
  3044 to 'converse' from 'inverse' (for compatibility with ZF and some
  3045 literature);
  3046 
  3047 * HOL/recdef can now declare non-recursive functions, with {} supplied as
  3048 the well-founded relation;
  3049 
  3050 * HOL/Set INCOMPATIBILITY: the complement of set A is now written -A instead of
  3051     Compl A.  The "Compl" syntax remains available as input syntax for this
  3052     release ONLY.
  3053 
  3054 * HOL/Update: new theory of function updates:
  3055     f(a:=b) == %x. if x=a then b else f x
  3056 may also be iterated as in f(a:=b,c:=d,...);
  3057 
  3058 * HOL/Vimage: new theory for inverse image of a function, syntax f-``B;
  3059 
  3060 * HOL/List:
  3061   - new function list_update written xs[i:=v] that updates the i-th
  3062     list position. May also be iterated as in xs[i:=a,j:=b,...].
  3063   - new function `upt' written [i..j(] which generates the list
  3064     [i,i+1,...,j-1], i.e. the upper bound is excluded. To include the upper
  3065     bound write [i..j], which is a shorthand for [i..j+1(].
  3066   - new lexicographic orderings and corresponding wellfoundedness theorems.
  3067 
  3068 * HOL/Arith:
  3069   - removed 'pred' (predecessor) function;
  3070   - generalized some theorems about n-1;
  3071   - many new laws about "div" and "mod";
  3072   - new laws about greatest common divisors (see theory ex/Primes);
  3073 
  3074 * HOL/Relation: renamed the relational operator r^-1 "converse"
  3075 instead of "inverse";
  3076 
  3077 * HOL/Induct/Multiset: a theory of multisets, including the wellfoundedness
  3078   of the multiset ordering;
  3079 
  3080 * directory HOL/Real: a construction of the reals using Dedekind cuts
  3081   (not included by default);
  3082 
  3083 * directory HOL/UNITY: Chandy and Misra's UNITY formalism;
  3084 
  3085 * directory HOL/Hoare: a new version of Hoare logic which permits many-sorted
  3086   programs, i.e. different program variables may have different types.
  3087 
  3088 * calling (stac rew i) now fails if "rew" has no effect on the goal
  3089   [previously, this check worked only if the rewrite rule was unconditional]
  3090   Now rew can involve either definitions or equalities (either == or =).
  3091 
  3092 
  3093 *** ZF ***
  3094 
  3095 * theory Main includes everything; INCOMPATIBILITY: theory ZF.thy contains
  3096   only the theorems proved on ZF.ML;
  3097 
  3098 * ZF INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
  3099   It and 'sym RS equals0D' are now in the default  claset, giving automatic
  3100   disjointness reasoning but breaking a few old proofs.
  3101 
  3102 * ZF/Update: new theory of function updates
  3103     with default rewrite rule  f(x:=y) ` z = if(z=x, y, f`z)
  3104   may also be iterated as in f(a:=b,c:=d,...);
  3105 
  3106 * in  let x=t in u(x), neither t nor u(x) has to be an FOL term.
  3107 
  3108 * calling (stac rew i) now fails if "rew" has no effect on the goal
  3109   [previously, this check worked only if the rewrite rule was unconditional]
  3110   Now rew can involve either definitions or equalities (either == or =).
  3111 
  3112 * case_tac provided for compatibility with HOL
  3113     (like the old excluded_middle_tac, but with subgoals swapped)
  3114 
  3115 
  3116 *** Internal programming interfaces ***
  3117 
  3118 * Pure: several new basic modules made available for general use, see
  3119 also src/Pure/README;
  3120 
  3121 * improved the theory data mechanism to support encapsulation (data
  3122 kind name replaced by private Object.kind, acting as authorization
  3123 key); new type-safe user interface via functor TheoryDataFun; generic
  3124 print_data function becomes basically useless;
  3125 
  3126 * removed global_names compatibility flag -- all theory declarations
  3127 are qualified by default;
  3128 
  3129 * module Pure/Syntax now offers quote / antiquote translation
  3130 functions (useful for Hoare logic etc. with implicit dependencies);
  3131 see HOL/ex/Antiquote for an example use;
  3132 
  3133 * Simplifier now offers conversions (asm_)(full_)rewrite: simpset ->
  3134 cterm -> thm;
  3135 
  3136 * new tactical CHANGED_GOAL for checking that a tactic modifies a
  3137 subgoal;
  3138 
  3139 * Display.print_goals function moved to Locale.print_goals;
  3140 
  3141 * standard print function for goals supports current_goals_markers
  3142 variable for marking begin of proof, end of proof, start of goal; the
  3143 default is ("", "", ""); setting current_goals_markers := ("<proof>",
  3144 "</proof>", "<goal>") causes SGML like tagged proof state printing,
  3145 for example;
  3146 
  3147 
  3148 
  3149 New in Isabelle98 (January 1998)
  3150 --------------------------------
  3151 
  3152 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  3153 
  3154 * changed lexical syntax of terms / types: dots made part of long
  3155 identifiers, e.g. "%x.x" no longer possible, should be "%x. x";
  3156 
  3157 * simpset (and claset) reference variable replaced by functions
  3158 simpset / simpset_ref;
  3159 
  3160 * no longer supports theory aliases (via merge) and non-trivial
  3161 implicit merge of thms' signatures;
  3162 
  3163 * most internal names of constants changed due to qualified names;
  3164 
  3165 * changed Pure/Sequence interface (see Pure/seq.ML);
  3166 
  3167 
  3168 *** General Changes ***
  3169 
  3170 * hierachically structured name spaces (for consts, types, axms, thms
  3171 etc.); new lexical class 'longid' (e.g. Foo.bar.x) may render much of
  3172 old input syntactically incorrect (e.g. "%x.x"); COMPATIBILITY:
  3173 isatool fixdots ensures space after dots (e.g. "%x. x"); set
  3174 long_names for fully qualified output names; NOTE: ML programs
  3175 (special tactics, packages etc.) referring to internal names may have
  3176 to be adapted to cope with fully qualified names; in case of severe
  3177 backward campatibility problems try setting 'global_names' at compile
  3178 time to have enrything declared within a flat name space; one may also
  3179 fine tune name declarations in theories via the 'global' and 'local'
  3180 section;
  3181 
  3182 * reimplemented the implicit simpset and claset using the new anytype
  3183 data filed in signatures; references simpset:simpset ref etc. are
  3184 replaced by functions simpset:unit->simpset and
  3185 simpset_ref:unit->simpset ref; COMPATIBILITY: use isatool fixclasimp
  3186 to patch your ML files accordingly;
  3187 
  3188 * HTML output now includes theory graph data for display with Java
  3189 applet or isatool browser; data generated automatically via isatool
  3190 usedir (see -i option, ISABELLE_USEDIR_OPTIONS);
  3191 
  3192 * defs may now be conditional; improved rewrite_goals_tac to handle
  3193 conditional equations;
  3194 
  3195 * defs now admits additional type arguments, using TYPE('a) syntax;
  3196 
  3197 * theory aliases via merge (e.g. M=A+B+C) no longer supported, always
  3198 creates a new theory node; implicit merge of thms' signatures is
  3199 restricted to 'trivial' ones; COMPATIBILITY: one may have to use
  3200 transfer:theory->thm->thm in (rare) cases;
  3201 
  3202 * improved handling of draft signatures / theories; draft thms (and
  3203 ctyps, cterms) are automatically promoted to real ones;
  3204 
  3205 * slightly changed interfaces for oracles: admit many per theory, named
  3206 (e.g. oracle foo = mlfun), additional name argument for invoke_oracle;
  3207 
  3208 * print_goals: optional output of const types (set show_consts and
  3209 show_types);
  3210 
  3211 * improved output of warnings (###) and errors (***);
  3212 
  3213 * subgoal_tac displays a warning if the new subgoal has type variables;
  3214 
  3215 * removed old README and Makefiles;
  3216 
  3217 * replaced print_goals_ref hook by print_current_goals_fn and result_error_fn;
  3218 
  3219 * removed obsolete init_pps and init_database;
  3220 
  3221 * deleted the obsolete tactical STATE, which was declared by
  3222     fun STATE tacfun st = tacfun st st;
  3223 
  3224 * cd and use now support path variables, e.g. $ISABELLE_HOME, or ~
  3225 (which abbreviates $HOME);
  3226 
  3227 * changed Pure/Sequence interface (see Pure/seq.ML); COMPATIBILITY:
  3228 use isatool fixseq to adapt your ML programs (this works for fully
  3229 qualified references to the Sequence structure only!);
  3230 
  3231 * use_thy no longer requires writable current directory; it always
  3232 reloads .ML *and* .thy file, if either one is out of date;
  3233 
  3234 
  3235 *** Classical Reasoner ***
  3236 
  3237 * Clarify_tac, clarify_tac, clarify_step_tac, Clarify_step_tac: new
  3238 tactics that use classical reasoning to simplify a subgoal without
  3239 splitting it into several subgoals;
  3240 
  3241 * Safe_tac: like safe_tac but uses the default claset;
  3242 
  3243 
  3244 *** Simplifier ***
  3245 
  3246 * added simplification meta rules:
  3247     (asm_)(full_)simplify: simpset -> thm -> thm;
  3248 
  3249 * simplifier.ML no longer part of Pure -- has to be loaded by object
  3250 logics (again);
  3251 
  3252 * added prems argument to simplification procedures;
  3253 
  3254 * HOL, FOL, ZF: added infix function `addsplits':
  3255   instead of `<simpset> setloop (split_tac <thms>)'
  3256   you can simply write `<simpset> addsplits <thms>'
  3257 
  3258 
  3259 *** Syntax ***
  3260 
  3261 * TYPE('a) syntax for type reflection terms;
  3262 
  3263 * no longer handles consts with name "" -- declare as 'syntax' instead;
  3264 
  3265 * pretty printer: changed order of mixfix annotation preference (again!);
  3266 
  3267 * Pure: fixed idt/idts vs. pttrn/pttrns syntactic categories;
  3268 
  3269 
  3270 *** HOL ***
  3271 
  3272 * HOL: there is a new splitter `split_asm_tac' that can be used e.g.
  3273   with `addloop' of the simplifier to faciliate case splitting in premises.
  3274 
  3275 * HOL/TLA: Stephan Merz's formalization of Lamport's Temporal Logic of Actions;
  3276 
  3277 * HOL/Auth: new protocol proofs including some for the Internet
  3278   protocol TLS;
  3279 
  3280 * HOL/Map: new theory of `maps' a la VDM;
  3281 
  3282 * HOL/simplifier: simplification procedures nat_cancel_sums for
  3283 cancelling out common nat summands from =, <, <= (in)equalities, or
  3284 differences; simplification procedures nat_cancel_factor for
  3285 cancelling common factor from =, <, <= (in)equalities over natural
  3286 sums; nat_cancel contains both kinds of procedures, it is installed by
  3287 default in Arith.thy -- this COULD MAKE EXISTING PROOFS FAIL;
  3288 
  3289 * HOL/simplifier: terms of the form
  3290   `? x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x)'  (or t=x)
  3291   are rewritten to
  3292   `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t)',
  3293   and those of the form
  3294   `! x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x) --> R(x)'  (or t=x)
  3295   are rewritten to
  3296   `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t) --> R(t)',
  3297 
  3298 * HOL/datatype
  3299   Each datatype `t' now comes with a theorem `split_t_case' of the form
  3300 
  3301   P(t_case f1 ... fn x) =
  3302      ( (!y1 ... ym1. x = C1 y1 ... ym1 --> P(f1 y1 ... ym1)) &
  3303         ...
  3304        (!y1 ... ymn. x = Cn y1 ... ymn --> P(f1 y1 ... ymn))
  3305      )
  3306 
  3307   and a theorem `split_t_case_asm' of the form
  3308 
  3309   P(t_case f1 ... fn x) =
  3310     ~( (? y1 ... ym1. x = C1 y1 ... ym1 & ~P(f1 y1 ... ym1)) |
  3311         ...
  3312        (? y1 ... ymn. x = Cn y1 ... ymn & ~P(f1 y1 ... ymn))
  3313      )
  3314   which can be added to a simpset via `addsplits'. The existing theorems
  3315   expand_list_case and expand_option_case have been renamed to
  3316   split_list_case and split_option_case.
  3317 
  3318 * HOL/Arithmetic:
  3319   - `pred n' is automatically converted to `n-1'.
  3320     Users are strongly encouraged not to use `pred' any longer,
  3321     because it will disappear altogether at some point.
  3322   - Users are strongly encouraged to write "0 < n" rather than
  3323     "n ~= 0". Theorems and proof tools have been modified towards this
  3324     `standard'.
  3325 
  3326 * HOL/Lists:
  3327   the function "set_of_list" has been renamed "set" (and its theorems too);
  3328   the function "nth" now takes its arguments in the reverse order and
  3329   has acquired the infix notation "!" as in "xs!n".
  3330 
  3331 * HOL/Set: UNIV is now a constant and is no longer translated to Compl{};
  3332 
  3333 * HOL/Set: The operator (UN x.B x) now abbreviates (UN x:UNIV. B x) and its
  3334   specialist theorems (like UN1_I) are gone.  Similarly for (INT x.B x);
  3335 
  3336 * HOL/record: extensible records with schematic structural subtyping
  3337 (single inheritance); EXPERIMENTAL version demonstrating the encoding,
  3338 still lacks various theorems and concrete record syntax;
  3339 
  3340 
  3341 *** HOLCF ***
  3342 
  3343 * removed "axioms" and "generated by" sections;
  3344 
  3345 * replaced "ops" section by extended "consts" section, which is capable of
  3346   handling the continuous function space "->" directly;
  3347 
  3348 * domain package:
  3349   . proves theorems immediately and stores them in the theory,
  3350   . creates hierachical name space,
  3351   . now uses normal mixfix annotations (instead of cinfix...),
  3352   . minor changes to some names and values (for consistency),
  3353   . e.g. cases -> casedist, dists_eq -> dist_eqs, [take_lemma] -> take_lemmas,
  3354   . separator between mutual domain defs: changed "," to "and",
  3355   . improved handling of sort constraints;  now they have to
  3356     appear on the left-hand side of the equations only;
  3357 
  3358 * fixed LAM <x,y,zs>.b syntax;
  3359 
  3360 * added extended adm_tac to simplifier in HOLCF -- can now discharge
  3361 adm (%x. P (t x)), where P is chainfinite and t continuous;
  3362 
  3363 
  3364 *** FOL and ZF ***
  3365 
  3366 * FOL: there is a new splitter `split_asm_tac' that can be used e.g.
  3367   with `addloop' of the simplifier to faciliate case splitting in premises.
  3368 
  3369 * qed_spec_mp, qed_goal_spec_mp, qed_goalw_spec_mp are available, as
  3370 in HOL, they strip ALL and --> from proved theorems;
  3371 
  3372 
  3373 
  3374 New in Isabelle94-8 (May 1997)
  3375 ------------------------------
  3376 
  3377 *** General Changes ***
  3378 
  3379 * new utilities to build / run / maintain Isabelle etc. (in parts
  3380 still somewhat experimental); old Makefiles etc. still functional;
  3381 
  3382 * new 'Isabelle System Manual';
  3383 
  3384 * INSTALL text, together with ./configure and ./build scripts;
  3385 
  3386 * reimplemented type inference for greater efficiency, better error
  3387 messages and clean internal interface;
  3388 
  3389 * prlim command for dealing with lots of subgoals (an easier way of
  3390 setting goals_limit);
  3391 
  3392 
  3393 *** Syntax ***
  3394 
  3395 * supports alternative (named) syntax tables (parser and pretty
  3396 printer); internal interface is provided by add_modesyntax(_i);
  3397 
  3398 * Pure, FOL, ZF, HOL, HOLCF now support symbolic input and output; to
  3399 be used in conjunction with the Isabelle symbol font; uses the
  3400 "symbols" syntax table;
  3401 
  3402 * added token_translation interface (may translate name tokens in
  3403 arbitrary ways, dependent on their type (free, bound, tfree, ...) and
  3404 the current print_mode); IMPORTANT: user print translation functions
  3405 are responsible for marking newly introduced bounds
  3406 (Syntax.mark_boundT);
  3407 
  3408 * token translations for modes "xterm" and "xterm_color" that display
  3409 names in bold, underline etc. or colors (which requires a color
  3410 version of xterm);
  3411 
  3412 * infixes may now be declared with names independent of their syntax;
  3413 
  3414 * added typed_print_translation (like print_translation, but may
  3415 access type of constant);
  3416 
  3417 
  3418 *** Classical Reasoner ***
  3419 
  3420 Blast_tac: a new tactic!  It is often more powerful than fast_tac, but has
  3421 some limitations.  Blast_tac...
  3422   + ignores addss, addbefore, addafter; this restriction is intrinsic
  3423   + ignores elimination rules that don't have the correct format
  3424         (the conclusion MUST be a formula variable)
  3425   + ignores types, which can make HOL proofs fail
  3426   + rules must not require higher-order unification, e.g. apply_type in ZF
  3427     [message "Function Var's argument not a bound variable" relates to this]
  3428   + its proof strategy is more general but can actually be slower
  3429 
  3430 * substitution with equality assumptions no longer permutes other
  3431 assumptions;
  3432 
  3433 * minor changes in semantics of addafter (now called addaltern); renamed
  3434 setwrapper to setWrapper and compwrapper to compWrapper; added safe wrapper
  3435 (and access functions for it);
  3436 
  3437 * improved combination of classical reasoner and simplifier:
  3438   + functions for handling clasimpsets
  3439   + improvement of addss: now the simplifier is called _after_ the
  3440     safe steps.
  3441   + safe variant of addss called addSss: uses safe simplifications
  3442     _during_ the safe steps. It is more complete as it allows multiple
  3443     instantiations of unknowns (e.g. with slow_tac).
  3444 
  3445 *** Simplifier ***
  3446 
  3447 * added interface for simplification procedures (functions that
  3448 produce *proven* rewrite rules on the fly, depending on current
  3449 redex);
  3450 
  3451 * ordering on terms as parameter (used for ordered rewriting);
  3452 
  3453 * new functions delcongs, deleqcongs, and Delcongs. richer rep_ss;
  3454 
  3455 * the solver is now split into a safe and an unsafe part.
  3456 This should be invisible for the normal user, except that the
  3457 functions setsolver and addsolver have been renamed to setSolver and
  3458 addSolver; added safe_asm_full_simp_tac;
  3459 
  3460 
  3461 *** HOL ***
  3462 
  3463 * a generic induction tactic `induct_tac' which works for all datatypes and
  3464 also for type `nat';
  3465 
  3466 * a generic case distinction tactic `exhaust_tac' which works for all
  3467 datatypes and also for type `nat';
  3468 
  3469 * each datatype comes with a function `size';
  3470 
  3471 * patterns in case expressions allow tuple patterns as arguments to
  3472 constructors, for example `case x of [] => ... | (x,y,z)#ps => ...';
  3473 
  3474 * primrec now also works with type nat;
  3475 
  3476 * recdef: a new declaration form, allows general recursive functions to be
  3477 defined in theory files.  See HOL/ex/Fib, HOL/ex/Primes, HOL/Subst/Unify.
  3478 
  3479 * the constant for negation has been renamed from "not" to "Not" to
  3480 harmonize with FOL, ZF, LK, etc.;
  3481 
  3482 * HOL/ex/LFilter theory of a corecursive "filter" functional for
  3483 infinite lists;
  3484 
  3485 * HOL/Modelcheck demonstrates invocation of model checker oracle;
  3486 
  3487 * HOL/ex/Ring.thy declares cring_simp, which solves equational
  3488 problems in commutative rings, using axiomatic type classes for + and *;
  3489 
  3490 * more examples in HOL/MiniML and HOL/Auth;
  3491 
  3492 * more default rewrite rules for quantifiers, union/intersection;
  3493 
  3494 * a new constant `arbitrary == @x.False';
  3495 
  3496 * HOLCF/IOA replaces old HOL/IOA;
  3497 
  3498 * HOLCF changes: derived all rules and arities
  3499   + axiomatic type classes instead of classes
  3500   + typedef instead of faking type definitions
  3501   + eliminated the internal constants less_fun, less_cfun, UU_fun, UU_cfun etc.
  3502   + new axclasses cpo, chfin, flat with flat < chfin < pcpo < cpo < po
  3503   + eliminated the types void, one, tr
  3504   + use unit lift and bool lift (with translations) instead of one and tr
  3505   + eliminated blift from Lift3.thy (use Def instead of blift)
  3506   all eliminated rules are derived as theorems --> no visible changes ;
  3507 
  3508 
  3509 *** ZF ***
  3510 
  3511 * ZF now has Fast_tac, Simp_tac and Auto_tac.  Union_iff is a now a default
  3512 rewrite rule; this may affect some proofs.  eq_cs is gone but can be put back
  3513 as ZF_cs addSIs [equalityI];
  3514 
  3515 
  3516 
  3517 New in Isabelle94-7 (November 96)
  3518 ---------------------------------
  3519 
  3520 * allowing negative levels (as offsets) in prlev and choplev;
  3521 
  3522 * super-linear speedup for large simplifications;
  3523 
  3524 * FOL, ZF and HOL now use miniscoping: rewriting pushes
  3525 quantifications in as far as possible (COULD MAKE EXISTING PROOFS
  3526 FAIL); can suppress it using the command Delsimps (ex_simps @
  3527 all_simps); De Morgan laws are also now included, by default;
  3528 
  3529 * improved printing of ==>  :  ~:
  3530 
  3531 * new object-logic "Sequents" adds linear logic, while replacing LK
  3532 and Modal (thanks to Sara Kalvala);
  3533 
  3534 * HOL/Auth: correctness proofs for authentication protocols;
  3535 
  3536 * HOL: new auto_tac combines rewriting and classical reasoning (many
  3537 examples on HOL/Auth);
  3538 
  3539 * HOL: new command AddIffs for declaring theorems of the form P=Q to
  3540 the rewriter and classical reasoner simultaneously;
  3541 
  3542 * function uresult no longer returns theorems in "standard" format;
  3543 regain previous version by: val uresult = standard o uresult;
  3544 
  3545 
  3546 
  3547 New in Isabelle94-6
  3548 -------------------
  3549 
  3550 * oracles -- these establish an interface between Isabelle and trusted
  3551 external reasoners, which may deliver results as theorems;
  3552 
  3553 * proof objects (in particular record all uses of oracles);
  3554 
  3555 * Simp_tac, Fast_tac, etc. that refer to implicit simpset / claset;
  3556 
  3557 * "constdefs" section in theory files;
  3558 
  3559 * "primrec" section (HOL) no longer requires names;
  3560 
  3561 * internal type "tactic" now simply "thm -> thm Sequence.seq";
  3562 
  3563 
  3564 
  3565 New in Isabelle94-5
  3566 -------------------
  3567 
  3568 * reduced space requirements;
  3569 
  3570 * automatic HTML generation from theories;
  3571 
  3572 * theory files no longer require "..." (quotes) around most types;
  3573 
  3574 * new examples, including two proofs of the Church-Rosser theorem;
  3575 
  3576 * non-curried (1994) version of HOL is no longer distributed;
  3577 
  3578 
  3579 
  3580 New in Isabelle94-4
  3581 -------------------
  3582 
  3583 * greatly reduced space requirements;
  3584 
  3585 * theory files (.thy) no longer require \...\ escapes at line breaks;
  3586 
  3587 * searchable theorem database (see the section "Retrieving theorems" on
  3588 page 8 of the Reference Manual);
  3589 
  3590 * new examples, including Grabczewski's monumental case study of the
  3591 Axiom of Choice;
  3592 
  3593 * The previous version of HOL renamed to Old_HOL;
  3594 
  3595 * The new version of HOL (previously called CHOL) uses a curried syntax
  3596 for functions.  Application looks like f a b instead of f(a,b);
  3597 
  3598 * Mutually recursive inductive definitions finally work in HOL;
  3599 
  3600 * In ZF, pattern-matching on tuples is now available in all abstractions and
  3601 translates to the operator "split";
  3602 
  3603 
  3604 
  3605 New in Isabelle94-3
  3606 -------------------
  3607 
  3608 * new infix operator, addss, allowing the classical reasoner to
  3609 perform simplification at each step of its search.  Example:
  3610         fast_tac (cs addss ss)
  3611 
  3612 * a new logic, CHOL, the same as HOL, but with a curried syntax
  3613 for functions.  Application looks like f a b instead of f(a,b).  Also pairs
  3614 look like (a,b) instead of <a,b>;
  3615 
  3616 * PLEASE NOTE: CHOL will eventually replace HOL!
  3617 
  3618 * In CHOL, pattern-matching on tuples is now available in all abstractions.
  3619 It translates to the operator "split".  A new theory of integers is available;
  3620 
  3621 * In ZF, integer numerals now denote two's-complement binary integers.
  3622 Arithmetic operations can be performed by rewriting.  See ZF/ex/Bin.ML;
  3623 
  3624 * Many new examples: I/O automata, Church-Rosser theorem, equivalents
  3625 of the Axiom of Choice;
  3626 
  3627 
  3628 
  3629 New in Isabelle94-2
  3630 -------------------
  3631 
  3632 * Significantly faster resolution;
  3633 
  3634 * the different sections in a .thy file can now be mixed and repeated
  3635 freely;
  3636 
  3637 * Database of theorems for FOL, HOL and ZF.  New
  3638 commands including qed, qed_goal and bind_thm store theorems in the database.
  3639 
  3640 * Simple database queries: return a named theorem (get_thm) or all theorems of
  3641 a given theory (thms_of), or find out what theory a theorem was proved in
  3642 (theory_of_thm);
  3643 
  3644 * Bugs fixed in the inductive definition and datatype packages;
  3645 
  3646 * The classical reasoner provides deepen_tac and depth_tac, making FOL_dup_cs
  3647 and HOL_dup_cs obsolete;
  3648 
  3649 * Syntactic ambiguities caused by the new treatment of syntax in Isabelle94-1
  3650 have been removed;
  3651 
  3652 * Simpler definition of function space in ZF;
  3653 
  3654 * new results about cardinal and ordinal arithmetic in ZF;
  3655 
  3656 * 'subtype' facility in HOL for introducing new types as subsets of existing
  3657 types;
  3658 
  3659 
  3660 $Id$