NEWS
author wenzelm
Wed, 25 Jul 2007 22:20:54 +0200
changeset 23992 bf352c4c499b
parent 23977 5a3ec03c825b
child 24032 b3d7eb6f535f
permissions -rw-r--r--
updated;
     1 Isabelle NEWS -- history user-relevant changes
     2 ==============================================
     3 
     4 New in this Isabelle version
     5 ----------------------------
     6 
     7 *** General ***
     8 
     9 * More uniform information about legacy features, notably a
    10 warning/error of "Legacy feature: ...", depending on the state of the
    11 tolerate_legacy_features flag (default true). FUTURE INCOMPATIBILITY:
    12 legacy features will disappear eventually.
    13 
    14 * Theory syntax: the header format ``theory A = B + C:'' has been
    15 discontinued in favour of ``theory A imports B C begin''.  Use isatool
    16 fixheaders to convert existing theory files.  INCOMPATIBILITY.
    17 
    18 * Theory syntax: the old non-Isar theory file format has been
    19 discontinued altogether.  Note that ML proof scripts may still be used
    20 with Isar theories; migration is usually quite simple with the ML
    21 function use_legacy_bindings.  INCOMPATIBILITY.
    22 
    23 * Theory syntax: some popular names (e.g. 'class', 'declaration',
    24 'fun', 'help', 'if') are now keywords.  INCOMPATIBILITY, use double
    25 quotes.
    26 
    27 * Theory loader: be more serious about observing the static theory
    28 header specifications (including optional directories), but not the
    29 accidental file locations of previously successful loads.  Potential
    30 INCOMPATIBILITY, may need to refine theory headers.
    31 
    32 * Theory loader: optional support for content-based file
    33 identification, instead of the traditional scheme of full physical
    34 path plus date stamp; configured by the ISABELLE_FILE_IDENT setting
    35 (cf. the system manual).  The new scheme allows to work with
    36 non-finished theories in persistent session images, such that source
    37 files may be moved later on without requiring reloads.
    38 
    39 * Legacy goal package: reduced interface to the bare minimum required
    40 to keep existing proof scripts running.  Most other user-level
    41 functions are now part of the OldGoals structure, which is *not* open
    42 by default (consider isatool expandshort before open OldGoals).
    43 Removed top_sg, prin, printyp, pprint_term/typ altogether, because
    44 these tend to cause confusion about the actual goal (!) context being
    45 used here, which is not necessarily the same as the_context().
    46 
    47 * Command 'find_theorems': supports "*" wild-card in "name:"
    48 criterion; "with_dups" option.  Certain ProofGeneral versions might
    49 support a specific search form (see ProofGeneral/CHANGES).
    50 
    51 * The ``prems limit'' option (cf. ProofContext.prems_limit) is now -1
    52 by default, which means that "prems" (and also "fixed variables") are
    53 suppressed from proof state output.  Note that the ProofGeneral
    54 settings mechanism allows to change and save options persistently, but
    55 older versions of Isabelle will fail to start up if a negative prems
    56 limit is imposed.
    57 
    58 * Local theory targets may be specified by non-nested blocks of
    59 ``context/locale/class ... begin'' followed by ``end''.  The body may
    60 contain definitions, theorems etc., including any derived mechanism
    61 that has been implemented on top of these primitives.  This concept
    62 generalizes the existing ``theorem (in ...)'' towards more versatility
    63 and scalability.
    64 
    65 * Proof General interface: proper undo of final 'end' command;
    66 discontinued Isabelle/classic mode (ML proof scripts).
    67 
    68 
    69 *** Document preparation ***
    70 
    71 * Added antiquotation @{theory name} which prints the given name,
    72 after checking that it refers to a valid ancestor theory in the
    73 current context.
    74 
    75 * Added antiquotations @{ML_type text} and @{ML_struct text} which
    76 check the given source text as ML type/structure, printing verbatim.
    77 
    78 * Added antiquotation @{abbrev "c args"} which prints the abbreviation
    79 "c args == rhs" given in the current context.  (Any number of
    80 arguments may be given on the LHS.)
    81 
    82 
    83 
    84 *** Pure ***
    85 
    86 * code generator: consts in 'consts_code' Isar commands are now referred
    87   to by usual term syntax (including optional type annotations).
    88 
    89 * code generator: 
    90   - Isar 'definition's, 'constdef's and primitive instance definitions are added
    91     explicitly to the table of defining equations
    92   - primitive definitions are not used as defining equations by default any longer
    93   - defining equations are now definitly restricted to meta "==" and object
    94         equality "="
    95   - HOL theories have been adopted accordingly
    96 
    97 * class_package.ML offers a combination of axclasses and locales to
    98 achieve Haskell-like type classes in Isabelle.  See
    99 HOL/ex/Classpackage.thy for examples.
   100 
   101 * Yet another code generator framework allows to generate executable
   102 code for ML and Haskell (including "class"es).  A short usage sketch:
   103 
   104     internal compilation:
   105         code_gen <list of constants (term syntax)> in SML
   106     writing SML code to a file:
   107         code_gen <list of constants (term syntax)> in SML <filename>
   108     writing OCaml code to a file:
   109         code_gen <list of constants (term syntax)> in OCaml <filename>
   110     writing Haskell code to a bunch of files:
   111         code_gen <list of constants (term syntax)> in Haskell <filename>
   112 
   113 Reasonable default setup of framework in HOL/Main.
   114 
   115 Theorem attributs for selecting and transforming function equations theorems:
   116 
   117     [code fun]:        select a theorem as function equation for a specific constant
   118     [code fun del]:    deselect a theorem as function equation for a specific constant
   119     [code inline]:     select an equation theorem for unfolding (inlining) in place
   120     [code inline del]: deselect an equation theorem for unfolding (inlining) in place
   121 
   122 User-defined serializations (target in {SML, OCaml, Haskell}):
   123 
   124     code_const <and-list of constants (term syntax)>
   125       {(target) <and-list of const target syntax>}+
   126 
   127     code_type <and-list of type constructors>
   128       {(target) <and-list of type target syntax>}+
   129 
   130     code_instance <and-list of instances>
   131       {(target)}+
   132         where instance ::= <type constructor> :: <class>
   133 
   134     code_class <and_list of classes>
   135       {(target) <and-list of class target syntax>}+
   136         where class target syntax ::= <class name> {where {<classop> == <target syntax>}+}?
   137 
   138 code_instance and code_class only apply to target Haskell.
   139 
   140 See HOL theories and HOL/ex/Codegenerator*.thy for usage examples.
   141 Doc/Isar/Advanced/Codegen/ provides a tutorial.
   142 
   143 * Command 'no_translations' removes translation rules from theory
   144 syntax.
   145 
   146 * Overloaded definitions are now actually checked for acyclic
   147 dependencies.  The overloading scheme is slightly more general than
   148 that of Haskell98, although Isabelle does not demand an exact
   149 correspondence to type class and instance declarations.
   150 INCOMPATIBILITY, use ``defs (unchecked overloaded)'' to admit more
   151 exotic versions of overloading -- at the discretion of the user!
   152 
   153 Polymorphic constants are represented via type arguments, i.e. the
   154 instantiation that matches an instance against the most general
   155 declaration given in the signature.  For example, with the declaration
   156 c :: 'a => 'a => 'a, an instance c :: nat => nat => nat is represented
   157 as c(nat).  Overloading is essentially simultaneous structural
   158 recursion over such type arguments.  Incomplete specification patterns
   159 impose global constraints on all occurrences, e.g. c('a * 'a) on the
   160 LHS means that more general c('a * 'b) will be disallowed on any RHS.
   161 Command 'print_theory' outputs the normalized system of recursive
   162 equations, see section "definitions".
   163 
   164 * Isar: method "assumption" (and implicit closing of subproofs) now
   165 takes simple non-atomic goal assumptions into account: after applying
   166 an assumption as a rule the resulting subgoals are solved by atomic
   167 assumption steps.  This is particularly useful to finish 'obtain'
   168 goals, such as "!!x. (!!x. P x ==> thesis) ==> P x ==> thesis",
   169 without referring to the original premise "!!x. P x ==> thesis" in the
   170 Isar proof context.  POTENTIAL INCOMPATIBILITY: method "assumption" is
   171 more permissive.
   172 
   173 * Isar: implicit use of prems from the Isar proof context is
   174 considered a legacy feature.  Common applications like ``have A .''
   175 may be replaced by ``have A by fact'' or ``note `A`''.  In general,
   176 referencing facts explicitly here improves readability and
   177 maintainability of proof texts.
   178 
   179 * Isar: improper proof element 'guess' is like 'obtain', but derives
   180 the obtained context from the course of reasoning!  For example:
   181 
   182   assume "EX x y. A x & B y"   -- "any previous fact"
   183   then guess x and y by clarify
   184 
   185 This technique is potentially adventurous, depending on the facts and
   186 proof tools being involved here.
   187 
   188 * Isar: known facts from the proof context may be specified as literal
   189 propositions, using ASCII back-quote syntax.  This works wherever
   190 named facts used to be allowed so far, in proof commands, proof
   191 methods, attributes etc.  Literal facts are retrieved from the context
   192 according to unification of type and term parameters.  For example,
   193 provided that "A" and "A ==> B" and "!!x. P x ==> Q x" are known
   194 theorems in the current context, then these are valid literal facts:
   195 `A` and `A ==> B` and `!!x. P x ==> Q x" as well as `P a ==> Q a` etc.
   196 
   197 There is also a proof method "fact" which does the same composition
   198 for explicit goal states, e.g. the following proof texts coincide with
   199 certain special cases of literal facts:
   200 
   201   have "A" by fact                 ==  note `A`
   202   have "A ==> B" by fact           ==  note `A ==> B`
   203   have "!!x. P x ==> Q x" by fact  ==  note `!!x. P x ==> Q x`
   204   have "P a ==> Q a" by fact       ==  note `P a ==> Q a`
   205 
   206 * Isar: ":" (colon) is no longer a symbolic identifier character in
   207 outer syntax.  Thus symbolic identifiers may be used without
   208 additional white space in declarations like this: ``assume *: A''.
   209 
   210 * Isar: 'print_facts' prints all local facts of the current context,
   211 both named and unnamed ones.
   212 
   213 * Isar: 'def' now admits simultaneous definitions, e.g.:
   214 
   215   def x == "t" and y == "u"
   216 
   217 * Isar: added command 'unfolding', which is structurally similar to
   218 'using', but affects both the goal state and facts by unfolding given
   219 rewrite rules.  Thus many occurrences of the 'unfold' method or
   220 'unfolded' attribute may be replaced by first-class proof text.
   221 
   222 * Isar: methods 'unfold' / 'fold', attributes 'unfolded' / 'folded',
   223 and command 'unfolding' now all support object-level equalities
   224 (potentially conditional).  The underlying notion of rewrite rule is
   225 analogous to the 'rule_format' attribute, but *not* that of the
   226 Simplifier (which is usually more generous).
   227 
   228 * Isar: the goal restriction operator [N] (default N = 1) evaluates a
   229 method expression within a sandbox consisting of the first N
   230 sub-goals, which need to exist.  For example, ``simp_all [3]''
   231 simplifies the first three sub-goals, while (rule foo, simp_all)[]
   232 simplifies all new goals that emerge from applying rule foo to the
   233 originally first one.
   234 
   235 * Isar: schematic goals are no longer restricted to higher-order
   236 patterns; e.g. ``lemma "?P(?x)" by (rule TrueI)'' now works as
   237 expected.
   238 
   239 * Isar: the conclusion of a long theorem statement is now either
   240 'shows' (a simultaneous conjunction, as before), or 'obtains'
   241 (essentially a disjunction of cases with local parameters and
   242 assumptions).  The latter allows to express general elimination rules
   243 adequately; in this notation common elimination rules look like this:
   244 
   245   lemma exE:    -- "EX x. P x ==> (!!x. P x ==> thesis) ==> thesis"
   246     assumes "EX x. P x"
   247     obtains x where "P x"
   248 
   249   lemma conjE:  -- "A & B ==> (A ==> B ==> thesis) ==> thesis"
   250     assumes "A & B"
   251     obtains A and B
   252 
   253   lemma disjE:  -- "A | B ==> (A ==> thesis) ==> (B ==> thesis) ==> thesis"
   254     assumes "A | B"
   255     obtains
   256       A
   257     | B
   258 
   259 The subsequent classical rules even refer to the formal "thesis"
   260 explicitly:
   261 
   262   lemma classical:     -- "(~ thesis ==> thesis) ==> thesis"
   263     obtains "~ thesis"
   264 
   265   lemma Peirce's_Law:  -- "((thesis ==> something) ==> thesis) ==> thesis"
   266     obtains "thesis ==> something"
   267 
   268 The actual proof of an 'obtains' statement is analogous to that of the
   269 Isar proof element 'obtain', only that there may be several cases.
   270 Optional case names may be specified in parentheses; these will be
   271 available both in the present proof and as annotations in the
   272 resulting rule, for later use with the 'cases' method (cf. attribute
   273 case_names).
   274 
   275 * Isar: the assumptions of a long theorem statement are available as
   276 "assms" fact in the proof context.  This is more appropriate than the
   277 (historical) "prems", which refers to all assumptions of the current
   278 context, including those from the target locale, proof body etc.
   279 
   280 * Isar: 'print_statement' prints theorems from the current theory or
   281 proof context in long statement form, according to the syntax of a
   282 top-level lemma.
   283 
   284 * Isar: 'obtain' takes an optional case name for the local context
   285 introduction rule (default "that").
   286 
   287 * Isar: removed obsolete 'concl is' patterns.  INCOMPATIBILITY, use
   288 explicit (is "_ ==> ?foo") in the rare cases where this still happens
   289 to occur.
   290 
   291 * Pure: syntax "CONST name" produces a fully internalized constant
   292 according to the current context.  This is particularly useful for
   293 syntax translations that should refer to internal constant
   294 representations independently of name spaces.
   295 
   296 * Pure: syntax constant for foo (binder "FOO ") is called "foo_binder"
   297 instead of "FOO ". This allows multiple binder declarations to coexist
   298 in the same context.  INCOMPATIBILITY.
   299 
   300 * Isar/locales: 'notation' provides a robust interface to the 'syntax'
   301 primitive that also works in a locale context (both for constants and
   302 fixed variables).  Type declaration and internal syntactic
   303 representation of given constants retrieved from the context.
   304 
   305 * Isar/locales: new derived specification elements 'axiomatization',
   306 'definition', 'abbreviation', which support type-inference, admit
   307 object-level specifications (equality, equivalence).  See also the
   308 isar-ref manual.  Examples:
   309 
   310   axiomatization
   311     eq  (infix "===" 50) where
   312     eq_refl: "x === x" and eq_subst: "x === y ==> P x ==> P y"
   313 
   314   definition "f x y = x + y + 1"
   315   definition g where "g x = f x x"
   316 
   317   abbreviation
   318     neq  (infix "=!=" 50) where
   319     "x =!= y == ~ (x === y)"
   320 
   321 These specifications may be also used in a locale context.  Then the
   322 constants being introduced depend on certain fixed parameters, and the
   323 constant name is qualified by the locale base name.  An internal
   324 abbreviation takes care for convenient input and output, making the
   325 parameters implicit and using the original short name.  See also
   326 HOL/ex/Abstract_NAT.thy for an example of deriving polymorphic
   327 entities from a monomorphic theory.
   328 
   329 Presently, abbreviations are only available 'in' a target locale, but
   330 not inherited by general import expressions.  Also note that
   331 'abbreviation' may be used as a type-safe replacement for 'syntax' +
   332 'translations' in common applications.
   333 
   334 Concrete syntax is attached to specified constants in internal form,
   335 independently of name spaces.  The parse tree representation is
   336 slightly different -- use 'notation' instead of raw 'syntax', and
   337 'translations' with explicit "CONST" markup to accommodate this.
   338 
   339 * Pure: command 'print_abbrevs' prints all constant abbreviations of
   340 the current context.  Print mode "no_abbrevs" prevents inversion of
   341 abbreviations on output.
   342 
   343 * Isar/locales: improved parameter handling:
   344 - use of locales "var" and "struct" no longer necessary;
   345 - parameter renamings are no longer required to be injective.
   346   This enables, for example, to define a locale for endomorphisms thus:
   347   locale endom = homom mult mult h.
   348 
   349 * Isar/locales: changed the way locales with predicates are defined.
   350 Instead of accumulating the specification, the imported expression is
   351 now an interpretation.  INCOMPATIBILITY: different normal form of
   352 locale expressions.  In particular, in interpretations of locales with
   353 predicates, goals repesenting already interpreted fragments are not
   354 removed automatically.  Use methods `intro_locales' and
   355 `unfold_locales'; see below.
   356 
   357 * Isar/locales: new methods `intro_locales' and `unfold_locales'
   358 provide backward reasoning on locales predicates.  The methods are
   359 aware of interpretations and discharge corresponding goals.
   360 `intro_locales' is less aggressive then `unfold_locales' and does not
   361 unfold predicates to assumptions.
   362 
   363 * Isar/locales: the order in which locale fragments are accumulated
   364 has changed.  This enables to override declarations from fragments due
   365 to interpretations -- for example, unwanted simp rules.
   366 
   367 * Isar/locales: interpretation in theories and proof contexts has been
   368 extended.  One may now specify (and prove) equations, which are
   369 unfolded in interpreted theorems.  This is useful for replacing
   370 defined concepts (constants depending on locale parameters) by
   371 concepts already existing in the target context.  Example:
   372 
   373   interpretation partial_order ["op <= :: [int, int] => bool"]
   374     where "partial_order.less (op <=) (x::int) y = (x < y)"
   375 
   376 Typically, the constant `partial_order.less' is created by a definition
   377 specification element in the context of locale partial_order.
   378 
   379 * Provers/induct: improved internal context management to support
   380 local fixes and defines on-the-fly.  Thus explicit meta-level
   381 connectives !! and ==> are rarely required anymore in inductive goals
   382 (using object-logic connectives for this purpose has been long
   383 obsolete anyway).  The subsequent proof patterns illustrate advanced
   384 techniques of natural induction; general datatypes and inductive sets
   385 work analogously (see also src/HOL/Lambda for realistic examples).
   386 
   387 (1) This is how to ``strengthen'' an inductive goal wrt. certain
   388 parameters:
   389 
   390   lemma
   391     fixes n :: nat and x :: 'a
   392     assumes a: "A n x"
   393     shows "P n x"
   394     using a                     -- {* make induct insert fact a *}
   395   proof (induct n arbitrary: x) -- {* generalize goal to "!!x. A n x ==> P n x" *}
   396     case 0
   397     show ?case sorry
   398   next
   399     case (Suc n)
   400     note `!!x. A n x ==> P n x` -- {* induction hypothesis, according to induction rule *}
   401     note `A (Suc n) x`          -- {* induction premise, stemming from fact a *}
   402     show ?case sorry
   403   qed
   404 
   405 (2) This is how to perform induction over ``expressions of a certain
   406 form'', using a locally defined inductive parameter n == "a x"
   407 together with strengthening (the latter is usually required to get
   408 sufficiently flexible induction hypotheses):
   409 
   410   lemma
   411     fixes a :: "'a => nat"
   412     assumes a: "A (a x)"
   413     shows "P (a x)"
   414     using a
   415   proof (induct n == "a x" arbitrary: x)
   416     ...
   417 
   418 See also HOL/Isar_examples/Puzzle.thy for an application of the this
   419 particular technique.
   420 
   421 (3) This is how to perform existential reasoning ('obtains' or
   422 'obtain') by induction, while avoiding explicit object-logic
   423 encodings:
   424 
   425   lemma
   426     fixes n :: nat
   427     obtains x :: 'a where "P n x" and "Q n x"
   428   proof (induct n arbitrary: thesis)
   429     case 0
   430     obtain x where "P 0 x" and "Q 0 x" sorry
   431     then show thesis by (rule 0)
   432   next
   433     case (Suc n)
   434     obtain x where "P n x" and "Q n x" by (rule Suc.hyps)
   435     obtain x where "P (Suc n) x" and "Q (Suc n) x" sorry
   436     then show thesis by (rule Suc.prems)
   437   qed
   438 
   439 Here the 'arbitrary: thesis' specification essentially modifies the
   440 scope of the formal thesis parameter, in order to the get the whole
   441 existence statement through the induction as expected.
   442 
   443 * Provers/induct: mutual induction rules are now specified as a list
   444 of rule sharing the same induction cases.  HOL packages usually
   445 provide foo_bar.inducts for mutually defined items foo and bar
   446 (e.g. inductive sets or datatypes).  INCOMPATIBILITY, users need to
   447 specify mutual induction rules differently, i.e. like this:
   448 
   449   (induct rule: foo_bar.inducts)
   450   (induct set: foo bar)
   451   (induct type: foo bar)
   452 
   453 The ML function ProjectRule.projections turns old-style rules into the
   454 new format.
   455 
   456 * Provers/induct: improved handling of simultaneous goals.  Instead of
   457 introducing object-level conjunction, the statement is now split into
   458 several conclusions, while the corresponding symbolic cases are
   459 nested accordingly.  INCOMPATIBILITY, proofs need to be structured
   460 explicitly.  For example:
   461 
   462   lemma
   463     fixes n :: nat
   464     shows "P n" and "Q n"
   465   proof (induct n)
   466     case 0 case 1
   467     show "P 0" sorry
   468   next
   469     case 0 case 2
   470     show "Q 0" sorry
   471   next
   472     case (Suc n) case 1
   473     note `P n` and `Q n`
   474     show "P (Suc n)" sorry
   475   next
   476     case (Suc n) case 2
   477     note `P n` and `Q n`
   478     show "Q (Suc n)" sorry
   479   qed
   480 
   481 The split into subcases may be deferred as follows -- this is
   482 particularly relevant for goal statements with local premises.
   483 
   484   lemma
   485     fixes n :: nat
   486     shows "A n ==> P n" and "B n ==> Q n"
   487   proof (induct n)
   488     case 0
   489     {
   490       case 1
   491       note `A 0`
   492       show "P 0" sorry
   493     next
   494       case 2
   495       note `B 0`
   496       show "Q 0" sorry
   497     }
   498   next
   499     case (Suc n)
   500     note `A n ==> P n` and `B n ==> Q n`
   501     {
   502       case 1
   503       note `A (Suc n)`
   504       show "P (Suc n)" sorry
   505     next
   506       case 2
   507       note `B (Suc n)`
   508       show "Q (Suc n)" sorry
   509     }
   510   qed
   511 
   512 If simultaneous goals are to be used with mutual rules, the statement
   513 needs to be structured carefully as a two-level conjunction, using
   514 lists of propositions separated by 'and':
   515 
   516   lemma
   517     shows "a : A ==> P1 a"
   518           "a : A ==> P2 a"
   519       and "b : B ==> Q1 b"
   520           "b : B ==> Q2 b"
   521           "b : B ==> Q3 b"
   522   proof (induct set: A B)
   523 
   524 * Provers/induct: support coinduction as well.  See
   525 src/HOL/Library/Coinductive_List.thy for various examples.
   526 
   527 * Attribute "symmetric" produces result with standardized schematic
   528 variables (index 0).  Potential INCOMPATIBILITY.
   529 
   530 * Simplifier: by default the simplifier trace only shows top level
   531 rewrites now. That is, trace_simp_depth_limit is set to 1 by
   532 default. Thus there is less danger of being flooded by the trace. The
   533 trace indicates where parts have been suppressed.
   534   
   535 * Provers/classical: removed obsolete classical version of elim_format
   536 attribute; classical elim/dest rules are now treated uniformly when
   537 manipulating the claset.
   538 
   539 * Provers/classical: stricter checks to ensure that supplied intro,
   540 dest and elim rules are well-formed; dest and elim rules must have at
   541 least one premise.
   542 
   543 * Provers/classical: attributes dest/elim/intro take an optional
   544 weight argument for the rule (just as the Pure versions).  Weights are
   545 ignored by automated tools, but determine the search order of single
   546 rule steps.
   547 
   548 * Syntax: input syntax now supports dummy variable binding "%_. b",
   549 where the body does not mention the bound variable.  Note that dummy
   550 patterns implicitly depend on their context of bounds, which makes
   551 "{_. _}" match any set comprehension as expected.  Potential
   552 INCOMPATIBILITY -- parse translations need to cope with syntactic
   553 constant "_idtdummy" in the binding position.
   554 
   555 * Syntax: removed obsolete syntactic constant "_K" and its associated
   556 parse translation.  INCOMPATIBILITY -- use dummy abstraction instead,
   557 for example "A -> B" => "Pi A (%_. B)".
   558 
   559 * Pure: 'class_deps' command visualizes the subclass relation, using
   560 the graph browser tool.
   561 
   562 * Pure: 'print_theory' now suppresses entities with internal name
   563 (trailing "_") by default; use '!' option for full details.
   564 
   565 
   566 *** HOL ***
   567 
   568 * Code generator library theories:
   569   * Pretty_Int represents HOL integers by big integer literals in target
   570     languages.
   571   * Pretty_Char represents HOL characters by character literals in target
   572     languages.
   573   * Pretty_Char_chr like Pretty_Char, but also offers treatment of character
   574     codes; includes Pretty_Int.
   575   * Executable_Set allows to generate code for finite sets using lists.
   576   * Executable_Rat implements rational numbers as triples (sign, enumerator,
   577     denominator).
   578   * Executable_Real implements a subset of real numbers, namly those
   579     representable by rational numbers.
   580   * Efficient_Nat implements natural numbers by integers, which in general will
   581     result in higher efficency; pattern matching with 0/Suc is eliminated;
   582     includes Pretty_Int.
   583   * ML_String provides an additional datatype ml_string; in the HOL default
   584     setup, strings in HOL are mapped to lists of HOL characters in SML; values
   585     of type ml_string are mapped to strings in SML.
   586   * ML_Int provides an additional datatype ml_int which is mapped to to SML
   587     built-in integers.
   588 
   589 * New package for inductive predicates
   590 
   591   An n-ary predicate p with m parameters z_1, ..., z_m can now be defined via
   592 
   593     inductive
   594       p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool"
   595       for z_1 :: U_1 and ... and z_n :: U_m
   596     where
   597       rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n"
   598     | ...
   599 
   600   rather than
   601 
   602     consts s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set"
   603 
   604     abbreviation p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool"
   605     where "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m"
   606 
   607     inductive "s z_1 ... z_m"
   608     intros
   609       rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m"
   610       ...
   611 
   612   For backward compatibility, there is a wrapper allowing inductive
   613   sets to be defined with the new package via
   614 
   615     inductive_set
   616       s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set"
   617       for z_1 :: U_1 and ... and z_n :: U_m
   618     where
   619       rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m"
   620     | ...
   621 
   622   or
   623 
   624     inductive_set
   625       s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set"
   626       and p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool"
   627       for z_1 :: U_1 and ... and z_n :: U_m
   628     where
   629       "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m"
   630     | rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n"
   631     | ...
   632 
   633   if the additional syntax "p ..." is required.
   634 
   635   Many examples can be found in the subdirectories Auth, Bali, Induct,
   636   or MicroJava.
   637 
   638   INCOMPATIBILITIES:
   639 
   640   - Since declaration and definition of inductive sets or predicates
   641     is no longer separated, abbreviations involving the newly introduced
   642     sets or predicates must be specified together with the introduction
   643     rules after the "where" keyword (see example above), rather than before
   644     the actual inductive definition.
   645 
   646   - The variables in induction and elimination rules are now quantified
   647     in the order of their occurrence in the introduction rules, rather than
   648     in alphabetical order. Since this may break some proofs, these proofs
   649     either have to be repaired, e.g. by reordering the variables
   650     a_i_1 ... a_i_{k_i} in Isar "case" statements of the form
   651 
   652       case (rule_i a_i_1 ... a_i_{k_i})
   653 
   654     or the old order of quantification has to be restored by explicitly adding
   655     meta-level quantifiers in the introduction rules, i.e.
   656 
   657       | rule_i: "!!a_i_1 ... a_i_{k_i}. ... ==> p z_1 ... z_m t_i_1 ... t_i_n"
   658 
   659   - The format of the elimination rules is now
   660 
   661       p z_1 ... z_m x_1 ... x_n ==>
   662         (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P)
   663         ==> ... ==> P
   664 
   665     for predicates and
   666 
   667       (x_1, ..., x_n) : s z_1 ... z_m ==>
   668         (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P)
   669         ==> ... ==> P
   670 
   671     for sets rather than
   672 
   673       x : s z_1 ... z_m ==>
   674         (!!a_1_1 ... a_1_{k_1}. x = (t_1_1, ..., t_1_n) ==> ... ==> P)
   675         ==> ... ==> P
   676 
   677     This may require terms in goals to be expanded to n-tuples (e.g. using case_tac
   678     or simplification with the split_paired_all rule) before the above elimination
   679     rule is applicable.
   680 
   681   - The elimination or case analysis rules for (mutually) inductive sets or
   682     predicates are now called "p_1.cases" ... "p_k.cases". The list of rules
   683     "p_1_..._p_k.elims" is no longer available.
   684 
   685 * Method "metis" proves goals by applying the Metis general-purpose
   686 resolution prover.  Examples are in the directory MetisExamples.  See
   687 also http://gilith.com/software/metis/
   688   
   689 * Command 'sledgehammer' invokes external automatic theorem provers as
   690 background processes.  It generates calls to the "metis" method if
   691 successful. These can be pasted into the proof.  Users do not have to
   692 wait for the automatic provers to return.
   693 
   694 * Case-expressions allow arbitrary constructor-patterns (including "_") and
   695   take their order into account, like in functional programming.
   696   Internally, this is translated into nested case-expressions; missing cases
   697   are added and mapped to the predefined constant "undefined". In complicated
   698   cases printing may no longer show the original input but the internal
   699   form. Lambda-abstractions allow the same form of pattern matching:
   700   "% pat1 => e1 | ..." is an abbreviation for
   701   "%x. case x of pat1 => e1 | ..." where x is a new variable.
   702 
   703 * IntDef: The constant "int :: nat => int" has been removed; now "int"
   704   is an abbreviation for "of_nat :: nat => int". The simplification rules
   705   for "of_nat" have been changed to work like "int" did previously.
   706   (potential INCOMPATIBILITY)
   707   - "of_nat (Suc m)" simplifies to "1 + of_nat m" instead of "of_nat m + 1"
   708   - of_nat_diff and of_nat_mult are no longer default simp rules
   709 
   710 * Method "algebra" solves polynomial equations over (semi)rings using
   711   Groebner bases. The (semi)ring structure is defined by locales and
   712   the tool setup depends on that generic context. Installing the
   713   method for a specific type involves instantiating the locale and
   714   possibly adding declarations for computation on the coefficients.
   715   The method is already instantiated for natural numbers and for the
   716   axiomatic class of idoms with numerals.  See also the paper by
   717   Chaieb and Wenzel at CALCULEMUS 2007 for the general principles
   718   underlying this architecture of context-aware proof-tools.
   719 
   720 * constant "List.op @" now named "List.append".  Use ML antiquotations
   721 @{const_name List.append} or @{term " ... @ ... "} to circumvent
   722 possible incompatibilities when working on ML level.
   723 
   724 * Constant renames due to introduction of canonical name prefixing for
   725   class package:
   726 
   727     HOL.abs ~> HOL.minus_class.abs
   728     HOL.divide ~> HOL.divide_class.divide
   729     Nat.power ~> Nat.power_class.power
   730     Nat.size ~> Nat.size_class.size
   731     Numeral.number_of ~> Numeral.number_class.number_of
   732     FixedPoint.Inf ~> FixedPoint.complete_lattice_class.Inf
   733 
   734 * Rudimentary class target mechanism involves constant renames:
   735 
   736     Orderings.min ~> Orderings.ord_class.min
   737     Orderings.max ~> Orderings.ord_class.max
   738     FixedPoint.Sup ~> FixedPoint.complete_lattice_class.Sup
   739 
   740 * primrec: missing cases mapped to "undefined" instead of "arbitrary"
   741 
   742 * new constant "undefined" with axiom "undefined x = undefined"
   743 
   744 * new class "default" with associated constant "default"
   745 
   746 * new function listsum :: 'a list => 'a for arbitrary monoids.
   747   Special syntax: "SUM x <- xs. f x" (and latex variants)
   748 
   749 * new (input only) syntax for Haskell-like list comprehension, eg
   750   [(x,y). x <- xs, y <- ys, x ~= y]
   751   For details see List.thy.
   752 
   753 * The special syntax for function "filter" has changed from [x : xs. P] to
   754   [x <- xs. P] to avoid an ambiguity caused by list comprehension syntax,
   755   and for uniformity. INCOMPATIBILITY
   756 
   757 * Lemma "set_take_whileD" renamed to "set_takeWhileD"
   758 
   759 * New lemma collection field_simps (an extension of ring_simps)
   760   for manipulating (in)equations involving division. Multiplies
   761   with all denominators that can be proved to be non-zero (in equations)
   762   or positive/negative (in inequations).
   763 
   764 * Lemma collections ring_eq_simps, group_eq_simps and ring_distrib
   765   have been improved and renamed to ring_simps, group_simps and ring_distribs.
   766   Removed lemmas field_xyz in Ring_and_Field
   767   because they were subsumed by lemmas xyz.
   768 INCOMPATIBILITY.
   769 
   770 * Library/Pretty_Int.thy: maps HOL numerals on target language integer literals
   771   when generating code.
   772 
   773 * Library/Pretty_Char.thy: maps HOL characters on target language character literals
   774   when generating code.
   775 
   776 * Library/Commutative_Ring.thy: switched from recdef to function package;
   777   constants add, mul, pow now curried.  Infix syntax for algebraic operations.
   778 
   779 * Some steps towards more uniform lattice theory development in HOL.
   780 
   781     constants "meet" and "join" now named "inf" and "sup"
   782     constant "Meet" now named "Inf"
   783 
   784     classes "meet_semilorder" and "join_semilorder" now named
   785       "lower_semilattice" and "upper_semilattice"
   786     class "lorder" now named "lattice"
   787     class "comp_lat" now named "complete_lattice"
   788 
   789     Instantiation of lattice classes allows explicit definitions
   790     for "inf" and "sup" operations.
   791 
   792   INCOMPATIBILITY.  Theorem renames:
   793 
   794     meet_left_le            ~> inf_le1
   795     meet_right_le           ~> inf_le2
   796     join_left_le            ~> sup_ge1
   797     join_right_le           ~> sup_ge2
   798     meet_join_le            ~> inf_sup_ord
   799     le_meetI                ~> le_infI
   800     join_leI                ~> le_supI
   801     le_meet                 ~> le_inf_iff
   802     le_join                 ~> ge_sup_conv
   803     meet_idempotent         ~> inf_idem
   804     join_idempotent         ~> sup_idem
   805     meet_comm               ~> inf_commute
   806     join_comm               ~> sup_commute
   807     meet_leI1               ~> le_infI1
   808     meet_leI2               ~> le_infI2
   809     le_joinI1               ~> le_supI1
   810     le_joinI2               ~> le_supI2
   811     meet_assoc              ~> inf_assoc
   812     join_assoc              ~> sup_assoc
   813     meet_left_comm          ~> inf_left_commute
   814     meet_left_idempotent    ~> inf_left_idem
   815     join_left_comm          ~> sup_left_commute
   816     join_left_idempotent    ~> sup_left_idem
   817     meet_aci                ~> inf_aci
   818     join_aci                ~> sup_aci
   819     le_def_meet             ~> le_iff_inf
   820     le_def_join             ~> le_iff_sup
   821     join_absorp2            ~> sup_absorb2
   822     join_absorp1            ~> sup_absorb1
   823     meet_absorp1            ~> inf_absorb1
   824     meet_absorp2            ~> inf_absorb2
   825     meet_join_absorp        ~> inf_sup_absorb
   826     join_meet_absorp        ~> sup_inf_absorb
   827     distrib_join_le         ~> distrib_sup_le
   828     distrib_meet_le         ~> distrib_inf_le
   829 
   830     add_meet_distrib_left   ~> add_inf_distrib_left
   831     add_join_distrib_left   ~> add_sup_distrib_left
   832     is_join_neg_meet        ~> is_join_neg_inf
   833     is_meet_neg_join        ~> is_meet_neg_sup
   834     add_meet_distrib_right  ~> add_inf_distrib_right
   835     add_join_distrib_right  ~> add_sup_distrib_right
   836     add_meet_join_distribs  ~> add_sup_inf_distribs
   837     join_eq_neg_meet        ~> sup_eq_neg_inf
   838     meet_eq_neg_join        ~> inf_eq_neg_sup
   839     add_eq_meet_join        ~> add_eq_inf_sup
   840     meet_0_imp_0            ~> inf_0_imp_0
   841     join_0_imp_0            ~> sup_0_imp_0
   842     meet_0_eq_0             ~> inf_0_eq_0
   843     join_0_eq_0             ~> sup_0_eq_0
   844     neg_meet_eq_join        ~> neg_inf_eq_sup
   845     neg_join_eq_meet        ~> neg_sup_eq_inf
   846     join_eq_if              ~> sup_eq_if
   847 
   848     mono_meet               ~> mono_inf
   849     mono_join               ~> mono_sup
   850     meet_bool_eq            ~> inf_bool_eq
   851     join_bool_eq            ~> sup_bool_eq
   852     meet_fun_eq             ~> inf_fun_eq
   853     join_fun_eq             ~> sup_fun_eq
   854     meet_set_eq             ~> inf_set_eq
   855     join_set_eq             ~> sup_set_eq
   856     meet1_iff               ~> inf1_iff
   857     meet2_iff               ~> inf2_iff
   858     meet1I                  ~> inf1I
   859     meet2I                  ~> inf2I
   860     meet1D1                 ~> inf1D1
   861     meet2D1                 ~> inf2D1
   862     meet1D2                 ~> inf1D2
   863     meet2D2                 ~> inf2D2
   864     meet1E                  ~> inf1E
   865     meet2E                  ~> inf2E
   866     join1_iff               ~> sup1_iff
   867     join2_iff               ~> sup2_iff
   868     join1I1                 ~> sup1I1
   869     join2I1                 ~> sup2I1
   870     join1I1                 ~> sup1I1
   871     join2I2                 ~> sup1I2
   872     join1CI                 ~> sup1CI
   873     join2CI                 ~> sup2CI
   874     join1E                  ~> sup1E
   875     join2E                  ~> sup2E
   876 
   877     is_meet_Meet            ~> is_meet_Inf
   878     Meet_bool_def           ~> Inf_bool_def
   879     Meet_fun_def            ~> Inf_fun_def
   880     Meet_greatest           ~> Inf_greatest
   881     Meet_lower              ~> Inf_lower
   882     Meet_set_def            ~> Inf_set_def
   883 
   884     listsp_meetI            ~> listsp_infI
   885     listsp_meet_eq          ~> listsp_inf_eq
   886 
   887     meet_min                ~> inf_min
   888     join_max                ~> sup_max
   889 
   890 * Classes "order" and "linorder": facts "refl", "trans" and
   891 "cases" renamed ro "order_refl", "order_trans" and "linorder_cases", to
   892 avoid clashes with HOL "refl" and "trans". INCOMPATIBILITY.
   893 
   894 * Classes "order" and "linorder": 
   895 potential INCOMPATIBILITY: order of proof goals in order/linorder instance
   896 proofs changed.
   897 
   898 * Dropped lemma duplicate def_imp_eq in favor of meta_eq_to_obj_eq.
   899 INCOMPATIBILITY.
   900 
   901 * Dropped lemma duplicate if_def2 in favor of if_bool_eq_conj.
   902 INCOMPATIBILITY.
   903 
   904 * Added syntactic class "size"; overloaded constant "size" now has
   905 type "'a::size ==> bool"
   906 
   907 * Renamed constants "Divides.op div", "Divides.op mod" and "Divides.op
   908 dvd" to "Divides.div_class.div", "Divides.div_class.mod" and "Divides.dvd". INCOMPATIBILITY.
   909 
   910 * Added method "lexicographic_order" automatically synthesizes
   911 termination relations as lexicographic combinations of size measures
   912 -- 'function' package.
   913 
   914 * HOL/records: generalised field-update to take a function on the
   915 field rather than the new value: r(|A := x|) is translated to A_update
   916 (K x) r The K-combinator that is internally used is called K_record.
   917 INCOMPATIBILITY: Usage of the plain update functions has to be
   918 adapted.
   919  
   920 * axclass "semiring_0" now contains annihilation axioms x * 0 = 0 and
   921 0 * x = 0, which are required for a semiring.  Richer structures do
   922 not inherit from semiring_0 anymore, because this property is a
   923 theorem there, not an axiom.  INCOMPATIBILITY: In instances of
   924 semiring_0, there is more to prove, but this is mostly trivial.
   925 
   926 * axclass "recpower" was generalized to arbitrary monoids, not just
   927 commutative semirings.  INCOMPATIBILITY: If you use recpower and need
   928 commutativity or a semiring property, add the corresponding classes.
   929 
   930 * Unified locale partial_order with class definition (cf. theory
   931 Orderings), added parameter ``less''.  INCOMPATIBILITY.
   932 
   933 * Constant "List.list_all2" in List.thy now uses authentic syntax.
   934 INCOMPATIBILITY: translations containing list_all2 may go wrong.  On
   935 Isar level, use abbreviations instead.
   936 
   937 * Renamed constant "List.op mem" to "List.memberl" INCOMPATIBILITY:
   938 rarely occuring name references (e.g. ``List.op mem.simps'') require
   939 renaming (e.g. ``List.memberl.simps'').
   940 
   941 * Renamed constants "0" to "HOL.zero_class.zero" and "1" to "HOL.one_class.one".
   942 INCOMPATIBILITY.
   943 
   944 * Added class "HOL.eq", allowing for code generation with polymorphic equality.
   945 
   946 * Numeral syntax: type 'bin' which was a mere type copy of 'int' has
   947 been abandoned in favour of plain 'int'. INCOMPATIBILITY --
   948 significant changes for setting up numeral syntax for types:
   949 
   950   - new constants Numeral.pred and Numeral.succ instead
   951       of former Numeral.bin_pred and Numeral.bin_succ.
   952   - Use integer operations instead of bin_add, bin_mult and so on.
   953   - Numeral simplification theorems named Numeral.numeral_simps instead of Bin_simps.
   954   - ML structure Bin_Simprocs now named Int_Numeral_Base_Simprocs.
   955 
   956 See HOL/Integ/IntArith.thy for an example setup.
   957 
   958 * New top level command 'normal_form' computes the normal form of a
   959 term that may contain free variables. For example ``normal_form
   960 "rev[a,b,c]"'' produces ``[b,c,a]'' (without proof).  This command is
   961 suitable for heavy-duty computations because the functions are
   962 compiled to ML first.
   963 
   964 * Alternative iff syntax "A <-> B" for equality on bool (with priority
   965 25 like -->); output depends on the "iff" print_mode, the default is
   966 "A = B" (with priority 50).
   967 
   968 * Renamed constants in HOL.thy and Orderings.thy:
   969     op +   ~> HOL.plus_class.plus
   970     op -   ~> HOL.minus_class.minus
   971     uminus ~> HOL.minus_class.uminus
   972     abs    ~> HOL.abs_class.abs
   973     op *   ~> HOL.times_class.times
   974     op <   ~> HOL.ord_class.less
   975     op <=  ~> HOL.ord_class.less_eq
   976 
   977 Adaptions may be required in the following cases:
   978 
   979 a) User-defined constants using any of the names "plus", "minus", "times",
   980 "less" or "less_eq". The standard syntax translations for "+", "-" and "*"
   981 may go wrong.
   982 INCOMPATIBILITY: use more specific names.
   983 
   984 b) Variables named "plus", "minus", "times", "less", "less_eq"
   985 INCOMPATIBILITY: use more specific names.
   986 
   987 c) Permutative equations (e.g. "a + b = b + a")
   988 Since the change of names also changes the order of terms, permutative
   989 rewrite rules may get applied in a different order. Experience shows that
   990 this is rarely the case (only two adaptions in the whole Isabelle
   991 distribution).
   992 INCOMPATIBILITY: rewrite proofs
   993 
   994 d) ML code directly refering to constant names
   995 This in general only affects hand-written proof tactics, simprocs and so on.
   996 INCOMPATIBILITY: grep your sourcecode and replace names.  Consider use
   997 of const_name ML antiquotations.
   998 
   999 * Relations less (<) and less_eq (<=) are also available on type bool.
  1000 Modified syntax to disallow nesting without explicit parentheses,
  1001 e.g. "(x < y) < z" or "x < (y < z)", but NOT "x < y < z".
  1002 
  1003 * "LEAST x:A. P" expands to "LEAST x. x:A & P" (input only).
  1004 
  1005 * Relation composition operator "op O" now has precedence 75 and binds
  1006 stronger than union and intersection. INCOMPATIBILITY.
  1007 
  1008 * The old set interval syntax "{m..n(}" (and relatives) has been
  1009 removed.  Use "{m..<n}" (and relatives) instead.
  1010 
  1011 * In the context of the assumption "~(s = t)" the Simplifier rewrites
  1012 "t = s" to False (by simproc "neq_simproc").  For backward
  1013 compatibility this can be disabled by ML "reset use_neq_simproc".
  1014 
  1015 * "m dvd n" where m and n are numbers is evaluated to True/False by
  1016 simp.
  1017 
  1018 * Theorem Cons_eq_map_conv no longer declared as ``simp''.
  1019 
  1020 * Theorem setsum_mult renamed to setsum_right_distrib.
  1021 
  1022 * Prefer ex1I over ex_ex1I in single-step reasoning, e.g. by the
  1023 ``rule'' method.
  1024 
  1025 * Reimplemented methods ``sat'' and ``satx'', with several
  1026 improvements: goals no longer need to be stated as "<prems> ==>
  1027 False", equivalences (i.e. "=" on type bool) are handled, variable
  1028 names of the form "lit_<n>" are no longer reserved, significant
  1029 speedup.
  1030 
  1031 * Methods ``sat'' and ``satx'' can now replay MiniSat proof traces.
  1032 zChaff is still supported as well.
  1033 
  1034 * 'inductive' and 'datatype': provide projections of mutual rules,
  1035 bundled as foo_bar.inducts;
  1036 
  1037 * Library: moved theories Parity, GCD, Binomial, Infinite_Set to
  1038 Library.
  1039 
  1040 * Library: moved theory Accessible_Part to main HOL.
  1041 
  1042 * Library: added theory Coinductive_List of potentially infinite lists
  1043 as greatest fixed-point.
  1044 
  1045 * Library: added theory AssocList which implements (finite) maps as
  1046 association lists.
  1047 
  1048 * Added proof method ``evaluation'' for efficiently solving a goal
  1049 (i.e. a boolean expression) by compiling it to ML. The goal is
  1050 "proved" (via an oracle) if it evaluates to True.
  1051 
  1052 * Linear arithmetic now splits certain operators (e.g. min, max, abs)
  1053 also when invoked by the simplifier.  This results in the simplifier
  1054 being more powerful on arithmetic goals.  INCOMPATIBILITY.  Set
  1055 fast_arith_split_limit to 0 to obtain the old behavior.
  1056 
  1057 * Support for hex (0x20) and binary (0b1001) numerals.
  1058 
  1059 * New method: reify eqs (t), where eqs are equations for an
  1060 interpretation I :: 'a list => 'b => 'c and t::'c is an optional
  1061 parameter, computes a term s::'b and a list xs::'a list and proves the
  1062 theorem I xs s = t. This is also known as reification or quoting. The
  1063 resulting theorem is applied to the subgoal to substitute t with I xs
  1064 s.  If t is omitted, the subgoal itself is reified.
  1065 
  1066 * New method: reflection corr_thm eqs (t). The parameters eqs and (t)
  1067 are as explained above. corr_thm is a theorem for I vs (f t) = I vs t,
  1068 where f is supposed to be a computable function (in the sense of code
  1069 generattion). The method uses reify to compute s and xs as above then
  1070 applies corr_thm and uses normalization by evaluation to "prove" f s =
  1071 r and finally gets the theorem t = r, which is again applied to the
  1072 subgoal. An Example is available in HOL/ex/ReflectionEx.thy.
  1073 
  1074 * Reflection: Automatic reification now handels binding, an example
  1075 is available in HOL/ex/ReflectionEx.thy
  1076 
  1077 
  1078 *** HOL-Algebra ***
  1079 
  1080 * Formalisation of ideals and the quotient construction over rings.
  1081 
  1082 * Order and lattice theory no longer based on records.
  1083 INCOMPATIBILITY.
  1084 
  1085 * Renamed lemmas least_carrier -> least_closed and greatest_carrier ->
  1086 greatest_closed.  INCOMPATIBILITY.
  1087 
  1088 * Method algebra is now set up via an attribute.  For examples see
  1089 Ring.thy.  INCOMPATIBILITY: the method is now weaker on combinations
  1090 of algebraic structures.
  1091 
  1092 * Renamed theory CRing to Ring.
  1093 
  1094 
  1095 *** HOL-Complex ***
  1096 
  1097 * Theory Real: new method ferrack implements quantifier elimination
  1098 for linear arithmetic over the reals. The quantifier elimination
  1099 feature is used only for decision, for compatibility with arith. This
  1100 means a goal is either solved or left unchanged, no simplification.
  1101 
  1102 * Hyperreal: Functions root and sqrt are now defined on negative real
  1103 inputs so that root n (- x) = - root n x and sqrt (- x) = - sqrt x.
  1104 Nonnegativity side conditions have been removed from many lemmas, so
  1105 that more subgoals may now be solved by simplification; potential
  1106 INCOMPATIBILITY.
  1107 
  1108 * Real: New axiomatic classes formalize real normed vector spaces and
  1109 algebras, using new overloaded constants scaleR :: real => 'a => 'a
  1110 and norm :: 'a => real.
  1111 
  1112 * Real: New constant of_real :: real => 'a::real_algebra_1 injects
  1113 from reals into other types. The overloaded constant Reals :: 'a set
  1114 is now defined as range of_real; potential INCOMPATIBILITY.
  1115 
  1116 * Real: ML code generation is supported now and hence also quickcheck.
  1117 Reals are implemented as arbitrary precision rationals.
  1118 
  1119 * Hyperreal: Several constants that previously worked only for the
  1120 reals have been generalized, so they now work over arbitrary vector
  1121 spaces. Type annotations may need to be added in some cases; potential
  1122 INCOMPATIBILITY.
  1123 
  1124   Infinitesimal  :: ('a::real_normed_vector) star set
  1125   HFinite        :: ('a::real_normed_vector) star set
  1126   HInfinite      :: ('a::real_normed_vector) star set
  1127   approx         :: ('a::real_normed_vector) star => 'a star => bool
  1128   monad          :: ('a::real_normed_vector) star => 'a star set
  1129   galaxy         :: ('a::real_normed_vector) star => 'a star set
  1130   (NS)LIMSEQ     :: [nat => 'a::real_normed_vector, 'a] => bool
  1131   (NS)convergent :: (nat => 'a::real_normed_vector) => bool
  1132   (NS)Bseq       :: (nat => 'a::real_normed_vector) => bool
  1133   (NS)Cauchy     :: (nat => 'a::real_normed_vector) => bool
  1134   (NS)LIM        :: ['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool
  1135   is(NS)Cont     :: ['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool
  1136   deriv          :: ['a::real_normed_field => 'a, 'a, 'a] => bool
  1137   sgn            :: 'a::real_normed_vector => 'a
  1138   exp            :: 'a::{recpower,real_normed_field,banach} => 'a
  1139 
  1140 * Complex: Some complex-specific constants are now abbreviations for
  1141 overloaded ones: complex_of_real = of_real, cmod = norm, hcmod =
  1142 hnorm.  Other constants have been entirely removed in favor of the
  1143 polymorphic versions (INCOMPATIBILITY):
  1144 
  1145   approx        <-- capprox
  1146   HFinite       <-- CFinite
  1147   HInfinite     <-- CInfinite
  1148   Infinitesimal <-- CInfinitesimal
  1149   monad         <-- cmonad
  1150   galaxy        <-- cgalaxy
  1151   (NS)LIM       <-- (NS)CLIM, (NS)CRLIM
  1152   is(NS)Cont    <-- is(NS)Contc, is(NS)contCR
  1153   (ns)deriv     <-- (ns)cderiv
  1154 
  1155 
  1156 *** ML ***
  1157 
  1158 * Generic arithmetic modules: Tools/integer.ML, Tools/rat.ML, Tools/float.ML
  1159 
  1160 * Context data interfaces (Theory/Proof/GenericDataFun): removed
  1161 name/print, uninitialized data defaults to ad-hoc copy of empty value,
  1162 init only required for impure data.  INCOMPATIBILITY: empty really
  1163 need to be empty (no dependencies on theory content!)
  1164 
  1165 * ML within Isar: antiquotations allow to embed statically-checked
  1166 formal entities in the source, referring to the context available at
  1167 compile-time.  For example:
  1168 
  1169 ML {* @{typ "'a => 'b"} *}
  1170 ML {* @{term "%x. x"} *}
  1171 ML {* @{prop "x == y"} *}
  1172 ML {* @{ctyp "'a => 'b"} *}
  1173 ML {* @{cterm "%x. x"} *}
  1174 ML {* @{cprop "x == y"} *}
  1175 ML {* @{thm asm_rl} *}
  1176 ML {* @{thms asm_rl} *}
  1177 ML {* @{const_name c} *}
  1178 ML {* @{const_syntax c} *}
  1179 ML {* @{context} *}
  1180 ML {* @{theory} *}
  1181 ML {* @{theory Pure} *}
  1182 ML {* @{simpset} *}
  1183 ML {* @{claset} *}
  1184 ML {* @{clasimpset} *}
  1185 
  1186 The same works for sources being ``used'' within an Isar context.
  1187 
  1188 * ML in Isar: improved error reporting; extra verbosity with
  1189 Toplevel.debug enabled.
  1190 
  1191 * Pure/library:
  1192 
  1193   val burrow: ('a list -> 'b list) -> 'a list list -> 'b list list
  1194   val fold_burrow: ('a list -> 'c -> 'b list * 'd) -> 'a list list -> 'c -> 'b list list * 'd
  1195 
  1196 The semantics of "burrow" is: "take a function with *simulatanously*
  1197 transforms a list of value, and apply it *simulatanously* to a list of
  1198 list of values of the appropriate type". Compare this with "map" which
  1199 would *not* apply its argument function simulatanously but in
  1200 sequence; "fold_burrow" has an additional context.
  1201 
  1202 * Pure/library: functions map2 and fold2 with curried syntax for
  1203 simultanous mapping and folding:
  1204 
  1205     val map2: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
  1206     val fold2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
  1207 
  1208 * Pure/library: indexed lists - some functions in the Isabelle library
  1209 treating lists over 'a as finite mappings from [0...n] to 'a have been
  1210 given more convenient names and signatures reminiscent of similar
  1211 functions for alists, tables, etc:
  1212 
  1213   val nth: 'a list -> int -> 'a 
  1214   val nth_map: int -> ('a -> 'a) -> 'a list -> 'a list
  1215   val fold_index: (int * 'a -> 'b -> 'b) -> 'a list -> 'b -> 'b
  1216 
  1217 Note that fold_index starts counting at index 0, not 1 like foldln
  1218 used to.
  1219 
  1220 * Pure/library: added general ``divide_and_conquer'' combinator on
  1221 lists.
  1222 
  1223 * Pure/General/table.ML: the join operations now works via exceptions
  1224 DUP/SAME instead of type option.  This is simpler in simple cases, and
  1225 admits slightly more efficient complex applications.
  1226 
  1227 * Pure: datatype Context.generic joins theory/Proof.context and
  1228 provides some facilities for code that works in either kind of
  1229 context, notably GenericDataFun for uniform theory and proof data.
  1230 
  1231 * Pure: 'advanced' translation functions (parse_translation etc.) now
  1232 use Context.generic instead of just theory.
  1233 
  1234 * Pure: simplified internal attribute type, which is now always
  1235 Context.generic * thm -> Context.generic * thm.  Global (theory)
  1236 vs. local (Proof.context) attributes have been discontinued, while
  1237 minimizing code duplication.  Thm.rule_attribute and
  1238 Thm.declaration_attribute build canonical attributes; see also
  1239 structure Context for further operations on Context.generic, notably
  1240 GenericDataFun.  INCOMPATIBILITY, need to adapt attribute type
  1241 declarations and definitions.
  1242 
  1243 * Pure/kernel: consts certification ignores sort constraints given in
  1244 signature declarations.  (This information is not relevant to the
  1245 logic, but only for type inference.)  IMPORTANT INTERNAL CHANGE,
  1246 potential INCOMPATIBILITY.
  1247 
  1248 * Pure: axiomatic type classes are now purely definitional, with
  1249 explicit proofs of class axioms and super class relations performed
  1250 internally.  See Pure/axclass.ML for the main internal interfaces --
  1251 notably AxClass.define_class supercedes AxClass.add_axclass, and
  1252 AxClass.axiomatize_class/classrel/arity supercede
  1253 Sign.add_classes/classrel/arities.
  1254 
  1255 * Pure/Isar: Args/Attrib parsers operate on Context.generic --
  1256 global/local versions on theory vs. Proof.context have been
  1257 discontinued; Attrib.syntax and Method.syntax have been adapted
  1258 accordingly.  INCOMPATIBILITY, need to adapt parser expressions for
  1259 attributes, methods, etc.
  1260 
  1261 * Pure: several functions of signature "... -> theory -> theory * ..."
  1262 have been reoriented to "... -> theory -> ... * theory" in order to
  1263 allow natural usage in combination with the ||>, ||>>, |-> and
  1264 fold_map combinators.
  1265 
  1266 * Pure: official theorem names (closed derivations) and additional
  1267 comments (tags) are now strictly separate.  Name hints -- which are
  1268 maintained as tags -- may be attached any time without affecting the
  1269 derivation.
  1270 
  1271 * Pure: primitive rule lift_rule now takes goal cterm instead of an
  1272 actual goal state (thm).  Use Thm.lift_rule (Thm.cprem_of st i) to
  1273 achieve the old behaviour.
  1274 
  1275 * Pure: the "Goal" constant is now called "prop", supporting a
  1276 slightly more general idea of ``protecting'' meta-level rule
  1277 statements.
  1278 
  1279 * Pure: Logic.(un)varify only works in a global context, which is now
  1280 enforced instead of silently assumed.  INCOMPATIBILITY, may use
  1281 Logic.legacy_(un)varify as temporary workaround.
  1282 
  1283 * Pure: structure Name provides scalable operations for generating
  1284 internal variable names, notably Name.variants etc.  This replaces
  1285 some popular functions from term.ML:
  1286 
  1287   Term.variant		->  Name.variant
  1288   Term.variantlist	->  Name.variant_list  (*canonical argument order*)
  1289   Term.invent_names	->  Name.invent_list
  1290 
  1291 Note that low-level renaming rarely occurs in new code -- operations
  1292 from structure Variable are used instead (see below).
  1293 
  1294 * Pure: structure Variable provides fundamental operations for proper
  1295 treatment of fixed/schematic variables in a context.  For example,
  1296 Variable.import introduces fixes for schematics of given facts and
  1297 Variable.export reverses the effect (up to renaming) -- this replaces
  1298 various freeze_thaw operations.
  1299 
  1300 * Pure: structure Goal provides simple interfaces for
  1301 init/conclude/finish and tactical prove operations (replacing former
  1302 Tactic.prove).  Goal.prove is the canonical way to prove results
  1303 within a given context; Goal.prove_global is a degraded version for
  1304 theory level goals, including a global Drule.standard.  Note that
  1305 OldGoals.prove_goalw_cterm has long been obsolete, since it is
  1306 ill-behaved in a local proof context (e.g. with local fixes/assumes or
  1307 in a locale context).
  1308 
  1309 * Isar: simplified treatment of user-level errors, using exception
  1310 ERROR of string uniformly.  Function error now merely raises ERROR,
  1311 without any side effect on output channels.  The Isar toplevel takes
  1312 care of proper display of ERROR exceptions.  ML code may use plain
  1313 handle/can/try; cat_error may be used to concatenate errors like this:
  1314 
  1315   ... handle ERROR msg => cat_error msg "..."
  1316 
  1317 Toplevel ML code (run directly or through the Isar toplevel) may be
  1318 embedded into the Isar toplevel with exception display/debug like
  1319 this:
  1320 
  1321   Isar.toplevel (fn () => ...)
  1322 
  1323 INCOMPATIBILITY, removed special transform_error facilities, removed
  1324 obsolete variants of user-level exceptions (ERROR_MESSAGE,
  1325 Context.PROOF, ProofContext.CONTEXT, Proof.STATE, ProofHistory.FAIL)
  1326 -- use plain ERROR instead.
  1327 
  1328 * Isar: theory setup now has type (theory -> theory), instead of a
  1329 list.  INCOMPATIBILITY, may use #> to compose setup functions.
  1330 
  1331 * Isar: installed ML toplevel pretty printer for type Proof.context,
  1332 subject to ProofContext.debug/verbose flags.
  1333 
  1334 * Isar: Toplevel.theory_to_proof admits transactions that modify the
  1335 theory before entering a proof state.  Transactions now always see a
  1336 quasi-functional intermediate checkpoint, both in interactive and
  1337 batch mode.
  1338 
  1339 * Simplifier: the simpset of a running simplification process now
  1340 contains a proof context (cf. Simplifier.the_context), which is the
  1341 very context that the initial simpset has been retrieved from (by
  1342 simpset_of/local_simpset_of).  Consequently, all plug-in components
  1343 (solver, looper etc.) may depend on arbitrary proof data.
  1344 
  1345 * Simplifier.inherit_context inherits the proof context (plus the
  1346 local bounds) of the current simplification process; any simproc
  1347 etc. that calls the Simplifier recursively should do this!  Removed
  1348 former Simplifier.inherit_bounds, which is already included here --
  1349 INCOMPATIBILITY.  Tools based on low-level rewriting may even have to
  1350 specify an explicit context using Simplifier.context/theory_context.
  1351 
  1352 * Simplifier/Classical Reasoner: more abstract interfaces
  1353 change_simpset/claset for modifying the simpset/claset reference of a
  1354 theory; raw versions simpset/claset_ref etc. have been discontinued --
  1355 INCOMPATIBILITY.
  1356 
  1357 * Provers: more generic wrt. syntax of object-logics, avoid hardwired
  1358 "Trueprop" etc.
  1359 
  1360 
  1361 *** System ***
  1362 
  1363 * settings: ML_IDENTIFIER -- which is appended to user specific heap
  1364 locations -- now includes the Isabelle version identifier as well.
  1365 This simplifies use of multiple Isabelle installations.
  1366 
  1367 * isabelle-process: option -S (secure mode) disables some critical
  1368 operations, notably runtime compilation and evaluation of ML source
  1369 code.
  1370 
  1371 
  1372 New in Isabelle2005 (October 2005)
  1373 ----------------------------------
  1374 
  1375 *** General ***
  1376 
  1377 * Theory headers: the new header syntax for Isar theories is
  1378 
  1379   theory <name>
  1380   imports <theory1> ... <theoryN>
  1381   uses <file1> ... <fileM>
  1382   begin
  1383 
  1384 where the 'uses' part is optional.  The previous syntax
  1385 
  1386   theory <name> = <theory1> + ... + <theoryN>:
  1387 
  1388 will disappear in the next release.  Use isatool fixheaders to convert
  1389 existing theory files.  Note that there is no change in ancient
  1390 non-Isar theories now, but these will disappear soon.
  1391 
  1392 * Theory loader: parent theories can now also be referred to via
  1393 relative and absolute paths.
  1394 
  1395 * Command 'find_theorems' searches for a list of criteria instead of a
  1396 list of constants. Known criteria are: intro, elim, dest, name:string,
  1397 simp:term, and any term. Criteria can be preceded by '-' to select
  1398 theorems that do not match. Intro, elim, dest select theorems that
  1399 match the current goal, name:s selects theorems whose fully qualified
  1400 name contain s, and simp:term selects all simplification rules whose
  1401 lhs match term.  Any other term is interpreted as pattern and selects
  1402 all theorems matching the pattern. Available in ProofGeneral under
  1403 'ProofGeneral -> Find Theorems' or C-c C-f.  Example:
  1404 
  1405   C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL."
  1406 
  1407 prints the last 100 theorems matching the pattern "(_::nat) + _ + _",
  1408 matching the current goal as introduction rule and not having "HOL."
  1409 in their name (i.e. not being defined in theory HOL).
  1410 
  1411 * Command 'thms_containing' has been discontinued in favour of
  1412 'find_theorems'; INCOMPATIBILITY.
  1413 
  1414 * Communication with Proof General is now 8bit clean, which means that
  1415 Unicode text in UTF-8 encoding may be used within theory texts (both
  1416 formal and informal parts).  Cf. option -U of the Isabelle Proof
  1417 General interface.  Here are some simple examples (cf. src/HOL/ex):
  1418 
  1419   http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html
  1420   http://isabelle.in.tum.de/library/HOL/ex/Chinese.html
  1421 
  1422 * Improved efficiency of the Simplifier and, to a lesser degree, the
  1423 Classical Reasoner.  Typical big applications run around 2 times
  1424 faster.
  1425 
  1426 
  1427 *** Document preparation ***
  1428 
  1429 * Commands 'display_drafts' and 'print_drafts' perform simple output
  1430 of raw sources.  Only those symbols that do not require additional
  1431 LaTeX packages (depending on comments in isabellesym.sty) are
  1432 displayed properly, everything else is left verbatim.  isatool display
  1433 and isatool print are used as front ends (these are subject to the
  1434 DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively).
  1435 
  1436 * Command tags control specific markup of certain regions of text,
  1437 notably folding and hiding.  Predefined tags include "theory" (for
  1438 theory begin and end), "proof" for proof commands, and "ML" for
  1439 commands involving ML code; the additional tags "visible" and
  1440 "invisible" are unused by default.  Users may give explicit tag
  1441 specifications in the text, e.g. ''by %invisible (auto)''.  The
  1442 interpretation of tags is determined by the LaTeX job during document
  1443 preparation: see option -V of isatool usedir, or options -n and -t of
  1444 isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag,
  1445 \isadroptag.
  1446 
  1447 Several document versions may be produced at the same time via isatool
  1448 usedir (the generated index.html will link all of them).  Typical
  1449 specifications include ''-V document=theory,proof,ML'' to present
  1450 theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold
  1451 proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit
  1452 these parts without any formal replacement text.  The Isabelle site
  1453 default settings produce ''document'' and ''outline'' versions as
  1454 specified above.
  1455 
  1456 * Several new antiquotations:
  1457 
  1458   @{term_type term} prints a term with its type annotated;
  1459 
  1460   @{typeof term} prints the type of a term;
  1461 
  1462   @{const const} is the same as @{term const}, but checks that the
  1463   argument is a known logical constant;
  1464 
  1465   @{term_style style term} and @{thm_style style thm} print a term or
  1466   theorem applying a "style" to it
  1467 
  1468   @{ML text}
  1469 
  1470 Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of
  1471 definitions, equations, inequations etc., 'concl' printing only the
  1472 conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19'
  1473 to print the specified premise.  TermStyle.add_style provides an ML
  1474 interface for introducing further styles.  See also the "LaTeX Sugar"
  1475 document practical applications.  The ML antiquotation prints
  1476 type-checked ML expressions verbatim.
  1477 
  1478 * Markup commands 'chapter', 'section', 'subsection', 'subsubsection',
  1479 and 'text' support optional locale specification '(in loc)', which
  1480 specifies the default context for interpreting antiquotations.  For
  1481 example: 'text (in lattice) {* @{thm inf_assoc}*}'.
  1482 
  1483 * Option 'locale=NAME' of antiquotations specifies an alternative
  1484 context interpreting the subsequent argument.  For example: @{thm
  1485 [locale=lattice] inf_assoc}.
  1486 
  1487 * Proper output of proof terms (@{prf ...} and @{full_prf ...}) within
  1488 a proof context.
  1489 
  1490 * Proper output of antiquotations for theory commands involving a
  1491 proof context (such as 'locale' or 'theorem (in loc) ...').
  1492 
  1493 * Delimiters of outer tokens (string etc.) now produce separate LaTeX
  1494 macros (\isachardoublequoteopen, isachardoublequoteclose etc.).
  1495 
  1496 * isatool usedir: new option -C (default true) controls whether option
  1497 -D should include a copy of the original document directory; -C false
  1498 prevents unwanted effects such as copying of administrative CVS data.
  1499 
  1500 
  1501 *** Pure ***
  1502 
  1503 * Considerably improved version of 'constdefs' command.  Now performs
  1504 automatic type-inference of declared constants; additional support for
  1505 local structure declarations (cf. locales and HOL records), see also
  1506 isar-ref manual.  Potential INCOMPATIBILITY: need to observe strictly
  1507 sequential dependencies of definitions within a single 'constdefs'
  1508 section; moreover, the declared name needs to be an identifier.  If
  1509 all fails, consider to fall back on 'consts' and 'defs' separately.
  1510 
  1511 * Improved indexed syntax and implicit structures.  First of all,
  1512 indexed syntax provides a notational device for subscripted
  1513 application, using the new syntax \<^bsub>term\<^esub> for arbitrary
  1514 expressions.  Secondly, in a local context with structure
  1515 declarations, number indexes \<^sub>n or the empty index (default
  1516 number 1) refer to a certain fixed variable implicitly; option
  1517 show_structs controls printing of implicit structures.  Typical
  1518 applications of these concepts involve record types and locales.
  1519 
  1520 * New command 'no_syntax' removes grammar declarations (and
  1521 translations) resulting from the given syntax specification, which is
  1522 interpreted in the same manner as for the 'syntax' command.
  1523 
  1524 * 'Advanced' translation functions (parse_translation etc.) may depend
  1525 on the signature of the theory context being presently used for
  1526 parsing/printing, see also isar-ref manual.
  1527 
  1528 * Improved 'oracle' command provides a type-safe interface to turn an
  1529 ML expression of type theory -> T -> term into a primitive rule of
  1530 type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle
  1531 is already included here); see also FOL/ex/IffExample.thy;
  1532 INCOMPATIBILITY.
  1533 
  1534 * axclass: name space prefix for class "c" is now "c_class" (was "c"
  1535 before); "cI" is no longer bound, use "c.intro" instead.
  1536 INCOMPATIBILITY.  This change avoids clashes of fact bindings for
  1537 axclasses vs. locales.
  1538 
  1539 * Improved internal renaming of symbolic identifiers -- attach primes
  1540 instead of base 26 numbers.
  1541 
  1542 * New flag show_question_marks controls printing of leading question
  1543 marks in schematic variable names.
  1544 
  1545 * In schematic variable names, *any* symbol following \<^isub> or
  1546 \<^isup> is now treated as part of the base name.  For example, the
  1547 following works without printing of awkward ".0" indexes:
  1548 
  1549   lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1"
  1550     by simp
  1551 
  1552 * Inner syntax includes (*(*nested*) comments*).
  1553 
  1554 * Pretty printer now supports unbreakable blocks, specified in mixfix
  1555 annotations as "(00...)".
  1556 
  1557 * Clear separation of logical types and nonterminals, where the latter
  1558 may only occur in 'syntax' specifications or type abbreviations.
  1559 Before that distinction was only partially implemented via type class
  1560 "logic" vs. "{}".  Potential INCOMPATIBILITY in rare cases of improper
  1561 use of 'types'/'consts' instead of 'nonterminals'/'syntax'.  Some very
  1562 exotic syntax specifications may require further adaption
  1563 (e.g. Cube/Cube.thy).
  1564 
  1565 * Removed obsolete type class "logic", use the top sort {} instead.
  1566 Note that non-logical types should be declared as 'nonterminals'
  1567 rather than 'types'.  INCOMPATIBILITY for new object-logic
  1568 specifications.
  1569 
  1570 * Attributes 'induct' and 'cases': type or set names may now be
  1571 locally fixed variables as well.
  1572 
  1573 * Simplifier: can now control the depth to which conditional rewriting
  1574 is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth
  1575 Limit.
  1576 
  1577 * Simplifier: simplification procedures may now take the current
  1578 simpset into account (cf. Simplifier.simproc(_i) / mk_simproc
  1579 interface), which is very useful for calling the Simplifier
  1580 recursively.  Minor INCOMPATIBILITY: the 'prems' argument of simprocs
  1581 is gone -- use prems_of_ss on the simpset instead.  Moreover, the
  1582 low-level mk_simproc no longer applies Logic.varify internally, to
  1583 allow for use in a context of fixed variables.
  1584 
  1585 * thin_tac now works even if the assumption being deleted contains !!
  1586 or ==>.  More generally, erule now works even if the major premise of
  1587 the elimination rule contains !! or ==>.
  1588 
  1589 * Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY.
  1590 
  1591 * Reorganized bootstrapping of the Pure theories; CPure is now derived
  1592 from Pure, which contains all common declarations already.  Both
  1593 theories are defined via plain Isabelle/Isar .thy files.
  1594 INCOMPATIBILITY: elements of CPure (such as the CPure.intro /
  1595 CPure.elim / CPure.dest attributes) now appear in the Pure name space;
  1596 use isatool fixcpure to adapt your theory and ML sources.
  1597 
  1598 * New syntax 'name(i-j, i-, i, ...)' for referring to specific
  1599 selections of theorems in named facts via index ranges.
  1600 
  1601 * 'print_theorems': in theory mode, really print the difference
  1602 wrt. the last state (works for interactive theory development only),
  1603 in proof mode print all local facts (cf. 'print_facts');
  1604 
  1605 * 'hide': option '(open)' hides only base names.
  1606 
  1607 * More efficient treatment of intermediate checkpoints in interactive
  1608 theory development.
  1609 
  1610 * Code generator is now invoked via code_module (incremental code
  1611 generation) and code_library (modular code generation, ML structures
  1612 for each theory).  INCOMPATIBILITY: new keywords 'file' and 'contains'
  1613 must be quoted when used as identifiers.
  1614 
  1615 * New 'value' command for reading, evaluating and printing terms using
  1616 the code generator.  INCOMPATIBILITY: command keyword 'value' must be
  1617 quoted when used as identifier.
  1618 
  1619 
  1620 *** Locales ***
  1621 
  1622 * New commands for the interpretation of locale expressions in
  1623 theories (1), locales (2) and proof contexts (3).  These generate
  1624 proof obligations from the expression specification.  After the
  1625 obligations have been discharged, theorems of the expression are added
  1626 to the theory, target locale or proof context.  The synopsis of the
  1627 commands is a follows:
  1628 
  1629   (1) interpretation expr inst
  1630   (2) interpretation target < expr
  1631   (3) interpret expr inst
  1632 
  1633 Interpretation in theories and proof contexts require a parameter
  1634 instantiation of terms from the current context.  This is applied to
  1635 specifications and theorems of the interpreted expression.
  1636 Interpretation in locales only permits parameter renaming through the
  1637 locale expression.  Interpretation is smart in that interpretations
  1638 that are active already do not occur in proof obligations, neither are
  1639 instantiated theorems stored in duplicate.  Use 'print_interps' to
  1640 inspect active interpretations of a particular locale.  For details,
  1641 see the Isar Reference manual.  Examples can be found in
  1642 HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy.
  1643 
  1644 INCOMPATIBILITY: former 'instantiate' has been withdrawn, use
  1645 'interpret' instead.
  1646 
  1647 * New context element 'constrains' for adding type constraints to
  1648 parameters.
  1649 
  1650 * Context expressions: renaming of parameters with syntax
  1651 redeclaration.
  1652 
  1653 * Locale declaration: 'includes' disallowed.
  1654 
  1655 * Proper static binding of attribute syntax -- i.e. types / terms /
  1656 facts mentioned as arguments are always those of the locale definition
  1657 context, independently of the context of later invocations.  Moreover,
  1658 locale operations (renaming and type / term instantiation) are applied
  1659 to attribute arguments as expected.
  1660 
  1661 INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of
  1662 actual attributes; rare situations may require Attrib.attribute to
  1663 embed those attributes into Attrib.src that lack concrete syntax.
  1664 Attribute implementations need to cooperate properly with the static
  1665 binding mechanism.  Basic parsers Args.XXX_typ/term/prop and
  1666 Attrib.XXX_thm etc. already do the right thing without further
  1667 intervention.  Only unusual applications -- such as "where" or "of"
  1668 (cf. src/Pure/Isar/attrib.ML), which process arguments depending both
  1669 on the context and the facts involved -- may have to assign parsed
  1670 values to argument tokens explicitly.
  1671 
  1672 * Changed parameter management in theorem generation for long goal
  1673 statements with 'includes'.  INCOMPATIBILITY: produces a different
  1674 theorem statement in rare situations.
  1675 
  1676 * Locale inspection command 'print_locale' omits notes elements.  Use
  1677 'print_locale!' to have them included in the output.
  1678 
  1679 
  1680 *** Provers ***
  1681 
  1682 * Provers/hypsubst.ML: improved version of the subst method, for
  1683 single-step rewriting: it now works in bound variable contexts. New is
  1684 'subst (asm)', for rewriting an assumption.  INCOMPATIBILITY: may
  1685 rewrite a different subterm than the original subst method, which is
  1686 still available as 'simplesubst'.
  1687 
  1688 * Provers/quasi.ML: new transitivity reasoners for transitivity only
  1689 and quasi orders.
  1690 
  1691 * Provers/trancl.ML: new transitivity reasoner for transitive and
  1692 reflexive-transitive closure of relations.
  1693 
  1694 * Provers/blast.ML: new reference depth_limit to make blast's depth
  1695 limit (previously hard-coded with a value of 20) user-definable.
  1696 
  1697 * Provers/simplifier.ML has been moved to Pure, where Simplifier.setup
  1698 is peformed already.  Object-logics merely need to finish their
  1699 initial simpset configuration as before.  INCOMPATIBILITY.
  1700 
  1701 
  1702 *** HOL ***
  1703 
  1704 * Symbolic syntax of Hilbert Choice Operator is now as follows:
  1705 
  1706   syntax (epsilon)
  1707     "_Eps" :: "[pttrn, bool] => 'a"    ("(3\<some>_./ _)" [0, 10] 10)
  1708 
  1709 The symbol \<some> is displayed as the alternative epsilon of LaTeX
  1710 and x-symbol; use option '-m epsilon' to get it actually printed.
  1711 Moreover, the mathematically important symbolic identifier \<epsilon>
  1712 becomes available as variable, constant etc.  INCOMPATIBILITY,
  1713 
  1714 * "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x".
  1715 Similarly for all quantifiers: "ALL x > y" etc.  The x-symbol for >=
  1716 is \<ge>. New transitivity rules have been added to HOL/Orderings.thy to
  1717 support corresponding Isar calculations.
  1718 
  1719 * "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\<in>"
  1720 instead of ":".
  1721 
  1722 * theory SetInterval: changed the syntax for open intervals:
  1723 
  1724   Old       New
  1725   {..n(}    {..<n}
  1726   {)n..}    {n<..}
  1727   {m..n(}   {m..<n}
  1728   {)m..n}   {m<..n}
  1729   {)m..n(}  {m<..<n}
  1730 
  1731 The old syntax is still supported but will disappear in the next
  1732 release.  For conversion use the following Emacs search and replace
  1733 patterns (these are not perfect but work quite well):
  1734 
  1735   {)\([^\.]*\)\.\.  ->  {\1<\.\.}
  1736   \.\.\([^(}]*\)(}  ->  \.\.<\1}
  1737 
  1738 * Theory Commutative_Ring (in Library): method comm_ring for proving
  1739 equalities in commutative rings; method 'algebra' provides a generic
  1740 interface.
  1741 
  1742 * Theory Finite_Set: changed the syntax for 'setsum', summation over
  1743 finite sets: "setsum (%x. e) A", which used to be "\<Sum>x:A. e", is
  1744 now either "SUM x:A. e" or "\<Sum>x \<in> A. e". The bound variable can
  1745 be a tuple pattern.
  1746 
  1747 Some new syntax forms are available:
  1748 
  1749   "\<Sum>x | P. e"      for     "setsum (%x. e) {x. P}"
  1750   "\<Sum>x = a..b. e"   for     "setsum (%x. e) {a..b}"
  1751   "\<Sum>x = a..<b. e"  for     "setsum (%x. e) {a..<b}"
  1752   "\<Sum>x < k. e"      for     "setsum (%x. e) {..<k}"
  1753 
  1754 The latter form "\<Sum>x < k. e" used to be based on a separate
  1755 function "Summation", which has been discontinued.
  1756 
  1757 * theory Finite_Set: in structured induction proofs, the insert case
  1758 is now 'case (insert x F)' instead of the old counterintuitive 'case
  1759 (insert F x)'.
  1760 
  1761 * The 'refute' command has been extended to support a much larger
  1762 fragment of HOL, including axiomatic type classes, constdefs and
  1763 typedefs, inductive datatypes and recursion.
  1764 
  1765 * New tactics 'sat' and 'satx' to prove propositional tautologies.
  1766 Requires zChaff with proof generation to be installed.  See
  1767 HOL/ex/SAT_Examples.thy for examples.
  1768 
  1769 * Datatype induction via method 'induct' now preserves the name of the
  1770 induction variable. For example, when proving P(xs::'a list) by
  1771 induction on xs, the induction step is now P(xs) ==> P(a#xs) rather
  1772 than P(list) ==> P(a#list) as previously.  Potential INCOMPATIBILITY
  1773 in unstructured proof scripts.
  1774 
  1775 * Reworked implementation of records.  Improved scalability for
  1776 records with many fields, avoiding performance problems for type
  1777 inference. Records are no longer composed of nested field types, but
  1778 of nested extension types. Therefore the record type only grows linear
  1779 in the number of extensions and not in the number of fields.  The
  1780 top-level (users) view on records is preserved.  Potential
  1781 INCOMPATIBILITY only in strange cases, where the theory depends on the
  1782 old record representation. The type generated for a record is called
  1783 <record_name>_ext_type.
  1784 
  1785 Flag record_quick_and_dirty_sensitive can be enabled to skip the
  1786 proofs triggered by a record definition or a simproc (if
  1787 quick_and_dirty is enabled).  Definitions of large records can take
  1788 quite long.
  1789 
  1790 New simproc record_upd_simproc for simplification of multiple record
  1791 updates enabled by default.  Moreover, trivial updates are also
  1792 removed: r(|x := x r|) = r.  INCOMPATIBILITY: old proofs break
  1793 occasionally, since simplification is more powerful by default.
  1794 
  1795 * typedef: proper support for polymorphic sets, which contain extra
  1796 type-variables in the term.
  1797 
  1798 * Simplifier: automatically reasons about transitivity chains
  1799 involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics
  1800 provided by Provers/trancl.ML as additional solvers.  INCOMPATIBILITY:
  1801 old proofs break occasionally as simplification may now solve more
  1802 goals than previously.
  1803 
  1804 * Simplifier: converts x <= y into x = y if assumption y <= x is
  1805 present.  Works for all partial orders (class "order"), in particular
  1806 numbers and sets.  For linear orders (e.g. numbers) it treats ~ x < y
  1807 just like y <= x.
  1808 
  1809 * Simplifier: new simproc for "let x = a in f x".  If a is a free or
  1810 bound variable or a constant then the let is unfolded.  Otherwise
  1811 first a is simplified to b, and then f b is simplified to g. If
  1812 possible we abstract b from g arriving at "let x = b in h x",
  1813 otherwise we unfold the let and arrive at g.  The simproc can be
  1814 enabled/disabled by the reference use_let_simproc.  Potential
  1815 INCOMPATIBILITY since simplification is more powerful by default.
  1816 
  1817 * Classical reasoning: the meson method now accepts theorems as arguments.
  1818 
  1819 * Prover support: pre-release of the Isabelle-ATP linkup, which runs background
  1820 jobs to provide advice on the provability of subgoals.
  1821 
  1822 * Theory OrderedGroup and Ring_and_Field: various additions and
  1823 improvements to faciliate calculations involving equalities and
  1824 inequalities.
  1825 
  1826 The following theorems have been eliminated or modified
  1827 (INCOMPATIBILITY):
  1828 
  1829   abs_eq             now named abs_of_nonneg
  1830   abs_of_ge_0        now named abs_of_nonneg
  1831   abs_minus_eq       now named abs_of_nonpos
  1832   imp_abs_id         now named abs_of_nonneg
  1833   imp_abs_neg_id     now named abs_of_nonpos
  1834   mult_pos           now named mult_pos_pos
  1835   mult_pos_le        now named mult_nonneg_nonneg
  1836   mult_pos_neg_le    now named mult_nonneg_nonpos
  1837   mult_pos_neg2_le   now named mult_nonneg_nonpos2
  1838   mult_neg           now named mult_neg_neg
  1839   mult_neg_le        now named mult_nonpos_nonpos
  1840 
  1841 * The following lemmas in Ring_and_Field have been added to the simplifier:
  1842      
  1843      zero_le_square
  1844      not_square_less_zero 
  1845 
  1846   The following lemmas have been deleted from Real/RealPow:
  1847   
  1848      realpow_zero_zero
  1849      realpow_two
  1850      realpow_less
  1851      zero_le_power
  1852      realpow_two_le
  1853      abs_realpow_two
  1854      realpow_two_abs     
  1855 
  1856 * Theory Parity: added rules for simplifying exponents.
  1857 
  1858 * Theory List:
  1859 
  1860 The following theorems have been eliminated or modified
  1861 (INCOMPATIBILITY):
  1862 
  1863   list_all_Nil       now named list_all.simps(1)
  1864   list_all_Cons      now named list_all.simps(2)
  1865   list_all_conv      now named list_all_iff
  1866   set_mem_eq         now named mem_iff
  1867 
  1868 * Theories SetsAndFunctions and BigO (see HOL/Library) support
  1869 asymptotic "big O" calculations.  See the notes in BigO.thy.
  1870 
  1871 
  1872 *** HOL-Complex ***
  1873 
  1874 * Theory RealDef: better support for embedding natural numbers and
  1875 integers in the reals.
  1876 
  1877 The following theorems have been eliminated or modified
  1878 (INCOMPATIBILITY):
  1879 
  1880   exp_ge_add_one_self  now requires no hypotheses
  1881   real_of_int_add      reversed direction of equality (use [symmetric])
  1882   real_of_int_minus    reversed direction of equality (use [symmetric])
  1883   real_of_int_diff     reversed direction of equality (use [symmetric])
  1884   real_of_int_mult     reversed direction of equality (use [symmetric])
  1885 
  1886 * Theory RComplete: expanded support for floor and ceiling functions.
  1887 
  1888 * Theory Ln is new, with properties of the natural logarithm
  1889 
  1890 * Hyperreal: There is a new type constructor "star" for making
  1891 nonstandard types.  The old type names are now type synonyms:
  1892 
  1893   hypreal = real star
  1894   hypnat = nat star
  1895   hcomplex = complex star
  1896 
  1897 * Hyperreal: Many groups of similarly-defined constants have been
  1898 replaced by polymorphic versions (INCOMPATIBILITY):
  1899 
  1900   star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex
  1901 
  1902   starset      <-- starsetNat, starsetC
  1903   *s*          <-- *sNat*, *sc*
  1904   starset_n    <-- starsetNat_n, starsetC_n
  1905   *sn*         <-- *sNatn*, *scn*
  1906   InternalSets <-- InternalNatSets, InternalCSets
  1907 
  1908   starfun      <-- starfun{Nat,Nat2,C,RC,CR}
  1909   *f*          <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR*
  1910   starfun_n    <-- starfun{Nat,Nat2,C,RC,CR}_n
  1911   *fn*         <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn*
  1912   InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs
  1913 
  1914 * Hyperreal: Many type-specific theorems have been removed in favor of
  1915 theorems specific to various axiomatic type classes (INCOMPATIBILITY):
  1916 
  1917   add_commute <-- {hypreal,hypnat,hcomplex}_add_commute
  1918   add_assoc   <-- {hypreal,hypnat,hcomplex}_add_assocs
  1919   OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left
  1920   OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right
  1921   right_minus <-- hypreal_add_minus
  1922   left_minus <-- {hypreal,hcomplex}_add_minus_left
  1923   mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute
  1924   mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc
  1925   mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left
  1926   mult_1_right <-- hcomplex_mult_one_right
  1927   mult_zero_left <-- hcomplex_mult_zero_left
  1928   left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib
  1929   right_distrib <-- hypnat_add_mult_distrib2
  1930   zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one
  1931   right_inverse <-- hypreal_mult_inverse
  1932   left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left
  1933   order_refl <-- {hypreal,hypnat}_le_refl
  1934   order_trans <-- {hypreal,hypnat}_le_trans
  1935   order_antisym <-- {hypreal,hypnat}_le_anti_sym
  1936   order_less_le <-- {hypreal,hypnat}_less_le
  1937   linorder_linear <-- {hypreal,hypnat}_le_linear
  1938   add_left_mono <-- {hypreal,hypnat}_add_left_mono
  1939   mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2
  1940   add_nonneg_nonneg <-- hypreal_le_add_order
  1941 
  1942 * Hyperreal: Separate theorems having to do with type-specific
  1943 versions of constants have been merged into theorems that apply to the
  1944 new polymorphic constants (INCOMPATIBILITY):
  1945 
  1946   STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set
  1947   STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set
  1948   STAR_Un <-- {STAR,NatStar,STARC}_Un
  1949   STAR_Int <-- {STAR,NatStar,STARC}_Int
  1950   STAR_Compl <-- {STAR,NatStar,STARC}_Compl
  1951   STAR_subset <-- {STAR,NatStar,STARC}_subset
  1952   STAR_mem <-- {STAR,NatStar,STARC}_mem
  1953   STAR_mem_Compl <-- {STAR,STARC}_mem_Compl
  1954   STAR_diff <-- {STAR,STARC}_diff
  1955   STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real,
  1956     STARC_hcomplex_of_complex}_image_subset
  1957   starset_n_Un <-- starset{Nat,C}_n_Un
  1958   starset_n_Int <-- starset{Nat,C}_n_Int
  1959   starset_n_Compl <-- starset{Nat,C}_n_Compl
  1960   starset_n_diff <-- starset{Nat,C}_n_diff
  1961   InternalSets_Un <-- Internal{Nat,C}Sets_Un
  1962   InternalSets_Int <-- Internal{Nat,C}Sets_Int
  1963   InternalSets_Compl <-- Internal{Nat,C}Sets_Compl
  1964   InternalSets_diff <-- Internal{Nat,C}Sets_diff
  1965   InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff
  1966   InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n
  1967   starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq
  1968   starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C}
  1969   starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR}
  1970   starfun <-- starfun{Nat,Nat2,C,RC,CR}
  1971   starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult
  1972   starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add
  1973   starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus
  1974   starfun_diff <-- starfun{C,RC,CR}_diff
  1975   starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o
  1976   starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2
  1977   starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun
  1978   starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse
  1979   starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq
  1980   starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff
  1981   starfun_Id <-- starfunC_Id
  1982   starfun_approx <-- starfun{Nat,CR}_approx
  1983   starfun_capprox <-- starfun{C,RC}_capprox
  1984   starfun_abs <-- starfunNat_rabs
  1985   starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel
  1986   starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2
  1987   starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox
  1988   starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox
  1989   starfun_add_capprox <-- starfun{C,RC}_add_capprox
  1990   starfun_add_approx <-- starfunCR_add_approx
  1991   starfun_inverse_inverse <-- starfunC_inverse_inverse
  1992   starfun_divide <-- starfun{C,CR,RC}_divide
  1993   starfun_n <-- starfun{Nat,C}_n
  1994   starfun_n_mult <-- starfun{Nat,C}_n_mult
  1995   starfun_n_add <-- starfun{Nat,C}_n_add
  1996   starfun_n_add_minus <-- starfunNat_n_add_minus
  1997   starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun
  1998   starfun_n_minus <-- starfun{Nat,C}_n_minus
  1999   starfun_n_eq <-- starfun{Nat,C}_n_eq
  2000 
  2001   star_n_add <-- {hypreal,hypnat,hcomplex}_add
  2002   star_n_minus <-- {hypreal,hcomplex}_minus
  2003   star_n_diff <-- {hypreal,hcomplex}_diff
  2004   star_n_mult <-- {hypreal,hcomplex}_mult
  2005   star_n_inverse <-- {hypreal,hcomplex}_inverse
  2006   star_n_le <-- {hypreal,hypnat}_le
  2007   star_n_less <-- {hypreal,hypnat}_less
  2008   star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num
  2009   star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num
  2010   star_n_abs <-- hypreal_hrabs
  2011   star_n_divide <-- hcomplex_divide
  2012 
  2013   star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add
  2014   star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus
  2015   star_of_diff <-- hypreal_of_real_diff
  2016   star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult
  2017   star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one
  2018   star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero
  2019   star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff
  2020   star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff
  2021   star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff
  2022   star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse
  2023   star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide
  2024   star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat
  2025   star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int
  2026   star_of_number_of <-- {hypreal,hcomplex}_number_of
  2027   star_of_number_less <-- number_of_less_hypreal_of_real_iff
  2028   star_of_number_le <-- number_of_le_hypreal_of_real_iff
  2029   star_of_eq_number <-- hypreal_of_real_eq_number_of_iff
  2030   star_of_less_number <-- hypreal_of_real_less_number_of_iff
  2031   star_of_le_number <-- hypreal_of_real_le_number_of_iff
  2032   star_of_power <-- hypreal_of_real_power
  2033   star_of_eq_0 <-- hcomplex_of_complex_zero_iff
  2034 
  2035 * Hyperreal: new method "transfer" that implements the transfer
  2036 principle of nonstandard analysis. With a subgoal that mentions
  2037 nonstandard types like "'a star", the command "apply transfer"
  2038 replaces it with an equivalent one that mentions only standard types.
  2039 To be successful, all free variables must have standard types; non-
  2040 standard variables must have explicit universal quantifiers.
  2041 
  2042 * Hyperreal: A theory of Taylor series.
  2043 
  2044 
  2045 *** HOLCF ***
  2046 
  2047 * Discontinued special version of 'constdefs' (which used to support
  2048 continuous functions) in favor of the general Pure one with full
  2049 type-inference.
  2050 
  2051 * New simplification procedure for solving continuity conditions; it
  2052 is much faster on terms with many nested lambda abstractions (cubic
  2053 instead of exponential time).
  2054 
  2055 * New syntax for domain package: selector names are now optional.
  2056 Parentheses should be omitted unless argument is lazy, for example:
  2057 
  2058   domain 'a stream = cons "'a" (lazy "'a stream")
  2059 
  2060 * New command 'fixrec' for defining recursive functions with pattern
  2061 matching; defining multiple functions with mutual recursion is also
  2062 supported.  Patterns may include the constants cpair, spair, up, sinl,
  2063 sinr, or any data constructor defined by the domain package. The given
  2064 equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for
  2065 syntax and examples.
  2066 
  2067 * New commands 'cpodef' and 'pcpodef' for defining predicate subtypes
  2068 of cpo and pcpo types. Syntax is exactly like the 'typedef' command,
  2069 but the proof obligation additionally includes an admissibility
  2070 requirement. The packages generate instances of class cpo or pcpo,
  2071 with continuity and strictness theorems for Rep and Abs.
  2072 
  2073 * HOLCF: Many theorems have been renamed according to a more standard naming
  2074 scheme (INCOMPATIBILITY):
  2075 
  2076   foo_inject:  "foo$x = foo$y ==> x = y"
  2077   foo_eq:      "(foo$x = foo$y) = (x = y)"
  2078   foo_less:    "(foo$x << foo$y) = (x << y)"
  2079   foo_strict:  "foo$UU = UU"
  2080   foo_defined: "... ==> foo$x ~= UU"
  2081   foo_defined_iff: "(foo$x = UU) = (x = UU)"
  2082 
  2083 
  2084 *** ZF ***
  2085 
  2086 * ZF/ex: theories Group and Ring provide examples in abstract algebra,
  2087 including the First Isomorphism Theorem (on quotienting by the kernel
  2088 of a homomorphism).
  2089 
  2090 * ZF/Simplifier: install second copy of type solver that actually
  2091 makes use of TC rules declared to Isar proof contexts (or locales);
  2092 the old version is still required for ML proof scripts.
  2093 
  2094 
  2095 *** Cube ***
  2096 
  2097 * Converted to Isar theory format; use locales instead of axiomatic
  2098 theories.
  2099 
  2100 
  2101 *** ML ***
  2102 
  2103 * Pure/library.ML: added ##>, ##>>, #>> -- higher-order counterparts
  2104 for ||>, ||>>, |>>,
  2105 
  2106 * Pure/library.ML no longer defines its own option datatype, but uses
  2107 that of the SML basis, which has constructors NONE and SOME instead of
  2108 None and Some, as well as exception Option.Option instead of OPTION.
  2109 The functions the, if_none, is_some, is_none have been adapted
  2110 accordingly, while Option.map replaces apsome.
  2111 
  2112 * Pure/library.ML: the exception LIST has been given up in favour of
  2113 the standard exceptions Empty and Subscript, as well as
  2114 Library.UnequalLengths.  Function like Library.hd and Library.tl are
  2115 superceded by the standard hd and tl functions etc.
  2116 
  2117 A number of basic list functions are no longer exported to the ML
  2118 toplevel, as they are variants of predefined functions.  The following
  2119 suggests how one can translate existing code:
  2120 
  2121     rev_append xs ys = List.revAppend (xs, ys)
  2122     nth_elem (i, xs) = List.nth (xs, i)
  2123     last_elem xs = List.last xs
  2124     flat xss = List.concat xss
  2125     seq fs = List.app fs
  2126     partition P xs = List.partition P xs
  2127     mapfilter f xs = List.mapPartial f xs
  2128 
  2129 * Pure/library.ML: several combinators for linear functional
  2130 transformations, notably reverse application and composition:
  2131 
  2132   x |> f                f #> g
  2133   (x, y) |-> f          f #-> g
  2134 
  2135 * Pure/library.ML: introduced/changed precedence of infix operators:
  2136 
  2137   infix 1 |> |-> ||> ||>> |>> |>>> #> #->;
  2138   infix 2 ?;
  2139   infix 3 o oo ooo oooo;
  2140   infix 4 ~~ upto downto;
  2141 
  2142 Maybe INCOMPATIBILITY when any of those is used in conjunction with other
  2143 infix operators.
  2144 
  2145 * Pure/library.ML: natural list combinators fold, fold_rev, and
  2146 fold_map support linear functional transformations and nesting.  For
  2147 example:
  2148 
  2149   fold f [x1, ..., xN] y =
  2150     y |> f x1 |> ... |> f xN
  2151 
  2152   (fold o fold) f [xs1, ..., xsN] y =
  2153     y |> fold f xs1 |> ... |> fold f xsN
  2154 
  2155   fold f [x1, ..., xN] =
  2156     f x1 #> ... #> f xN
  2157 
  2158   (fold o fold) f [xs1, ..., xsN] =
  2159     fold f xs1 #> ... #> fold f xsN
  2160 
  2161 * Pure/library.ML: the following selectors on type 'a option are
  2162 available:
  2163 
  2164   the:               'a option -> 'a  (*partial*)
  2165   these:             'a option -> 'a  where 'a = 'b list
  2166   the_default: 'a -> 'a option -> 'a
  2167   the_list:          'a option -> 'a list
  2168 
  2169 * Pure/General: structure AList (cf. Pure/General/alist.ML) provides
  2170 basic operations for association lists, following natural argument
  2171 order; moreover the explicit equality predicate passed here avoids
  2172 potentially expensive polymorphic runtime equality checks.
  2173 The old functions may be expressed as follows:
  2174 
  2175   assoc = uncurry (AList.lookup (op =))
  2176   assocs = these oo AList.lookup (op =)
  2177   overwrite = uncurry (AList.update (op =)) o swap
  2178 
  2179 * Pure/General: structure AList (cf. Pure/General/alist.ML) provides
  2180 
  2181   val make: ('a -> 'b) -> 'a list -> ('a * 'b) list
  2182   val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list
  2183 
  2184 replacing make_keylist and keyfilter (occassionally used)
  2185 Naive rewrites:
  2186 
  2187   make_keylist = AList.make
  2188   keyfilter = AList.find (op =)
  2189 
  2190 * eq_fst and eq_snd now take explicit equality parameter, thus
  2191   avoiding eqtypes. Naive rewrites:
  2192 
  2193     eq_fst = eq_fst (op =)
  2194     eq_snd = eq_snd (op =)
  2195 
  2196 * Removed deprecated apl and apr (rarely used).
  2197   Naive rewrites:
  2198 
  2199     apl (n, op) =>>= curry op n
  2200     apr (op, m) =>>= fn n => op (n, m)
  2201 
  2202 * Pure/General: structure OrdList (cf. Pure/General/ord_list.ML)
  2203 provides a reasonably efficient light-weight implementation of sets as
  2204 lists.
  2205 
  2206 * Pure/General: generic tables (cf. Pure/General/table.ML) provide a
  2207 few new operations; existing lookup and update are now curried to
  2208 follow natural argument order (for use with fold etc.);
  2209 INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort.
  2210 
  2211 * Pure/General: output via the Isabelle channels of
  2212 writeln/warning/error etc. is now passed through Output.output, with a
  2213 hook for arbitrary transformations depending on the print_mode
  2214 (cf. Output.add_mode -- the first active mode that provides a output
  2215 function wins).  Already formatted output may be embedded into further
  2216 text via Output.raw; the result of Pretty.string_of/str_of and derived
  2217 functions (string_of_term/cterm/thm etc.) is already marked raw to
  2218 accommodate easy composition of diagnostic messages etc.  Programmers
  2219 rarely need to care about Output.output or Output.raw at all, with
  2220 some notable exceptions: Output.output is required when bypassing the
  2221 standard channels (writeln etc.), or in token translations to produce
  2222 properly formatted results; Output.raw is required when capturing
  2223 already output material that will eventually be presented to the user
  2224 a second time.  For the default print mode, both Output.output and
  2225 Output.raw have no effect.
  2226 
  2227 * Pure/General: Output.time_accumulator NAME creates an operator ('a
  2228 -> 'b) -> 'a -> 'b to measure runtime and count invocations; the
  2229 cumulative results are displayed at the end of a batch session.
  2230 
  2231 * Pure/General: File.sysify_path and File.quote_sysify path have been
  2232 replaced by File.platform_path and File.shell_path (with appropriate
  2233 hooks).  This provides a clean interface for unusual systems where the
  2234 internal and external process view of file names are different.
  2235 
  2236 * Pure: more efficient orders for basic syntactic entities: added
  2237 fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord
  2238 and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is
  2239 NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast
  2240 orders now -- potential INCOMPATIBILITY for code that depends on a
  2241 particular order for Symtab.keys, Symtab.dest, etc. (consider using
  2242 Library.sort_strings on result).
  2243 
  2244 * Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types,
  2245 fold_types traverse types/terms from left to right, observing natural
  2246 argument order.  Supercedes previous foldl_XXX versions, add_frees,
  2247 add_vars etc. have been adapted as well: INCOMPATIBILITY.
  2248 
  2249 * Pure: name spaces have been refined, with significant changes of the
  2250 internal interfaces -- INCOMPATIBILITY.  Renamed cond_extern(_table)
  2251 to extern(_table).  The plain name entry path is superceded by a
  2252 general 'naming' context, which also includes the 'policy' to produce
  2253 a fully qualified name and external accesses of a fully qualified
  2254 name; NameSpace.extend is superceded by context dependent
  2255 Sign.declare_name.  Several theory and proof context operations modify
  2256 the naming context.  Especially note Theory.restore_naming and
  2257 ProofContext.restore_naming to get back to a sane state; note that
  2258 Theory.add_path is no longer sufficient to recover from
  2259 Theory.absolute_path in particular.
  2260 
  2261 * Pure: new flags short_names (default false) and unique_names
  2262 (default true) for controlling output of qualified names.  If
  2263 short_names is set, names are printed unqualified.  If unique_names is
  2264 reset, the name prefix is reduced to the minimum required to achieve
  2265 the original result when interning again, even if there is an overlap
  2266 with earlier declarations.
  2267 
  2268 * Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is
  2269 now 'extend', and 'merge' gets an additional Pretty.pp argument
  2270 (useful for printing error messages).  INCOMPATIBILITY.
  2271 
  2272 * Pure: major reorganization of the theory context.  Type Sign.sg and
  2273 Theory.theory are now identified, referring to the universal
  2274 Context.theory (see Pure/context.ML).  Actual signature and theory
  2275 content is managed as theory data.  The old code and interfaces were
  2276 spread over many files and structures; the new arrangement introduces
  2277 considerable INCOMPATIBILITY to gain more clarity:
  2278 
  2279   Context -- theory management operations (name, identity, inclusion,
  2280     parents, ancestors, merge, etc.), plus generic theory data;
  2281 
  2282   Sign -- logical signature and syntax operations (declaring consts,
  2283     types, etc.), plus certify/read for common entities;
  2284 
  2285   Theory -- logical theory operations (stating axioms, definitions,
  2286     oracles), plus a copy of logical signature operations (consts,
  2287     types, etc.); also a few basic management operations (Theory.copy,
  2288     Theory.merge, etc.)
  2289 
  2290 The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm
  2291 etc.) as well as the sign field in Thm.rep_thm etc. have been retained
  2292 for convenience -- they merely return the theory.
  2293 
  2294 * Pure: type Type.tsig is superceded by theory in most interfaces.
  2295 
  2296 * Pure: the Isar proof context type is already defined early in Pure
  2297 as Context.proof (note that ProofContext.context and Proof.context are
  2298 aliases, where the latter is the preferred name).  This enables other
  2299 Isabelle components to refer to that type even before Isar is present.
  2300 
  2301 * Pure/sign/theory: discontinued named name spaces (i.e. classK,
  2302 typeK, constK, axiomK, oracleK), but provide explicit operations for
  2303 any of these kinds.  For example, Sign.intern typeK is now
  2304 Sign.intern_type, Theory.hide_space Sign.typeK is now
  2305 Theory.hide_types.  Also note that former
  2306 Theory.hide_classes/types/consts are now
  2307 Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions
  2308 internalize their arguments!  INCOMPATIBILITY.
  2309 
  2310 * Pure: get_thm interface (of PureThy and ProofContext) expects
  2311 datatype thmref (with constructors Name and NameSelection) instead of
  2312 plain string -- INCOMPATIBILITY;
  2313 
  2314 * Pure: cases produced by proof methods specify options, where NONE
  2315 means to remove case bindings -- INCOMPATIBILITY in
  2316 (RAW_)METHOD_CASES.
  2317 
  2318 * Pure: the following operations retrieve axioms or theorems from a
  2319 theory node or theory hierarchy, respectively:
  2320 
  2321   Theory.axioms_of: theory -> (string * term) list
  2322   Theory.all_axioms_of: theory -> (string * term) list
  2323   PureThy.thms_of: theory -> (string * thm) list
  2324   PureThy.all_thms_of: theory -> (string * thm) list
  2325 
  2326 * Pure: print_tac now outputs the goal through the trace channel.
  2327 
  2328 * Isar toplevel: improved diagnostics, mostly for Poly/ML only.
  2329 Reference Toplevel.debug (default false) controls detailed printing
  2330 and tracing of low-level exceptions; Toplevel.profiling (default 0)
  2331 controls execution profiling -- set to 1 for time and 2 for space
  2332 (both increase the runtime).
  2333 
  2334 * Isar session: The initial use of ROOT.ML is now always timed,
  2335 i.e. the log will show the actual process times, in contrast to the
  2336 elapsed wall-clock time that the outer shell wrapper produces.
  2337 
  2338 * Simplifier: improved handling of bound variables (nameless
  2339 representation, avoid allocating new strings).  Simprocs that invoke
  2340 the Simplifier recursively should use Simplifier.inherit_bounds to
  2341 avoid local name clashes.  Failure to do so produces warnings
  2342 "Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds
  2343 for further details.
  2344 
  2345 * ML functions legacy_bindings and use_legacy_bindings produce ML fact
  2346 bindings for all theorems stored within a given theory; this may help
  2347 in porting non-Isar theories to Isar ones, while keeping ML proof
  2348 scripts for the time being.
  2349 
  2350 * ML operator HTML.with_charset specifies the charset begin used for
  2351 generated HTML files.  For example:
  2352 
  2353   HTML.with_charset "utf-8" use_thy "Hebrew";
  2354   HTML.with_charset "utf-8" use_thy "Chinese";
  2355 
  2356 
  2357 *** System ***
  2358 
  2359 * Allow symlinks to all proper Isabelle executables (Isabelle,
  2360 isabelle, isatool etc.).
  2361 
  2362 * ISABELLE_DOC_FORMAT setting specifies preferred document format (for
  2363 isatool doc, isatool mkdir, display_drafts etc.).
  2364 
  2365 * isatool usedir: option -f allows specification of the ML file to be
  2366 used by Isabelle; default is ROOT.ML.
  2367 
  2368 * New isatool version outputs the version identifier of the Isabelle
  2369 distribution being used.
  2370 
  2371 * HOL: new isatool dimacs2hol converts files in DIMACS CNF format
  2372 (containing Boolean satisfiability problems) into Isabelle/HOL
  2373 theories.
  2374 
  2375 
  2376 
  2377 New in Isabelle2004 (April 2004)
  2378 --------------------------------
  2379 
  2380 *** General ***
  2381 
  2382 * Provers/order.ML:  new efficient reasoner for partial and linear orders.
  2383   Replaces linorder.ML.
  2384 
  2385 * Pure: Greek letters (except small lambda, \<lambda>), as well as Gothic
  2386   (\<aa>...\<zz>\<AA>...\<ZZ>), calligraphic (\<A>...\<Z>), and Euler
  2387   (\<a>...\<z>), are now considered normal letters, and can therefore
  2388   be used anywhere where an ASCII letter (a...zA...Z) has until
  2389   now. COMPATIBILITY: This obviously changes the parsing of some
  2390   terms, especially where a symbol has been used as a binder, say
  2391   '\<Pi>x. ...', which is now a type error since \<Pi>x will be parsed
  2392   as an identifier.  Fix it by inserting a space around former
  2393   symbols.  Call 'isatool fixgreek' to try to fix parsing errors in
  2394   existing theory and ML files.
  2395 
  2396 * Pure: Macintosh and Windows line-breaks are now allowed in theory files.
  2397 
  2398 * Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now
  2399   allowed in identifiers. Similar to Greek letters \<^isub> is now considered
  2400   a normal (but invisible) letter. For multiple letter subscripts repeat
  2401   \<^isub> like this: x\<^isub>1\<^isub>2.
  2402 
  2403 * Pure: There are now sub-/superscripts that can span more than one
  2404   character. Text between \<^bsub> and \<^esub> is set in subscript in
  2405   ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in
  2406   superscript. The new control characters are not identifier parts.
  2407 
  2408 * Pure: Control-symbols of the form \<^raw:...> will literally print the
  2409   content of "..." to the latex file instead of \isacntrl... . The "..."
  2410   may consist of any printable characters excluding the end bracket >.
  2411 
  2412 * Pure: Using new Isar command "finalconsts" (or the ML functions
  2413   Theory.add_finals or Theory.add_finals_i) it is now possible to
  2414   declare constants "final", which prevents their being given a definition
  2415   later.  It is useful for constants whose behaviour is fixed axiomatically
  2416   rather than definitionally, such as the meta-logic connectives.
  2417 
  2418 * Pure: 'instance' now handles general arities with general sorts
  2419   (i.e. intersections of classes),
  2420 
  2421 * Presentation: generated HTML now uses a CSS style sheet to make layout
  2422   (somewhat) independent of content. It is copied from lib/html/isabelle.css.
  2423   It can be changed to alter the colors/layout of generated pages.
  2424 
  2425 
  2426 *** Isar ***
  2427 
  2428 * Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac,
  2429   cut_tac, subgoal_tac and thin_tac:
  2430   - Now understand static (Isar) contexts.  As a consequence, users of Isar
  2431     locales are no longer forced to write Isar proof scripts.
  2432     For details see Isar Reference Manual, paragraph 4.3.2: Further tactic
  2433     emulations.
  2434   - INCOMPATIBILITY: names of variables to be instantiated may no
  2435     longer be enclosed in quotes.  Instead, precede variable name with `?'.
  2436     This is consistent with the instantiation attribute "where".
  2437 
  2438 * Attributes "where" and "of":
  2439   - Now take type variables of instantiated theorem into account when reading
  2440     the instantiation string.  This fixes a bug that caused instantiated
  2441     theorems to have too special types in some circumstances.
  2442   - "where" permits explicit instantiations of type variables.
  2443 
  2444 * Calculation commands "moreover" and "also" no longer interfere with
  2445   current facts ("this"), admitting arbitrary combinations with "then"
  2446   and derived forms.
  2447 
  2448 * Locales:
  2449   - Goal statements involving the context element "includes" no longer
  2450     generate theorems with internal delta predicates (those ending on
  2451     "_axioms") in the premise.
  2452     Resolve particular premise with <locale>.intro to obtain old form.
  2453   - Fixed bug in type inference ("unify_frozen") that prevented mix of target
  2454     specification and "includes" elements in goal statement.
  2455   - Rule sets <locale>.intro and <locale>.axioms no longer declared as
  2456     [intro?] and [elim?] (respectively) by default.
  2457   - Experimental command for instantiation of locales in proof contexts:
  2458         instantiate <label>[<attrs>]: <loc>
  2459     Instantiates locale <loc> and adds all its theorems to the current context
  2460     taking into account their attributes.  Label and attrs are optional
  2461     modifiers, like in theorem declarations.  If present, names of
  2462     instantiated theorems are qualified with <label>, and the attributes
  2463     <attrs> are applied after any attributes these theorems might have already.
  2464       If the locale has assumptions, a chained fact of the form
  2465     "<loc> t1 ... tn" is expected from which instantiations of the parameters
  2466     are derived.  The command does not support old-style locales declared
  2467     with "locale (open)".
  2468       A few (very simple) examples can be found in FOL/ex/LocaleInst.thy.
  2469 
  2470 * HOL: Tactic emulation methods induct_tac and case_tac understand static
  2471   (Isar) contexts.
  2472 
  2473 
  2474 *** HOL ***
  2475 
  2476 * Proof import: new image HOL4 contains the imported library from
  2477   the HOL4 system with about 2500 theorems. It is imported by
  2478   replaying proof terms produced by HOL4 in Isabelle. The HOL4 image
  2479   can be used like any other Isabelle image.  See
  2480   HOL/Import/HOL/README for more information.
  2481 
  2482 * Simplifier:
  2483   - Much improved handling of linear and partial orders.
  2484     Reasoners for linear and partial orders are set up for type classes
  2485     "linorder" and "order" respectively, and are added to the default simpset
  2486     as solvers.  This means that the simplifier can build transitivity chains
  2487     to solve goals from the assumptions.
  2488   - INCOMPATIBILITY: old proofs break occasionally.  Typically, applications
  2489     of blast or auto after simplification become unnecessary because the goal
  2490     is solved by simplification already.
  2491 
  2492 * Numerics: new theory Ring_and_Field contains over 250 basic numerical laws,
  2493     all proved in axiomatic type classes for semirings, rings and fields.
  2494 
  2495 * Numerics:
  2496   - Numeric types (nat, int, and in HOL-Complex rat, real, complex, etc.) are
  2497     now formalized using the Ring_and_Field theory mentioned above.
  2498   - INCOMPATIBILITY: simplification and arithmetic behaves somewhat differently
  2499     than before, because now they are set up once in a generic manner.
  2500   - INCOMPATIBILITY: many type-specific arithmetic laws have gone.
  2501     Look for the general versions in Ring_and_Field (and Power if they concern
  2502     exponentiation).
  2503 
  2504 * Type "rat" of the rational numbers is now available in HOL-Complex.
  2505 
  2506 * Records:
  2507   - Record types are now by default printed with their type abbreviation
  2508     instead of the list of all field types. This can be configured via
  2509     the reference "print_record_type_abbr".
  2510   - Simproc "record_upd_simproc" for simplification of multiple updates added
  2511     (not enabled by default).
  2512   - Simproc "record_ex_sel_eq_simproc" to simplify EX x. sel r = x resp.
  2513     EX x. x = sel r to True (not enabled by default).
  2514   - Tactic "record_split_simp_tac" to split and simplify records added.
  2515 
  2516 * 'specification' command added, allowing for definition by
  2517   specification.  There is also an 'ax_specification' command that
  2518   introduces the new constants axiomatically.
  2519 
  2520 * arith(_tac) is now able to generate counterexamples for reals as well.
  2521 
  2522 * HOL-Algebra: new locale "ring" for non-commutative rings.
  2523 
  2524 * HOL-ex: InductiveInvariant_examples illustrates advanced recursive function
  2525   definitions, thanks to Sava Krsti\'{c} and John Matthews.
  2526 
  2527 * HOL-Matrix: a first theory for matrices in HOL with an application of
  2528   matrix theory to linear programming.
  2529 
  2530 * Unions and Intersections:
  2531   The latex output syntax of UN and INT has been changed
  2532   from "\Union x \in A. B" to "\Union_{x \in A} B"
  2533   i.e. the index formulae has become a subscript.
  2534   Similarly for "\Union x. B", and for \Inter instead of \Union.
  2535 
  2536 * Unions and Intersections over Intervals:
  2537   There is new short syntax "UN i<=n. A" for "UN i:{0..n}. A". There is
  2538   also an x-symbol version with subscripts "\<Union>\<^bsub>i <= n\<^esub>. A"
  2539   like in normal math, and corresponding versions for < and for intersection.
  2540 
  2541 * HOL/List: Ordering "lexico" is renamed "lenlex" and the standard
  2542   lexicographic dictonary ordering has been added as "lexord".
  2543 
  2544 * ML: the legacy theory structures Int and List have been removed. They had
  2545   conflicted with ML Basis Library structures having the same names.
  2546 
  2547 * 'refute' command added to search for (finite) countermodels.  Only works
  2548   for a fragment of HOL.  The installation of an external SAT solver is
  2549   highly recommended.  See "HOL/Refute.thy" for details.
  2550 
  2551 * 'quickcheck' command: Allows to find counterexamples by evaluating
  2552   formulae under an assignment of free variables to random values.
  2553   In contrast to 'refute', it can deal with inductive datatypes,
  2554   but cannot handle quantifiers. See "HOL/ex/Quickcheck_Examples.thy"
  2555   for examples.
  2556 
  2557 
  2558 *** HOLCF ***
  2559 
  2560 * Streams now come with concatenation and are part of the HOLCF image
  2561 
  2562 
  2563 
  2564 New in Isabelle2003 (May 2003)
  2565 ------------------------------
  2566 
  2567 *** General ***
  2568 
  2569 * Provers/simplifier:
  2570 
  2571   - Completely reimplemented method simp (ML: Asm_full_simp_tac):
  2572     Assumptions are now subject to complete mutual simplification,
  2573     not just from left to right. The simplifier now preserves
  2574     the order of assumptions.
  2575 
  2576     Potential INCOMPATIBILITY:
  2577 
  2578     -- simp sometimes diverges where the old version did
  2579        not, e.g. invoking simp on the goal
  2580 
  2581         [| P (f x); y = x; f x = f y |] ==> Q
  2582 
  2583        now gives rise to the infinite reduction sequence
  2584 
  2585         P(f x) --(f x = f y)--> P(f y) --(y = x)--> P(f x) --(f x = f y)--> ...
  2586 
  2587        Using "simp (asm_lr)" (ML: Asm_lr_simp_tac) instead often solves this
  2588        kind of problem.
  2589 
  2590     -- Tactics combining classical reasoner and simplification (such as auto)
  2591        are also affected by this change, because many of them rely on
  2592        simp. They may sometimes diverge as well or yield a different numbers
  2593        of subgoals. Try to use e.g. force, fastsimp, or safe instead of auto
  2594        in case of problems. Sometimes subsequent calls to the classical
  2595        reasoner will fail because a preceeding call to the simplifier too
  2596        eagerly simplified the goal, e.g. deleted redundant premises.
  2597 
  2598   - The simplifier trace now shows the names of the applied rewrite rules
  2599 
  2600   - You can limit the number of recursive invocations of the simplifier
  2601     during conditional rewriting (where the simplifie tries to solve the
  2602     conditions before applying the rewrite rule):
  2603     ML "simp_depth_limit := n"
  2604     where n is an integer. Thus you can force termination where previously
  2605     the simplifier would diverge.
  2606 
  2607   - Accepts free variables as head terms in congruence rules.  Useful in Isar.
  2608 
  2609   - No longer aborts on failed congruence proof.  Instead, the
  2610     congruence is ignored.
  2611 
  2612 * Pure: New generic framework for extracting programs from constructive
  2613   proofs. See HOL/Extraction.thy for an example instantiation, as well
  2614   as HOL/Extraction for some case studies.
  2615 
  2616 * Pure: The main goal of the proof state is no longer shown by default, only
  2617 the subgoals. This behaviour is controlled by a new flag.
  2618    PG menu: Isabelle/Isar -> Settings -> Show Main Goal
  2619 (ML: Proof.show_main_goal).
  2620 
  2621 * Pure: You can find all matching introduction rules for subgoal 1, i.e. all
  2622 rules whose conclusion matches subgoal 1:
  2623       PG menu: Isabelle/Isar -> Show me -> matching rules
  2624 The rules are ordered by how closely they match the subgoal.
  2625 In particular, rules that solve a subgoal outright are displayed first
  2626 (or rather last, the way they are printed).
  2627 (ML: ProofGeneral.print_intros())
  2628 
  2629 * Pure: New flag trace_unify_fail causes unification to print
  2630 diagnostic information (PG: in trace buffer) when it fails. This is
  2631 useful for figuring out why single step proofs like rule, erule or
  2632 assumption failed.
  2633 
  2634 * Pure: Locale specifications now produce predicate definitions
  2635 according to the body of text (covering assumptions modulo local
  2636 definitions); predicate "loc_axioms" covers newly introduced text,
  2637 while "loc" is cumulative wrt. all included locale expressions; the
  2638 latter view is presented only on export into the global theory
  2639 context; potential INCOMPATIBILITY, use "(open)" option to fall back
  2640 on the old view without predicates;
  2641 
  2642 * Pure: predefined locales "var" and "struct" are useful for sharing
  2643 parameters (as in CASL, for example); just specify something like
  2644 ``var x + var y + struct M'' as import;
  2645 
  2646 * Pure: improved thms_containing: proper indexing of facts instead of
  2647 raw theorems; check validity of results wrt. current name space;
  2648 include local facts of proof configuration (also covers active
  2649 locales), cover fixed variables in index; may use "_" in term
  2650 specification; an optional limit for the number of printed facts may
  2651 be given (the default is 40);
  2652 
  2653 * Pure: disallow duplicate fact bindings within new-style theory files
  2654 (batch-mode only);
  2655 
  2656 * Provers: improved induct method: assumptions introduced by case
  2657 "foo" are split into "foo.hyps" (from the rule) and "foo.prems" (from
  2658 the goal statement); "foo" still refers to all facts collectively;
  2659 
  2660 * Provers: the function blast.overloaded has been removed: all constants
  2661 are regarded as potentially overloaded, which improves robustness in exchange
  2662 for slight decrease in efficiency;
  2663 
  2664 * Provers/linorder: New generic prover for transitivity reasoning over
  2665 linear orders.  Note: this prover is not efficient!
  2666 
  2667 * Isar: preview of problems to finish 'show' now produce an error
  2668 rather than just a warning (in interactive mode);
  2669 
  2670 
  2671 *** HOL ***
  2672 
  2673 * arith(_tac)
  2674 
  2675  - Produces a counter example if it cannot prove a goal.
  2676    Note that the counter example may be spurious if the goal is not a formula
  2677    of quantifier-free linear arithmetic.
  2678    In ProofGeneral the counter example appears in the trace buffer.
  2679 
  2680  - Knows about div k and mod k where k is a numeral of type nat or int.
  2681 
  2682  - Calls full Presburger arithmetic (by Amine Chaieb) if quantifier-free
  2683    linear arithmetic fails. This takes account of quantifiers and divisibility.
  2684    Presburger arithmetic can also be called explicitly via presburger(_tac).
  2685 
  2686 * simp's arithmetic capabilities have been enhanced a bit: it now
  2687 takes ~= in premises into account (by performing a case split);
  2688 
  2689 * simp reduces "m*(n div m) + n mod m" to n, even if the two summands
  2690 are distributed over a sum of terms;
  2691 
  2692 * New tactic "trans_tac" and method "trans" instantiate
  2693 Provers/linorder.ML for axclasses "order" and "linorder" (predicates
  2694 "<=", "<" and "=").
  2695 
  2696 * function INCOMPATIBILITIES: Pi-sets have been redefined and moved from main
  2697 HOL to Library/FuncSet; constant "Fun.op o" is now called "Fun.comp";
  2698 
  2699 * 'typedef' command has new option "open" to suppress the set
  2700 definition;
  2701 
  2702 * functions Min and Max on finite sets have been introduced (theory
  2703 Finite_Set);
  2704 
  2705 * attribute [symmetric] now works for relations as well; it turns
  2706 (x,y) : R^-1 into (y,x) : R, and vice versa;
  2707 
  2708 * induct over a !!-quantified statement (say !!x1..xn):
  2709   each "case" automatically performs "fix x1 .. xn" with exactly those names.
  2710 
  2711 * Map: `empty' is no longer a constant but a syntactic abbreviation for
  2712 %x. None. Warning: empty_def now refers to the previously hidden definition
  2713 of the empty set.
  2714 
  2715 * Algebra: formalization of classical algebra.  Intended as base for
  2716 any algebraic development in Isabelle.  Currently covers group theory
  2717 (up to Sylow's theorem) and ring theory (Universal Property of
  2718 Univariate Polynomials).  Contributions welcome;
  2719 
  2720 * GroupTheory: deleted, since its material has been moved to Algebra;
  2721 
  2722 * Complex: new directory of the complex numbers with numeric constants,
  2723 nonstandard complex numbers, and some complex analysis, standard and
  2724 nonstandard (Jacques Fleuriot);
  2725 
  2726 * HOL-Complex: new image for analysis, replacing HOL-Real and HOL-Hyperreal;
  2727 
  2728 * Hyperreal: introduced Gauge integration and hyperreal logarithms (Jacques
  2729 Fleuriot);
  2730 
  2731 * Real/HahnBanach: updated and adapted to locales;
  2732 
  2733 * NumberTheory: added Gauss's law of quadratic reciprocity (by Avigad,
  2734 Gray and Kramer);
  2735 
  2736 * UNITY: added the Meier-Sanders theory of progress sets;
  2737 
  2738 * MicroJava: bytecode verifier and lightweight bytecode verifier
  2739 as abstract algorithms, instantiated to the JVM;
  2740 
  2741 * Bali: Java source language formalization. Type system, operational
  2742 semantics, axiomatic semantics. Supported language features:
  2743 classes, interfaces, objects,virtual methods, static methods,
  2744 static/instance fields, arrays, access modifiers, definite
  2745 assignment, exceptions.
  2746 
  2747 
  2748 *** ZF ***
  2749 
  2750 * ZF/Constructible: consistency proof for AC (Gdel's constructible
  2751 universe, etc.);
  2752 
  2753 * Main ZF: virtually all theories converted to new-style format;
  2754 
  2755 
  2756 *** ML ***
  2757 
  2758 * Pure: Tactic.prove provides sane interface for internal proofs;
  2759 omits the infamous "standard" operation, so this is more appropriate
  2760 than prove_goalw_cterm in many situations (e.g. in simprocs);
  2761 
  2762 * Pure: improved error reporting of simprocs;
  2763 
  2764 * Provers: Simplifier.simproc(_i) provides sane interface for setting
  2765 up simprocs;
  2766 
  2767 
  2768 *** Document preparation ***
  2769 
  2770 * uses \par instead of \\ for line breaks in theory text. This may
  2771 shift some page breaks in large documents. To get the old behaviour
  2772 use \renewcommand{\isanewline}{\mbox{}\\\mbox{}} in root.tex.
  2773 
  2774 * minimized dependencies of isabelle.sty and isabellesym.sty on
  2775 other packages
  2776 
  2777 * \<euro> now needs package babel/greek instead of marvosym (which
  2778 broke \Rightarrow)
  2779 
  2780 * normal size for \<zero>...\<nine> (uses \mathbf instead of
  2781 textcomp package)
  2782 
  2783 
  2784 
  2785 New in Isabelle2002 (March 2002)
  2786 --------------------------------
  2787 
  2788 *** Document preparation ***
  2789 
  2790 * greatly simplified document preparation setup, including more
  2791 graceful interpretation of isatool usedir -i/-d/-D options, and more
  2792 instructive isatool mkdir; users should basically be able to get
  2793 started with "isatool mkdir HOL Test && isatool make"; alternatively,
  2794 users may run a separate document processing stage manually like this:
  2795 "isatool usedir -D output HOL Test && isatool document Test/output";
  2796 
  2797 * theory dependency graph may now be incorporated into documents;
  2798 isatool usedir -g true will produce session_graph.eps/.pdf for use
  2799 with \includegraphics of LaTeX;
  2800 
  2801 * proper spacing of consecutive markup elements, especially text
  2802 blocks after section headings;
  2803 
  2804 * support bold style (for single symbols only), input syntax is like
  2805 this: "\<^bold>\<alpha>" or "\<^bold>A";
  2806 
  2807 * \<bullet> is now output as bold \cdot by default, which looks much
  2808 better in printed text;
  2809 
  2810 * added default LaTeX bindings for \<tturnstile> and \<TTurnstile>;
  2811 note that these symbols are currently unavailable in Proof General /
  2812 X-Symbol; new symbols \<zero>, \<one>, ..., \<nine>, and \<euro>;
  2813 
  2814 * isatool latex no longer depends on changed TEXINPUTS, instead
  2815 isatool document copies the Isabelle style files to the target
  2816 location;
  2817 
  2818 
  2819 *** Isar ***
  2820 
  2821 * Pure/Provers: improved proof by cases and induction;
  2822   - 'case' command admits impromptu naming of parameters (such as
  2823     "case (Suc n)");
  2824   - 'induct' method divinates rule instantiation from the inductive
  2825     claim; no longer requires excessive ?P bindings for proper
  2826     instantiation of cases;
  2827   - 'induct' method properly enumerates all possibilities of set/type
  2828     rules; as a consequence facts may be also passed through *type*
  2829     rules without further ado;
  2830   - 'induct' method now derives symbolic cases from the *rulified*
  2831     rule (before it used to rulify cases stemming from the internal
  2832     atomized version); this means that the context of a non-atomic
  2833     statement becomes is included in the hypothesis, avoiding the
  2834     slightly cumbersome show "PROP ?case" form;
  2835   - 'induct' may now use elim-style induction rules without chaining
  2836     facts, using ``missing'' premises from the goal state; this allows
  2837     rules stemming from inductive sets to be applied in unstructured
  2838     scripts, while still benefitting from proper handling of non-atomic
  2839     statements; NB: major inductive premises need to be put first, all
  2840     the rest of the goal is passed through the induction;
  2841   - 'induct' proper support for mutual induction involving non-atomic
  2842     rule statements (uses the new concept of simultaneous goals, see
  2843     below);
  2844   - append all possible rule selections, but only use the first
  2845     success (no backtracking);
  2846   - removed obsolete "(simplified)" and "(stripped)" options of methods;
  2847   - undeclared rule case names default to numbers 1, 2, 3, ...;
  2848   - added 'print_induct_rules' (covered by help item in recent Proof
  2849     General versions);
  2850   - moved induct/cases attributes to Pure, methods to Provers;
  2851   - generic method setup instantiated for FOL and HOL;
  2852 
  2853 * Pure: support multiple simultaneous goal statements, for example
  2854 "have a: A and b: B" (same for 'theorem' etc.); being a pure
  2855 meta-level mechanism, this acts as if several individual goals had
  2856 been stated separately; in particular common proof methods need to be
  2857 repeated in order to cover all claims; note that a single elimination
  2858 step is *not* sufficient to establish the two conjunctions, so this
  2859 fails:
  2860 
  2861   assume "A & B" then have A and B ..   (*".." fails*)
  2862 
  2863 better use "obtain" in situations as above; alternative refer to
  2864 multi-step methods like 'auto', 'simp_all', 'blast+' etc.;
  2865 
  2866 * Pure: proper integration with ``locales''; unlike the original
  2867 version by Florian Kammller, Isar locales package high-level proof
  2868 contexts rather than raw logical ones (e.g. we admit to include
  2869 attributes everywhere); operations on locales include merge and
  2870 rename; support for implicit arguments (``structures''); simultaneous
  2871 type-inference over imports and text; see also HOL/ex/Locales.thy for
  2872 some examples;
  2873 
  2874 * Pure: the following commands have been ``localized'', supporting a
  2875 target locale specification "(in name)": 'lemma', 'theorem',
  2876 'corollary', 'lemmas', 'theorems', 'declare'; the results will be
  2877 stored both within the locale and at the theory level (exported and
  2878 qualified by the locale name);
  2879 
  2880 * Pure: theory goals may now be specified in ``long'' form, with
  2881 ad-hoc contexts consisting of arbitrary locale elements. for example
  2882 ``lemma foo: fixes x assumes "A x" shows "B x"'' (local syntax and
  2883 definitions may be given, too); the result is a meta-level rule with
  2884 the context elements being discharged in the obvious way;
  2885 
  2886 * Pure: new proof command 'using' allows to augment currently used
  2887 facts after a goal statement ('using' is syntactically analogous to
  2888 'apply', but acts on the goal's facts only); this allows chained facts
  2889 to be separated into parts given before and after a claim, as in
  2890 ``from a and b have C using d and e <proof>'';
  2891 
  2892 * Pure: renamed "antecedent" case to "rule_context";
  2893 
  2894 * Pure: new 'judgment' command records explicit information about the
  2895 object-logic embedding (used by several tools internally); no longer
  2896 use hard-wired "Trueprop";
  2897 
  2898 * Pure: added 'corollary' command;
  2899 
  2900 * Pure: fixed 'token_translation' command;
  2901 
  2902 * Pure: removed obsolete 'exported' attribute;
  2903 
  2904 * Pure: dummy pattern "_" in is/let is now automatically lifted over
  2905 bound variables: "ALL x. P x --> Q x" (is "ALL x. _ --> ?C x")
  2906 supersedes more cumbersome ... (is "ALL x. _ x --> ?C x");
  2907 
  2908 * Pure: method 'atomize' presents local goal premises as object-level
  2909 statements (atomic meta-level propositions); setup controlled via
  2910 rewrite rules declarations of 'atomize' attribute; example
  2911 application: 'induct' method with proper rule statements in improper
  2912 proof *scripts*;
  2913 
  2914 * Pure: emulation of instantiation tactics (rule_tac, cut_tac, etc.)
  2915 now consider the syntactic context of assumptions, giving a better
  2916 chance to get type-inference of the arguments right (this is
  2917 especially important for locales);
  2918 
  2919 * Pure: "sorry" no longer requires quick_and_dirty in interactive
  2920 mode;
  2921 
  2922 * Pure/obtain: the formal conclusion "thesis", being marked as
  2923 ``internal'', may no longer be reference directly in the text;
  2924 potential INCOMPATIBILITY, may need to use "?thesis" in rare
  2925 situations;
  2926 
  2927 * Pure: generic 'sym' attribute which declares a rule both as pure
  2928 'elim?' and for the 'symmetric' operation;
  2929 
  2930 * Pure: marginal comments ``--'' may now occur just anywhere in the
  2931 text; the fixed correlation with particular command syntax has been
  2932 discontinued;
  2933 
  2934 * Pure: new method 'rules' is particularly well-suited for proof
  2935 search in intuitionistic logic; a bit slower than 'blast' or 'fast',
  2936 but often produces more compact proof terms with less detours;
  2937 
  2938 * Pure/Provers/classical: simplified integration with pure rule
  2939 attributes and methods; the classical "intro?/elim?/dest?"
  2940 declarations coincide with the pure ones; the "rule" method no longer
  2941 includes classically swapped intros; "intro" and "elim" methods no
  2942 longer pick rules from the context; also got rid of ML declarations
  2943 AddXIs/AddXEs/AddXDs; all of this has some potential for
  2944 INCOMPATIBILITY;
  2945 
  2946 * Provers/classical: attribute 'swapped' produces classical inversions
  2947 of introduction rules;
  2948 
  2949 * Provers/simplifier: 'simplified' attribute may refer to explicit
  2950 rules instead of full simplifier context; 'iff' attribute handles
  2951 conditional rules;
  2952 
  2953 * HOL: 'typedef' now allows alternative names for Rep/Abs morphisms;
  2954 
  2955 * HOL: 'recdef' now fails on unfinished automated proofs, use
  2956 "(permissive)" option to recover old behavior;
  2957 
  2958 * HOL: 'inductive' no longer features separate (collective) attributes
  2959 for 'intros' (was found too confusing);
  2960 
  2961 * HOL: properly declared induction rules less_induct and
  2962 wf_induct_rule;
  2963 
  2964 
  2965 *** HOL ***
  2966 
  2967 * HOL: moved over to sane numeral syntax; the new policy is as
  2968 follows:
  2969 
  2970   - 0 and 1 are polymorphic constants, which are defined on any
  2971   numeric type (nat, int, real etc.);
  2972 
  2973   - 2, 3, 4, ... and -1, -2, -3, ... are polymorphic numerals, based
  2974   binary representation internally;
  2975 
  2976   - type nat has special constructor Suc, and generally prefers Suc 0
  2977   over 1::nat and Suc (Suc 0) over 2::nat;
  2978 
  2979 This change may cause significant problems of INCOMPATIBILITY; here
  2980 are some hints on converting existing sources:
  2981 
  2982   - due to the new "num" token, "-0" and "-1" etc. are now atomic
  2983   entities, so expressions involving "-" (unary or binary minus) need
  2984   to be spaced properly;
  2985 
  2986   - existing occurrences of "1" may need to be constraint "1::nat" or
  2987   even replaced by Suc 0; similar for old "2";
  2988 
  2989   - replace "#nnn" by "nnn", and "#-nnn" by "-nnn";
  2990 
  2991   - remove all special provisions on numerals in proofs;
  2992 
  2993 * HOL: simp rules nat_number expand numerals on nat to Suc/0
  2994 representation (depends on bin_arith_simps in the default context);
  2995 
  2996 * HOL: symbolic syntax for x^2 (numeral 2);
  2997 
  2998 * HOL: the class of all HOL types is now called "type" rather than
  2999 "term"; INCOMPATIBILITY, need to adapt references to this type class
  3000 in axclass/classes, instance/arities, and (usually rare) occurrences
  3001 in typings (of consts etc.); internally the class is called
  3002 "HOL.type", ML programs should refer to HOLogic.typeS;
  3003 
  3004 * HOL/record package improvements:
  3005   - new derived operations "fields" to build a partial record section,
  3006     "extend" to promote a fixed record to a record scheme, and
  3007     "truncate" for the reverse; cf. theorems "xxx.defs", which are *not*
  3008     declared as simp by default;
  3009   - shared operations ("more", "fields", etc.) now need to be always
  3010     qualified) --- potential INCOMPATIBILITY;
  3011   - removed "make_scheme" operations (use "make" with "extend") --
  3012     INCOMPATIBILITY;
  3013   - removed "more" class (simply use "term") -- INCOMPATIBILITY;
  3014   - provides cases/induct rules for use with corresponding Isar
  3015     methods (for concrete records, record schemes, concrete more
  3016     parts, and schematic more parts -- in that order);
  3017   - internal definitions directly based on a light-weight abstract
  3018     theory of product types over typedef rather than datatype;
  3019 
  3020 * HOL: generic code generator for generating executable ML code from
  3021 specifications; specific support for HOL constructs such as inductive
  3022 datatypes and sets, as well as recursive functions; can be invoked
  3023 via 'generate_code' theory section;
  3024 
  3025 * HOL: canonical cases/induct rules for n-tuples (n = 3..7);
  3026 
  3027 * HOL: consolidated and renamed several theories.  In particular:
  3028         Ord.thy has been absorbed into HOL.thy
  3029         String.thy has been absorbed into List.thy
  3030 
  3031 * HOL: concrete setsum syntax "\<Sum>i:A. b" == "setsum (%i. b) A"
  3032 (beware of argument permutation!);
  3033 
  3034 * HOL: linorder_less_split superseded by linorder_cases;
  3035 
  3036 * HOL/List: "nodups" renamed to "distinct";
  3037 
  3038 * HOL: added "The" definite description operator; move Hilbert's "Eps"
  3039 to peripheral theory "Hilbert_Choice"; some INCOMPATIBILITIES:
  3040   - Ex_def has changed, now need to use some_eq_ex
  3041 
  3042 * HOL: made split_all_tac safe; EXISTING PROOFS MAY FAIL OR LOOP, so
  3043 in this (rare) case use:
  3044 
  3045   delSWrapper "split_all_tac"
  3046   addSbefore ("unsafe_split_all_tac", unsafe_split_all_tac)
  3047 
  3048 * HOL: added safe wrapper "split_conv_tac" to claset; EXISTING PROOFS
  3049 MAY FAIL;
  3050 
  3051 * HOL: introduced f^n = f o ... o f; warning: due to the limits of
  3052 Isabelle's type classes, ^ on functions and relations has too general
  3053 a domain, namely ('a * 'b) set and 'a => 'b; this means that it may be
  3054 necessary to attach explicit type constraints;
  3055 
  3056 * HOL/Relation: the prefix name of the infix "O" has been changed from
  3057 "comp" to "rel_comp"; INCOMPATIBILITY: a few theorems have been
  3058 renamed accordingly (eg "compI" -> "rel_compI").
  3059 
  3060 * HOL: syntax translations now work properly with numerals and records
  3061 expressions;
  3062 
  3063 * HOL: bounded abstraction now uses syntax "%" / "\<lambda>" instead
  3064 of "lam" -- INCOMPATIBILITY;
  3065 
  3066 * HOL: got rid of some global declarations (potential INCOMPATIBILITY
  3067 for ML tools): const "()" renamed "Product_Type.Unity", type "unit"
  3068 renamed "Product_Type.unit";
  3069 
  3070 * HOL: renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
  3071 
  3072 * HOL: removed obsolete theorem "optionE" (use "option.exhaust", or
  3073 the "cases" method);
  3074 
  3075 * HOL/GroupTheory: group theory examples including Sylow's theorem (by
  3076 Florian Kammller);
  3077 
  3078 * HOL/IMP: updated and converted to new-style theory format; several
  3079 parts turned into readable document, with proper Isar proof texts and
  3080 some explanations (by Gerwin Klein);
  3081 
  3082 * HOL-Real: added Complex_Numbers (by Gertrud Bauer);
  3083 
  3084 * HOL-Hyperreal is now a logic image;
  3085 
  3086 
  3087 *** HOLCF ***
  3088 
  3089 * Isar: consts/constdefs supports mixfix syntax for continuous
  3090 operations;
  3091 
  3092 * Isar: domain package adapted to new-style theory format, e.g. see
  3093 HOLCF/ex/Dnat.thy;
  3094 
  3095 * theory Lift: proper use of rep_datatype lift instead of ML hacks --
  3096 potential INCOMPATIBILITY; now use plain induct_tac instead of former
  3097 lift.induct_tac, always use UU instead of Undef;
  3098 
  3099 * HOLCF/IMP: updated and converted to new-style theory;
  3100 
  3101 
  3102 *** ZF ***
  3103 
  3104 * Isar: proper integration of logic-specific tools and packages,
  3105 including theory commands '(co)inductive', '(co)datatype',
  3106 'rep_datatype', 'inductive_cases', as well as methods 'ind_cases',
  3107 'induct_tac', 'case_tac', and 'typecheck' (with attribute 'TC');
  3108 
  3109 * theory Main no longer includes AC; for the Axiom of Choice, base
  3110 your theory on Main_ZFC;
  3111 
  3112 * the integer library now covers quotients and remainders, with many
  3113 laws relating division to addition, multiplication, etc.;
  3114 
  3115 * ZF/UNITY: Chandy and Misra's UNITY is now available in ZF, giving a
  3116 typeless version of the formalism;
  3117 
  3118 * ZF/AC, Coind, IMP, Resid: updated and converted to new-style theory
  3119 format;
  3120 
  3121 * ZF/Induct: new directory for examples of inductive definitions,
  3122 including theory Multiset for multiset orderings; converted to
  3123 new-style theory format;
  3124 
  3125 * ZF: many new theorems about lists, ordinals, etc.;
  3126 
  3127 
  3128 *** General ***
  3129 
  3130 * Pure/kernel: meta-level proof terms (by Stefan Berghofer); reference
  3131 variable proof controls level of detail: 0 = no proofs (only oracle
  3132 dependencies), 1 = lemma dependencies, 2 = compact proof terms; see
  3133 also ref manual for further ML interfaces;
  3134 
  3135 * Pure/axclass: removed obsolete ML interface
  3136 goal_subclass/goal_arity;
  3137 
  3138 * Pure/syntax: new token syntax "num" for plain numerals (without "#"
  3139 of "xnum"); potential INCOMPATIBILITY, since -0, -1 etc. are now
  3140 separate tokens, so expressions involving minus need to be spaced
  3141 properly;
  3142 
  3143 * Pure/syntax: support non-oriented infixes, using keyword "infix"
  3144 rather than "infixl" or "infixr";
  3145 
  3146 * Pure/syntax: concrete syntax for dummy type variables admits genuine
  3147 sort constraint specifications in type inference; e.g. "x::_::foo"
  3148 ensures that the type of "x" is of sort "foo" (but not necessarily a
  3149 type variable);
  3150 
  3151 * Pure/syntax: print modes "type_brackets" and "no_type_brackets"
  3152 control output of nested => (types); the default behavior is
  3153 "type_brackets";
  3154 
  3155 * Pure/syntax: builtin parse translation for "_constify" turns valued
  3156 tokens into AST constants;
  3157 
  3158 * Pure/syntax: prefer later declarations of translations and print
  3159 translation functions; potential INCOMPATIBILITY: need to reverse
  3160 multiple declarations for same syntax element constant;
  3161 
  3162 * Pure/show_hyps reset by default (in accordance to existing Isar
  3163 practice);
  3164 
  3165 * Provers/classical: renamed addaltern to addafter, addSaltern to
  3166 addSafter;
  3167 
  3168 * Provers/clasimp: ``iff'' declarations now handle conditional rules
  3169 as well;
  3170 
  3171 * system: tested support for MacOS X; should be able to get Isabelle +
  3172 Proof General to work in a plain Terminal after installing Poly/ML
  3173 (e.g. from the Isabelle distribution area) and GNU bash alone
  3174 (e.g. from http://www.apple.com); full X11, XEmacs and X-Symbol
  3175 support requires further installations, e.g. from
  3176 http://fink.sourceforge.net/);
  3177 
  3178 * system: support Poly/ML 4.1.1 (able to manage larger heaps);
  3179 
  3180 * system: reduced base memory usage by Poly/ML (approx. 20 MB instead
  3181 of 40 MB), cf. ML_OPTIONS;
  3182 
  3183 * system: Proof General keywords specification is now part of the
  3184 Isabelle distribution (see etc/isar-keywords.el);
  3185 
  3186 * system: support for persistent Proof General sessions (refrain from
  3187 outdating all loaded theories on startup); user may create writable
  3188 logic images like this: ``isabelle -q HOL Test'';
  3189 
  3190 * system: smart selection of Isabelle process versus Isabelle
  3191 interface, accommodates case-insensitive file systems (e.g. HFS+); may
  3192 run both "isabelle" and "Isabelle" even if file names are badly
  3193 damaged (executable inspects the case of the first letter of its own
  3194 name); added separate "isabelle-process" and "isabelle-interface";
  3195 
  3196 * system: refrain from any attempt at filtering input streams; no
  3197 longer support ``8bit'' encoding of old isabelle font, instead proper
  3198 iso-latin characters may now be used; the related isatools
  3199 "symbolinput" and "nonascii" have disappeared as well;
  3200 
  3201 * system: removed old "xterm" interface (the print modes "xterm" and
  3202 "xterm_color" are still available for direct use in a suitable
  3203 terminal);
  3204 
  3205 
  3206 
  3207 New in Isabelle99-2 (February 2001)
  3208 -----------------------------------
  3209 
  3210 *** Overview of INCOMPATIBILITIES ***
  3211 
  3212 * HOL: please note that theories in the Library and elsewhere often use the
  3213 new-style (Isar) format; to refer to their theorems in an ML script you must
  3214 bind them to ML identifers by e.g.      val thm_name = thm "thm_name";
  3215 
  3216 * HOL: inductive package no longer splits induction rule aggressively,
  3217 but only as far as specified by the introductions given; the old
  3218 format may be recovered via ML function complete_split_rule or attribute
  3219 'split_rule (complete)';
  3220 
  3221 * HOL: induct renamed to lfp_induct, lfp_Tarski to lfp_unfold,
  3222 gfp_Tarski to gfp_unfold;
  3223 
  3224 * HOL: contrapos, contrapos2 renamed to contrapos_nn, contrapos_pp;
  3225 
  3226 * HOL: infix "dvd" now has priority 50 rather than 70 (because it is a
  3227 relation); infix "^^" has been renamed "``"; infix "``" has been
  3228 renamed "`"; "univalent" has been renamed "single_valued";
  3229 
  3230 * HOL/Real: "rinv" and "hrinv" replaced by overloaded "inverse"
  3231 operation;
  3232 
  3233 * HOLCF: infix "`" has been renamed "$"; the symbol syntax is \<cdot>;
  3234 
  3235 * Isar: 'obtain' no longer declares "that" fact as simp/intro;
  3236 
  3237 * Isar/HOL: method 'induct' now handles non-atomic goals; as a
  3238 consequence, it is no longer monotonic wrt. the local goal context
  3239 (which is now passed through the inductive cases);
  3240 
  3241 * Document preparation: renamed standard symbols \<ll> to \<lless> and
  3242 \<gg> to \<ggreater>;
  3243 
  3244 
  3245 *** Document preparation ***
  3246 
  3247 * \isabellestyle{NAME} selects version of Isabelle output (currently
  3248 available: are "it" for near math-mode best-style output, "sl" for
  3249 slanted text style, and "tt" for plain type-writer; if no
  3250 \isabellestyle command is given, output is according to slanted
  3251 type-writer);
  3252 
  3253 * support sub/super scripts (for single symbols only), input syntax is
  3254 like this: "A\<^sup>*" or "A\<^sup>\<star>";
  3255 
  3256 * some more standard symbols; see Appendix A of the system manual for
  3257 the complete list of symbols defined in isabellesym.sty;
  3258 
  3259 * improved isabelle style files; more abstract symbol implementation
  3260 (should now use \isamath{...} and \isatext{...} in custom symbol
  3261 definitions);
  3262 
  3263 * antiquotation @{goals} and @{subgoals} for output of *dynamic* goals
  3264 state; Note that presentation of goal states does not conform to
  3265 actual human-readable proof documents.  Please do not include goal
  3266 states into document output unless you really know what you are doing!
  3267 
  3268 * proper indentation of antiquoted output with proportional LaTeX
  3269 fonts;
  3270 
  3271 * no_document ML operator temporarily disables LaTeX document
  3272 generation;
  3273 
  3274 * isatool unsymbolize tunes sources for plain ASCII communication;
  3275 
  3276 
  3277 *** Isar ***
  3278 
  3279 * Pure: Isar now suffers initial goal statements to contain unbound
  3280 schematic variables (this does not conform to actual readable proof
  3281 documents, due to unpredictable outcome and non-compositional proof
  3282 checking); users who know what they are doing may use schematic goals
  3283 for Prolog-style synthesis of proven results;
  3284 
  3285 * Pure: assumption method (an implicit finishing) now handles actual
  3286 rules as well;
  3287 
  3288 * Pure: improved 'obtain' --- moved to Pure, insert "that" into
  3289 initial goal, declare "that" only as Pure intro (only for single
  3290 steps); the "that" rule assumption may now be involved in implicit
  3291 finishing, thus ".." becomes a feasible for trivial obtains;
  3292 
  3293 * Pure: default proof step now includes 'intro_classes'; thus trivial
  3294 instance proofs may be performed by "..";
  3295 
  3296 * Pure: ?thesis / ?this / "..." now work for pure meta-level
  3297 statements as well;
  3298 
  3299 * Pure: more robust selection of calculational rules;
  3300 
  3301 * Pure: the builtin notion of 'finished' goal now includes the ==-refl
  3302 rule (as well as the assumption rule);
  3303 
  3304 * Pure: 'thm_deps' command visualizes dependencies of theorems and
  3305 lemmas, using the graph browser tool;
  3306 
  3307 * Pure: predict failure of "show" in interactive mode;
  3308 
  3309 * Pure: 'thms_containing' now takes actual terms as arguments;
  3310 
  3311 * HOL: improved method 'induct' --- now handles non-atomic goals
  3312 (potential INCOMPATIBILITY); tuned error handling;
  3313 
  3314 * HOL: cases and induct rules now provide explicit hints about the
  3315 number of facts to be consumed (0 for "type" and 1 for "set" rules);
  3316 any remaining facts are inserted into the goal verbatim;
  3317 
  3318 * HOL: local contexts (aka cases) may now contain term bindings as
  3319 well; the 'cases' and 'induct' methods new provide a ?case binding for
  3320 the result to be shown in each case;
  3321 
  3322 * HOL: added 'recdef_tc' command;
  3323 
  3324 * isatool convert assists in eliminating legacy ML scripts;
  3325 
  3326 
  3327 *** HOL ***
  3328 
  3329 * HOL/Library: a collection of generic theories to be used together
  3330 with main HOL; the theory loader path already includes this directory
  3331 by default; the following existing theories have been moved here:
  3332 HOL/Induct/Multiset, HOL/Induct/Acc (as Accessible_Part), HOL/While
  3333 (as While_Combinator), HOL/Lex/Prefix (as List_Prefix);
  3334 
  3335 * HOL/Unix: "Some aspects of Unix file-system security", a typical
  3336 modelling and verification task performed in Isabelle/HOL +
  3337 Isabelle/Isar + Isabelle document preparation (by Markus Wenzel).
  3338 
  3339 * HOL/Algebra: special summation operator SUM no longer exists, it has
  3340 been replaced by setsum; infix 'assoc' now has priority 50 (like
  3341 'dvd'); axiom 'one_not_zero' has been moved from axclass 'ring' to
  3342 'domain', this makes the theory consistent with mathematical
  3343 literature;
  3344 
  3345 * HOL basics: added overloaded operations "inverse" and "divide"
  3346 (infix "/"), syntax for generic "abs" operation, generic summation
  3347 operator \<Sum>;
  3348 
  3349 * HOL/typedef: simplified package, provide more useful rules (see also
  3350 HOL/subset.thy);
  3351 
  3352 * HOL/datatype: induction rule for arbitrarily branching datatypes is
  3353 now expressed as a proper nested rule (old-style tactic scripts may
  3354 require atomize_strip_tac to cope with non-atomic premises);
  3355 
  3356 * HOL: renamed theory "Prod" to "Product_Type", renamed "split" rule
  3357 to "split_conv" (old name still available for compatibility);
  3358 
  3359 * HOL: improved concrete syntax for strings (e.g. allows translation
  3360 rules with string literals);
  3361 
  3362 * HOL-Real-Hyperreal: this extends HOL-Real with the hyperreals
  3363  and Fleuriot's mechanization of analysis, including the transcendental
  3364  functions for the reals;
  3365 
  3366 * HOL/Real, HOL/Hyperreal: improved arithmetic simplification;
  3367 
  3368 
  3369 *** CTT ***
  3370 
  3371 * CTT: x-symbol support for Pi, Sigma, -->, : (membership); note that
  3372 "lam" is displayed as TWO lambda-symbols
  3373 
  3374 * CTT: theory Main now available, containing everything (that is, Bool
  3375 and Arith);
  3376 
  3377 
  3378 *** General ***
  3379 
  3380 * Pure: the Simplifier has been implemented properly as a derived rule
  3381 outside of the actual kernel (at last!); the overall performance
  3382 penalty in practical applications is about 50%, while reliability of
  3383 the Isabelle inference kernel has been greatly improved;
  3384 
  3385 * print modes "brackets" and "no_brackets" control output of nested =>
  3386 (types) and ==> (props); the default behaviour is "brackets";
  3387 
  3388 * Provers: fast_tac (and friends) now handle actual object-logic rules
  3389 as assumptions as well;
  3390 
  3391 * system: support Poly/ML 4.0;
  3392 
  3393 * system: isatool install handles KDE version 1 or 2;
  3394 
  3395 
  3396 
  3397 New in Isabelle99-1 (October 2000)
  3398 ----------------------------------
  3399 
  3400 *** Overview of INCOMPATIBILITIES ***
  3401 
  3402 * HOL: simplification of natural numbers is much changed; to partly
  3403 recover the old behaviour (e.g. to prevent n+n rewriting to #2*n)
  3404 issue the following ML commands:
  3405 
  3406   Delsimprocs Nat_Numeral_Simprocs.cancel_numerals;
  3407   Delsimprocs [Nat_Numeral_Simprocs.combine_numerals];
  3408 
  3409 * HOL: simplification no longer dives into case-expressions; this is
  3410 controlled by "t.weak_case_cong" for each datatype t;
  3411 
  3412 * HOL: nat_less_induct renamed to less_induct;
  3413 
  3414 * HOL: systematic renaming of the SOME (Eps) rules, may use isatool
  3415 fixsome to patch .thy and .ML sources automatically;
  3416 
  3417   select_equality  -> some_equality
  3418   select_eq_Ex     -> some_eq_ex
  3419   selectI2EX       -> someI2_ex
  3420   selectI2         -> someI2
  3421   selectI          -> someI
  3422   select1_equality -> some1_equality
  3423   Eps_sym_eq       -> some_sym_eq_trivial
  3424   Eps_eq           -> some_eq_trivial
  3425 
  3426 * HOL: exhaust_tac on datatypes superceded by new generic case_tac;
  3427 
  3428 * HOL: removed obsolete theorem binding expand_if (refer to split_if
  3429 instead);
  3430 
  3431 * HOL: the recursion equations generated by 'recdef' are now called
  3432 f.simps instead of f.rules;
  3433 
  3434 * HOL: qed_spec_mp now also handles bounded ALL as well;
  3435 
  3436 * HOL: 0 is now overloaded, so the type constraint ":: nat" may
  3437 sometimes be needed;
  3438 
  3439 * HOL: the constant for "f``x" is now "image" rather than "op ``";
  3440 
  3441 * HOL: the constant for "f-``x" is now "vimage" rather than "op -``";
  3442 
  3443 * HOL: the disjoint sum is now "<+>" instead of "Plus"; the cartesian
  3444 product is now "<*>" instead of "Times"; the lexicographic product is
  3445 now "<*lex*>" instead of "**";
  3446 
  3447 * HOL: theory Sexp is now in HOL/Induct examples (it used to be part
  3448 of main HOL, but was unused); better use HOL's datatype package;
  3449 
  3450 * HOL: removed "symbols" syntax for constant "override" of theory Map;
  3451 the old syntax may be recovered as follows:
  3452 
  3453   syntax (symbols)
  3454     override  :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)"
  3455       (infixl "\\<oplus>" 100)
  3456 
  3457 * HOL/Real: "rabs" replaced by overloaded "abs" function;
  3458 
  3459 * HOL/ML: even fewer consts are declared as global (see theories Ord,
  3460 Lfp, Gfp, WF); this only affects ML packages that refer to const names
  3461 internally;
  3462 
  3463 * HOL and ZF: syntax for quotienting wrt an equivalence relation
  3464 changed from A/r to A//r;
  3465 
  3466 * ZF: new treatment of arithmetic (nat & int) may break some old
  3467 proofs;
  3468 
  3469 * Isar: renamed some attributes (RS -> THEN, simplify -> simplified,
  3470 rulify -> rule_format, elimify -> elim_format, ...);
  3471 
  3472 * Isar/Provers: intro/elim/dest attributes changed; renamed
  3473 intro/intro!/intro!! flags to intro!/intro/intro? (in most cases, one
  3474 should have to change intro!! to intro? only); replaced "delrule" by
  3475 "rule del";
  3476 
  3477 * Isar/HOL: renamed "intrs" to "intros" in inductive definitions;
  3478 
  3479 * Provers: strengthened force_tac by using new first_best_tac;
  3480 
  3481 * LaTeX document preparation: several changes of isabelle.sty (see
  3482 lib/texinputs);
  3483 
  3484 
  3485 *** Document preparation ***
  3486 
  3487 * formal comments (text blocks etc.) in new-style theories may now
  3488 contain antiquotations of thm/prop/term/typ/text to be presented
  3489 according to latex print mode; concrete syntax is like this:
  3490 @{term[show_types] "f(x) = a + x"};
  3491 
  3492 * isatool mkdir provides easy setup of Isabelle session directories,
  3493 including proper document sources;
  3494 
  3495 * generated LaTeX sources are now deleted after successful run
  3496 (isatool document -c); may retain a copy somewhere else via -D option
  3497 of isatool usedir;
  3498 
  3499 * isatool usedir -D now lets isatool latex -o sty update the Isabelle
  3500 style files, achieving self-contained LaTeX sources and simplifying
  3501 LaTeX debugging;
  3502 
  3503 * old-style theories now produce (crude) LaTeX output as well;
  3504 
  3505 * browser info session directories are now self-contained (may be put
  3506 on WWW server seperately); improved graphs of nested sessions; removed
  3507 graph for 'all sessions';
  3508 
  3509 * several improvements in isabelle style files; \isabellestyle{it}
  3510 produces fake math mode output; \isamarkupheader is now \section by
  3511 default; see lib/texinputs/isabelle.sty etc.;
  3512 
  3513 
  3514 *** Isar ***
  3515 
  3516 * Isar/Pure: local results and corresponding term bindings are now
  3517 subject to Hindley-Milner polymorphism (similar to ML); this
  3518 accommodates incremental type-inference very nicely;
  3519 
  3520 * Isar/Pure: new derived language element 'obtain' supports
  3521 generalized existence reasoning;
  3522 
  3523 * Isar/Pure: new calculational elements 'moreover' and 'ultimately'
  3524 support accumulation of results, without applying any rules yet;
  3525 useful to collect intermediate results without explicit name
  3526 references, and for use with transitivity rules with more than 2
  3527 premises;
  3528 
  3529 * Isar/Pure: scalable support for case-analysis type proofs: new
  3530 'case' language element refers to local contexts symbolically, as
  3531 produced by certain proof methods; internally, case names are attached
  3532 to theorems as "tags";
  3533 
  3534 * Isar/Pure: theory command 'hide' removes declarations from
  3535 class/type/const name spaces;
  3536 
  3537 * Isar/Pure: theory command 'defs' supports option "(overloaded)" to
  3538 indicate potential overloading;
  3539 
  3540 * Isar/Pure: changed syntax of local blocks from {{ }} to { };
  3541 
  3542 * Isar/Pure: syntax of sorts made 'inner', i.e. have to write
  3543 "{a,b,c}" instead of {a,b,c};
  3544 
  3545 * Isar/Pure now provides its own version of intro/elim/dest
  3546 attributes; useful for building new logics, but beware of confusion
  3547 with the version in Provers/classical;
  3548 
  3549 * Isar/Pure: the local context of (non-atomic) goals is provided via
  3550 case name 'antecedent';
  3551 
  3552 * Isar/Pure: removed obsolete 'transfer' attribute (transfer of thms
  3553 to the current context is now done automatically);
  3554 
  3555 * Isar/Pure: theory command 'method_setup' provides a simple interface
  3556 for definining proof methods in ML;
  3557 
  3558 * Isar/Provers: intro/elim/dest attributes changed; renamed
  3559 intro/intro!/intro!! flags to intro!/intro/intro? (INCOMPATIBILITY, in
  3560 most cases, one should have to change intro!! to intro? only);
  3561 replaced "delrule" by "rule del";
  3562 
  3563 * Isar/Provers: new 'hypsubst' method, plain 'subst' method and
  3564 'symmetric' attribute (the latter supercedes [RS sym]);
  3565 
  3566 * Isar/Provers: splitter support (via 'split' attribute and 'simp'
  3567 method modifier); 'simp' method: 'only:' modifier removes loopers as
  3568 well (including splits);
  3569 
  3570 * Isar/Provers: Simplifier and Classical methods now support all kind
  3571 of modifiers used in the past, including 'cong', 'iff', etc.
  3572 
  3573 * Isar/Provers: added 'fastsimp' and 'clarsimp' methods (combination
  3574 of Simplifier and Classical reasoner);
  3575 
  3576 * Isar/HOL: new proof method 'cases' and improved version of 'induct'
  3577 now support named cases; major packages (inductive, datatype, primrec,
  3578 recdef) support case names and properly name parameters;
  3579 
  3580 * Isar/HOL: new transitivity rules for substitution in inequalities --
  3581 monotonicity conditions are extracted to be proven at end of
  3582 calculations;
  3583 
  3584 * Isar/HOL: removed 'case_split' thm binding, should use 'cases' proof
  3585 method anyway;
  3586 
  3587 * Isar/HOL: removed old expand_if = split_if; theorems if_splits =
  3588 split_if split_if_asm; datatype package provides theorems foo.splits =
  3589 foo.split foo.split_asm for each datatype;
  3590 
  3591 * Isar/HOL: tuned inductive package, rename "intrs" to "intros"
  3592 (potential INCOMPATIBILITY), emulation of mk_cases feature for proof
  3593 scripts: new 'inductive_cases' command and 'ind_cases' method; (Note:
  3594 use "(cases (simplified))" method in proper proof texts);
  3595 
  3596 * Isar/HOL: added global 'arith_split' attribute for 'arith' method;
  3597 
  3598 * Isar: names of theorems etc. may be natural numbers as well;
  3599 
  3600 * Isar: 'pr' command: optional arguments for goals_limit and
  3601 ProofContext.prems_limit; no longer prints theory contexts, but only
  3602 proof states;
  3603 
  3604 * Isar: diagnostic commands 'pr', 'thm', 'prop', 'term', 'typ' admit
  3605 additional print modes to be specified; e.g. "pr(latex)" will print
  3606 proof state according to the Isabelle LaTeX style;
  3607 
  3608 * Isar: improved support for emulating tactic scripts, including proof
  3609 methods 'rule_tac' etc., 'cut_tac', 'thin_tac', 'subgoal_tac',
  3610 'rename_tac', 'rotate_tac', 'tactic', and 'case_tac' / 'induct_tac'
  3611 (for HOL datatypes);
  3612 
  3613 * Isar: simplified (more robust) goal selection of proof methods: 1st
  3614 goal, all goals, or explicit goal specifier (tactic emulation); thus
  3615 'proof method scripts' have to be in depth-first order;
  3616 
  3617 * Isar: tuned 'let' syntax: replaced 'as' keyword by 'and';
  3618 
  3619 * Isar: removed 'help' command, which hasn't been too helpful anyway;
  3620 should instead use individual commands for printing items
  3621 (print_commands, print_methods etc.);
  3622 
  3623 * Isar: added 'nothing' --- the empty list of theorems;
  3624 
  3625 
  3626 *** HOL ***
  3627 
  3628 * HOL/MicroJava: formalization of a fragment of Java, together with a
  3629 corresponding virtual machine and a specification of its bytecode
  3630 verifier and a lightweight bytecode verifier, including proofs of
  3631 type-safety; by Gerwin Klein, Tobias Nipkow, David von Oheimb, and
  3632 Cornelia Pusch (see also the homepage of project Bali at
  3633 http://isabelle.in.tum.de/Bali/);
  3634 
  3635 * HOL/Algebra: new theory of rings and univariate polynomials, by
  3636 Clemens Ballarin;
  3637 
  3638 * HOL/NumberTheory: fundamental Theorem of Arithmetic, Chinese
  3639 Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, by Thomas M
  3640 Rasmussen;
  3641 
  3642 * HOL/Lattice: fundamental concepts of lattice theory and order
  3643 structures, including duals, properties of bounds versus algebraic
  3644 laws, lattice operations versus set-theoretic ones, the Knaster-Tarski
  3645 Theorem for complete lattices etc.; may also serve as a demonstration
  3646 for abstract algebraic reasoning using axiomatic type classes, and
  3647 mathematics-style proof in Isabelle/Isar; by Markus Wenzel;
  3648 
  3649 * HOL/Prolog: a (bare-bones) implementation of Lambda-Prolog, by David
  3650 von Oheimb;
  3651 
  3652 * HOL/IMPP: extension of IMP with local variables and mutually
  3653 recursive procedures, by David von Oheimb;
  3654 
  3655 * HOL/Lambda: converted into new-style theory and document;
  3656 
  3657 * HOL/ex/Multiquote: example of multiple nested quotations and
  3658 anti-quotations -- basically a generalized version of de-Bruijn
  3659 representation; very useful in avoiding lifting of operations;
  3660 
  3661 * HOL/record: added general record equality rule to simpset; fixed
  3662 select-update simplification procedure to handle extended records as
  3663 well; admit "r" as field name;
  3664 
  3665 * HOL: 0 is now overloaded over the new sort "zero", allowing its use with
  3666 other numeric types and also as the identity of groups, rings, etc.;
  3667 
  3668 * HOL: new axclass plus_ac0 for addition with the AC-laws and 0 as identity.
  3669 Types nat and int belong to this axclass;
  3670 
  3671 * HOL: greatly improved simplification involving numerals of type nat, int, real:
  3672    (i + #8 + j) = Suc k simplifies to  #7 + (i + j) = k
  3673    i*j + k + j*#3*i     simplifies to  #4*(i*j) + k
  3674   two terms #m*u and #n*u are replaced by #(m+n)*u
  3675     (where #m, #n and u can implicitly be 1; this is simproc combine_numerals)
  3676   and the term/formula #m*u+x ~~ #n*u+y simplifies simplifies to #(m-n)+x ~~ y
  3677     or x ~~ #(n-m)+y, where ~~ is one of = < <= or - (simproc cancel_numerals);
  3678 
  3679 * HOL: meson_tac is available (previously in ex/meson.ML); it is a
  3680 powerful prover for predicate logic but knows nothing of clasets; see
  3681 ex/mesontest.ML and ex/mesontest2.ML for example applications;
  3682 
  3683 * HOL: new version of "case_tac" subsumes both boolean case split and
  3684 "exhaust_tac" on datatypes; INCOMPATIBILITY: exhaust_tac no longer
  3685 exists, may define val exhaust_tac = case_tac for ad-hoc portability;
  3686 
  3687 * HOL: simplification no longer dives into case-expressions: only the
  3688 selector expression is simplified, but not the remaining arms; to
  3689 enable full simplification of case-expressions for datatype t, you may
  3690 remove t.weak_case_cong from the simpset, either globally (Delcongs
  3691 [thm"t.weak_case_cong"];) or locally (delcongs [...]).
  3692 
  3693 * HOL/recdef: the recursion equations generated by 'recdef' for
  3694 function 'f' are now called f.simps instead of f.rules; if all
  3695 termination conditions are proved automatically, these simplification
  3696 rules are added to the simpset, as in primrec; rules may be named
  3697 individually as well, resulting in a separate list of theorems for
  3698 each equation;
  3699 
  3700 * HOL/While is a new theory that provides a while-combinator. It
  3701 permits the definition of tail-recursive functions without the
  3702 provision of a termination measure. The latter is necessary once the
  3703 invariant proof rule for while is applied.
  3704 
  3705 * HOL: new (overloaded) notation for the set of elements below/above
  3706 some element: {..u}, {..u(}, {l..}, {)l..}. See theory SetInterval.
  3707 
  3708 * HOL: theorems impI, allI, ballI bound as "strip";
  3709 
  3710 * HOL: new tactic induct_thm_tac: thm -> string -> int -> tactic
  3711 induct_tac th "x1 ... xn" expects th to have a conclusion of the form
  3712 P v1 ... vn and abbreviates res_inst_tac [("v1","x1"),...,("vn","xn")] th;
  3713 
  3714 * HOL/Real: "rabs" replaced by overloaded "abs" function;
  3715 
  3716 * HOL: theory Sexp now in HOL/Induct examples (it used to be part of
  3717 main HOL, but was unused);
  3718 
  3719 * HOL: fewer consts declared as global (e.g. have to refer to
  3720 "Lfp.lfp" instead of "lfp" internally; affects ML packages only);
  3721 
  3722 * HOL: tuned AST representation of nested pairs, avoiding bogus output
  3723 in case of overlap with user translations (e.g. judgements over
  3724 tuples); (note that the underlying logical represenation is still
  3725 bogus);
  3726 
  3727 
  3728 *** ZF ***
  3729 
  3730 * ZF: simplification automatically cancels common terms in arithmetic
  3731 expressions over nat and int;
  3732 
  3733 * ZF: new treatment of nat to minimize type-checking: all operators
  3734 coerce their operands to a natural number using the function natify,
  3735 making the algebraic laws unconditional;
  3736 
  3737 * ZF: as above, for int: operators coerce their operands to an integer
  3738 using the function intify;
  3739 
  3740 * ZF: the integer library now contains many of the usual laws for the
  3741 orderings, including $<=, and monotonicity laws for $+ and $*;
  3742 
  3743 * ZF: new example ZF/ex/NatSum to demonstrate integer arithmetic
  3744 simplification;
  3745 
  3746 * FOL and ZF: AddIffs now available, giving theorems of the form P<->Q
  3747 to the simplifier and classical reasoner simultaneously;
  3748 
  3749 
  3750 *** General ***
  3751 
  3752 * Provers: blast_tac now handles actual object-logic rules as
  3753 assumptions; note that auto_tac uses blast_tac internally as well;
  3754 
  3755 * Provers: new functions rulify/rulify_no_asm: thm -> thm for turning
  3756 outer -->/All/Ball into ==>/!!; qed_spec_mp now uses rulify_no_asm;
  3757 
  3758 * Provers: delrules now handles destruct rules as well (no longer need
  3759 explicit make_elim);
  3760 
  3761 * Provers: Blast_tac now warns of and ignores "weak elimination rules" e.g.
  3762   [| inj ?f;          ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
  3763 use instead the strong form,
  3764   [| inj ?f; ~ ?W ==> ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
  3765 in HOL, FOL and ZF the function cla_make_elim will create such rules
  3766 from destruct-rules;
  3767 
  3768 * Provers: Simplifier.easy_setup provides a fast path to basic
  3769 Simplifier setup for new object-logics;
  3770 
  3771 * Pure: AST translation rules no longer require constant head on LHS;
  3772 
  3773 * Pure: improved name spaces: ambiguous output is qualified; support
  3774 for hiding of names;
  3775 
  3776 * system: smart setup of canonical ML_HOME, ISABELLE_INTERFACE, and
  3777 XSYMBOL_HOME; no longer need to do manual configuration in most
  3778 situations;
  3779 
  3780 * system: compression of ML heaps images may now be controlled via -c
  3781 option of isabelle and isatool usedir (currently only observed by
  3782 Poly/ML);
  3783 
  3784 * system: isatool installfonts may handle X-Symbol fonts as well (very
  3785 useful for remote X11);
  3786 
  3787 * system: provide TAGS file for Isabelle sources;
  3788 
  3789 * ML: infix 'OF' is a version of 'MRS' with more appropriate argument
  3790 order;
  3791 
  3792 * ML: renamed flags Syntax.trace_norm_ast to Syntax.trace_ast; global
  3793 timing flag supersedes proof_timing and Toplevel.trace;
  3794 
  3795 * ML: new combinators |>> and |>>> for incremental transformations
  3796 with secondary results (e.g. certain theory extensions):
  3797 
  3798 * ML: PureThy.add_defs gets additional argument to indicate potential
  3799 overloading (usually false);
  3800 
  3801 * ML: PureThy.add_thms/add_axioms/add_defs now return theorems as
  3802 results;
  3803 
  3804 
  3805 
  3806 New in Isabelle99 (October 1999)
  3807 --------------------------------
  3808 
  3809 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  3810 
  3811 * HOL: The THEN and ELSE parts of conditional expressions (if P then x else y)
  3812 are no longer simplified.  (This allows the simplifier to unfold recursive
  3813 functional programs.)  To restore the old behaviour, declare
  3814 
  3815     Delcongs [if_weak_cong];
  3816 
  3817 * HOL: Removed the obsolete syntax "Compl A"; use -A for set
  3818 complement;
  3819 
  3820 * HOL: the predicate "inj" is now defined by translation to "inj_on";
  3821 
  3822 * HOL/datatype: mutual_induct_tac no longer exists --
  3823   use induct_tac "x_1 ... x_n" instead of mutual_induct_tac ["x_1", ..., "x_n"]
  3824 
  3825 * HOL/typedef: fixed type inference for representing set; type
  3826 arguments now have to occur explicitly on the rhs as type constraints;
  3827 
  3828 * ZF: The con_defs part of an inductive definition may no longer refer
  3829 to constants declared in the same theory;
  3830 
  3831 * HOL, ZF: the function mk_cases, generated by the inductive
  3832 definition package, has lost an argument.  To simplify its result, it
  3833 uses the default simpset instead of a supplied list of theorems.
  3834 
  3835 * HOL/List: the constructors of type list are now Nil and Cons;
  3836 
  3837 * Simplifier: the type of the infix ML functions
  3838         setSSolver addSSolver setSolver addSolver
  3839 is now  simpset * solver -> simpset  where `solver' is a new abstract type
  3840 for packaging solvers. A solver is created via
  3841         mk_solver: string -> (thm list -> int -> tactic) -> solver
  3842 where the string argument is only a comment.
  3843 
  3844 
  3845 *** Proof tools ***
  3846 
  3847 * Provers/Arith/fast_lin_arith.ML contains a functor for creating a
  3848 decision procedure for linear arithmetic. Currently it is used for
  3849 types `nat', `int', and `real' in HOL (see below); it can, should and
  3850 will be instantiated for other types and logics as well.
  3851 
  3852 * The simplifier now accepts rewrite rules with flexible heads, eg
  3853      hom ?f ==> ?f(?x+?y) = ?f ?x + ?f ?y
  3854   They are applied like any rule with a non-pattern lhs, i.e. by first-order
  3855   matching.
  3856 
  3857 
  3858 *** General ***
  3859 
  3860 * New Isabelle/Isar subsystem provides an alternative to traditional
  3861 tactical theorem proving; together with the ProofGeneral/isar user
  3862 interface it offers an interactive environment for developing human
  3863 readable proof documents (Isar == Intelligible semi-automated
  3864 reasoning); for further information see isatool doc isar-ref,
  3865 src/HOL/Isar_examples and http://isabelle.in.tum.de/Isar/
  3866 
  3867 * improved and simplified presentation of theories: better HTML markup
  3868 (including colors), graph views in several sizes; isatool usedir now
  3869 provides a proper interface for user theories (via -P option); actual
  3870 document preparation based on (PDF)LaTeX is available as well (for
  3871 new-style theories only); see isatool doc system for more information;
  3872 
  3873 * native support for Proof General, both for classic Isabelle and
  3874 Isabelle/Isar;
  3875 
  3876 * ML function thm_deps visualizes dependencies of theorems and lemmas,
  3877 using the graph browser tool;
  3878 
  3879 * Isabelle manuals now also available as PDF;
  3880 
  3881 * theory loader rewritten from scratch (may not be fully
  3882 bug-compatible); old loadpath variable has been replaced by show_path,
  3883 add_path, del_path, reset_path functions; new operations such as
  3884 update_thy, touch_thy, remove_thy, use/update_thy_only (see also
  3885 isatool doc ref);
  3886 
  3887 * improved isatool install: option -k creates KDE application icon,
  3888 option -p DIR installs standalone binaries;
  3889 
  3890 * added ML_PLATFORM setting (useful for cross-platform installations);
  3891 more robust handling of platform specific ML images for SML/NJ;
  3892 
  3893 * the settings environment is now statically scoped, i.e. it is never
  3894 created again in sub-processes invoked from isabelle, isatool, or
  3895 Isabelle;
  3896 
  3897 * path element specification '~~' refers to '$ISABELLE_HOME';
  3898 
  3899 * in locales, the "assumes" and "defines" parts may be omitted if
  3900 empty;
  3901 
  3902 * new print_mode "xsymbols" for extended symbol support (e.g. genuine
  3903 long arrows);
  3904 
  3905 * new print_mode "HTML";
  3906 
  3907 * new flag show_tags controls display of tags of theorems (which are
  3908 basically just comments that may be attached by some tools);
  3909 
  3910 * Isamode 2.6 requires patch to accomodate change of Isabelle font
  3911 mode and goal output format:
  3912 
  3913 diff -r Isamode-2.6/elisp/isa-load.el Isamode/elisp/isa-load.el
  3914 244c244
  3915 <       (list (isa-getenv "ISABELLE") "-msymbols" logic-name)
  3916 ---
  3917 >       (list (isa-getenv "ISABELLE") "-misabelle_font" "-msymbols" logic-name)
  3918 diff -r Isabelle-2.6/elisp/isa-proofstate.el Isamode/elisp/isa-proofstate.el
  3919 181c181
  3920 < (defconst proofstate-proofstart-regexp "^Level [0-9]+$"
  3921 ---
  3922 > (defconst proofstate-proofstart-regexp "^Level [0-9]+"
  3923 
  3924 * function bind_thms stores lists of theorems (cf. bind_thm);
  3925 
  3926 * new shorthand tactics ftac, eatac, datac, fatac;
  3927 
  3928 * qed (and friends) now accept "" as result name; in that case the
  3929 theorem is not stored, but proper checks and presentation of the
  3930 result still apply;
  3931 
  3932 * theorem database now also indexes constants "Trueprop", "all",
  3933 "==>", "=="; thus thms_containing, findI etc. may retrieve more rules;
  3934 
  3935 
  3936 *** HOL ***
  3937 
  3938 ** HOL arithmetic **
  3939 
  3940 * There are now decision procedures for linear arithmetic over nat and
  3941 int:
  3942 
  3943 1. arith_tac copes with arbitrary formulae involving `=', `<', `<=',
  3944 `+', `-', `Suc', `min', `max' and numerical constants; other subterms
  3945 are treated as atomic; subformulae not involving type `nat' or `int'
  3946 are ignored; quantified subformulae are ignored unless they are
  3947 positive universal or negative existential. The tactic has to be
  3948 invoked by hand and can be a little bit slow. In particular, the
  3949 running time is exponential in the number of occurrences of `min' and
  3950 `max', and `-' on `nat'.
  3951 
  3952 2. fast_arith_tac is a cut-down version of arith_tac: it only takes
  3953 (negated) (in)equalities among the premises and the conclusion into
  3954 account (i.e. no compound formulae) and does not know about `min' and
  3955 `max', and `-' on `nat'. It is fast and is used automatically by the
  3956 simplifier.
  3957 
  3958 NB: At the moment, these decision procedures do not cope with mixed
  3959 nat/int formulae where the two parts interact, such as `m < n ==>
  3960 int(m) < int(n)'.
  3961 
  3962 * HOL/Numeral provides a generic theory of numerals (encoded
  3963 efficiently as bit strings); setup for types nat/int/real is in place;
  3964 INCOMPATIBILITY: since numeral syntax is now polymorphic, rather than
  3965 int, existing theories and proof scripts may require a few additional
  3966 type constraints;
  3967 
  3968 * integer division and remainder can now be performed on constant
  3969 arguments;
  3970 
  3971 * many properties of integer multiplication, division and remainder
  3972 are now available;
  3973 
  3974 * An interface to the Stanford Validity Checker (SVC) is available through the
  3975 tactic svc_tac.  Propositional tautologies and theorems of linear arithmetic
  3976 are proved automatically.  SVC must be installed separately, and its results
  3977 must be TAKEN ON TRUST (Isabelle does not check the proofs, but tags any
  3978 invocation of the underlying oracle).  For SVC see
  3979   http://verify.stanford.edu/SVC
  3980 
  3981 * IsaMakefile: the HOL-Real target now builds an actual image;
  3982 
  3983 
  3984 ** HOL misc **
  3985 
  3986 * HOL/Real/HahnBanach: the Hahn-Banach theorem for real vector spaces
  3987 (in Isabelle/Isar) -- by Gertrud Bauer;
  3988 
  3989 * HOL/BCV: generic model of bytecode verification, i.e. data-flow
  3990 analysis for assembly languages with subtypes;
  3991 
  3992 * HOL/TLA (Lamport's Temporal Logic of Actions): major reorganization
  3993 -- avoids syntactic ambiguities and treats state, transition, and
  3994 temporal levels more uniformly; introduces INCOMPATIBILITIES due to
  3995 changed syntax and (many) tactics;
  3996 
  3997 * HOL/inductive: Now also handles more general introduction rules such
  3998   as "ALL y. (y, x) : r --> y : acc r ==> x : acc r"; monotonicity
  3999   theorems are now maintained within the theory (maintained via the
  4000   "mono" attribute);
  4001 
  4002 * HOL/datatype: Now also handles arbitrarily branching datatypes
  4003   (using function types) such as
  4004 
  4005   datatype 'a tree = Atom 'a | Branch "nat => 'a tree"
  4006 
  4007 * HOL/record: record_simproc (part of the default simpset) takes care
  4008 of selectors applied to updated records; record_split_tac is no longer
  4009 part of the default claset; update_defs may now be removed from the
  4010 simpset in many cases; COMPATIBILITY: old behavior achieved by
  4011 
  4012   claset_ref () := claset() addSWrapper record_split_wrapper;
  4013   Delsimprocs [record_simproc]
  4014 
  4015 * HOL/typedef: fixed type inference for representing set; type
  4016 arguments now have to occur explicitly on the rhs as type constraints;
  4017 
  4018 * HOL/recdef (TFL): 'congs' syntax now expects comma separated list of theorem
  4019 names rather than an ML expression;
  4020 
  4021 * HOL/defer_recdef (TFL): like recdef but the well-founded relation can be
  4022 supplied later.  Program schemes can be defined, such as
  4023     "While B C s = (if B s then While B C (C s) else s)"
  4024 where the well-founded relation can be chosen after B and C have been given.
  4025 
  4026 * HOL/List: the constructors of type list are now Nil and Cons;
  4027 INCOMPATIBILITY: while [] and infix # syntax is still there, of
  4028 course, ML tools referring to List.list.op # etc. have to be adapted;
  4029 
  4030 * HOL_quantifiers flag superseded by "HOL" print mode, which is
  4031 disabled by default; run isabelle with option -m HOL to get back to
  4032 the original Gordon/HOL-style output;
  4033 
  4034 * HOL/Ord.thy: new bounded quantifier syntax (input only): ALL x<y. P,
  4035 ALL x<=y. P, EX x<y. P, EX x<=y. P;
  4036 
  4037 * HOL basic syntax simplified (more orthogonal): all variants of
  4038 All/Ex now support plain / symbolic / HOL notation; plain syntax for
  4039 Eps operator is provided as well: "SOME x. P[x]";
  4040 
  4041 * HOL/Sum.thy: sum_case has been moved to HOL/Datatype;
  4042 
  4043 * HOL/Univ.thy: infix syntax <*>, <+>, <**>, <+> eliminated and made
  4044 thus available for user theories;
  4045 
  4046 * HOLCF/IOA/Sequents: renamed 'Cons' to 'Consq' to avoid clash with
  4047 HOL/List; hardly an INCOMPATIBILITY since '>>' syntax is used all the
  4048 time;
  4049 
  4050 * HOL: new tactic smp_tac: int -> int -> tactic, which applies spec
  4051 several times and then mp;
  4052 
  4053 
  4054 *** LK ***
  4055 
  4056 * the notation <<...>> is now available as a notation for sequences of
  4057 formulas;
  4058 
  4059 * the simplifier is now installed
  4060 
  4061 * the axiom system has been generalized (thanks to Soren Heilmann)
  4062 
  4063 * the classical reasoner now has a default rule database
  4064 
  4065 
  4066 *** ZF ***
  4067 
  4068 * new primrec section allows primitive recursive functions to be given
  4069 directly (as in HOL) over datatypes and the natural numbers;
  4070 
  4071 * new tactics induct_tac and exhaust_tac for induction (or case
  4072 analysis) over datatypes and the natural numbers;
  4073 
  4074 * the datatype declaration of type T now defines the recursor T_rec;
  4075 
  4076 * simplification automatically does freeness reasoning for datatype
  4077 constructors;
  4078 
  4079 * automatic type-inference, with AddTCs command to insert new
  4080 type-checking rules;
  4081 
  4082 * datatype introduction rules are now added as Safe Introduction rules
  4083 to the claset;
  4084 
  4085 * the syntax "if P then x else y" is now available in addition to
  4086 if(P,x,y);
  4087 
  4088 
  4089 *** Internal programming interfaces ***
  4090 
  4091 * tuned simplifier trace output; new flag debug_simp;
  4092 
  4093 * structures Vartab / Termtab (instances of TableFun) offer efficient
  4094 tables indexed by indexname_ord / term_ord (compatible with aconv);
  4095 
  4096 * AxClass.axclass_tac lost the theory argument;
  4097 
  4098 * tuned current_goals_markers semantics: begin / end goal avoids
  4099 printing empty lines;
  4100 
  4101 * removed prs and prs_fn hook, which was broken because it did not
  4102 include \n in its semantics, forcing writeln to add one
  4103 uncoditionally; replaced prs_fn by writeln_fn; consider std_output:
  4104 string -> unit if you really want to output text without newline;
  4105 
  4106 * Symbol.output subject to print mode; INCOMPATIBILITY: defaults to
  4107 plain output, interface builders may have to enable 'isabelle_font'
  4108 mode to get Isabelle font glyphs as before;
  4109 
  4110 * refined token_translation interface; INCOMPATIBILITY: output length
  4111 now of type real instead of int;
  4112 
  4113 * theory loader actions may be traced via new ThyInfo.add_hook
  4114 interface (see src/Pure/Thy/thy_info.ML); example application: keep
  4115 your own database of information attached to *whole* theories -- as
  4116 opposed to intra-theory data slots offered via TheoryDataFun;
  4117 
  4118 * proper handling of dangling sort hypotheses (at last!);
  4119 Thm.strip_shyps and Drule.strip_shyps_warning take care of removing
  4120 extra sort hypotheses that can be witnessed from the type signature;
  4121 the force_strip_shyps flag is gone, any remaining shyps are simply
  4122 left in the theorem (with a warning issued by strip_shyps_warning);
  4123 
  4124 
  4125 
  4126 New in Isabelle98-1 (October 1998)
  4127 ----------------------------------
  4128 
  4129 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  4130 
  4131 * several changes of automated proof tools;
  4132 
  4133 * HOL: major changes to the inductive and datatype packages, including
  4134 some minor incompatibilities of theory syntax;
  4135 
  4136 * HOL: renamed r^-1 to 'converse' from 'inverse'; 'inj_onto' is now
  4137 called `inj_on';
  4138 
  4139 * HOL: removed duplicate thms in Arith:
  4140   less_imp_add_less  should be replaced by  trans_less_add1
  4141   le_imp_add_le      should be replaced by  trans_le_add1
  4142 
  4143 * HOL: unary minus is now overloaded (new type constraints may be
  4144 required);
  4145 
  4146 * HOL and ZF: unary minus for integers is now #- instead of #~.  In
  4147 ZF, expressions such as n#-1 must be changed to n#- 1, since #-1 is
  4148 now taken as an integer constant.
  4149 
  4150 * Pure: ML function 'theory_of' renamed to 'theory';
  4151 
  4152 
  4153 *** Proof tools ***
  4154 
  4155 * Simplifier:
  4156   1. Asm_full_simp_tac is now more aggressive.
  4157      1. It will sometimes reorient premises if that increases their power to
  4158         simplify.
  4159      2. It does no longer proceed strictly from left to right but may also
  4160         rotate premises to achieve further simplification.
  4161      For compatibility reasons there is now Asm_lr_simp_tac which is like the
  4162      old Asm_full_simp_tac in that it does not rotate premises.
  4163   2. The simplifier now knows a little bit about nat-arithmetic.
  4164 
  4165 * Classical reasoner: wrapper mechanism for the classical reasoner now
  4166 allows for selected deletion of wrappers, by introduction of names for
  4167 wrapper functionals.  This implies that addbefore, addSbefore,
  4168 addaltern, and addSaltern now take a pair (name, tactic) as argument,
  4169 and that adding two tactics with the same name overwrites the first
  4170 one (emitting a warning).
  4171   type wrapper = (int -> tactic) -> (int -> tactic)
  4172   setWrapper, setSWrapper, compWrapper and compSWrapper are replaced by
  4173   addWrapper, addSWrapper: claset * (string * wrapper) -> claset
  4174   delWrapper, delSWrapper: claset *  string            -> claset
  4175   getWrapper is renamed to appWrappers, getSWrapper to appSWrappers;
  4176 
  4177 * Classical reasoner: addbefore/addSbefore now have APPEND/ORELSE
  4178 semantics; addbefore now affects only the unsafe part of step_tac
  4179 etc.; this affects addss/auto_tac/force_tac, so EXISTING PROOFS MAY
  4180 FAIL, but proofs should be fixable easily, e.g. by replacing Auto_tac
  4181 by Force_tac;
  4182 
  4183 * Classical reasoner: setwrapper to setWrapper and compwrapper to
  4184 compWrapper; added safe wrapper (and access functions for it);
  4185 
  4186 * HOL/split_all_tac is now much faster and fails if there is nothing
  4187 to split.  Some EXISTING PROOFS MAY REQUIRE ADAPTION because the order
  4188 and the names of the automatically generated variables have changed.
  4189 split_all_tac has moved within claset() from unsafe wrappers to safe
  4190 wrappers, which means that !!-bound variables are split much more
  4191 aggressively, and safe_tac and clarify_tac now split such variables.
  4192 If this splitting is not appropriate, use delSWrapper "split_all_tac".
  4193 Note: the same holds for record_split_tac, which does the job of
  4194 split_all_tac for record fields.
  4195 
  4196 * HOL/Simplifier: Rewrite rules for case distinctions can now be added
  4197 permanently to the default simpset using Addsplits just like
  4198 Addsimps. They can be removed via Delsplits just like
  4199 Delsimps. Lower-case versions are also available.
  4200 
  4201 * HOL/Simplifier: The rule split_if is now part of the default
  4202 simpset. This means that the simplifier will eliminate all occurrences
  4203 of if-then-else in the conclusion of a goal. To prevent this, you can
  4204 either remove split_if completely from the default simpset by
  4205 `Delsplits [split_if]' or remove it in a specific call of the
  4206 simplifier using `... delsplits [split_if]'.  You can also add/delete
  4207 other case splitting rules to/from the default simpset: every datatype
  4208 generates suitable rules `split_t_case' and `split_t_case_asm' (where
  4209 t is the name of the datatype).
  4210 
  4211 * Classical reasoner / Simplifier combination: new force_tac (and
  4212 derivatives Force_tac, force) combines rewriting and classical
  4213 reasoning (and whatever other tools) similarly to auto_tac, but is
  4214 aimed to solve the given subgoal completely.
  4215 
  4216 
  4217 *** General ***
  4218 
  4219 * new top-level commands `Goal' and `Goalw' that improve upon `goal'
  4220 and `goalw': the theory is no longer needed as an explicit argument -
  4221 the current theory context is used; assumptions are no longer returned
  4222 at the ML-level unless one of them starts with ==> or !!; it is
  4223 recommended to convert to these new commands using isatool fixgoal
  4224 (backup your sources first!);
  4225 
  4226 * new top-level commands 'thm' and 'thms' for retrieving theorems from
  4227 the current theory context, and 'theory' to lookup stored theories;
  4228 
  4229 * new theory section 'locale' for declaring constants, assumptions and
  4230 definitions that have local scope;
  4231 
  4232 * new theory section 'nonterminals' for purely syntactic types;
  4233 
  4234 * new theory section 'setup' for generic ML setup functions
  4235 (e.g. package initialization);
  4236 
  4237 * the distribution now includes Isabelle icons: see
  4238 lib/logo/isabelle-{small,tiny}.xpm;
  4239 
  4240 * isatool install - install binaries with absolute references to
  4241 ISABELLE_HOME/bin;
  4242 
  4243 * isatool logo -- create instances of the Isabelle logo (as EPS);
  4244 
  4245 * print mode 'emacs' reserved for Isamode;
  4246 
  4247 * support multiple print (ast) translations per constant name;
  4248 
  4249 * theorems involving oracles are now printed with a suffixed [!];
  4250 
  4251 
  4252 *** HOL ***
  4253 
  4254 * there is now a tutorial on Isabelle/HOL (do 'isatool doc tutorial');
  4255 
  4256 * HOL/inductive package reorganized and improved: now supports mutual
  4257 definitions such as
  4258 
  4259   inductive EVEN ODD
  4260     intrs
  4261       null "0 : EVEN"
  4262       oddI "n : EVEN ==> Suc n : ODD"
  4263       evenI "n : ODD ==> Suc n : EVEN"
  4264 
  4265 new theorem list "elims" contains an elimination rule for each of the
  4266 recursive sets; inductive definitions now handle disjunctive premises
  4267 correctly (also ZF);
  4268 
  4269 INCOMPATIBILITIES: requires Inductive as an ancestor; component
  4270 "mutual_induct" no longer exists - the induction rule is always
  4271 contained in "induct";
  4272 
  4273 
  4274 * HOL/datatype package re-implemented and greatly improved: now
  4275 supports mutually recursive datatypes such as
  4276 
  4277   datatype
  4278     'a aexp = IF_THEN_ELSE ('a bexp) ('a aexp) ('a aexp)
  4279             | SUM ('a aexp) ('a aexp)
  4280             | DIFF ('a aexp) ('a aexp)
  4281             | NUM 'a
  4282   and
  4283     'a bexp = LESS ('a aexp) ('a aexp)
  4284             | AND ('a bexp) ('a bexp)
  4285             | OR ('a bexp) ('a bexp)
  4286 
  4287 as well as indirectly recursive datatypes such as
  4288 
  4289   datatype
  4290     ('a, 'b) term = Var 'a
  4291                   | App 'b ((('a, 'b) term) list)
  4292 
  4293 The new tactic  mutual_induct_tac [<var_1>, ..., <var_n>] i  performs
  4294 induction on mutually / indirectly recursive datatypes.
  4295 
  4296 Primrec equations are now stored in theory and can be accessed via
  4297 <function_name>.simps.
  4298 
  4299 INCOMPATIBILITIES:
  4300 
  4301   - Theories using datatypes must now have theory Datatype as an
  4302     ancestor.
  4303   - The specific <typename>.induct_tac no longer exists - use the
  4304     generic induct_tac instead.
  4305   - natE has been renamed to nat.exhaust - use exhaust_tac
  4306     instead of res_inst_tac ... natE. Note that the variable
  4307     names in nat.exhaust differ from the names in natE, this
  4308     may cause some "fragile" proofs to fail.
  4309   - The theorems split_<typename>_case and split_<typename>_case_asm
  4310     have been renamed to <typename>.split and <typename>.split_asm.
  4311   - Since default sorts of type variables are now handled correctly,
  4312     some datatype definitions may have to be annotated with explicit
  4313     sort constraints.
  4314   - Primrec definitions no longer require function name and type
  4315     of recursive argument.
  4316 
  4317 Consider using isatool fixdatatype to adapt your theories and proof
  4318 scripts to the new package (backup your sources first!).
  4319 
  4320 
  4321 * HOL/record package: considerably improved implementation; now
  4322 includes concrete syntax for record types, terms, updates; theorems
  4323 for surjective pairing and splitting !!-bound record variables; proof
  4324 support is as follows:
  4325 
  4326   1) standard conversions (selectors or updates applied to record
  4327 constructor terms) are part of the standard simpset;
  4328 
  4329   2) inject equations of the form ((x, y) = (x', y')) == x=x' & y=y' are
  4330 made part of standard simpset and claset via addIffs;
  4331 
  4332   3) a tactic for record field splitting (record_split_tac) is part of
  4333 the standard claset (addSWrapper);
  4334 
  4335 To get a better idea about these rules you may retrieve them via
  4336 something like 'thms "foo.simps"' or 'thms "foo.iffs"', where "foo" is
  4337 the name of your record type.
  4338 
  4339 The split tactic 3) conceptually simplifies by the following rule:
  4340 
  4341   "(!!x. PROP ?P x) == (!!a b. PROP ?P (a, b))"
  4342 
  4343 Thus any record variable that is bound by meta-all will automatically
  4344 blow up into some record constructor term, consequently the
  4345 simplifications of 1), 2) apply.  Thus force_tac, auto_tac etc. shall
  4346 solve record problems automatically.
  4347 
  4348 
  4349 * reorganized the main HOL image: HOL/Integ and String loaded by
  4350 default; theory Main includes everything;
  4351 
  4352 * automatic simplification of integer sums and comparisons, using cancellation;
  4353 
  4354 * added option_map_eq_Some and not_Some_eq to the default simpset and claset;
  4355 
  4356 * added disj_not1 = "(~P | Q) = (P --> Q)" to the default simpset;
  4357 
  4358 * many new identities for unions, intersections, set difference, etc.;
  4359 
  4360 * expand_if, expand_split, expand_sum_case and expand_nat_case are now
  4361 called split_if, split_split, split_sum_case and split_nat_case (to go
  4362 with add/delsplits);
  4363 
  4364 * HOL/Prod introduces simplification procedure unit_eq_proc rewriting
  4365 (?x::unit) = (); this is made part of the default simpset, which COULD
  4366 MAKE EXISTING PROOFS FAIL under rare circumstances (consider
  4367 'Delsimprocs [unit_eq_proc];' as last resort); also note that
  4368 unit_abs_eta_conv is added in order to counter the effect of
  4369 unit_eq_proc on (%u::unit. f u), replacing it by f rather than by
  4370 %u.f();
  4371 
  4372 * HOL/Fun INCOMPATIBILITY: `inj_onto' is now called `inj_on' (which
  4373 makes more sense);
  4374 
  4375 * HOL/Set INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
  4376   It and 'sym RS equals0D' are now in the default  claset, giving automatic
  4377   disjointness reasoning but breaking a few old proofs.
  4378 
  4379 * HOL/Relation INCOMPATIBILITY: renamed the relational operator r^-1
  4380 to 'converse' from 'inverse' (for compatibility with ZF and some
  4381 literature);
  4382 
  4383 * HOL/recdef can now declare non-recursive functions, with {} supplied as
  4384 the well-founded relation;
  4385 
  4386 * HOL/Set INCOMPATIBILITY: the complement of set A is now written -A instead of
  4387     Compl A.  The "Compl" syntax remains available as input syntax for this
  4388     release ONLY.
  4389 
  4390 * HOL/Update: new theory of function updates:
  4391     f(a:=b) == %x. if x=a then b else f x
  4392 may also be iterated as in f(a:=b,c:=d,...);
  4393 
  4394 * HOL/Vimage: new theory for inverse image of a function, syntax f-``B;
  4395 
  4396 * HOL/List:
  4397   - new function list_update written xs[i:=v] that updates the i-th
  4398     list position. May also be iterated as in xs[i:=a,j:=b,...].
  4399   - new function `upt' written [i..j(] which generates the list
  4400     [i,i+1,...,j-1], i.e. the upper bound is excluded. To include the upper
  4401     bound write [i..j], which is a shorthand for [i..j+1(].
  4402   - new lexicographic orderings and corresponding wellfoundedness theorems.
  4403 
  4404 * HOL/Arith:
  4405   - removed 'pred' (predecessor) function;
  4406   - generalized some theorems about n-1;
  4407   - many new laws about "div" and "mod";
  4408   - new laws about greatest common divisors (see theory ex/Primes);
  4409 
  4410 * HOL/Relation: renamed the relational operator r^-1 "converse"
  4411 instead of "inverse";
  4412 
  4413 * HOL/Induct/Multiset: a theory of multisets, including the wellfoundedness
  4414   of the multiset ordering;
  4415 
  4416 * directory HOL/Real: a construction of the reals using Dedekind cuts
  4417   (not included by default);
  4418 
  4419 * directory HOL/UNITY: Chandy and Misra's UNITY formalism;
  4420 
  4421 * directory HOL/Hoare: a new version of Hoare logic which permits many-sorted
  4422   programs, i.e. different program variables may have different types.
  4423 
  4424 * calling (stac rew i) now fails if "rew" has no effect on the goal
  4425   [previously, this check worked only if the rewrite rule was unconditional]
  4426   Now rew can involve either definitions or equalities (either == or =).
  4427 
  4428 
  4429 *** ZF ***
  4430 
  4431 * theory Main includes everything; INCOMPATIBILITY: theory ZF.thy contains
  4432   only the theorems proved on ZF.ML;
  4433 
  4434 * ZF INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
  4435   It and 'sym RS equals0D' are now in the default  claset, giving automatic
  4436   disjointness reasoning but breaking a few old proofs.
  4437 
  4438 * ZF/Update: new theory of function updates
  4439     with default rewrite rule  f(x:=y) ` z = if(z=x, y, f`z)
  4440   may also be iterated as in f(a:=b,c:=d,...);
  4441 
  4442 * in  let x=t in u(x), neither t nor u(x) has to be an FOL term.
  4443 
  4444 * calling (stac rew i) now fails if "rew" has no effect on the goal
  4445   [previously, this check worked only if the rewrite rule was unconditional]
  4446   Now rew can involve either definitions or equalities (either == or =).
  4447 
  4448 * case_tac provided for compatibility with HOL
  4449     (like the old excluded_middle_tac, but with subgoals swapped)
  4450 
  4451 
  4452 *** Internal programming interfaces ***
  4453 
  4454 * Pure: several new basic modules made available for general use, see
  4455 also src/Pure/README;
  4456 
  4457 * improved the theory data mechanism to support encapsulation (data
  4458 kind name replaced by private Object.kind, acting as authorization
  4459 key); new type-safe user interface via functor TheoryDataFun; generic
  4460 print_data function becomes basically useless;
  4461 
  4462 * removed global_names compatibility flag -- all theory declarations
  4463 are qualified by default;
  4464 
  4465 * module Pure/Syntax now offers quote / antiquote translation
  4466 functions (useful for Hoare logic etc. with implicit dependencies);
  4467 see HOL/ex/Antiquote for an example use;
  4468 
  4469 * Simplifier now offers conversions (asm_)(full_)rewrite: simpset ->
  4470 cterm -> thm;
  4471 
  4472 * new tactical CHANGED_GOAL for checking that a tactic modifies a
  4473 subgoal;
  4474 
  4475 * Display.print_goals function moved to Locale.print_goals;
  4476 
  4477 * standard print function for goals supports current_goals_markers
  4478 variable for marking begin of proof, end of proof, start of goal; the
  4479 default is ("", "", ""); setting current_goals_markers := ("<proof>",
  4480 "</proof>", "<goal>") causes SGML like tagged proof state printing,
  4481 for example;
  4482 
  4483 
  4484 
  4485 New in Isabelle98 (January 1998)
  4486 --------------------------------
  4487 
  4488 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  4489 
  4490 * changed lexical syntax of terms / types: dots made part of long
  4491 identifiers, e.g. "%x.x" no longer possible, should be "%x. x";
  4492 
  4493 * simpset (and claset) reference variable replaced by functions
  4494 simpset / simpset_ref;
  4495 
  4496 * no longer supports theory aliases (via merge) and non-trivial
  4497 implicit merge of thms' signatures;
  4498 
  4499 * most internal names of constants changed due to qualified names;
  4500 
  4501 * changed Pure/Sequence interface (see Pure/seq.ML);
  4502 
  4503 
  4504 *** General Changes ***
  4505 
  4506 * hierachically structured name spaces (for consts, types, axms, thms
  4507 etc.); new lexical class 'longid' (e.g. Foo.bar.x) may render much of
  4508 old input syntactically incorrect (e.g. "%x.x"); COMPATIBILITY:
  4509 isatool fixdots ensures space after dots (e.g. "%x. x"); set
  4510 long_names for fully qualified output names; NOTE: ML programs
  4511 (special tactics, packages etc.) referring to internal names may have
  4512 to be adapted to cope with fully qualified names; in case of severe
  4513 backward campatibility problems try setting 'global_names' at compile
  4514 time to have enrything declared within a flat name space; one may also
  4515 fine tune name declarations in theories via the 'global' and 'local'
  4516 section;
  4517 
  4518 * reimplemented the implicit simpset and claset using the new anytype
  4519 data filed in signatures; references simpset:simpset ref etc. are
  4520 replaced by functions simpset:unit->simpset and
  4521 simpset_ref:unit->simpset ref; COMPATIBILITY: use isatool fixclasimp
  4522 to patch your ML files accordingly;
  4523 
  4524 * HTML output now includes theory graph data for display with Java
  4525 applet or isatool browser; data generated automatically via isatool
  4526 usedir (see -i option, ISABELLE_USEDIR_OPTIONS);
  4527 
  4528 * defs may now be conditional; improved rewrite_goals_tac to handle
  4529 conditional equations;
  4530 
  4531 * defs now admits additional type arguments, using TYPE('a) syntax;
  4532 
  4533 * theory aliases via merge (e.g. M=A+B+C) no longer supported, always
  4534 creates a new theory node; implicit merge of thms' signatures is
  4535 restricted to 'trivial' ones; COMPATIBILITY: one may have to use
  4536 transfer:theory->thm->thm in (rare) cases;
  4537 
  4538 * improved handling of draft signatures / theories; draft thms (and
  4539 ctyps, cterms) are automatically promoted to real ones;
  4540 
  4541 * slightly changed interfaces for oracles: admit many per theory, named
  4542 (e.g. oracle foo = mlfun), additional name argument for invoke_oracle;
  4543 
  4544 * print_goals: optional output of const types (set show_consts and
  4545 show_types);
  4546 
  4547 * improved output of warnings (###) and errors (***);
  4548 
  4549 * subgoal_tac displays a warning if the new subgoal has type variables;
  4550 
  4551 * removed old README and Makefiles;
  4552 
  4553 * replaced print_goals_ref hook by print_current_goals_fn and result_error_fn;
  4554 
  4555 * removed obsolete init_pps and init_database;
  4556 
  4557 * deleted the obsolete tactical STATE, which was declared by
  4558     fun STATE tacfun st = tacfun st st;
  4559 
  4560 * cd and use now support path variables, e.g. $ISABELLE_HOME, or ~
  4561 (which abbreviates $HOME);
  4562 
  4563 * changed Pure/Sequence interface (see Pure/seq.ML); COMPATIBILITY:
  4564 use isatool fixseq to adapt your ML programs (this works for fully
  4565 qualified references to the Sequence structure only!);
  4566 
  4567 * use_thy no longer requires writable current directory; it always
  4568 reloads .ML *and* .thy file, if either one is out of date;
  4569 
  4570 
  4571 *** Classical Reasoner ***
  4572 
  4573 * Clarify_tac, clarify_tac, clarify_step_tac, Clarify_step_tac: new
  4574 tactics that use classical reasoning to simplify a subgoal without
  4575 splitting it into several subgoals;
  4576 
  4577 * Safe_tac: like safe_tac but uses the default claset;
  4578 
  4579 
  4580 *** Simplifier ***
  4581 
  4582 * added simplification meta rules:
  4583     (asm_)(full_)simplify: simpset -> thm -> thm;
  4584 
  4585 * simplifier.ML no longer part of Pure -- has to be loaded by object
  4586 logics (again);
  4587 
  4588 * added prems argument to simplification procedures;
  4589 
  4590 * HOL, FOL, ZF: added infix function `addsplits':
  4591   instead of `<simpset> setloop (split_tac <thms>)'
  4592   you can simply write `<simpset> addsplits <thms>'
  4593 
  4594 
  4595 *** Syntax ***
  4596 
  4597 * TYPE('a) syntax for type reflection terms;
  4598 
  4599 * no longer handles consts with name "" -- declare as 'syntax' instead;
  4600 
  4601 * pretty printer: changed order of mixfix annotation preference (again!);
  4602 
  4603 * Pure: fixed idt/idts vs. pttrn/pttrns syntactic categories;
  4604 
  4605 
  4606 *** HOL ***
  4607 
  4608 * HOL: there is a new splitter `split_asm_tac' that can be used e.g.
  4609   with `addloop' of the simplifier to faciliate case splitting in premises.
  4610 
  4611 * HOL/TLA: Stephan Merz's formalization of Lamport's Temporal Logic of Actions;
  4612 
  4613 * HOL/Auth: new protocol proofs including some for the Internet
  4614   protocol TLS;
  4615 
  4616 * HOL/Map: new theory of `maps' a la VDM;
  4617 
  4618 * HOL/simplifier: simplification procedures nat_cancel_sums for
  4619 cancelling out common nat summands from =, <, <= (in)equalities, or
  4620 differences; simplification procedures nat_cancel_factor for
  4621 cancelling common factor from =, <, <= (in)equalities over natural
  4622 sums; nat_cancel contains both kinds of procedures, it is installed by
  4623 default in Arith.thy -- this COULD MAKE EXISTING PROOFS FAIL;
  4624 
  4625 * HOL/simplifier: terms of the form
  4626   `? x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x)'  (or t=x)
  4627   are rewritten to
  4628   `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t)',
  4629   and those of the form
  4630   `! x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x) --> R(x)'  (or t=x)
  4631   are rewritten to
  4632   `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t) --> R(t)',
  4633 
  4634 * HOL/datatype
  4635   Each datatype `t' now comes with a theorem `split_t_case' of the form
  4636 
  4637   P(t_case f1 ... fn x) =
  4638      ( (!y1 ... ym1. x = C1 y1 ... ym1 --> P(f1 y1 ... ym1)) &
  4639         ...
  4640        (!y1 ... ymn. x = Cn y1 ... ymn --> P(f1 y1 ... ymn))
  4641      )
  4642 
  4643   and a theorem `split_t_case_asm' of the form
  4644 
  4645   P(t_case f1 ... fn x) =
  4646     ~( (? y1 ... ym1. x = C1 y1 ... ym1 & ~P(f1 y1 ... ym1)) |
  4647         ...
  4648        (? y1 ... ymn. x = Cn y1 ... ymn & ~P(f1 y1 ... ymn))
  4649      )
  4650   which can be added to a simpset via `addsplits'. The existing theorems
  4651   expand_list_case and expand_option_case have been renamed to
  4652   split_list_case and split_option_case.
  4653 
  4654 * HOL/Arithmetic:
  4655   - `pred n' is automatically converted to `n-1'.
  4656     Users are strongly encouraged not to use `pred' any longer,
  4657     because it will disappear altogether at some point.
  4658   - Users are strongly encouraged to write "0 < n" rather than
  4659     "n ~= 0". Theorems and proof tools have been modified towards this
  4660     `standard'.
  4661 
  4662 * HOL/Lists:
  4663   the function "set_of_list" has been renamed "set" (and its theorems too);
  4664   the function "nth" now takes its arguments in the reverse order and
  4665   has acquired the infix notation "!" as in "xs!n".
  4666 
  4667 * HOL/Set: UNIV is now a constant and is no longer translated to Compl{};
  4668 
  4669 * HOL/Set: The operator (UN x.B x) now abbreviates (UN x:UNIV. B x) and its
  4670   specialist theorems (like UN1_I) are gone.  Similarly for (INT x.B x);
  4671 
  4672 * HOL/record: extensible records with schematic structural subtyping
  4673 (single inheritance); EXPERIMENTAL version demonstrating the encoding,
  4674 still lacks various theorems and concrete record syntax;
  4675 
  4676 
  4677 *** HOLCF ***
  4678 
  4679 * removed "axioms" and "generated by" sections;
  4680 
  4681 * replaced "ops" section by extended "consts" section, which is capable of
  4682   handling the continuous function space "->" directly;
  4683 
  4684 * domain package:
  4685   . proves theorems immediately and stores them in the theory,
  4686   . creates hierachical name space,
  4687   . now uses normal mixfix annotations (instead of cinfix...),
  4688   . minor changes to some names and values (for consistency),
  4689   . e.g. cases -> casedist, dists_eq -> dist_eqs, [take_lemma] -> take_lemmas,
  4690   . separator between mutual domain defs: changed "," to "and",
  4691   . improved handling of sort constraints;  now they have to
  4692     appear on the left-hand side of the equations only;
  4693 
  4694 * fixed LAM <x,y,zs>.b syntax;
  4695 
  4696 * added extended adm_tac to simplifier in HOLCF -- can now discharge
  4697 adm (%x. P (t x)), where P is chainfinite and t continuous;
  4698 
  4699 
  4700 *** FOL and ZF ***
  4701 
  4702 * FOL: there is a new splitter `split_asm_tac' that can be used e.g.
  4703   with `addloop' of the simplifier to faciliate case splitting in premises.
  4704 
  4705 * qed_spec_mp, qed_goal_spec_mp, qed_goalw_spec_mp are available, as
  4706 in HOL, they strip ALL and --> from proved theorems;
  4707 
  4708 
  4709 
  4710 New in Isabelle94-8 (May 1997)
  4711 ------------------------------
  4712 
  4713 *** General Changes ***
  4714 
  4715 * new utilities to build / run / maintain Isabelle etc. (in parts
  4716 still somewhat experimental); old Makefiles etc. still functional;
  4717 
  4718 * new 'Isabelle System Manual';
  4719 
  4720 * INSTALL text, together with ./configure and ./build scripts;
  4721 
  4722 * reimplemented type inference for greater efficiency, better error
  4723 messages and clean internal interface;
  4724 
  4725 * prlim command for dealing with lots of subgoals (an easier way of
  4726 setting goals_limit);
  4727 
  4728 
  4729 *** Syntax ***
  4730 
  4731 * supports alternative (named) syntax tables (parser and pretty
  4732 printer); internal interface is provided by add_modesyntax(_i);
  4733 
  4734 * Pure, FOL, ZF, HOL, HOLCF now support symbolic input and output; to
  4735 be used in conjunction with the Isabelle symbol font; uses the
  4736 "symbols" syntax table;
  4737 
  4738 * added token_translation interface (may translate name tokens in
  4739 arbitrary ways, dependent on their type (free, bound, tfree, ...) and
  4740 the current print_mode); IMPORTANT: user print translation functions
  4741 are responsible for marking newly introduced bounds
  4742 (Syntax.mark_boundT);
  4743 
  4744 * token translations for modes "xterm" and "xterm_color" that display
  4745 names in bold, underline etc. or colors (which requires a color
  4746 version of xterm);
  4747 
  4748 * infixes may now be declared with names independent of their syntax;
  4749 
  4750 * added typed_print_translation (like print_translation, but may
  4751 access type of constant);
  4752 
  4753 
  4754 *** Classical Reasoner ***
  4755 
  4756 Blast_tac: a new tactic!  It is often more powerful than fast_tac, but has
  4757 some limitations.  Blast_tac...
  4758   + ignores addss, addbefore, addafter; this restriction is intrinsic
  4759   + ignores elimination rules that don't have the correct format
  4760         (the conclusion MUST be a formula variable)
  4761   + ignores types, which can make HOL proofs fail
  4762   + rules must not require higher-order unification, e.g. apply_type in ZF
  4763     [message "Function Var's argument not a bound variable" relates to this]
  4764   + its proof strategy is more general but can actually be slower
  4765 
  4766 * substitution with equality assumptions no longer permutes other
  4767 assumptions;
  4768 
  4769 * minor changes in semantics of addafter (now called addaltern); renamed
  4770 setwrapper to setWrapper and compwrapper to compWrapper; added safe wrapper
  4771 (and access functions for it);
  4772 
  4773 * improved combination of classical reasoner and simplifier:
  4774   + functions for handling clasimpsets
  4775   + improvement of addss: now the simplifier is called _after_ the
  4776     safe steps.
  4777   + safe variant of addss called addSss: uses safe simplifications
  4778     _during_ the safe steps. It is more complete as it allows multiple
  4779     instantiations of unknowns (e.g. with slow_tac).
  4780 
  4781 *** Simplifier ***
  4782 
  4783 * added interface for simplification procedures (functions that
  4784 produce *proven* rewrite rules on the fly, depending on current
  4785 redex);
  4786 
  4787 * ordering on terms as parameter (used for ordered rewriting);
  4788 
  4789 * new functions delcongs, deleqcongs, and Delcongs. richer rep_ss;
  4790 
  4791 * the solver is now split into a safe and an unsafe part.
  4792 This should be invisible for the normal user, except that the
  4793 functions setsolver and addsolver have been renamed to setSolver and
  4794 addSolver; added safe_asm_full_simp_tac;
  4795 
  4796 
  4797 *** HOL ***
  4798 
  4799 * a generic induction tactic `induct_tac' which works for all datatypes and
  4800 also for type `nat';
  4801 
  4802 * a generic case distinction tactic `exhaust_tac' which works for all
  4803 datatypes and also for type `nat';
  4804 
  4805 * each datatype comes with a function `size';
  4806 
  4807 * patterns in case expressions allow tuple patterns as arguments to
  4808 constructors, for example `case x of [] => ... | (x,y,z)#ps => ...';
  4809 
  4810 * primrec now also works with type nat;
  4811 
  4812 * recdef: a new declaration form, allows general recursive functions to be
  4813 defined in theory files.  See HOL/ex/Fib, HOL/ex/Primes, HOL/Subst/Unify.
  4814 
  4815 * the constant for negation has been renamed from "not" to "Not" to
  4816 harmonize with FOL, ZF, LK, etc.;
  4817 
  4818 * HOL/ex/LFilter theory of a corecursive "filter" functional for
  4819 infinite lists;
  4820 
  4821 * HOL/Modelcheck demonstrates invocation of model checker oracle;
  4822 
  4823 * HOL/ex/Ring.thy declares cring_simp, which solves equational
  4824 problems in commutative rings, using axiomatic type classes for + and *;
  4825 
  4826 * more examples in HOL/MiniML and HOL/Auth;
  4827 
  4828 * more default rewrite rules for quantifiers, union/intersection;
  4829 
  4830 * a new constant `arbitrary == @x.False';
  4831 
  4832 * HOLCF/IOA replaces old HOL/IOA;
  4833 
  4834 * HOLCF changes: derived all rules and arities
  4835   + axiomatic type classes instead of classes
  4836   + typedef instead of faking type definitions
  4837   + eliminated the internal constants less_fun, less_cfun, UU_fun, UU_cfun etc.
  4838   + new axclasses cpo, chfin, flat with flat < chfin < pcpo < cpo < po
  4839   + eliminated the types void, one, tr
  4840   + use unit lift and bool lift (with translations) instead of one and tr
  4841   + eliminated blift from Lift3.thy (use Def instead of blift)
  4842   all eliminated rules are derived as theorems --> no visible changes ;
  4843 
  4844 
  4845 *** ZF ***
  4846 
  4847 * ZF now has Fast_tac, Simp_tac and Auto_tac.  Union_iff is a now a default
  4848 rewrite rule; this may affect some proofs.  eq_cs is gone but can be put back
  4849 as ZF_cs addSIs [equalityI];
  4850 
  4851 
  4852 
  4853 New in Isabelle94-7 (November 96)
  4854 ---------------------------------
  4855 
  4856 * allowing negative levels (as offsets) in prlev and choplev;
  4857 
  4858 * super-linear speedup for large simplifications;
  4859 
  4860 * FOL, ZF and HOL now use miniscoping: rewriting pushes
  4861 quantifications in as far as possible (COULD MAKE EXISTING PROOFS
  4862 FAIL); can suppress it using the command Delsimps (ex_simps @
  4863 all_simps); De Morgan laws are also now included, by default;
  4864 
  4865 * improved printing of ==>  :  ~:
  4866 
  4867 * new object-logic "Sequents" adds linear logic, while replacing LK
  4868 and Modal (thanks to Sara Kalvala);
  4869 
  4870 * HOL/Auth: correctness proofs for authentication protocols;
  4871 
  4872 * HOL: new auto_tac combines rewriting and classical reasoning (many
  4873 examples on HOL/Auth);
  4874 
  4875 * HOL: new command AddIffs for declaring theorems of the form P=Q to
  4876 the rewriter and classical reasoner simultaneously;
  4877 
  4878 * function uresult no longer returns theorems in "standard" format;
  4879 regain previous version by: val uresult = standard o uresult;
  4880 
  4881 
  4882 
  4883 New in Isabelle94-6
  4884 -------------------
  4885 
  4886 * oracles -- these establish an interface between Isabelle and trusted
  4887 external reasoners, which may deliver results as theorems;
  4888 
  4889 * proof objects (in particular record all uses of oracles);
  4890 
  4891 * Simp_tac, Fast_tac, etc. that refer to implicit simpset / claset;
  4892 
  4893 * "constdefs" section in theory files;
  4894 
  4895 * "primrec" section (HOL) no longer requires names;
  4896 
  4897 * internal type "tactic" now simply "thm -> thm Sequence.seq";
  4898 
  4899 
  4900 
  4901 New in Isabelle94-5
  4902 -------------------
  4903 
  4904 * reduced space requirements;
  4905 
  4906 * automatic HTML generation from theories;
  4907 
  4908 * theory files no longer require "..." (quotes) around most types;
  4909 
  4910 * new examples, including two proofs of the Church-Rosser theorem;
  4911 
  4912 * non-curried (1994) version of HOL is no longer distributed;
  4913 
  4914 
  4915 
  4916 New in Isabelle94-4
  4917 -------------------
  4918 
  4919 * greatly reduced space requirements;
  4920 
  4921 * theory files (.thy) no longer require \...\ escapes at line breaks;
  4922 
  4923 * searchable theorem database (see the section "Retrieving theorems" on
  4924 page 8 of the Reference Manual);
  4925 
  4926 * new examples, including Grabczewski's monumental case study of the
  4927 Axiom of Choice;
  4928 
  4929 * The previous version of HOL renamed to Old_HOL;
  4930 
  4931 * The new version of HOL (previously called CHOL) uses a curried syntax
  4932 for functions.  Application looks like f a b instead of f(a,b);
  4933 
  4934 * Mutually recursive inductive definitions finally work in HOL;
  4935 
  4936 * In ZF, pattern-matching on tuples is now available in all abstractions and
  4937 translates to the operator "split";
  4938 
  4939 
  4940 
  4941 New in Isabelle94-3
  4942 -------------------
  4943 
  4944 * new infix operator, addss, allowing the classical reasoner to
  4945 perform simplification at each step of its search.  Example:
  4946         fast_tac (cs addss ss)
  4947 
  4948 * a new logic, CHOL, the same as HOL, but with a curried syntax
  4949 for functions.  Application looks like f a b instead of f(a,b).  Also pairs
  4950 look like (a,b) instead of <a,b>;
  4951 
  4952 * PLEASE NOTE: CHOL will eventually replace HOL!
  4953 
  4954 * In CHOL, pattern-matching on tuples is now available in all abstractions.
  4955 It translates to the operator "split".  A new theory of integers is available;
  4956 
  4957 * In ZF, integer numerals now denote two's-complement binary integers.
  4958 Arithmetic operations can be performed by rewriting.  See ZF/ex/Bin.ML;
  4959 
  4960 * Many new examples: I/O automata, Church-Rosser theorem, equivalents
  4961 of the Axiom of Choice;
  4962 
  4963 
  4964 
  4965 New in Isabelle94-2
  4966 -------------------
  4967 
  4968 * Significantly faster resolution;
  4969 
  4970 * the different sections in a .thy file can now be mixed and repeated
  4971 freely;
  4972 
  4973 * Database of theorems for FOL, HOL and ZF.  New
  4974 commands including qed, qed_goal and bind_thm store theorems in the database.
  4975 
  4976 * Simple database queries: return a named theorem (get_thm) or all theorems of
  4977 a given theory (thms_of), or find out what theory a theorem was proved in
  4978 (theory_of_thm);
  4979 
  4980 * Bugs fixed in the inductive definition and datatype packages;
  4981 
  4982 * The classical reasoner provides deepen_tac and depth_tac, making FOL_dup_cs
  4983 and HOL_dup_cs obsolete;
  4984 
  4985 * Syntactic ambiguities caused by the new treatment of syntax in Isabelle94-1
  4986 have been removed;
  4987 
  4988 * Simpler definition of function space in ZF;
  4989 
  4990 * new results about cardinal and ordinal arithmetic in ZF;
  4991 
  4992 * 'subtype' facility in HOL for introducing new types as subsets of existing
  4993 types;
  4994 
  4995 
  4996 $Id$