Isabelle NEWS -- history user-relevant changes ============================================== New in this Isabelle version ---------------------------- *** General *** * Discontinued old form of "escaped symbols" such as \\. Only one backslash should be used, even in ML sources. *** Pure *** * On instantiation of classes, remaining undefined class parameters are formally declared. INCOMPATIBILITY. *** HOL *** * Class semiring_div requires superclass no_zero_divisors and proof of div_mult_mult1; theorems div_mult_mult1, div_mult_mult2, div_mult_mult1_if, div_mult_mult1 and div_mult_mult2 have been generalized to class semiring_div, subsuming former theorems zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and zdiv_zmult_zmult2. div_mult_mult1 is now [simp] by default. INCOMPATIBILITY. * Power operations on relations and functions are now one dedicate constant compow with infix syntax "^^". Power operations on multiplicative monoids retains syntax "^" and is now defined generic in class power. INCOMPATIBILITY. * ML antiquotation @{code_datatype} inserts definition of a datatype generated by the code generator; see Predicate.thy for an example. * New method "linarith" invokes existing linear arithmetic decision procedure only. *** ML *** * Eliminated old Attrib.add_attributes, Method.add_methods and related cominators for "args". INCOMPATIBILITY, need to use simplified Attrib/Method.setup introduced in Isabelle2009. *** System *** * Discontinued support for Poly/ML 4.x versions. * Removed "compress" option from isabelle-process and isabelle usedir; this is always enabled. New in Isabelle2009 (April 2009) -------------------------------- *** General *** * Simplified main Isabelle executables, with less surprises on case-insensitive file-systems (such as Mac OS). - The main Isabelle tool wrapper is now called "isabelle" instead of "isatool." - The former "isabelle" alias for "isabelle-process" has been removed (should rarely occur to regular users). - The former "isabelle-interface" and its alias "Isabelle" have been removed (interfaces are now regular Isabelle tools). Within scripts and make files, the Isabelle environment variables ISABELLE_TOOL and ISABELLE_PROCESS replace old ISATOOL and ISABELLE, respectively. (The latter are still available as legacy feature.) The old isabelle-interface wrapper could react in confusing ways if the interface was uninstalled or changed otherwise. Individual interface tool configuration is now more explicit, see also the Isabelle system manual. In particular, Proof General is now available via "isabelle emacs". INCOMPATIBILITY, need to adapt derivative scripts. Users may need to purge installed copies of Isabelle executables and re-run "isabelle install -p ...", or use symlinks. * The default for ISABELLE_HOME_USER is now ~/.isabelle instead of the old ~/isabelle, which was slightly non-standard and apt to cause surprises on case-insensitive file-systems (such as Mac OS). INCOMPATIBILITY, need to move existing ~/isabelle/etc, ~/isabelle/heaps, ~/isabelle/browser_info to the new place. Special care is required when using older releases of Isabelle. Note that ISABELLE_HOME_USER can be changed in Isabelle/etc/settings of any Isabelle distribution, in order to use the new ~/.isabelle uniformly. * Proofs of fully specified statements are run in parallel on multi-core systems. A speedup factor of 2.5 to 3.2 can be expected on a regular 4-core machine, if the initial heap space is made reasonably large (cf. Poly/ML option -H). (Requires Poly/ML 5.2.1 or later.) * The main reference manuals ("isar-ref", "implementation", and "system") have been updated and extended. Formally checked references as hyperlinks are now available uniformly. *** Pure *** * Complete re-implementation of locales. INCOMPATIBILITY in several respects. The most important changes are listed below. See the Tutorial on Locales ("locales" manual) for details. - In locale expressions, instantiation replaces renaming. Parameters must be declared in a for clause. To aid compatibility with previous parameter inheritance, in locale declarations, parameters that are not 'touched' (instantiation position "_" or omitted) are implicitly added with their syntax at the beginning of the for clause. - Syntax from abbreviations and definitions in locales is available in locale expressions and context elements. The latter is particularly useful in locale declarations. - More flexible mechanisms to qualify names generated by locale expressions. Qualifiers (prefixes) may be specified in locale expressions, and can be marked as mandatory (syntax: "name!:") or optional (syntax "name?:"). The default depends for plain "name:" depends on the situation where a locale expression is used: in commands 'locale' and 'sublocale' prefixes are optional, in 'interpretation' and 'interpret' prefixes are mandatory. The old implicit qualifiers derived from the parameter names of a locale are no longer generated. - Command "sublocale l < e" replaces "interpretation l < e". The instantiation clause in "interpretation" and "interpret" (square brackets) is no longer available. Use locale expressions. - When converting proof scripts, mandatory qualifiers in 'interpretation' and 'interpret' should be retained by default, even if this is an INCOMPATIBILITY compared to former behavior. In the worst case, use the "name?:" form for non-mandatory ones. Qualifiers in locale expressions range over a single locale instance only. - Dropped locale element "includes". This is a major INCOMPATIBILITY. In existing theorem specifications replace the includes element by the respective context elements of the included locale, omitting those that are already present in the theorem specification. Multiple assume elements of a locale should be replaced by a single one involving the locale predicate. In the proof body, declarations (most notably theorems) may be regained by interpreting the respective locales in the proof context as required (command "interpret"). If using "includes" in replacement of a target solely because the parameter types in the theorem are not as general as in the target, consider declaring a new locale with additional type constraints on the parameters (context element "constrains"). - Discontinued "locale (open)". INCOMPATIBILITY. - Locale interpretation commands no longer attempt to simplify goal. INCOMPATIBILITY: in rare situations the generated goal differs. Use methods intro_locales and unfold_locales to clarify. - Locale interpretation commands no longer accept interpretation attributes. INCOMPATIBILITY. * Class declaration: so-called "base sort" must not be given in import list any longer, but is inferred from the specification. Particularly in HOL, write class foo = ... instead of class foo = type + ... * Class target: global versions of theorems stemming do not carry a parameter prefix any longer. INCOMPATIBILITY. * Class 'instance' command no longer accepts attached definitions. INCOMPATIBILITY, use proper 'instantiation' target instead. * Recovered hiding of consts, which was accidentally broken in Isabelle2007. Potential INCOMPATIBILITY, ``hide const c'' really makes c inaccessible; consider using ``hide (open) const c'' instead. * Slightly more coherent Pure syntax, with updated documentation in isar-ref manual. Removed locales meta_term_syntax and meta_conjunction_syntax: TERM and &&& (formerly &&) are now permanent, INCOMPATIBILITY in rare situations. Note that &&& should not be used directly in regular applications. * There is a new syntactic category "float_const" for signed decimal fractions (e.g. 123.45 or -123.45). * Removed exotic 'token_translation' command. INCOMPATIBILITY, use ML interface with 'setup' command instead. * Command 'local_setup' is similar to 'setup', but operates on a local theory context. * The 'axiomatization' command now only works within a global theory context. INCOMPATIBILITY. * Goal-directed proof now enforces strict proof irrelevance wrt. sort hypotheses. Sorts required in the course of reasoning need to be covered by the constraints in the initial statement, completed by the type instance information of the background theory. Non-trivial sort hypotheses, which rarely occur in practice, may be specified via vacuous propositions of the form SORT_CONSTRAINT('a::c). For example: lemma assumes "SORT_CONSTRAINT('a::empty)" shows False ... The result contains an implicit sort hypotheses as before -- SORT_CONSTRAINT premises are eliminated as part of the canonical rule normalization. * Generalized Isar history, with support for linear undo, direct state addressing etc. * Changed defaults for unify configuration options: unify_trace_bound = 50 (formerly 25) unify_search_bound = 60 (formerly 30) * Different bookkeeping for code equations (INCOMPATIBILITY): a) On theory merge, the last set of code equations for a particular constant is taken (in accordance with the policy applied by other parts of the code generator framework). b) Code equations stemming from explicit declarations (e.g. code attribute) gain priority over default code equations stemming from definition, primrec, fun etc. * Keyword 'code_exception' now named 'code_abort'. INCOMPATIBILITY. * Unified theorem tables for both code generators. Thus [code func] has disappeared and only [code] remains. INCOMPATIBILITY. * Command 'find_consts' searches for constants based on type and name patterns, e.g. find_consts "_ => bool" By default, matching is against subtypes, but it may be restricted to the whole type. Searching by name is possible. Multiple queries are conjunctive and queries may be negated by prefixing them with a hyphen: find_consts strict: "_ => bool" name: "Int" -"int => int" * New 'find_theorems' criterion "solves" matches theorems that directly solve the current goal (modulo higher-order unification). * Auto solve feature for main theorem statements: whenever a new goal is stated, "find_theorems solves" is called; any theorems that could solve the lemma directly are listed as part of the goal state. Cf. associated options in Proof General Isabelle settings menu, enabled by default, with reasonable timeout for pathological cases of higher-order unification. *** Document preparation *** * Antiquotation @{lemma} now imitates a regular terminal proof, demanding keyword 'by' and supporting the full method expression syntax just like the Isar command 'by'. *** HOL *** * Integrated main parts of former image HOL-Complex with HOL. Entry points Main and Complex_Main remain as before. * Logic image HOL-Plain provides a minimal HOL with the most important tools available (inductive, datatype, primrec, ...). This facilitates experimentation and tool development. Note that user applications (and library theories) should never refer to anything below theory Main, as before. * Logic image HOL-Main stops at theory Main, and thus facilitates experimentation due to shorter build times. * Logic image HOL-NSA contains theories of nonstandard analysis which were previously part of former HOL-Complex. Entry point Hyperreal remains valid, but theories formerly using Complex_Main should now use new entry point Hypercomplex. * Generic ATP manager for Sledgehammer, based on ML threads instead of Posix processes. Avoids potentially expensive forking of the ML process. New thread-based implementation also works on non-Unix platforms (Cygwin). Provers are no longer hardwired, but defined within the theory via plain ML wrapper functions. Basic Sledgehammer commands are covered in the isar-ref manual. * Wrapper scripts for remote SystemOnTPTP service allows to use sledgehammer without local ATP installation (Vampire etc.). Other provers may be included via suitable ML wrappers, see also src/HOL/ATP_Linkup.thy. * ATP selection (E/Vampire/Spass) is now via Proof General's settings menu. * The metis method no longer fails because the theorem is too trivial (contains the empty clause). * The metis method now fails in the usual manner, rather than raising an exception, if it determines that it cannot prove the theorem. * Method "coherent" implements a prover for coherent logic (see also src/Tools/coherent.ML). * Constants "undefined" and "default" replace "arbitrary". Usually "undefined" is the right choice to replace "arbitrary", though logically there is no difference. INCOMPATIBILITY. * Command "value" now integrates different evaluation mechanisms. The result of the first successful evaluation mechanism is printed. In square brackets a particular named evaluation mechanisms may be specified (currently, [SML], [code] or [nbe]). See further src/HOL/ex/Eval_Examples.thy. * Normalization by evaluation now allows non-leftlinear equations. Declare with attribute [code nbe]. * Methods "case_tac" and "induct_tac" now refer to the very same rules as the structured Isar versions "cases" and "induct", cf. the corresponding "cases" and "induct" attributes. Mutual induction rules are now presented as a list of individual projections (e.g. foo_bar.inducts for types foo and bar); the old format with explicit HOL conjunction is no longer supported. INCOMPATIBILITY, in rare situations a different rule is selected --- notably nested tuple elimination instead of former prod.exhaust: use explicit (case_tac t rule: prod.exhaust) here. * Attributes "cases", "induct", "coinduct" support "del" option. * Removed fact "case_split_thm", which duplicates "case_split". * The option datatype has been moved to a new theory Option. Renamed option_map to Option.map, and o2s to Option.set, INCOMPATIBILITY. * New predicate "strict_mono" classifies strict functions on partial orders. With strict functions on linear orders, reasoning about (in)equalities is facilitated by theorems "strict_mono_eq", "strict_mono_less_eq" and "strict_mono_less". * Some set operations are now proper qualified constants with authentic syntax. INCOMPATIBILITY: op Int ~> Set.Int op Un ~> Set.Un INTER ~> Set.INTER UNION ~> Set.UNION Inter ~> Set.Inter Union ~> Set.Union {} ~> Set.empty UNIV ~> Set.UNIV * Class complete_lattice with operations Inf, Sup, INFI, SUPR now in theory Set. * Auxiliary class "itself" has disappeared -- classes without any parameter are treated as expected by the 'class' command. * Leibnitz's Series for Pi and the arcus tangens and logarithm series. * Common decision procedures (Cooper, MIR, Ferrack, Approximation, Dense_Linear_Order) are now in directory HOL/Decision_Procs. * Theory src/HOL/Decision_Procs/Approximation provides the new proof method "approximation". It proves formulas on real values by using interval arithmetic. In the formulas are also the transcendental functions sin, cos, tan, atan, ln, exp and the constant pi are allowed. For examples see src/HOL/Descision_Procs/ex/Approximation_Ex.thy. * Theory "Reflection" now resides in HOL/Library. * Entry point to Word library now simply named "Word". INCOMPATIBILITY. * Made source layout more coherent with logical distribution structure: src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy src/HOL/Library/Code_Message.thy ~> src/HOL/ src/HOL/Library/GCD.thy ~> src/HOL/ src/HOL/Library/Order_Relation.thy ~> src/HOL/ src/HOL/Library/Parity.thy ~> src/HOL/ src/HOL/Library/Univ_Poly.thy ~> src/HOL/ src/HOL/Real/ContNotDenum.thy ~> src/HOL/Library/ src/HOL/Real/Lubs.thy ~> src/HOL/ src/HOL/Real/PReal.thy ~> src/HOL/ src/HOL/Real/Rational.thy ~> src/HOL/ src/HOL/Real/RComplete.thy ~> src/HOL/ src/HOL/Real/RealDef.thy ~> src/HOL/ src/HOL/Real/RealPow.thy ~> src/HOL/ src/HOL/Real/Real.thy ~> src/HOL/ src/HOL/Complex/Complex_Main.thy ~> src/HOL/ src/HOL/Complex/Complex.thy ~> src/HOL/ src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/Library/ src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/Library/ src/HOL/Hyperreal/Deriv.thy ~> src/HOL/ src/HOL/Hyperreal/Fact.thy ~> src/HOL/ src/HOL/Hyperreal/Integration.thy ~> src/HOL/ src/HOL/Hyperreal/Lim.thy ~> src/HOL/ src/HOL/Hyperreal/Ln.thy ~> src/HOL/ src/HOL/Hyperreal/Log.thy ~> src/HOL/ src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/ src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/ src/HOL/Hyperreal/Series.thy ~> src/HOL/ src/HOL/Hyperreal/SEQ.thy ~> src/HOL/ src/HOL/Hyperreal/Taylor.thy ~> src/HOL/ src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/ src/HOL/Real/Float ~> src/HOL/Library/ src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach src/HOL/Real/RealVector.thy ~> src/HOL/ src/HOL/arith_data.ML ~> src/HOL/Tools src/HOL/hologic.ML ~> src/HOL/Tools src/HOL/simpdata.ML ~> src/HOL/Tools src/HOL/int_arith1.ML ~> src/HOL/Tools/int_arith.ML src/HOL/int_factor_simprocs.ML ~> src/HOL/Tools src/HOL/nat_simprocs.ML ~> src/HOL/Tools src/HOL/Real/float_arith.ML ~> src/HOL/Tools src/HOL/Real/float_syntax.ML ~> src/HOL/Tools src/HOL/Real/rat_arith.ML ~> src/HOL/Tools src/HOL/Real/real_arith.ML ~> src/HOL/Tools src/HOL/Library/Array.thy ~> src/HOL/Imperative_HOL src/HOL/Library/Heap_Monad.thy ~> src/HOL/Imperative_HOL src/HOL/Library/Heap.thy ~> src/HOL/Imperative_HOL src/HOL/Library/Imperative_HOL.thy ~> src/HOL/Imperative_HOL src/HOL/Library/Ref.thy ~> src/HOL/Imperative_HOL src/HOL/Library/Relational.thy ~> src/HOL/Imperative_HOL * If methods "eval" and "evaluation" encounter a structured proof state with !!/==>, only the conclusion is evaluated to True (if possible), avoiding strange error messages. * Method "sizechange" automates termination proofs using (a modification of) the size-change principle. Requires SAT solver. See src/HOL/ex/Termination.thy for examples. * Simplifier: simproc for let expressions now unfolds if bound variable occurs at most once in let expression body. INCOMPATIBILITY. * Method "arith": Linear arithmetic now ignores all inequalities when fast_arith_neq_limit is exceeded, instead of giving up entirely. * New attribute "arith" for facts that should always be used automatically by arithmetic. It is intended to be used locally in proofs, e.g. assumes [arith]: "x > 0" Global usage is discouraged because of possible performance impact. * New classes "top" and "bot" with corresponding operations "top" and "bot" in theory Orderings; instantiation of class "complete_lattice" requires instantiation of classes "top" and "bot". INCOMPATIBILITY. * Changed definition lemma "less_fun_def" in order to provide an instance for preorders on functions; use lemma "less_le" instead. INCOMPATIBILITY. * Theory Orderings: class "wellorder" moved here, with explicit induction rule "less_induct" as assumption. For instantiation of "wellorder" by means of predicate "wf", use rule wf_wellorderI. INCOMPATIBILITY. * Theory Orderings: added class "preorder" as superclass of "order". INCOMPATIBILITY: Instantiation proofs for order, linorder etc. slightly changed. Some theorems named order_class.* now named preorder_class.*. * Theory Relation: renamed "refl" to "refl_on", "reflexive" to "refl, "diag" to "Id_on". * Theory Finite_Set: added a new fold combinator of type ('a => 'b => 'b) => 'b => 'a set => 'b Occasionally this is more convenient than the old fold combinator which is now defined in terms of the new one and renamed to fold_image. * Theories Ring_and_Field and OrderedGroup: The lemmas "group_simps" and "ring_simps" have been replaced by "algebra_simps" (which can be extended with further lemmas!). At the moment both still exist but the former will disappear at some point. * Theory Power: Lemma power_Suc is now declared as a simp rule in class recpower. Type-specific simp rules for various recpower types have been removed. INCOMPATIBILITY, rename old lemmas as follows: rat_power_0 -> power_0 rat_power_Suc -> power_Suc realpow_0 -> power_0 realpow_Suc -> power_Suc complexpow_0 -> power_0 complexpow_Suc -> power_Suc power_poly_0 -> power_0 power_poly_Suc -> power_Suc * Theories Ring_and_Field and Divides: Definition of "op dvd" has been moved to separate class dvd in Ring_and_Field; a couple of lemmas on dvd has been generalized to class comm_semiring_1. Likewise a bunch of lemmas from Divides has been generalized from nat to class semiring_div. INCOMPATIBILITY. This involves the following theorem renames resulting from duplicate elimination: dvd_def_mod ~> dvd_eq_mod_eq_0 zero_dvd_iff ~> dvd_0_left_iff dvd_0 ~> dvd_0_right DIVISION_BY_ZERO_DIV ~> div_by_0 DIVISION_BY_ZERO_MOD ~> mod_by_0 mult_div ~> div_mult_self2_is_id mult_mod ~> mod_mult_self2_is_0 * Theory IntDiv: removed many lemmas that are instances of class-based generalizations (from Divides and Ring_and_Field). INCOMPATIBILITY, rename old lemmas as follows: dvd_diff -> nat_dvd_diff dvd_zminus_iff -> dvd_minus_iff mod_add1_eq -> mod_add_eq mod_mult1_eq -> mod_mult_right_eq mod_mult1_eq' -> mod_mult_left_eq mod_mult_distrib_mod -> mod_mult_eq nat_mod_add_left_eq -> mod_add_left_eq nat_mod_add_right_eq -> mod_add_right_eq nat_mod_div_trivial -> mod_div_trivial nat_mod_mod_trivial -> mod_mod_trivial zdiv_zadd_self1 -> div_add_self1 zdiv_zadd_self2 -> div_add_self2 zdiv_zmult_self1 -> div_mult_self2_is_id zdiv_zmult_self2 -> div_mult_self1_is_id zdvd_triv_left -> dvd_triv_left zdvd_triv_right -> dvd_triv_right zdvd_zmult_cancel_disj -> dvd_mult_cancel_left zmod_eq0_zdvd_iff -> dvd_eq_mod_eq_0[symmetric] zmod_zadd_left_eq -> mod_add_left_eq zmod_zadd_right_eq -> mod_add_right_eq zmod_zadd_self1 -> mod_add_self1 zmod_zadd_self2 -> mod_add_self2 zmod_zadd1_eq -> mod_add_eq zmod_zdiff1_eq -> mod_diff_eq zmod_zdvd_zmod -> mod_mod_cancel zmod_zmod_cancel -> mod_mod_cancel zmod_zmult_self1 -> mod_mult_self2_is_0 zmod_zmult_self2 -> mod_mult_self1_is_0 zmod_1 -> mod_by_1 zdiv_1 -> div_by_1 zdvd_abs1 -> abs_dvd_iff zdvd_abs2 -> dvd_abs_iff zdvd_refl -> dvd_refl zdvd_trans -> dvd_trans zdvd_zadd -> dvd_add zdvd_zdiff -> dvd_diff zdvd_zminus_iff -> dvd_minus_iff zdvd_zminus2_iff -> minus_dvd_iff zdvd_zmultD -> dvd_mult_right zdvd_zmultD2 -> dvd_mult_left zdvd_zmult_mono -> mult_dvd_mono zdvd_0_right -> dvd_0_right zdvd_0_left -> dvd_0_left_iff zdvd_1_left -> one_dvd zminus_dvd_iff -> minus_dvd_iff * Theory Rational: 'Fract k 0' now equals '0'. INCOMPATIBILITY. * The real numbers offer decimal input syntax: 12.34 is translated into 1234/10^2. This translation is not reversed upon output. * Theory Library/Polynomial defines an abstract type 'a poly of univariate polynomials with coefficients of type 'a. In addition to the standard ring operations, it also supports div and mod. Code generation is also supported, using list-style constructors. * Theory Library/Inner_Product defines a class of real_inner for real inner product spaces, with an overloaded operation inner :: 'a => 'a => real. Class real_inner is a subclass of real_normed_vector from theory RealVector. * Theory Library/Product_Vector provides instances for the product type 'a * 'b of several classes from RealVector and Inner_Product. Definitions of addition, subtraction, scalar multiplication, norms, and inner products are included. * Theory Library/Bit defines the field "bit" of integers modulo 2. In addition to the field operations, numerals and case syntax are also supported. * Theory Library/Diagonalize provides constructive version of Cantor's first diagonalization argument. * Theory Library/GCD: Curried operations gcd, lcm (for nat) and zgcd, zlcm (for int); carried together from various gcd/lcm developements in the HOL Distribution. Constants zgcd and zlcm replace former igcd and ilcm; corresponding theorems renamed accordingly. INCOMPATIBILITY, may recover tupled syntax as follows: hide (open) const gcd abbreviation gcd where "gcd == (%(a, b). GCD.gcd a b)" notation (output) GCD.gcd ("gcd '(_, _')") The same works for lcm, zgcd, zlcm. * Theory Library/Nat_Infinity: added addition, numeral syntax and more instantiations for algebraic structures. Removed some duplicate theorems. Changes in simp rules. INCOMPATIBILITY. * ML antiquotation @{code} takes a constant as argument and generates corresponding code in background and inserts name of the corresponding resulting ML value/function/datatype constructor binding in place. All occurrences of @{code} with a single ML block are generated simultaneously. Provides a generic and safe interface for instrumentalizing code generation. See src/HOL/Decision_Procs/Ferrack.thy for a more ambitious application. In future you ought to refrain from ad-hoc compiling generated SML code on the ML toplevel. Note that (for technical reasons) @{code} cannot refer to constants for which user-defined serializations are set. Refer to the corresponding ML counterpart directly in that cases. * Command 'rep_datatype': instead of theorem names the command now takes a list of terms denoting the constructors of the type to be represented as datatype. The characteristic theorems have to be proven. INCOMPATIBILITY. Also observe that the following theorems have disappeared in favour of existing ones: unit_induct ~> unit.induct prod_induct ~> prod.induct sum_induct ~> sum.induct Suc_Suc_eq ~> nat.inject Suc_not_Zero Zero_not_Suc ~> nat.distinct *** HOL-Algebra *** * New locales for orders and lattices where the equivalence relation is not restricted to equality. INCOMPATIBILITY: all order and lattice locales use a record structure with field eq for the equivalence. * New theory of factorial domains. * Units_l_inv and Units_r_inv are now simp rules by default. INCOMPATIBILITY. Simplifier proof that require deletion of l_inv and/or r_inv will now also require deletion of these lemmas. * Renamed the following theorems, INCOMPATIBILITY: UpperD ~> Upper_memD LowerD ~> Lower_memD least_carrier ~> least_closed greatest_carrier ~> greatest_closed greatest_Lower_above ~> greatest_Lower_below one_zero ~> carrier_one_zero one_not_zero ~> carrier_one_not_zero (collision with assumption) *** HOL-Nominal *** * Nominal datatypes can now contain type-variables. * Commands 'nominal_inductive' and 'equivariance' work with local theory targets. * Nominal primrec can now works with local theory targets and its specification syntax now conforms to the general format as seen in 'inductive' etc. * Method "perm_simp" honours the standard simplifier attributes (no_asm), (no_asm_use) etc. * The new predicate #* is defined like freshness, except that on the left hand side can be a set or list of atoms. * Experimental command 'nominal_inductive2' derives strong induction principles for inductive definitions. In contrast to 'nominal_inductive', which can only deal with a fixed number of binders, it can deal with arbitrary expressions standing for sets of atoms to be avoided. The only inductive definition we have at the moment that needs this generalisation is the typing rule for Lets in the algorithm W: Gamma |- t1 : T1 (x,close Gamma T1)::Gamma |- t2 : T2 x#Gamma ----------------------------------------------------------------- Gamma |- Let x be t1 in t2 : T2 In this rule one wants to avoid all the binders that are introduced by "close Gamma T1". We are looking for other examples where this feature might be useful. Please let us know. *** HOLCF *** * Reimplemented the simplification procedure for proving continuity subgoals. The new simproc is extensible; users can declare additional continuity introduction rules with the attribute [cont2cont]. * The continuity simproc now uses a different introduction rule for solving continuity subgoals on terms with lambda abstractions. In some rare cases the new simproc may fail to solve subgoals that the old one could solve, and "simp add: cont2cont_LAM" may be necessary. Potential INCOMPATIBILITY. * Command 'fixrec': specification syntax now conforms to the general format as seen in 'inductive' etc. See src/HOLCF/ex/Fixrec_ex.thy for examples. INCOMPATIBILITY. *** ZF *** * Proof of Zorn's Lemma for partial orders. *** ML *** * Multithreading for Poly/ML 5.1/5.2 is no longer supported, only for Poly/ML 5.2.1 or later. Important note: the TimeLimit facility depends on multithreading, so timouts will not work before Poly/ML 5.2.1! * High-level support for concurrent ML programming, see src/Pure/Cuncurrent. The data-oriented model of "future values" is particularly convenient to organize independent functional computations. The concept of "synchronized variables" provides a higher-order interface for components with shared state, avoiding the delicate details of mutexes and condition variables. (Requires Poly/ML 5.2.1 or later.) * ML bindings produced via Isar commands are stored within the Isar context (theory or proof). Consequently, commands like 'use' and 'ML' become thread-safe and work with undo as expected (concerning top-level bindings, not side-effects on global references). INCOMPATIBILITY, need to provide proper Isar context when invoking the compiler at runtime; really global bindings need to be given outside a theory. (Requires Poly/ML 5.2 or later.) * Command 'ML_prf' is analogous to 'ML' but works within a proof context. Top-level ML bindings are stored within the proof context in a purely sequential fashion, disregarding the nested proof structure. ML bindings introduced by 'ML_prf' are discarded at the end of the proof. (Requires Poly/ML 5.2 or later.) * Simplified ML attribute and method setup, cf. functions Attrib.setup and Method.setup, as well as Isar commands 'attribute_setup' and 'method_setup'. INCOMPATIBILITY for 'method_setup', need to simplify existing code accordingly, or use plain 'setup' together with old Method.add_method. * Simplified ML oracle interface Thm.add_oracle promotes 'a -> cterm to 'a -> thm, while results are always tagged with an authentic oracle name. The Isar command 'oracle' is now polymorphic, no argument type is specified. INCOMPATIBILITY, need to simplify existing oracle code accordingly. Note that extra performance may be gained by producing the cterm carefully, avoiding slow Thm.cterm_of. * Simplified interface for defining document antiquotations via ThyOutput.antiquotation, ThyOutput.output, and optionally ThyOutput.maybe_pretty_source. INCOMPATIBILITY, need to simplify user antiquotations accordingly, see src/Pure/Thy/thy_output.ML for common examples. * More systematic treatment of long names, abstract name bindings, and name space operations. Basic operations on qualified names have been move from structure NameSpace to Long_Name, e.g. Long_Name.base_name, Long_Name.append. Old type bstring has been mostly replaced by abstract type binding (see structure Binding), which supports precise qualification by packages and local theory targets, as well as proper tracking of source positions. INCOMPATIBILITY, need to wrap old bstring values into Binding.name, or better pass through abstract bindings everywhere. See further src/Pure/General/long_name.ML, src/Pure/General/binding.ML and src/Pure/General/name_space.ML * Result facts (from PureThy.note_thms, ProofContext.note_thms, LocalTheory.note etc.) now refer to the *full* internal name, not the bstring as before. INCOMPATIBILITY, not detected by ML type-checking! * Disposed old type and term read functions (Sign.read_def_typ, Sign.read_typ, Sign.read_def_terms, Sign.read_term, Thm.read_def_cterms, Thm.read_cterm etc.). INCOMPATIBILITY, should use regular Syntax.read_typ, Syntax.read_term, Syntax.read_typ_global, Syntax.read_term_global etc.; see also OldGoals.read_term as last resort for legacy applications. * Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory (such as Addsimps, Simp_tac, SIMPSET). INCOMPATIBILITY, should use @{simpset} etc. in embedded ML text, or local_simpset_of with a proper context passed as explicit runtime argument. * Rules and tactics that read instantiations (read_instantiate, res_inst_tac, thin_tac, subgoal_tac etc.) now demand a proper proof context, which is required for parsing and type-checking. Moreover, the variables are specified as plain indexnames, not string encodings thereof. INCOMPATIBILITY. * Generic Toplevel.add_hook interface allows to analyze the result of transactions. E.g. see src/Pure/ProofGeneral/proof_general_pgip.ML for theorem dependency output of transactions resulting in a new theory state. * ML antiquotations: block-structured compilation context indicated by \ ... \; additional antiquotation forms: @{binding name} - basic name binding @{let ?pat = term} - term abbreviation (HO matching) @{note name = fact} - fact abbreviation @{thm fact} - singleton fact (with attributes) @{thms fact} - general fact (with attributes) @{lemma prop by method} - singleton goal @{lemma prop by meth1 meth2} - singleton goal @{lemma prop1 ... propN by method} - general goal @{lemma prop1 ... propN by meth1 meth2} - general goal @{lemma (open) ...} - open derivation *** System *** * The Isabelle "emacs" tool provides a specific interface to invoke Proof General / Emacs, with more explicit failure if that is not installed (the old isabelle-interface script silently falls back on isabelle-process). The PROOFGENERAL_HOME setting determines the installation location of the Proof General distribution. * Isabelle/lib/classes/Pure.jar provides basic support to integrate the Isabelle process into a JVM/Scala application. See Isabelle/lib/jedit/plugin for a minimal example. (The obsolete Java process wrapper has been discontinued.) * Added homegrown Isabelle font with unicode layout, see lib/fonts. * Various status messages (with exact source position information) are emitted, if proper markup print mode is enabled. This allows user-interface components to provide detailed feedback on internal prover operations. New in Isabelle2008 (June 2008) ------------------------------- *** General *** * The Isabelle/Isar Reference Manual (isar-ref) has been reorganized and updated, with formally checked references as hyperlinks. * Theory loader: use_thy (and similar operations) no longer set the implicit ML context, which was occasionally hard to predict and in conflict with concurrency. INCOMPATIBILITY, use ML within Isar which provides a proper context already. * Theory loader: old-style ML proof scripts being *attached* to a thy file are no longer supported. INCOMPATIBILITY, regular 'uses' and 'use' within a theory file will do the job. * Name space merge now observes canonical order, i.e. the second space is inserted into the first one, while existing entries in the first space take precedence. INCOMPATIBILITY in rare situations, may try to swap theory imports. * Syntax: symbol \ is now considered a letter. Potential INCOMPATIBILITY in identifier syntax etc. * Outer syntax: string tokens no longer admit escaped white space, which was an accidental (undocumented) feature. INCOMPATIBILITY, use white space without escapes. * Outer syntax: string tokens may contain arbitrary character codes specified via 3 decimal digits (as in SML). E.g. "foo\095bar" for "foo_bar". *** Pure *** * Context-dependent token translations. Default setup reverts locally fixed variables, and adds hilite markup for undeclared frees. * Unused theorems can be found using the new command 'unused_thms'. There are three ways of invoking it: (1) unused_thms Only finds unused theorems in the current theory. (2) unused_thms thy_1 ... thy_n - Finds unused theorems in the current theory and all of its ancestors, excluding the theories thy_1 ... thy_n and all of their ancestors. (3) unused_thms thy_1 ... thy_n - thy'_1 ... thy'_m Finds unused theorems in the theories thy'_1 ... thy'_m and all of their ancestors, excluding the theories thy_1 ... thy_n and all of their ancestors. In order to increase the readability of the list produced by unused_thms, theorems that have been created by a particular instance of a theory command such as 'inductive' or 'function' are considered to belong to the same "group", meaning that if at least one theorem in this group is used, the other theorems in the same group are no longer reported as unused. Moreover, if all theorems in the group are unused, only one theorem in the group is displayed. Note that proof objects have to be switched on in order for unused_thms to work properly (i.e. !proofs must be >= 1, which is usually the case when using Proof General with the default settings). * Authentic naming of facts disallows ad-hoc overwriting of previous theorems within the same name space. INCOMPATIBILITY, need to remove duplicate fact bindings, or even accidental fact duplications. Note that tools may maintain dynamically scoped facts systematically, using PureThy.add_thms_dynamic. * Command 'hide' now allows to hide from "fact" name space as well. * Eliminated destructive theorem database, simpset, claset, and clasimpset. Potential INCOMPATIBILITY, really need to observe linear update of theories within ML code. * Eliminated theory ProtoPure and CPure, leaving just one Pure theory. INCOMPATIBILITY, object-logics depending on former Pure require additional setup PureThy.old_appl_syntax_setup; object-logics depending on former CPure need to refer to Pure. * Commands 'use' and 'ML' are now purely functional, operating on theory/local_theory. Removed former 'ML_setup' (on theory), use 'ML' instead. Added 'ML_val' as mere diagnostic replacement for 'ML'. INCOMPATIBILITY. * Command 'setup': discontinued implicit version with ML reference. * Instantiation target allows for simultaneous specification of class instance operations together with an instantiation proof. Type-checking phase allows to refer to class operations uniformly. See src/HOL/Complex/Complex.thy for an Isar example and src/HOL/Library/Eval.thy for an ML example. * Indexing of literal facts: be more serious about including only facts from the visible specification/proof context, but not the background context (locale etc.). Affects `prop` notation and method "fact". INCOMPATIBILITY: need to name facts explicitly in rare situations. * Method "cases", "induct", "coinduct": removed obsolete/undocumented "(open)" option, which used to expose internal bound variables to the proof text. * Isar statements: removed obsolete case "rule_context". INCOMPATIBILITY, better use explicit fixes/assumes. * Locale proofs: default proof step now includes 'unfold_locales'; hence 'proof' without argument may be used to unfold locale predicates. *** Document preparation *** * Simplified pdfsetup.sty: color/hyperref is used unconditionally for both pdf and dvi (hyperlinks usually work in xdvi as well); removed obsolete thumbpdf setup (contemporary PDF viewers do this on the spot); renamed link color from "darkblue" to "linkcolor" (default value unchanged, can be redefined via \definecolor); no longer sets "a4paper" option (unnecessary or even intrusive). * Antiquotation @{lemma A method} proves proposition A by the given method (either a method name or a method name plus (optional) method arguments in parentheses) and prints A just like @{prop A}. *** HOL *** * New primrec package. Specification syntax conforms in style to definition/function/.... No separate induction rule is provided. The "primrec" command distinguishes old-style and new-style specifications by syntax. The former primrec package is now named OldPrimrecPackage. When adjusting theories, beware: constants stemming from new-style primrec specifications have authentic syntax. * Metis prover is now an order of magnitude faster, and also works with multithreading. * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed. * Sledgehammer no longer produces structured proofs by default. To enable, declare [[sledgehammer_full = true]]. Attributes reconstruction_modulus, reconstruction_sorts renamed sledgehammer_modulus, sledgehammer_sorts. INCOMPATIBILITY. * Method "induct_scheme" derives user-specified induction rules from well-founded induction and completeness of patterns. This factors out some operations that are done internally by the function package and makes them available separately. See src/HOL/ex/Induction_Scheme.thy for examples. * More flexible generation of measure functions for termination proofs: Measure functions can be declared by proving a rule of the form "is_measure f" and giving it the [measure_function] attribute. The "is_measure" predicate is logically meaningless (always true), and just guides the heuristic. To find suitable measure functions, the termination prover sets up the goal "is_measure ?f" of the appropriate type and generates all solutions by prolog-style backwards proof using the declared rules. This setup also deals with rules like "is_measure f ==> is_measure (list_size f)" which accommodates nested datatypes that recurse through lists. Similar rules are predeclared for products and option types. * Turned the type of sets "'a set" into an abbreviation for "'a => bool" INCOMPATIBILITIES: - Definitions of overloaded constants on sets have to be replaced by definitions on => and bool. - Some definitions of overloaded operators on sets can now be proved using the definitions of the operators on => and bool. Therefore, the following theorems have been renamed: subset_def -> subset_eq psubset_def -> psubset_eq set_diff_def -> set_diff_eq Compl_def -> Compl_eq Sup_set_def -> Sup_set_eq Inf_set_def -> Inf_set_eq sup_set_def -> sup_set_eq inf_set_def -> inf_set_eq - Due to the incompleteness of the HO unification algorithm, some rules such as subst may require manual instantiation, if some of the unknowns in the rule is a set. - Higher order unification and forward proofs: The proof pattern have "P (S::'a set)" <...> then have "EX S. P S" .. no longer works (due to the incompleteness of the HO unification algorithm) and must be replaced by the pattern have "EX S. P S" proof show "P S" <...> qed - Calculational reasoning with subst (or similar rules): The proof pattern have "P (S::'a set)" <...> also have "S = T" <...> finally have "P T" . no longer works (for similar reasons as the previous example) and must be replaced by something like have "P (S::'a set)" <...> moreover have "S = T" <...> ultimately have "P T" by simp - Tactics or packages written in ML code: Code performing pattern matching on types via Type ("set", [T]) => ... must be rewritten. Moreover, functions like strip_type or binder_types no longer return the right value when applied to a type of the form T1 => ... => Tn => U => bool rather than T1 => ... => Tn => U set * Merged theories Wellfounded_Recursion, Accessible_Part and Wellfounded_Relations to theory Wellfounded. * Explicit class "eq" for executable equality. INCOMPATIBILITY. * Class finite no longer treats UNIV as class parameter. Use class enum from theory Library/Enum instead to achieve a similar effect. INCOMPATIBILITY. * Theory List: rule list_induct2 now has explicitly named cases "Nil" and "Cons". INCOMPATIBILITY. * HOL (and FOL): renamed variables in rules imp_elim and swap. Potential INCOMPATIBILITY. * Theory Product_Type: duplicated lemmas split_Pair_apply and injective_fst_snd removed, use split_eta and prod_eqI instead. Renamed upd_fst to apfst and upd_snd to apsnd. INCOMPATIBILITY. * Theory Nat: removed redundant lemmas that merely duplicate lemmas of the same name in theory Orderings: less_trans less_linear le_imp_less_or_eq le_less_trans less_le_trans less_not_sym less_asym Renamed less_imp_le to less_imp_le_nat, and less_irrefl to less_irrefl_nat. Potential INCOMPATIBILITY due to more general types and different variable names. * Library/Option_ord.thy: Canonical order on option type. * Library/RBT.thy: Red-black trees, an efficient implementation of finite maps. * Library/Countable.thy: Type class for countable types. * Theory Int: The representation of numerals has changed. The infix operator BIT and the bit datatype with constructors B0 and B1 have disappeared. INCOMPATIBILITY, use "Int.Bit0 x" and "Int.Bit1 y" in place of "x BIT bit.B0" and "y BIT bit.B1", respectively. Theorems involving BIT, B0, or B1 have been renamed with "Bit0" or "Bit1" accordingly. * Theory Nat: definition of <= and < on natural numbers no longer depend on well-founded relations. INCOMPATIBILITY. Definitions le_def and less_def have disappeared. Consider lemmas not_less [symmetric, where ?'a = nat] and less_eq [symmetric] instead. * Theory Finite_Set: locales ACf, ACe, ACIf, ACIfSL and ACIfSLlin (whose purpose mainly is for various fold_set functionals) have been abandoned in favor of the existing algebraic classes ab_semigroup_mult, comm_monoid_mult, ab_semigroup_idem_mult, lower_semilattice (resp. upper_semilattice) and linorder. INCOMPATIBILITY. * Theory Transitive_Closure: induct and cases rules now declare proper case_names ("base" and "step"). INCOMPATIBILITY. * Theorem Inductive.lfp_ordinal_induct generalized to complete lattices. The form set-specific version is available as Inductive.lfp_ordinal_induct_set. * Renamed theorems "power.simps" to "power_int.simps". INCOMPATIBILITY. * Class semiring_div provides basic abstract properties of semirings with division and modulo operations. Subsumes former class dvd_mod. * Merged theories IntDef, Numeral and IntArith into unified theory Int. INCOMPATIBILITY. * Theory Library/Code_Index: type "index" now represents natural numbers rather than integers. INCOMPATIBILITY. * New class "uminus" with operation "uminus" (split of from class "minus" which now only has operation "minus", binary). INCOMPATIBILITY. * Constants "card", "internal_split", "option_map" now with authentic syntax. INCOMPATIBILITY. * Definitions subset_def, psubset_def, set_diff_def, Compl_def, le_bool_def, less_bool_def, le_fun_def, less_fun_def, inf_bool_def, sup_bool_def, Inf_bool_def, Sup_bool_def, inf_fun_def, sup_fun_def, Inf_fun_def, Sup_fun_def, inf_set_def, sup_set_def, Inf_set_def, Sup_set_def, le_def, less_def, option_map_def now with object equality. INCOMPATIBILITY. * Records. Removed K_record, and replaced it by pure lambda term %x. c. The simplifier setup is now more robust against eta expansion. INCOMPATIBILITY: in cases explicitly referring to K_record. * Library/Multiset: {#a, b, c#} abbreviates {#a#} + {#b#} + {#c#}. * Library/ListVector: new theory of arithmetic vector operations. * Library/Order_Relation: new theory of various orderings as sets of pairs. Defines preorders, partial orders, linear orders and well-orders on sets and on types. *** ZF *** * Renamed some theories to allow to loading both ZF and HOL in the same session: Datatype -> Datatype_ZF Inductive -> Inductive_ZF Int -> Int_ZF IntDiv -> IntDiv_ZF Nat -> Nat_ZF List -> List_ZF Main -> Main_ZF INCOMPATIBILITY: ZF theories that import individual theories below Main might need to be adapted. Regular theory Main is still available, as trivial extension of Main_ZF. *** ML *** * ML within Isar: antiquotation @{const name} or @{const name(typargs)} produces statically-checked Const term. * Functor NamedThmsFun: data is available to the user as dynamic fact (of the same name). Removed obsolete print command. * Removed obsolete "use_legacy_bindings" function. * The ``print mode'' is now a thread-local value derived from a global template (the former print_mode reference), thus access becomes non-critical. The global print_mode reference is for session management only; user-code should use print_mode_value, print_mode_active, PrintMode.setmp etc. INCOMPATIBILITY. * Functions system/system_out provide a robust way to invoke external shell commands, with propagation of interrupts (requires Poly/ML 5.2.1). Do not use OS.Process.system etc. from the basis library! *** System *** * Default settings: PROOFGENERAL_OPTIONS no longer impose xemacs --- in accordance with Proof General 3.7, which prefers GNU emacs. * isatool tty runs Isabelle process with plain tty interaction; optional line editor may be specified via ISABELLE_LINE_EDITOR setting, the default settings attempt to locate "ledit" and "rlwrap". * isatool browser now works with Cygwin as well, using general "javapath" function defined in Isabelle process environment. * YXML notation provides a simple and efficient alternative to standard XML transfer syntax. See src/Pure/General/yxml.ML and isatool yxml as described in the Isabelle system manual. * JVM class isabelle.IsabelleProcess (located in Isabelle/lib/classes) provides general wrapper for managing an Isabelle process in a robust fashion, with ``cooked'' output from stdin/stderr. * Rudimentary Isabelle plugin for jEdit (see Isabelle/lib/jedit), based on Isabelle/JVM process wrapper (see Isabelle/lib/classes). * Removed obsolete THIS_IS_ISABELLE_BUILD feature. NB: the documented way of changing the user's settings is via ISABELLE_HOME_USER/etc/settings, which is a fully featured bash script. * Multithreading.max_threads := 0 refers to the number of actual CPU cores of the underlying machine, which is a good starting point for optimal performance tuning. The corresponding usedir option -M allows "max" as an alias for "0". WARNING: does not work on certain versions of Mac OS (with Poly/ML 5.1). * isabelle-process: non-ML sessions are run with "nice", to reduce the adverse effect of Isabelle flooding interactive front-ends (notably ProofGeneral / XEmacs). New in Isabelle2007 (November 2007) ----------------------------------- *** General *** * More uniform information about legacy features, notably a warning/error of "Legacy feature: ...", depending on the state of the tolerate_legacy_features flag (default true). FUTURE INCOMPATIBILITY: legacy features will disappear eventually. * Theory syntax: the header format ``theory A = B + C:'' has been discontinued in favour of ``theory A imports B C begin''. Use isatool fixheaders to convert existing theory files. INCOMPATIBILITY. * Theory syntax: the old non-Isar theory file format has been discontinued altogether. Note that ML proof scripts may still be used with Isar theories; migration is usually quite simple with the ML function use_legacy_bindings. INCOMPATIBILITY. * Theory syntax: some popular names (e.g. 'class', 'declaration', 'fun', 'help', 'if') are now keywords. INCOMPATIBILITY, use double quotes. * Theory loader: be more serious about observing the static theory header specifications (including optional directories), but not the accidental file locations of previously successful loads. The strict update policy of former update_thy is now already performed by use_thy, so the former has been removed; use_thys updates several theories simultaneously, just as 'imports' within a theory header specification, but without merging the results. Potential INCOMPATIBILITY: may need to refine theory headers and commands ROOT.ML which depend on load order. * Theory loader: optional support for content-based file identification, instead of the traditional scheme of full physical path plus date stamp; configured by the ISABELLE_FILE_IDENT setting (cf. the system manual). The new scheme allows to work with non-finished theories in persistent session images, such that source files may be moved later on without requiring reloads. * Theory loader: old-style ML proof scripts being *attached* to a thy file (with the same base name as the theory) are considered a legacy feature, which will disappear eventually. Even now, the theory loader no longer maintains dependencies on such files. * Syntax: the scope for resolving ambiguities via type-inference is now limited to individual terms, instead of whole simultaneous specifications as before. This greatly reduces the complexity of the syntax module and improves flexibility by separating parsing and type-checking. INCOMPATIBILITY: additional type-constraints (explicit 'fixes' etc.) are required in rare situations. * Syntax: constants introduced by new-style packages ('definition', 'abbreviation' etc.) are passed through the syntax module in ``authentic mode''. This means that associated mixfix annotations really stick to such constants, independently of potential name space ambiguities introduced later on. INCOMPATIBILITY: constants in parse trees are represented slightly differently, may need to adapt syntax translations accordingly. Use CONST marker in 'translations' and @{const_syntax} antiquotation in 'parse_translation' etc. * Legacy goal package: reduced interface to the bare minimum required to keep existing proof scripts running. Most other user-level functions are now part of the OldGoals structure, which is *not* open by default (consider isatool expandshort before open OldGoals). Removed top_sg, prin, printyp, pprint_term/typ altogether, because these tend to cause confusion about the actual goal (!) context being used here, which is not necessarily the same as the_context(). * Command 'find_theorems': supports "*" wild-card in "name:" criterion; "with_dups" option. Certain ProofGeneral versions might support a specific search form (see ProofGeneral/CHANGES). * The ``prems limit'' option (cf. ProofContext.prems_limit) is now -1 by default, which means that "prems" (and also "fixed variables") are suppressed from proof state output. Note that the ProofGeneral settings mechanism allows to change and save options persistently, but older versions of Isabelle will fail to start up if a negative prems limit is imposed. * Local theory targets may be specified by non-nested blocks of ``context/locale/class ... begin'' followed by ``end''. The body may contain definitions, theorems etc., including any derived mechanism that has been implemented on top of these primitives. This concept generalizes the existing ``theorem (in ...)'' towards more versatility and scalability. * Proof General interface: proper undo of final 'end' command; discontinued Isabelle/classic mode (ML proof scripts). *** Document preparation *** * Added antiquotation @{theory name} which prints the given name, after checking that it refers to a valid ancestor theory in the current context. * Added antiquotations @{ML_type text} and @{ML_struct text} which check the given source text as ML type/structure, printing verbatim. * Added antiquotation @{abbrev "c args"} which prints the abbreviation "c args == rhs" given in the current context. (Any number of arguments may be given on the LHS.) *** Pure *** * The 'class' package offers a combination of axclass and locale to achieve Haskell-like type classes in Isabelle. Definitions and theorems within a class context produce both relative results (with implicit parameters according to the locale context), and polymorphic constants with qualified polymorphism (according to the class context). Within the body context of a 'class' target, a separate syntax layer ("user space type system") takes care of converting between global polymorphic consts and internal locale representation. See src/HOL/ex/Classpackage.thy for examples (as well as main HOL). "isatool doc classes" provides a tutorial. * Generic code generator framework allows to generate executable code for ML and Haskell (including Isabelle classes). A short usage sketch: internal compilation: export_code in SML writing SML code to a file: export_code in SML writing OCaml code to a file: export_code in OCaml writing Haskell code to a bunch of files: export_code in Haskell evaluating closed propositions to True/False using code generation: method ``eval'' Reasonable default setup of framework in HOL. Theorem attributs for selecting and transforming function equations theorems: [code fun]: select a theorem as function equation for a specific constant [code fun del]: deselect a theorem as function equation for a specific constant [code inline]: select an equation theorem for unfolding (inlining) in place [code inline del]: deselect an equation theorem for unfolding (inlining) in place User-defined serializations (target in {SML, OCaml, Haskell}): code_const {(target) }+ code_type {(target) }+ code_instance {(target)}+ where instance ::= :: code_class {(target) }+ where class target syntax ::= {where { == }+}? code_instance and code_class only are effective to target Haskell. For example usage see src/HOL/ex/Codegenerator.thy and src/HOL/ex/Codegenerator_Pretty.thy. A separate tutorial on code generation from Isabelle/HOL theories is available via "isatool doc codegen". * Code generator: consts in 'consts_code' Isar commands are now referred to by usual term syntax (including optional type annotations). * Command 'no_translations' removes translation rules from theory syntax. * Overloaded definitions are now actually checked for acyclic dependencies. The overloading scheme is slightly more general than that of Haskell98, although Isabelle does not demand an exact correspondence to type class and instance declarations. INCOMPATIBILITY, use ``defs (unchecked overloaded)'' to admit more exotic versions of overloading -- at the discretion of the user! Polymorphic constants are represented via type arguments, i.e. the instantiation that matches an instance against the most general declaration given in the signature. For example, with the declaration c :: 'a => 'a => 'a, an instance c :: nat => nat => nat is represented as c(nat). Overloading is essentially simultaneous structural recursion over such type arguments. Incomplete specification patterns impose global constraints on all occurrences, e.g. c('a * 'a) on the LHS means that more general c('a * 'b) will be disallowed on any RHS. Command 'print_theory' outputs the normalized system of recursive equations, see section "definitions". * Configuration options are maintained within the theory or proof context (with name and type bool/int/string), providing a very simple interface to a poor-man's version of general context data. Tools may declare options in ML (e.g. using Attrib.config_int) and then refer to these values using Config.get etc. Users may change options via an associated attribute of the same name. This form of context declaration works particularly well with commands 'declare' or 'using', for example ``declare [[foo = 42]]''. Thus it has become very easy to avoid global references, which would not observe Isar toplevel undo/redo and fail to work with multithreading. Various global ML references of Pure and HOL have been turned into configuration options: Unify.search_bound unify_search_bound Unify.trace_bound unify_trace_bound Unify.trace_simp unify_trace_simp Unify.trace_types unify_trace_types Simplifier.simp_depth_limit simp_depth_limit Blast.depth_limit blast_depth_limit DatatypeProp.dtK datatype_distinctness_limit fast_arith_neq_limit fast_arith_neq_limit fast_arith_split_limit fast_arith_split_limit * Named collections of theorems may be easily installed as context data using the functor NamedThmsFun (see also src/Pure/Tools/named_thms.ML). The user may add or delete facts via attributes; there is also a toplevel print command. This facility is just a common case of general context data, which is the preferred way for anything more complex than just a list of facts in canonical order. * Isar: command 'declaration' augments a local theory by generic declaration functions written in ML. This enables arbitrary content being added to the context, depending on a morphism that tells the difference of the original declaration context wrt. the application context encountered later on. * Isar: proper interfaces for simplification procedures. Command 'simproc_setup' declares named simprocs (with match patterns, and body text in ML). Attribute "simproc" adds/deletes simprocs in the current context. ML antiquotation @{simproc name} retrieves named simprocs. * Isar: an extra pair of brackets around attribute declarations abbreviates a theorem reference involving an internal dummy fact, which will be ignored later --- only the effect of the attribute on the background context will persist. This form of in-place declarations is particularly useful with commands like 'declare' and 'using', for example ``have A using [[simproc a]] by simp''. * Isar: method "assumption" (and implicit closing of subproofs) now takes simple non-atomic goal assumptions into account: after applying an assumption as a rule the resulting subgoals are solved by atomic assumption steps. This is particularly useful to finish 'obtain' goals, such as "!!x. (!!x. P x ==> thesis) ==> P x ==> thesis", without referring to the original premise "!!x. P x ==> thesis" in the Isar proof context. POTENTIAL INCOMPATIBILITY: method "assumption" is more permissive. * Isar: implicit use of prems from the Isar proof context is considered a legacy feature. Common applications like ``have A .'' may be replaced by ``have A by fact'' or ``note `A`''. In general, referencing facts explicitly here improves readability and maintainability of proof texts. * Isar: improper proof element 'guess' is like 'obtain', but derives the obtained context from the course of reasoning! For example: assume "EX x y. A x & B y" -- "any previous fact" then guess x and y by clarify This technique is potentially adventurous, depending on the facts and proof tools being involved here. * Isar: known facts from the proof context may be specified as literal propositions, using ASCII back-quote syntax. This works wherever named facts used to be allowed so far, in proof commands, proof methods, attributes etc. Literal facts are retrieved from the context according to unification of type and term parameters. For example, provided that "A" and "A ==> B" and "!!x. P x ==> Q x" are known theorems in the current context, then these are valid literal facts: `A` and `A ==> B` and `!!x. P x ==> Q x" as well as `P a ==> Q a` etc. There is also a proof method "fact" which does the same composition for explicit goal states, e.g. the following proof texts coincide with certain special cases of literal facts: have "A" by fact == note `A` have "A ==> B" by fact == note `A ==> B` have "!!x. P x ==> Q x" by fact == note `!!x. P x ==> Q x` have "P a ==> Q a" by fact == note `P a ==> Q a` * Isar: ":" (colon) is no longer a symbolic identifier character in outer syntax. Thus symbolic identifiers may be used without additional white space in declarations like this: ``assume *: A''. * Isar: 'print_facts' prints all local facts of the current context, both named and unnamed ones. * Isar: 'def' now admits simultaneous definitions, e.g.: def x == "t" and y == "u" * Isar: added command 'unfolding', which is structurally similar to 'using', but affects both the goal state and facts by unfolding given rewrite rules. Thus many occurrences of the 'unfold' method or 'unfolded' attribute may be replaced by first-class proof text. * Isar: methods 'unfold' / 'fold', attributes 'unfolded' / 'folded', and command 'unfolding' now all support object-level equalities (potentially conditional). The underlying notion of rewrite rule is analogous to the 'rule_format' attribute, but *not* that of the Simplifier (which is usually more generous). * Isar: the new attribute [rotated n] (default n = 1) rotates the premises of a theorem by n. Useful in conjunction with drule. * Isar: the goal restriction operator [N] (default N = 1) evaluates a method expression within a sandbox consisting of the first N sub-goals, which need to exist. For example, ``simp_all [3]'' simplifies the first three sub-goals, while (rule foo, simp_all)[] simplifies all new goals that emerge from applying rule foo to the originally first one. * Isar: schematic goals are no longer restricted to higher-order patterns; e.g. ``lemma "?P(?x)" by (rule TrueI)'' now works as expected. * Isar: the conclusion of a long theorem statement is now either 'shows' (a simultaneous conjunction, as before), or 'obtains' (essentially a disjunction of cases with local parameters and assumptions). The latter allows to express general elimination rules adequately; in this notation common elimination rules look like this: lemma exE: -- "EX x. P x ==> (!!x. P x ==> thesis) ==> thesis" assumes "EX x. P x" obtains x where "P x" lemma conjE: -- "A & B ==> (A ==> B ==> thesis) ==> thesis" assumes "A & B" obtains A and B lemma disjE: -- "A | B ==> (A ==> thesis) ==> (B ==> thesis) ==> thesis" assumes "A | B" obtains A | B The subsequent classical rules even refer to the formal "thesis" explicitly: lemma classical: -- "(~ thesis ==> thesis) ==> thesis" obtains "~ thesis" lemma Peirce's_Law: -- "((thesis ==> something) ==> thesis) ==> thesis" obtains "thesis ==> something" The actual proof of an 'obtains' statement is analogous to that of the Isar proof element 'obtain', only that there may be several cases. Optional case names may be specified in parentheses; these will be available both in the present proof and as annotations in the resulting rule, for later use with the 'cases' method (cf. attribute case_names). * Isar: the assumptions of a long theorem statement are available as "assms" fact in the proof context. This is more appropriate than the (historical) "prems", which refers to all assumptions of the current context, including those from the target locale, proof body etc. * Isar: 'print_statement' prints theorems from the current theory or proof context in long statement form, according to the syntax of a top-level lemma. * Isar: 'obtain' takes an optional case name for the local context introduction rule (default "that"). * Isar: removed obsolete 'concl is' patterns. INCOMPATIBILITY, use explicit (is "_ ==> ?foo") in the rare cases where this still happens to occur. * Pure: syntax "CONST name" produces a fully internalized constant according to the current context. This is particularly useful for syntax translations that should refer to internal constant representations independently of name spaces. * Pure: syntax constant for foo (binder "FOO ") is called "foo_binder" instead of "FOO ". This allows multiple binder declarations to coexist in the same context. INCOMPATIBILITY. * Isar/locales: 'notation' provides a robust interface to the 'syntax' primitive that also works in a locale context (both for constants and fixed variables). Type declaration and internal syntactic representation of given constants retrieved from the context. Likewise, the 'no_notation' command allows to remove given syntax annotations from the current context. * Isar/locales: new derived specification elements 'axiomatization', 'definition', 'abbreviation', which support type-inference, admit object-level specifications (equality, equivalence). See also the isar-ref manual. Examples: axiomatization eq (infix "===" 50) where eq_refl: "x === x" and eq_subst: "x === y ==> P x ==> P y" definition "f x y = x + y + 1" definition g where "g x = f x x" abbreviation neq (infix "=!=" 50) where "x =!= y == ~ (x === y)" These specifications may be also used in a locale context. Then the constants being introduced depend on certain fixed parameters, and the constant name is qualified by the locale base name. An internal abbreviation takes care for convenient input and output, making the parameters implicit and using the original short name. See also src/HOL/ex/Abstract_NAT.thy for an example of deriving polymorphic entities from a monomorphic theory. Presently, abbreviations are only available 'in' a target locale, but not inherited by general import expressions. Also note that 'abbreviation' may be used as a type-safe replacement for 'syntax' + 'translations' in common applications. The "no_abbrevs" print mode prevents folding of abbreviations in term output. Concrete syntax is attached to specified constants in internal form, independently of name spaces. The parse tree representation is slightly different -- use 'notation' instead of raw 'syntax', and 'translations' with explicit "CONST" markup to accommodate this. * Pure/Isar: unified syntax for new-style specification mechanisms (e.g. 'definition', 'abbreviation', or 'inductive' in HOL) admits full type inference and dummy patterns ("_"). For example: definition "K x _ = x" inductive conj for A B where "A ==> B ==> conj A B" * Pure: command 'print_abbrevs' prints all constant abbreviations of the current context. Print mode "no_abbrevs" prevents inversion of abbreviations on output. * Isar/locales: improved parameter handling: use of locales "var" and "struct" no longer necessary; - parameter renamings are no longer required to be injective. For example, this allows to define endomorphisms as locale endom = homom mult mult h. * Isar/locales: changed the way locales with predicates are defined. Instead of accumulating the specification, the imported expression is now an interpretation. INCOMPATIBILITY: different normal form of locale expressions. In particular, in interpretations of locales with predicates, goals repesenting already interpreted fragments are not removed automatically. Use methods `intro_locales' and `unfold_locales'; see below. * Isar/locales: new methods `intro_locales' and `unfold_locales' provide backward reasoning on locales predicates. The methods are aware of interpretations and discharge corresponding goals. `intro_locales' is less aggressive then `unfold_locales' and does not unfold predicates to assumptions. * Isar/locales: the order in which locale fragments are accumulated has changed. This enables to override declarations from fragments due to interpretations -- for example, unwanted simp rules. * Isar/locales: interpretation in theories and proof contexts has been extended. One may now specify (and prove) equations, which are unfolded in interpreted theorems. This is useful for replacing defined concepts (constants depending on locale parameters) by concepts already existing in the target context. Example: interpretation partial_order ["op <= :: [int, int] => bool"] where "partial_order.less (op <=) (x::int) y = (x < y)" Typically, the constant `partial_order.less' is created by a definition specification element in the context of locale partial_order. * Method "induct": improved internal context management to support local fixes and defines on-the-fly. Thus explicit meta-level connectives !! and ==> are rarely required anymore in inductive goals (using object-logic connectives for this purpose has been long obsolete anyway). Common proof patterns are explained in src/HOL/Induct/Common_Patterns.thy, see also src/HOL/Isar_examples/Puzzle.thy and src/HOL/Lambda for realistic examples. * Method "induct": improved handling of simultaneous goals. Instead of introducing object-level conjunction, the statement is now split into several conclusions, while the corresponding symbolic cases are nested accordingly. INCOMPATIBILITY, proofs need to be structured explicitly, see src/HOL/Induct/Common_Patterns.thy, for example. * Method "induct": mutual induction rules are now specified as a list of rule sharing the same induction cases. HOL packages usually provide foo_bar.inducts for mutually defined items foo and bar (e.g. inductive predicates/sets or datatypes). INCOMPATIBILITY, users need to specify mutual induction rules differently, i.e. like this: (induct rule: foo_bar.inducts) (induct set: foo bar) (induct pred: foo bar) (induct type: foo bar) The ML function ProjectRule.projections turns old-style rules into the new format. * Method "coinduct": dual of induction, see src/HOL/Library/Coinductive_List.thy for various examples. * Method "cases", "induct", "coinduct": the ``(open)'' option is considered a legacy feature. * Attribute "symmetric" produces result with standardized schematic variables (index 0). Potential INCOMPATIBILITY. * Simplifier: by default the simplifier trace only shows top level rewrites now. That is, trace_simp_depth_limit is set to 1 by default. Thus there is less danger of being flooded by the trace. The trace indicates where parts have been suppressed. * Provers/classical: removed obsolete classical version of elim_format attribute; classical elim/dest rules are now treated uniformly when manipulating the claset. * Provers/classical: stricter checks to ensure that supplied intro, dest and elim rules are well-formed; dest and elim rules must have at least one premise. * Provers/classical: attributes dest/elim/intro take an optional weight argument for the rule (just as the Pure versions). Weights are ignored by automated tools, but determine the search order of single rule steps. * Syntax: input syntax now supports dummy variable binding "%_. b", where the body does not mention the bound variable. Note that dummy patterns implicitly depend on their context of bounds, which makes "{_. _}" match any set comprehension as expected. Potential INCOMPATIBILITY -- parse translations need to cope with syntactic constant "_idtdummy" in the binding position. * Syntax: removed obsolete syntactic constant "_K" and its associated parse translation. INCOMPATIBILITY -- use dummy abstraction instead, for example "A -> B" => "Pi A (%_. B)". * Pure: 'class_deps' command visualizes the subclass relation, using the graph browser tool. * Pure: 'print_theory' now suppresses certain internal declarations by default; use '!' option for full details. *** HOL *** * Method "metis" proves goals by applying the Metis general-purpose resolution prover (see also http://gilith.com/software/metis/). Examples are in the directory MetisExamples. WARNING: the Isabelle/HOL-Metis integration does not yet work properly with multi-threading. * Command 'sledgehammer' invokes external automatic theorem provers as background processes. It generates calls to the "metis" method if successful. These can be pasted into the proof. Users do not have to wait for the automatic provers to return. WARNING: does not really work with multi-threading. * New "auto_quickcheck" feature tests outermost goal statements for potential counter-examples. Controlled by ML references auto_quickcheck (default true) and auto_quickcheck_time_limit (default 5000 milliseconds). Fails silently if statements is outside of executable fragment, or any other codgenerator problem occurs. * New constant "undefined" with axiom "undefined x = undefined". * Added class "HOL.eq", allowing for code generation with polymorphic equality. * Some renaming of class constants due to canonical name prefixing in the new 'class' package: HOL.abs ~> HOL.abs_class.abs HOL.divide ~> HOL.divide_class.divide 0 ~> HOL.zero_class.zero 1 ~> HOL.one_class.one op + ~> HOL.plus_class.plus op - ~> HOL.minus_class.minus uminus ~> HOL.minus_class.uminus op * ~> HOL.times_class.times op < ~> HOL.ord_class.less op <= > HOL.ord_class.less_eq Nat.power ~> Power.power_class.power Nat.size ~> Nat.size_class.size Numeral.number_of ~> Numeral.number_class.number_of FixedPoint.Inf ~> Lattices.complete_lattice_class.Inf FixedPoint.Sup ~> Lattices.complete_lattice_class.Sup Orderings.min ~> Orderings.ord_class.min Orderings.max ~> Orderings.ord_class.max Divides.op div ~> Divides.div_class.div Divides.op mod ~> Divides.div_class.mod Divides.op dvd ~> Divides.div_class.dvd INCOMPATIBILITY. Adaptions may be required in the following cases: a) User-defined constants using any of the names "plus", "minus", "times", "less" or "less_eq". The standard syntax translations for "+", "-" and "*" may go wrong. INCOMPATIBILITY: use more specific names. b) Variables named "plus", "minus", "times", "less", "less_eq" INCOMPATIBILITY: use more specific names. c) Permutative equations (e.g. "a + b = b + a") Since the change of names also changes the order of terms, permutative rewrite rules may get applied in a different order. Experience shows that this is rarely the case (only two adaptions in the whole Isabelle distribution). INCOMPATIBILITY: rewrite proofs d) ML code directly refering to constant names This in general only affects hand-written proof tactics, simprocs and so on. INCOMPATIBILITY: grep your sourcecode and replace names. Consider using @{const_name} antiquotation. * New class "default" with associated constant "default". * Function "sgn" is now overloaded and available on int, real, complex (and other numeric types), using class "sgn". Two possible defs of sgn are given as equational assumptions in the classes sgn_if and sgn_div_norm; ordered_idom now also inherits from sgn_if. INCOMPATIBILITY. * Locale "partial_order" now unified with class "order" (cf. theory Orderings), added parameter "less". INCOMPATIBILITY. * Renamings in classes "order" and "linorder": facts "refl", "trans" and "cases" to "order_refl", "order_trans" and "linorder_cases", to avoid clashes with HOL "refl" and "trans". INCOMPATIBILITY. * Classes "order" and "linorder": potential INCOMPATIBILITY due to changed order of proof goals in instance proofs. * The transitivity reasoner for partial and linear orders is set up for classes "order" and "linorder". Instances of the reasoner are available in all contexts importing or interpreting the corresponding locales. Method "order" invokes the reasoner separately; the reasoner is also integrated with the Simplifier as a solver. Diagnostic command 'print_orders' shows the available instances of the reasoner in the current context. * Localized monotonicity predicate in theory "Orderings"; integrated lemmas max_of_mono and min_of_mono with this predicate. INCOMPATIBILITY. * Formulation of theorem "dense" changed slightly due to integration with new class dense_linear_order. * Uniform lattice theory development in HOL. constants "meet" and "join" now named "inf" and "sup" constant "Meet" now named "Inf" classes "meet_semilorder" and "join_semilorder" now named "lower_semilattice" and "upper_semilattice" class "lorder" now named "lattice" class "comp_lat" now named "complete_lattice" Instantiation of lattice classes allows explicit definitions for "inf" and "sup" operations (or "Inf" and "Sup" for complete lattices). INCOMPATIBILITY. Theorem renames: meet_left_le ~> inf_le1 meet_right_le ~> inf_le2 join_left_le ~> sup_ge1 join_right_le ~> sup_ge2 meet_join_le ~> inf_sup_ord le_meetI ~> le_infI join_leI ~> le_supI le_meet ~> le_inf_iff le_join ~> ge_sup_conv meet_idempotent ~> inf_idem join_idempotent ~> sup_idem meet_comm ~> inf_commute join_comm ~> sup_commute meet_leI1 ~> le_infI1 meet_leI2 ~> le_infI2 le_joinI1 ~> le_supI1 le_joinI2 ~> le_supI2 meet_assoc ~> inf_assoc join_assoc ~> sup_assoc meet_left_comm ~> inf_left_commute meet_left_idempotent ~> inf_left_idem join_left_comm ~> sup_left_commute join_left_idempotent ~> sup_left_idem meet_aci ~> inf_aci join_aci ~> sup_aci le_def_meet ~> le_iff_inf le_def_join ~> le_iff_sup join_absorp2 ~> sup_absorb2 join_absorp1 ~> sup_absorb1 meet_absorp1 ~> inf_absorb1 meet_absorp2 ~> inf_absorb2 meet_join_absorp ~> inf_sup_absorb join_meet_absorp ~> sup_inf_absorb distrib_join_le ~> distrib_sup_le distrib_meet_le ~> distrib_inf_le add_meet_distrib_left ~> add_inf_distrib_left add_join_distrib_left ~> add_sup_distrib_left is_join_neg_meet ~> is_join_neg_inf is_meet_neg_join ~> is_meet_neg_sup add_meet_distrib_right ~> add_inf_distrib_right add_join_distrib_right ~> add_sup_distrib_right add_meet_join_distribs ~> add_sup_inf_distribs join_eq_neg_meet ~> sup_eq_neg_inf meet_eq_neg_join ~> inf_eq_neg_sup add_eq_meet_join ~> add_eq_inf_sup meet_0_imp_0 ~> inf_0_imp_0 join_0_imp_0 ~> sup_0_imp_0 meet_0_eq_0 ~> inf_0_eq_0 join_0_eq_0 ~> sup_0_eq_0 neg_meet_eq_join ~> neg_inf_eq_sup neg_join_eq_meet ~> neg_sup_eq_inf join_eq_if ~> sup_eq_if mono_meet ~> mono_inf mono_join ~> mono_sup meet_bool_eq ~> inf_bool_eq join_bool_eq ~> sup_bool_eq meet_fun_eq ~> inf_fun_eq join_fun_eq ~> sup_fun_eq meet_set_eq ~> inf_set_eq join_set_eq ~> sup_set_eq meet1_iff ~> inf1_iff meet2_iff ~> inf2_iff meet1I ~> inf1I meet2I ~> inf2I meet1D1 ~> inf1D1 meet2D1 ~> inf2D1 meet1D2 ~> inf1D2 meet2D2 ~> inf2D2 meet1E ~> inf1E meet2E ~> inf2E join1_iff ~> sup1_iff join2_iff ~> sup2_iff join1I1 ~> sup1I1 join2I1 ~> sup2I1 join1I1 ~> sup1I1 join2I2 ~> sup1I2 join1CI ~> sup1CI join2CI ~> sup2CI join1E ~> sup1E join2E ~> sup2E is_meet_Meet ~> is_meet_Inf Meet_bool_def ~> Inf_bool_def Meet_fun_def ~> Inf_fun_def Meet_greatest ~> Inf_greatest Meet_lower ~> Inf_lower Meet_set_def ~> Inf_set_def Sup_def ~> Sup_Inf Sup_bool_eq ~> Sup_bool_def Sup_fun_eq ~> Sup_fun_def Sup_set_eq ~> Sup_set_def listsp_meetI ~> listsp_infI listsp_meet_eq ~> listsp_inf_eq meet_min ~> inf_min join_max ~> sup_max * Added syntactic class "size"; overloaded constant "size" now has type "'a::size ==> bool" * Internal reorganisation of `size' of datatypes: size theorems "foo.size" are no longer subsumed by "foo.simps" (but are still simplification rules by default!); theorems "prod.size" now named "*.size". * Class "div" now inherits from class "times" rather than "type". INCOMPATIBILITY. * HOL/Finite_Set: "name-space" locales Lattice, Distrib_lattice, Linorder etc. have disappeared; operations defined in terms of fold_set now are named Inf_fin, Sup_fin. INCOMPATIBILITY. * HOL/Nat: neq0_conv no longer declared as iff. INCOMPATIBILITY. * HOL-Word: New extensive library and type for generic, fixed size machine words, with arithemtic, bit-wise, shifting and rotating operations, reflection into int, nat, and bool lists, automation for linear arithmetic (by automatic reflection into nat or int), including lemmas on overflow and monotonicity. Instantiated to all appropriate arithmetic type classes, supporting automatic simplification of numerals on all operations. * Library/Boolean_Algebra: locales for abstract boolean algebras. * Library/Numeral_Type: numbers as types, e.g. TYPE(32). * Code generator library theories: - Code_Integer represents HOL integers by big integer literals in target languages. - Code_Char represents HOL characters by character literals in target languages. - Code_Char_chr like Code_Char, but also offers treatment of character codes; includes Code_Integer. - Executable_Set allows to generate code for finite sets using lists. - Executable_Rat implements rational numbers as triples (sign, enumerator, denominator). - Executable_Real implements a subset of real numbers, namly those representable by rational numbers. - Efficient_Nat implements natural numbers by integers, which in general will result in higher efficency; pattern matching with 0/Suc is eliminated; includes Code_Integer. - Code_Index provides an additional datatype index which is mapped to target-language built-in integers. - Code_Message provides an additional datatype message_string which is isomorphic to strings; messages are mapped to target-language strings. * New package for inductive predicates An n-ary predicate p with m parameters z_1, ..., z_m can now be defined via inductive p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" for z_1 :: U_1 and ... and z_n :: U_m where rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n" | ... with full support for type-inference, rather than consts s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" abbreviation p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" where "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m" inductive "s z_1 ... z_m" intros rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m" ... For backward compatibility, there is a wrapper allowing inductive sets to be defined with the new package via inductive_set s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" for z_1 :: U_1 and ... and z_n :: U_m where rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m" | ... or inductive_set s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" and p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" for z_1 :: U_1 and ... and z_n :: U_m where "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m" | rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n" | ... if the additional syntax "p ..." is required. Numerous examples can be found in the subdirectories src/HOL/Auth, src/HOL/Bali, src/HOL/Induct, and src/HOL/MicroJava. INCOMPATIBILITIES: - Since declaration and definition of inductive sets or predicates is no longer separated, abbreviations involving the newly introduced sets or predicates must be specified together with the introduction rules after the 'where' keyword (see above), rather than before the actual inductive definition. - The variables in induction and elimination rules are now quantified in the order of their occurrence in the introduction rules, rather than in alphabetical order. Since this may break some proofs, these proofs either have to be repaired, e.g. by reordering the variables a_i_1 ... a_i_{k_i} in Isar 'case' statements of the form case (rule_i a_i_1 ... a_i_{k_i}) or the old order of quantification has to be restored by explicitly adding meta-level quantifiers in the introduction rules, i.e. | rule_i: "!!a_i_1 ... a_i_{k_i}. ... ==> p z_1 ... z_m t_i_1 ... t_i_n" - The format of the elimination rules is now p z_1 ... z_m x_1 ... x_n ==> (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P) ==> ... ==> P for predicates and (x_1, ..., x_n) : s z_1 ... z_m ==> (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P) ==> ... ==> P for sets rather than x : s z_1 ... z_m ==> (!!a_1_1 ... a_1_{k_1}. x = (t_1_1, ..., t_1_n) ==> ... ==> P) ==> ... ==> P This may require terms in goals to be expanded to n-tuples (e.g. using case_tac or simplification with the split_paired_all rule) before the above elimination rule is applicable. - The elimination or case analysis rules for (mutually) inductive sets or predicates are now called "p_1.cases" ... "p_k.cases". The list of rules "p_1_..._p_k.elims" is no longer available. * New package "function"/"fun" for general recursive functions, supporting mutual and nested recursion, definitions in local contexts, more general pattern matching and partiality. See HOL/ex/Fundefs.thy for small examples, and the separate tutorial on the function package. The old recdef "package" is still available as before, but users are encouraged to use the new package. * Method "lexicographic_order" automatically synthesizes termination relations as lexicographic combinations of size measures. * Case-expressions allow arbitrary constructor-patterns (including "_") and take their order into account, like in functional programming. Internally, this is translated into nested case-expressions; missing cases are added and mapped to the predefined constant "undefined". In complicated cases printing may no longer show the original input but the internal form. Lambda-abstractions allow the same form of pattern matching: "% pat1 => e1 | ..." is an abbreviation for "%x. case x of pat1 => e1 | ..." where x is a new variable. * IntDef: The constant "int :: nat => int" has been removed; now "int" is an abbreviation for "of_nat :: nat => int". The simplification rules for "of_nat" have been changed to work like "int" did previously. Potential INCOMPATIBILITY: - "of_nat (Suc m)" simplifies to "1 + of_nat m" instead of "of_nat m + 1" - of_nat_diff and of_nat_mult are no longer default simp rules * Method "algebra" solves polynomial equations over (semi)rings using Groebner bases. The (semi)ring structure is defined by locales and the tool setup depends on that generic context. Installing the method for a specific type involves instantiating the locale and possibly adding declarations for computation on the coefficients. The method is already instantiated for natural numbers and for the axiomatic class of idoms with numerals. See also the paper by Chaieb and Wenzel at CALCULEMUS 2007 for the general principles underlying this architecture of context-aware proof-tools. * Method "ferrack" implements quantifier elimination over special-purpose dense linear orders using locales (analogous to "algebra"). The method is already installed for class {ordered_field,recpower,number_ring} which subsumes real, hyperreal, rat, etc. * Former constant "List.op @" now named "List.append". Use ML antiquotations @{const_name List.append} or @{term " ... @ ... "} to circumvent possible incompatibilities when working on ML level. * primrec: missing cases mapped to "undefined" instead of "arbitrary". * New function listsum :: 'a list => 'a for arbitrary monoids. Special syntax: "SUM x <- xs. f x" (and latex variants) * New syntax for Haskell-like list comprehension (input only), eg. [(x,y). x <- xs, y <- ys, x ~= y], see also src/HOL/List.thy. * The special syntax for function "filter" has changed from [x : xs. P] to [x <- xs. P] to avoid an ambiguity caused by list comprehension syntax, and for uniformity. INCOMPATIBILITY. * [a..b] is now defined for arbitrary linear orders. It used to be defined on nat only, as an abbreviation for [a.. B" for equality on bool (with priority 25 like -->); output depends on the "iff" print_mode, the default is "A = B" (with priority 50). * Relations less (<) and less_eq (<=) are also available on type bool. Modified syntax to disallow nesting without explicit parentheses, e.g. "(x < y) < z" or "x < (y < z)", but NOT "x < y < z". Potential INCOMPATIBILITY. * "LEAST x:A. P" expands to "LEAST x. x:A & P" (input only). * Relation composition operator "op O" now has precedence 75 and binds stronger than union and intersection. INCOMPATIBILITY. * The old set interval syntax "{m..n(}" (and relatives) has been removed. Use "{m.. ==> False", equivalences (i.e. "=" on type bool) are handled, variable names of the form "lit_" are no longer reserved, significant speedup. * Methods "sat" and "satx" can now replay MiniSat proof traces. zChaff is still supported as well. * 'inductive' and 'datatype': provide projections of mutual rules, bundled as foo_bar.inducts; * Library: moved theories Parity, GCD, Binomial, Infinite_Set to Library. * Library: moved theory Accessible_Part to main HOL. * Library: added theory Coinductive_List of potentially infinite lists as greatest fixed-point. * Library: added theory AssocList which implements (finite) maps as association lists. * Method "evaluation" solves goals (i.e. a boolean expression) efficiently by compiling it to ML. The goal is "proved" (via an oracle) if it evaluates to True. * Linear arithmetic now splits certain operators (e.g. min, max, abs) also when invoked by the simplifier. This results in the Simplifier being more powerful on arithmetic goals. INCOMPATIBILITY. Configuration option fast_arith_split_limit=0 recovers the old behavior. * Support for hex (0x20) and binary (0b1001) numerals. * New method: reify eqs (t), where eqs are equations for an interpretation I :: 'a list => 'b => 'c and t::'c is an optional parameter, computes a term s::'b and a list xs::'a list and proves the theorem I xs s = t. This is also known as reification or quoting. The resulting theorem is applied to the subgoal to substitute t with I xs s. If t is omitted, the subgoal itself is reified. * New method: reflection corr_thm eqs (t). The parameters eqs and (t) are as explained above. corr_thm is a theorem for I vs (f t) = I vs t, where f is supposed to be a computable function (in the sense of code generattion). The method uses reify to compute s and xs as above then applies corr_thm and uses normalization by evaluation to "prove" f s = r and finally gets the theorem t = r, which is again applied to the subgoal. An Example is available in src/HOL/ex/ReflectionEx.thy. * Reflection: Automatic reification now handels binding, an example is available in src/HOL/ex/ReflectionEx.thy * HOL-Statespace: ``State Spaces: The Locale Way'' introduces a command 'statespace' that is similar to 'record', but introduces an abstract specification based on the locale infrastructure instead of HOL types. This leads to extra flexibility in composing state spaces, in particular multiple inheritance and renaming of components. *** HOL-Complex *** * Hyperreal: Functions root and sqrt are now defined on negative real inputs so that root n (- x) = - root n x and sqrt (- x) = - sqrt x. Nonnegativity side conditions have been removed from many lemmas, so that more subgoals may now be solved by simplification; potential INCOMPATIBILITY. * Real: new type classes formalize real normed vector spaces and algebras, using new overloaded constants scaleR :: real => 'a => 'a and norm :: 'a => real. * Real: constant of_real :: real => 'a::real_algebra_1 injects from reals into other types. The overloaded constant Reals :: 'a set is now defined as range of_real; potential INCOMPATIBILITY. * Real: proper support for ML code generation, including 'quickcheck'. Reals are implemented as arbitrary precision rationals. * Hyperreal: Several constants that previously worked only for the reals have been generalized, so they now work over arbitrary vector spaces. Type annotations may need to be added in some cases; potential INCOMPATIBILITY. Infinitesimal :: ('a::real_normed_vector) star set HFinite :: ('a::real_normed_vector) star set HInfinite :: ('a::real_normed_vector) star set approx :: ('a::real_normed_vector) star => 'a star => bool monad :: ('a::real_normed_vector) star => 'a star set galaxy :: ('a::real_normed_vector) star => 'a star set (NS)LIMSEQ :: [nat => 'a::real_normed_vector, 'a] => bool (NS)convergent :: (nat => 'a::real_normed_vector) => bool (NS)Bseq :: (nat => 'a::real_normed_vector) => bool (NS)Cauchy :: (nat => 'a::real_normed_vector) => bool (NS)LIM :: ['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool is(NS)Cont :: ['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool deriv :: ['a::real_normed_field => 'a, 'a, 'a] => bool sgn :: 'a::real_normed_vector => 'a exp :: 'a::{recpower,real_normed_field,banach} => 'a * Complex: Some complex-specific constants are now abbreviations for overloaded ones: complex_of_real = of_real, cmod = norm, hcmod = hnorm. Other constants have been entirely removed in favor of the polymorphic versions (INCOMPATIBILITY): approx <-- capprox HFinite <-- CFinite HInfinite <-- CInfinite Infinitesimal <-- CInfinitesimal monad <-- cmonad galaxy <-- cgalaxy (NS)LIM <-- (NS)CLIM, (NS)CRLIM is(NS)Cont <-- is(NS)Contc, is(NS)contCR (ns)deriv <-- (ns)cderiv *** HOL-Algebra *** * Formalisation of ideals and the quotient construction over rings. * Order and lattice theory no longer based on records. INCOMPATIBILITY. * Renamed lemmas least_carrier -> least_closed and greatest_carrier -> greatest_closed. INCOMPATIBILITY. * Method algebra is now set up via an attribute. For examples see Ring.thy. INCOMPATIBILITY: the method is now weaker on combinations of algebraic structures. * Renamed theory CRing to Ring. *** HOL-Nominal *** * Substantial, yet incomplete support for nominal datatypes (binding structures) based on HOL-Nominal logic. See src/HOL/Nominal and src/HOL/Nominal/Examples. Prospective users should consult http://isabelle.in.tum.de/nominal/ *** ML *** * ML basics: just one true type int, which coincides with IntInf.int (even on SML/NJ). * ML within Isar: antiquotations allow to embed statically-checked formal entities in the source, referring to the context available at compile-time. For example: ML {* @{sort "{zero,one}"} *} ML {* @{typ "'a => 'b"} *} ML {* @{term "%x. x"} *} ML {* @{prop "x == y"} *} ML {* @{ctyp "'a => 'b"} *} ML {* @{cterm "%x. x"} *} ML {* @{cprop "x == y"} *} ML {* @{thm asm_rl} *} ML {* @{thms asm_rl} *} ML {* @{type_name c} *} ML {* @{type_syntax c} *} ML {* @{const_name c} *} ML {* @{const_syntax c} *} ML {* @{context} *} ML {* @{theory} *} ML {* @{theory Pure} *} ML {* @{theory_ref} *} ML {* @{theory_ref Pure} *} ML {* @{simpset} *} ML {* @{claset} *} ML {* @{clasimpset} *} The same works for sources being ``used'' within an Isar context. * ML in Isar: improved error reporting; extra verbosity with ML_Context.trace enabled. * Pure/General/table.ML: the join operations now works via exceptions DUP/SAME instead of type option. This is simpler in simple cases, and admits slightly more efficient complex applications. * Pure: 'advanced' translation functions (parse_translation etc.) now use Context.generic instead of just theory. * Pure: datatype Context.generic joins theory/Proof.context and provides some facilities for code that works in either kind of context, notably GenericDataFun for uniform theory and proof data. * Pure: simplified internal attribute type, which is now always Context.generic * thm -> Context.generic * thm. Global (theory) vs. local (Proof.context) attributes have been discontinued, while minimizing code duplication. Thm.rule_attribute and Thm.declaration_attribute build canonical attributes; see also structure Context for further operations on Context.generic, notably GenericDataFun. INCOMPATIBILITY, need to adapt attribute type declarations and definitions. * Context data interfaces (Theory/Proof/GenericDataFun): removed name/print, uninitialized data defaults to ad-hoc copy of empty value, init only required for impure data. INCOMPATIBILITY: empty really need to be empty (no dependencies on theory content!) * Pure/kernel: consts certification ignores sort constraints given in signature declarations. (This information is not relevant to the logic, but only for type inference.) SIGNIFICANT INTERNAL CHANGE, potential INCOMPATIBILITY. * Pure: axiomatic type classes are now purely definitional, with explicit proofs of class axioms and super class relations performed internally. See Pure/axclass.ML for the main internal interfaces -- notably AxClass.define_class supercedes AxClass.add_axclass, and AxClass.axiomatize_class/classrel/arity supersede Sign.add_classes/classrel/arities. * Pure/Isar: Args/Attrib parsers operate on Context.generic -- global/local versions on theory vs. Proof.context have been discontinued; Attrib.syntax and Method.syntax have been adapted accordingly. INCOMPATIBILITY, need to adapt parser expressions for attributes, methods, etc. * Pure: several functions of signature "... -> theory -> theory * ..." have been reoriented to "... -> theory -> ... * theory" in order to allow natural usage in combination with the ||>, ||>>, |-> and fold_map combinators. * Pure: official theorem names (closed derivations) and additional comments (tags) are now strictly separate. Name hints -- which are maintained as tags -- may be attached any time without affecting the derivation. * Pure: primitive rule lift_rule now takes goal cterm instead of an actual goal state (thm). Use Thm.lift_rule (Thm.cprem_of st i) to achieve the old behaviour. * Pure: the "Goal" constant is now called "prop", supporting a slightly more general idea of ``protecting'' meta-level rule statements. * Pure: Logic.(un)varify only works in a global context, which is now enforced instead of silently assumed. INCOMPATIBILITY, may use Logic.legacy_(un)varify as temporary workaround. * Pure: structure Name provides scalable operations for generating internal variable names, notably Name.variants etc. This replaces some popular functions from term.ML: Term.variant -> Name.variant Term.variantlist -> Name.variant_list Term.invent_names -> Name.invent_list Note that low-level renaming rarely occurs in new code -- operations from structure Variable are used instead (see below). * Pure: structure Variable provides fundamental operations for proper treatment of fixed/schematic variables in a context. For example, Variable.import introduces fixes for schematics of given facts and Variable.export reverses the effect (up to renaming) -- this replaces various freeze_thaw operations. * Pure: structure Goal provides simple interfaces for init/conclude/finish and tactical prove operations (replacing former Tactic.prove). Goal.prove is the canonical way to prove results within a given context; Goal.prove_global is a degraded version for theory level goals, including a global Drule.standard. Note that OldGoals.prove_goalw_cterm has long been obsolete, since it is ill-behaved in a local proof context (e.g. with local fixes/assumes or in a locale context). * Pure/Syntax: generic interfaces for parsing (Syntax.parse_term etc.) and type checking (Syntax.check_term etc.), with common combinations (Syntax.read_term etc.). These supersede former Sign.read_term etc. which are considered legacy and await removal. * Pure/Syntax: generic interfaces for type unchecking (Syntax.uncheck_terms etc.) and unparsing (Syntax.unparse_term etc.), with common combinations (Syntax.pretty_term, Syntax.string_of_term etc.). Former Sign.pretty_term, Sign.string_of_term etc. are still available for convenience, but refer to the very same operations using a mere theory instead of a full context. * Isar: simplified treatment of user-level errors, using exception ERROR of string uniformly. Function error now merely raises ERROR, without any side effect on output channels. The Isar toplevel takes care of proper display of ERROR exceptions. ML code may use plain handle/can/try; cat_error may be used to concatenate errors like this: ... handle ERROR msg => cat_error msg "..." Toplevel ML code (run directly or through the Isar toplevel) may be embedded into the Isar toplevel with exception display/debug like this: Isar.toplevel (fn () => ...) INCOMPATIBILITY, removed special transform_error facilities, removed obsolete variants of user-level exceptions (ERROR_MESSAGE, Context.PROOF, ProofContext.CONTEXT, Proof.STATE, ProofHistory.FAIL) -- use plain ERROR instead. * Isar: theory setup now has type (theory -> theory), instead of a list. INCOMPATIBILITY, may use #> to compose setup functions. * Isar: ML toplevel pretty printer for type Proof.context, subject to ProofContext.debug/verbose flags. * Isar: Toplevel.theory_to_proof admits transactions that modify the theory before entering a proof state. Transactions now always see a quasi-functional intermediate checkpoint, both in interactive and batch mode. * Isar: simplified interfaces for outer syntax. Renamed OuterSyntax.add_keywords to OuterSyntax.keywords. Removed OuterSyntax.add_parsers -- this functionality is now included in OuterSyntax.command etc. INCOMPATIBILITY. * Simplifier: the simpset of a running simplification process now contains a proof context (cf. Simplifier.the_context), which is the very context that the initial simpset has been retrieved from (by simpset_of/local_simpset_of). Consequently, all plug-in components (solver, looper etc.) may depend on arbitrary proof data. * Simplifier.inherit_context inherits the proof context (plus the local bounds) of the current simplification process; any simproc etc. that calls the Simplifier recursively should do this! Removed former Simplifier.inherit_bounds, which is already included here -- INCOMPATIBILITY. Tools based on low-level rewriting may even have to specify an explicit context using Simplifier.context/theory_context. * Simplifier/Classical Reasoner: more abstract interfaces change_simpset/claset for modifying the simpset/claset reference of a theory; raw versions simpset/claset_ref etc. have been discontinued -- INCOMPATIBILITY. * Provers: more generic wrt. syntax of object-logics, avoid hardwired "Trueprop" etc. *** System *** * settings: the default heap location within ISABELLE_HOME_USER now includes ISABELLE_IDENTIFIER. This simplifies use of multiple Isabelle installations. * isabelle-process: option -S (secure mode) disables some critical operations, notably runtime compilation and evaluation of ML source code. * Basic Isabelle mode for jEdit, see Isabelle/lib/jedit/. * Support for parallel execution, using native multicore support of Poly/ML 5.1. The theory loader exploits parallelism when processing independent theories, according to the given theory header specifications. The maximum number of worker threads is specified via usedir option -M or the "max-threads" setting in Proof General. A speedup factor of 1.5--3.5 can be expected on a 4-core machine, and up to 6 on a 8-core machine. User-code needs to observe certain guidelines for thread-safe programming, see appendix A in the Isar Implementation manual. New in Isabelle2005 (October 2005) ---------------------------------- *** General *** * Theory headers: the new header syntax for Isar theories is theory imports ... uses ... begin where the 'uses' part is optional. The previous syntax theory = + ... + : will disappear in the next release. Use isatool fixheaders to convert existing theory files. Note that there is no change in ancient non-Isar theories now, but these will disappear soon. * Theory loader: parent theories can now also be referred to via relative and absolute paths. * Command 'find_theorems' searches for a list of criteria instead of a list of constants. Known criteria are: intro, elim, dest, name:string, simp:term, and any term. Criteria can be preceded by '-' to select theorems that do not match. Intro, elim, dest select theorems that match the current goal, name:s selects theorems whose fully qualified name contain s, and simp:term selects all simplification rules whose lhs match term. Any other term is interpreted as pattern and selects all theorems matching the pattern. Available in ProofGeneral under 'ProofGeneral -> Find Theorems' or C-c C-f. Example: C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL." prints the last 100 theorems matching the pattern "(_::nat) + _ + _", matching the current goal as introduction rule and not having "HOL." in their name (i.e. not being defined in theory HOL). * Command 'thms_containing' has been discontinued in favour of 'find_theorems'; INCOMPATIBILITY. * Communication with Proof General is now 8bit clean, which means that Unicode text in UTF-8 encoding may be used within theory texts (both formal and informal parts). Cf. option -U of the Isabelle Proof General interface. Here are some simple examples (cf. src/HOL/ex): http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html http://isabelle.in.tum.de/library/HOL/ex/Chinese.html * Improved efficiency of the Simplifier and, to a lesser degree, the Classical Reasoner. Typical big applications run around 2 times faster. *** Document preparation *** * Commands 'display_drafts' and 'print_drafts' perform simple output of raw sources. Only those symbols that do not require additional LaTeX packages (depending on comments in isabellesym.sty) are displayed properly, everything else is left verbatim. isatool display and isatool print are used as front ends (these are subject to the DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively). * Command tags control specific markup of certain regions of text, notably folding and hiding. Predefined tags include "theory" (for theory begin and end), "proof" for proof commands, and "ML" for commands involving ML code; the additional tags "visible" and "invisible" are unused by default. Users may give explicit tag specifications in the text, e.g. ''by %invisible (auto)''. The interpretation of tags is determined by the LaTeX job during document preparation: see option -V of isatool usedir, or options -n and -t of isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag, \isadroptag. Several document versions may be produced at the same time via isatool usedir (the generated index.html will link all of them). Typical specifications include ''-V document=theory,proof,ML'' to present theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit these parts without any formal replacement text. The Isabelle site default settings produce ''document'' and ''outline'' versions as specified above. * Several new antiquotations: @{term_type term} prints a term with its type annotated; @{typeof term} prints the type of a term; @{const const} is the same as @{term const}, but checks that the argument is a known logical constant; @{term_style style term} and @{thm_style style thm} print a term or theorem applying a "style" to it @{ML text} Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of definitions, equations, inequations etc., 'concl' printing only the conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19' to print the specified premise. TermStyle.add_style provides an ML interface for introducing further styles. See also the "LaTeX Sugar" document practical applications. The ML antiquotation prints type-checked ML expressions verbatim. * Markup commands 'chapter', 'section', 'subsection', 'subsubsection', and 'text' support optional locale specification '(in loc)', which specifies the default context for interpreting antiquotations. For example: 'text (in lattice) {* @{thm inf_assoc}*}'. * Option 'locale=NAME' of antiquotations specifies an alternative context interpreting the subsequent argument. For example: @{thm [locale=lattice] inf_assoc}. * Proper output of proof terms (@{prf ...} and @{full_prf ...}) within a proof context. * Proper output of antiquotations for theory commands involving a proof context (such as 'locale' or 'theorem (in loc) ...'). * Delimiters of outer tokens (string etc.) now produce separate LaTeX macros (\isachardoublequoteopen, isachardoublequoteclose etc.). * isatool usedir: new option -C (default true) controls whether option -D should include a copy of the original document directory; -C false prevents unwanted effects such as copying of administrative CVS data. *** Pure *** * Considerably improved version of 'constdefs' command. Now performs automatic type-inference of declared constants; additional support for local structure declarations (cf. locales and HOL records), see also isar-ref manual. Potential INCOMPATIBILITY: need to observe strictly sequential dependencies of definitions within a single 'constdefs' section; moreover, the declared name needs to be an identifier. If all fails, consider to fall back on 'consts' and 'defs' separately. * Improved indexed syntax and implicit structures. First of all, indexed syntax provides a notational device for subscripted application, using the new syntax \<^bsub>term\<^esub> for arbitrary expressions. Secondly, in a local context with structure declarations, number indexes \<^sub>n or the empty index (default number 1) refer to a certain fixed variable implicitly; option show_structs controls printing of implicit structures. Typical applications of these concepts involve record types and locales. * New command 'no_syntax' removes grammar declarations (and translations) resulting from the given syntax specification, which is interpreted in the same manner as for the 'syntax' command. * 'Advanced' translation functions (parse_translation etc.) may depend on the signature of the theory context being presently used for parsing/printing, see also isar-ref manual. * Improved 'oracle' command provides a type-safe interface to turn an ML expression of type theory -> T -> term into a primitive rule of type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle is already included here); see also FOL/ex/IffExample.thy; INCOMPATIBILITY. * axclass: name space prefix for class "c" is now "c_class" (was "c" before); "cI" is no longer bound, use "c.intro" instead. INCOMPATIBILITY. This change avoids clashes of fact bindings for axclasses vs. locales. * Improved internal renaming of symbolic identifiers -- attach primes instead of base 26 numbers. * New flag show_question_marks controls printing of leading question marks in schematic variable names. * In schematic variable names, *any* symbol following \<^isub> or \<^isup> is now treated as part of the base name. For example, the following works without printing of awkward ".0" indexes: lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1" by simp * Inner syntax includes (*(*nested*) comments*). * Pretty printer now supports unbreakable blocks, specified in mixfix annotations as "(00...)". * Clear separation of logical types and nonterminals, where the latter may only occur in 'syntax' specifications or type abbreviations. Before that distinction was only partially implemented via type class "logic" vs. "{}". Potential INCOMPATIBILITY in rare cases of improper use of 'types'/'consts' instead of 'nonterminals'/'syntax'. Some very exotic syntax specifications may require further adaption (e.g. Cube/Cube.thy). * Removed obsolete type class "logic", use the top sort {} instead. Note that non-logical types should be declared as 'nonterminals' rather than 'types'. INCOMPATIBILITY for new object-logic specifications. * Attributes 'induct' and 'cases': type or set names may now be locally fixed variables as well. * Simplifier: can now control the depth to which conditional rewriting is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth Limit. * Simplifier: simplification procedures may now take the current simpset into account (cf. Simplifier.simproc(_i) / mk_simproc interface), which is very useful for calling the Simplifier recursively. Minor INCOMPATIBILITY: the 'prems' argument of simprocs is gone -- use prems_of_ss on the simpset instead. Moreover, the low-level mk_simproc no longer applies Logic.varify internally, to allow for use in a context of fixed variables. * thin_tac now works even if the assumption being deleted contains !! or ==>. More generally, erule now works even if the major premise of the elimination rule contains !! or ==>. * Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY. * Reorganized bootstrapping of the Pure theories; CPure is now derived from Pure, which contains all common declarations already. Both theories are defined via plain Isabelle/Isar .thy files. INCOMPATIBILITY: elements of CPure (such as the CPure.intro / CPure.elim / CPure.dest attributes) now appear in the Pure name space; use isatool fixcpure to adapt your theory and ML sources. * New syntax 'name(i-j, i-, i, ...)' for referring to specific selections of theorems in named facts via index ranges. * 'print_theorems': in theory mode, really print the difference wrt. the last state (works for interactive theory development only), in proof mode print all local facts (cf. 'print_facts'); * 'hide': option '(open)' hides only base names. * More efficient treatment of intermediate checkpoints in interactive theory development. * Code generator is now invoked via code_module (incremental code generation) and code_library (modular code generation, ML structures for each theory). INCOMPATIBILITY: new keywords 'file' and 'contains' must be quoted when used as identifiers. * New 'value' command for reading, evaluating and printing terms using the code generator. INCOMPATIBILITY: command keyword 'value' must be quoted when used as identifier. *** Locales *** * New commands for the interpretation of locale expressions in theories (1), locales (2) and proof contexts (3). These generate proof obligations from the expression specification. After the obligations have been discharged, theorems of the expression are added to the theory, target locale or proof context. The synopsis of the commands is a follows: (1) interpretation expr inst (2) interpretation target < expr (3) interpret expr inst Interpretation in theories and proof contexts require a parameter instantiation of terms from the current context. This is applied to specifications and theorems of the interpreted expression. Interpretation in locales only permits parameter renaming through the locale expression. Interpretation is smart in that interpretations that are active already do not occur in proof obligations, neither are instantiated theorems stored in duplicate. Use 'print_interps' to inspect active interpretations of a particular locale. For details, see the Isar Reference manual. Examples can be found in HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy. INCOMPATIBILITY: former 'instantiate' has been withdrawn, use 'interpret' instead. * New context element 'constrains' for adding type constraints to parameters. * Context expressions: renaming of parameters with syntax redeclaration. * Locale declaration: 'includes' disallowed. * Proper static binding of attribute syntax -- i.e. types / terms / facts mentioned as arguments are always those of the locale definition context, independently of the context of later invocations. Moreover, locale operations (renaming and type / term instantiation) are applied to attribute arguments as expected. INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of actual attributes; rare situations may require Attrib.attribute to embed those attributes into Attrib.src that lack concrete syntax. Attribute implementations need to cooperate properly with the static binding mechanism. Basic parsers Args.XXX_typ/term/prop and Attrib.XXX_thm etc. already do the right thing without further intervention. Only unusual applications -- such as "where" or "of" (cf. src/Pure/Isar/attrib.ML), which process arguments depending both on the context and the facts involved -- may have to assign parsed values to argument tokens explicitly. * Changed parameter management in theorem generation for long goal statements with 'includes'. INCOMPATIBILITY: produces a different theorem statement in rare situations. * Locale inspection command 'print_locale' omits notes elements. Use 'print_locale!' to have them included in the output. *** Provers *** * Provers/hypsubst.ML: improved version of the subst method, for single-step rewriting: it now works in bound variable contexts. New is 'subst (asm)', for rewriting an assumption. INCOMPATIBILITY: may rewrite a different subterm than the original subst method, which is still available as 'simplesubst'. * Provers/quasi.ML: new transitivity reasoners for transitivity only and quasi orders. * Provers/trancl.ML: new transitivity reasoner for transitive and reflexive-transitive closure of relations. * Provers/blast.ML: new reference depth_limit to make blast's depth limit (previously hard-coded with a value of 20) user-definable. * Provers/simplifier.ML has been moved to Pure, where Simplifier.setup is peformed already. Object-logics merely need to finish their initial simpset configuration as before. INCOMPATIBILITY. *** HOL *** * Symbolic syntax of Hilbert Choice Operator is now as follows: syntax (epsilon) "_Eps" :: "[pttrn, bool] => 'a" ("(3\_./ _)" [0, 10] 10) The symbol \ is displayed as the alternative epsilon of LaTeX and x-symbol; use option '-m epsilon' to get it actually printed. Moreover, the mathematically important symbolic identifier \ becomes available as variable, constant etc. INCOMPATIBILITY, * "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x". Similarly for all quantifiers: "ALL x > y" etc. The x-symbol for >= is \. New transitivity rules have been added to HOL/Orderings.thy to support corresponding Isar calculations. * "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\" instead of ":". * theory SetInterval: changed the syntax for open intervals: Old New {..n(} {.. {\1<\.\.} \.\.\([^(}]*\)(} -> \.\.<\1} * Theory Commutative_Ring (in Library): method comm_ring for proving equalities in commutative rings; method 'algebra' provides a generic interface. * Theory Finite_Set: changed the syntax for 'setsum', summation over finite sets: "setsum (%x. e) A", which used to be "\x:A. e", is now either "SUM x:A. e" or "\x \ A. e". The bound variable can be a tuple pattern. Some new syntax forms are available: "\x | P. e" for "setsum (%x. e) {x. P}" "\x = a..b. e" for "setsum (%x. e) {a..b}" "\x = a..x < k. e" for "setsum (%x. e) {..x < k. e" used to be based on a separate function "Summation", which has been discontinued. * theory Finite_Set: in structured induction proofs, the insert case is now 'case (insert x F)' instead of the old counterintuitive 'case (insert F x)'. * The 'refute' command has been extended to support a much larger fragment of HOL, including axiomatic type classes, constdefs and typedefs, inductive datatypes and recursion. * New tactics 'sat' and 'satx' to prove propositional tautologies. Requires zChaff with proof generation to be installed. See HOL/ex/SAT_Examples.thy for examples. * Datatype induction via method 'induct' now preserves the name of the induction variable. For example, when proving P(xs::'a list) by induction on xs, the induction step is now P(xs) ==> P(a#xs) rather than P(list) ==> P(a#list) as previously. Potential INCOMPATIBILITY in unstructured proof scripts. * Reworked implementation of records. Improved scalability for records with many fields, avoiding performance problems for type inference. Records are no longer composed of nested field types, but of nested extension types. Therefore the record type only grows linear in the number of extensions and not in the number of fields. The top-level (users) view on records is preserved. Potential INCOMPATIBILITY only in strange cases, where the theory depends on the old record representation. The type generated for a record is called _ext_type. Flag record_quick_and_dirty_sensitive can be enabled to skip the proofs triggered by a record definition or a simproc (if quick_and_dirty is enabled). Definitions of large records can take quite long. New simproc record_upd_simproc for simplification of multiple record updates enabled by default. Moreover, trivial updates are also removed: r(|x := x r|) = r. INCOMPATIBILITY: old proofs break occasionally, since simplification is more powerful by default. * typedef: proper support for polymorphic sets, which contain extra type-variables in the term. * Simplifier: automatically reasons about transitivity chains involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics provided by Provers/trancl.ML as additional solvers. INCOMPATIBILITY: old proofs break occasionally as simplification may now solve more goals than previously. * Simplifier: converts x <= y into x = y if assumption y <= x is present. Works for all partial orders (class "order"), in particular numbers and sets. For linear orders (e.g. numbers) it treats ~ x < y just like y <= x. * Simplifier: new simproc for "let x = a in f x". If a is a free or bound variable or a constant then the let is unfolded. Otherwise first a is simplified to b, and then f b is simplified to g. If possible we abstract b from g arriving at "let x = b in h x", otherwise we unfold the let and arrive at g. The simproc can be enabled/disabled by the reference use_let_simproc. Potential INCOMPATIBILITY since simplification is more powerful by default. * Classical reasoning: the meson method now accepts theorems as arguments. * Prover support: pre-release of the Isabelle-ATP linkup, which runs background jobs to provide advice on the provability of subgoals. * Theory OrderedGroup and Ring_and_Field: various additions and improvements to faciliate calculations involving equalities and inequalities. The following theorems have been eliminated or modified (INCOMPATIBILITY): abs_eq now named abs_of_nonneg abs_of_ge_0 now named abs_of_nonneg abs_minus_eq now named abs_of_nonpos imp_abs_id now named abs_of_nonneg imp_abs_neg_id now named abs_of_nonpos mult_pos now named mult_pos_pos mult_pos_le now named mult_nonneg_nonneg mult_pos_neg_le now named mult_nonneg_nonpos mult_pos_neg2_le now named mult_nonneg_nonpos2 mult_neg now named mult_neg_neg mult_neg_le now named mult_nonpos_nonpos * The following lemmas in Ring_and_Field have been added to the simplifier: zero_le_square not_square_less_zero The following lemmas have been deleted from Real/RealPow: realpow_zero_zero realpow_two realpow_less zero_le_power realpow_two_le abs_realpow_two realpow_two_abs * Theory Parity: added rules for simplifying exponents. * Theory List: The following theorems have been eliminated or modified (INCOMPATIBILITY): list_all_Nil now named list_all.simps(1) list_all_Cons now named list_all.simps(2) list_all_conv now named list_all_iff set_mem_eq now named mem_iff * Theories SetsAndFunctions and BigO (see HOL/Library) support asymptotic "big O" calculations. See the notes in BigO.thy. *** HOL-Complex *** * Theory RealDef: better support for embedding natural numbers and integers in the reals. The following theorems have been eliminated or modified (INCOMPATIBILITY): exp_ge_add_one_self now requires no hypotheses real_of_int_add reversed direction of equality (use [symmetric]) real_of_int_minus reversed direction of equality (use [symmetric]) real_of_int_diff reversed direction of equality (use [symmetric]) real_of_int_mult reversed direction of equality (use [symmetric]) * Theory RComplete: expanded support for floor and ceiling functions. * Theory Ln is new, with properties of the natural logarithm * Hyperreal: There is a new type constructor "star" for making nonstandard types. The old type names are now type synonyms: hypreal = real star hypnat = nat star hcomplex = complex star * Hyperreal: Many groups of similarly-defined constants have been replaced by polymorphic versions (INCOMPATIBILITY): star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex starset <-- starsetNat, starsetC *s* <-- *sNat*, *sc* starset_n <-- starsetNat_n, starsetC_n *sn* <-- *sNatn*, *scn* InternalSets <-- InternalNatSets, InternalCSets starfun <-- starfun{Nat,Nat2,C,RC,CR} *f* <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR* starfun_n <-- starfun{Nat,Nat2,C,RC,CR}_n *fn* <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn* InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs * Hyperreal: Many type-specific theorems have been removed in favor of theorems specific to various axiomatic type classes (INCOMPATIBILITY): add_commute <-- {hypreal,hypnat,hcomplex}_add_commute add_assoc <-- {hypreal,hypnat,hcomplex}_add_assocs OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right right_minus <-- hypreal_add_minus left_minus <-- {hypreal,hcomplex}_add_minus_left mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left mult_1_right <-- hcomplex_mult_one_right mult_zero_left <-- hcomplex_mult_zero_left left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib right_distrib <-- hypnat_add_mult_distrib2 zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one right_inverse <-- hypreal_mult_inverse left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left order_refl <-- {hypreal,hypnat}_le_refl order_trans <-- {hypreal,hypnat}_le_trans order_antisym <-- {hypreal,hypnat}_le_anti_sym order_less_le <-- {hypreal,hypnat}_less_le linorder_linear <-- {hypreal,hypnat}_le_linear add_left_mono <-- {hypreal,hypnat}_add_left_mono mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2 add_nonneg_nonneg <-- hypreal_le_add_order * Hyperreal: Separate theorems having to do with type-specific versions of constants have been merged into theorems that apply to the new polymorphic constants (INCOMPATIBILITY): STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set STAR_Un <-- {STAR,NatStar,STARC}_Un STAR_Int <-- {STAR,NatStar,STARC}_Int STAR_Compl <-- {STAR,NatStar,STARC}_Compl STAR_subset <-- {STAR,NatStar,STARC}_subset STAR_mem <-- {STAR,NatStar,STARC}_mem STAR_mem_Compl <-- {STAR,STARC}_mem_Compl STAR_diff <-- {STAR,STARC}_diff STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real, STARC_hcomplex_of_complex}_image_subset starset_n_Un <-- starset{Nat,C}_n_Un starset_n_Int <-- starset{Nat,C}_n_Int starset_n_Compl <-- starset{Nat,C}_n_Compl starset_n_diff <-- starset{Nat,C}_n_diff InternalSets_Un <-- Internal{Nat,C}Sets_Un InternalSets_Int <-- Internal{Nat,C}Sets_Int InternalSets_Compl <-- Internal{Nat,C}Sets_Compl InternalSets_diff <-- Internal{Nat,C}Sets_diff InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C} starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR} starfun <-- starfun{Nat,Nat2,C,RC,CR} starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus starfun_diff <-- starfun{C,RC,CR}_diff starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2 starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff starfun_Id <-- starfunC_Id starfun_approx <-- starfun{Nat,CR}_approx starfun_capprox <-- starfun{C,RC}_capprox starfun_abs <-- starfunNat_rabs starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2 starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox starfun_add_capprox <-- starfun{C,RC}_add_capprox starfun_add_approx <-- starfunCR_add_approx starfun_inverse_inverse <-- starfunC_inverse_inverse starfun_divide <-- starfun{C,CR,RC}_divide starfun_n <-- starfun{Nat,C}_n starfun_n_mult <-- starfun{Nat,C}_n_mult starfun_n_add <-- starfun{Nat,C}_n_add starfun_n_add_minus <-- starfunNat_n_add_minus starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun starfun_n_minus <-- starfun{Nat,C}_n_minus starfun_n_eq <-- starfun{Nat,C}_n_eq star_n_add <-- {hypreal,hypnat,hcomplex}_add star_n_minus <-- {hypreal,hcomplex}_minus star_n_diff <-- {hypreal,hcomplex}_diff star_n_mult <-- {hypreal,hcomplex}_mult star_n_inverse <-- {hypreal,hcomplex}_inverse star_n_le <-- {hypreal,hypnat}_le star_n_less <-- {hypreal,hypnat}_less star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num star_n_abs <-- hypreal_hrabs star_n_divide <-- hcomplex_divide star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus star_of_diff <-- hypreal_of_real_diff star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int star_of_number_of <-- {hypreal,hcomplex}_number_of star_of_number_less <-- number_of_less_hypreal_of_real_iff star_of_number_le <-- number_of_le_hypreal_of_real_iff star_of_eq_number <-- hypreal_of_real_eq_number_of_iff star_of_less_number <-- hypreal_of_real_less_number_of_iff star_of_le_number <-- hypreal_of_real_le_number_of_iff star_of_power <-- hypreal_of_real_power star_of_eq_0 <-- hcomplex_of_complex_zero_iff * Hyperreal: new method "transfer" that implements the transfer principle of nonstandard analysis. With a subgoal that mentions nonstandard types like "'a star", the command "apply transfer" replaces it with an equivalent one that mentions only standard types. To be successful, all free variables must have standard types; non- standard variables must have explicit universal quantifiers. * Hyperreal: A theory of Taylor series. *** HOLCF *** * Discontinued special version of 'constdefs' (which used to support continuous functions) in favor of the general Pure one with full type-inference. * New simplification procedure for solving continuity conditions; it is much faster on terms with many nested lambda abstractions (cubic instead of exponential time). * New syntax for domain package: selector names are now optional. Parentheses should be omitted unless argument is lazy, for example: domain 'a stream = cons "'a" (lazy "'a stream") * New command 'fixrec' for defining recursive functions with pattern matching; defining multiple functions with mutual recursion is also supported. Patterns may include the constants cpair, spair, up, sinl, sinr, or any data constructor defined by the domain package. The given equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for syntax and examples. * New commands 'cpodef' and 'pcpodef' for defining predicate subtypes of cpo and pcpo types. Syntax is exactly like the 'typedef' command, but the proof obligation additionally includes an admissibility requirement. The packages generate instances of class cpo or pcpo, with continuity and strictness theorems for Rep and Abs. * HOLCF: Many theorems have been renamed according to a more standard naming scheme (INCOMPATIBILITY): foo_inject: "foo$x = foo$y ==> x = y" foo_eq: "(foo$x = foo$y) = (x = y)" foo_less: "(foo$x << foo$y) = (x << y)" foo_strict: "foo$UU = UU" foo_defined: "... ==> foo$x ~= UU" foo_defined_iff: "(foo$x = UU) = (x = UU)" *** ZF *** * ZF/ex: theories Group and Ring provide examples in abstract algebra, including the First Isomorphism Theorem (on quotienting by the kernel of a homomorphism). * ZF/Simplifier: install second copy of type solver that actually makes use of TC rules declared to Isar proof contexts (or locales); the old version is still required for ML proof scripts. *** Cube *** * Converted to Isar theory format; use locales instead of axiomatic theories. *** ML *** * Pure/library.ML: added ##>, ##>>, #>> -- higher-order counterparts for ||>, ||>>, |>>, * Pure/library.ML no longer defines its own option datatype, but uses that of the SML basis, which has constructors NONE and SOME instead of None and Some, as well as exception Option.Option instead of OPTION. The functions the, if_none, is_some, is_none have been adapted accordingly, while Option.map replaces apsome. * Pure/library.ML: the exception LIST has been given up in favour of the standard exceptions Empty and Subscript, as well as Library.UnequalLengths. Function like Library.hd and Library.tl are superceded by the standard hd and tl functions etc. A number of basic list functions are no longer exported to the ML toplevel, as they are variants of predefined functions. The following suggests how one can translate existing code: rev_append xs ys = List.revAppend (xs, ys) nth_elem (i, xs) = List.nth (xs, i) last_elem xs = List.last xs flat xss = List.concat xss seq fs = List.app fs partition P xs = List.partition P xs mapfilter f xs = List.mapPartial f xs * Pure/library.ML: several combinators for linear functional transformations, notably reverse application and composition: x |> f f #> g (x, y) |-> f f #-> g * Pure/library.ML: introduced/changed precedence of infix operators: infix 1 |> |-> ||> ||>> |>> |>>> #> #->; infix 2 ?; infix 3 o oo ooo oooo; infix 4 ~~ upto downto; Maybe INCOMPATIBILITY when any of those is used in conjunction with other infix operators. * Pure/library.ML: natural list combinators fold, fold_rev, and fold_map support linear functional transformations and nesting. For example: fold f [x1, ..., xN] y = y |> f x1 |> ... |> f xN (fold o fold) f [xs1, ..., xsN] y = y |> fold f xs1 |> ... |> fold f xsN fold f [x1, ..., xN] = f x1 #> ... #> f xN (fold o fold) f [xs1, ..., xsN] = fold f xs1 #> ... #> fold f xsN * Pure/library.ML: the following selectors on type 'a option are available: the: 'a option -> 'a (*partial*) these: 'a option -> 'a where 'a = 'b list the_default: 'a -> 'a option -> 'a the_list: 'a option -> 'a list * Pure/General: structure AList (cf. Pure/General/alist.ML) provides basic operations for association lists, following natural argument order; moreover the explicit equality predicate passed here avoids potentially expensive polymorphic runtime equality checks. The old functions may be expressed as follows: assoc = uncurry (AList.lookup (op =)) assocs = these oo AList.lookup (op =) overwrite = uncurry (AList.update (op =)) o swap * Pure/General: structure AList (cf. Pure/General/alist.ML) provides val make: ('a -> 'b) -> 'a list -> ('a * 'b) list val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list replacing make_keylist and keyfilter (occassionally used) Naive rewrites: make_keylist = AList.make keyfilter = AList.find (op =) * eq_fst and eq_snd now take explicit equality parameter, thus avoiding eqtypes. Naive rewrites: eq_fst = eq_fst (op =) eq_snd = eq_snd (op =) * Removed deprecated apl and apr (rarely used). Naive rewrites: apl (n, op) =>>= curry op n apr (op, m) =>>= fn n => op (n, m) * Pure/General: structure OrdList (cf. Pure/General/ord_list.ML) provides a reasonably efficient light-weight implementation of sets as lists. * Pure/General: generic tables (cf. Pure/General/table.ML) provide a few new operations; existing lookup and update are now curried to follow natural argument order (for use with fold etc.); INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort. * Pure/General: output via the Isabelle channels of writeln/warning/error etc. is now passed through Output.output, with a hook for arbitrary transformations depending on the print_mode (cf. Output.add_mode -- the first active mode that provides a output function wins). Already formatted output may be embedded into further text via Output.raw; the result of Pretty.string_of/str_of and derived functions (string_of_term/cterm/thm etc.) is already marked raw to accommodate easy composition of diagnostic messages etc. Programmers rarely need to care about Output.output or Output.raw at all, with some notable exceptions: Output.output is required when bypassing the standard channels (writeln etc.), or in token translations to produce properly formatted results; Output.raw is required when capturing already output material that will eventually be presented to the user a second time. For the default print mode, both Output.output and Output.raw have no effect. * Pure/General: Output.time_accumulator NAME creates an operator ('a -> 'b) -> 'a -> 'b to measure runtime and count invocations; the cumulative results are displayed at the end of a batch session. * Pure/General: File.sysify_path and File.quote_sysify path have been replaced by File.platform_path and File.shell_path (with appropriate hooks). This provides a clean interface for unusual systems where the internal and external process view of file names are different. * Pure: more efficient orders for basic syntactic entities: added fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast orders now -- potential INCOMPATIBILITY for code that depends on a particular order for Symtab.keys, Symtab.dest, etc. (consider using Library.sort_strings on result). * Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types, fold_types traverse types/terms from left to right, observing natural argument order. Supercedes previous foldl_XXX versions, add_frees, add_vars etc. have been adapted as well: INCOMPATIBILITY. * Pure: name spaces have been refined, with significant changes of the internal interfaces -- INCOMPATIBILITY. Renamed cond_extern(_table) to extern(_table). The plain name entry path is superceded by a general 'naming' context, which also includes the 'policy' to produce a fully qualified name and external accesses of a fully qualified name; NameSpace.extend is superceded by context dependent Sign.declare_name. Several theory and proof context operations modify the naming context. Especially note Theory.restore_naming and ProofContext.restore_naming to get back to a sane state; note that Theory.add_path is no longer sufficient to recover from Theory.absolute_path in particular. * Pure: new flags short_names (default false) and unique_names (default true) for controlling output of qualified names. If short_names is set, names are printed unqualified. If unique_names is reset, the name prefix is reduced to the minimum required to achieve the original result when interning again, even if there is an overlap with earlier declarations. * Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is now 'extend', and 'merge' gets an additional Pretty.pp argument (useful for printing error messages). INCOMPATIBILITY. * Pure: major reorganization of the theory context. Type Sign.sg and Theory.theory are now identified, referring to the universal Context.theory (see Pure/context.ML). Actual signature and theory content is managed as theory data. The old code and interfaces were spread over many files and structures; the new arrangement introduces considerable INCOMPATIBILITY to gain more clarity: Context -- theory management operations (name, identity, inclusion, parents, ancestors, merge, etc.), plus generic theory data; Sign -- logical signature and syntax operations (declaring consts, types, etc.), plus certify/read for common entities; Theory -- logical theory operations (stating axioms, definitions, oracles), plus a copy of logical signature operations (consts, types, etc.); also a few basic management operations (Theory.copy, Theory.merge, etc.) The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm etc.) as well as the sign field in Thm.rep_thm etc. have been retained for convenience -- they merely return the theory. * Pure: type Type.tsig is superceded by theory in most interfaces. * Pure: the Isar proof context type is already defined early in Pure as Context.proof (note that ProofContext.context and Proof.context are aliases, where the latter is the preferred name). This enables other Isabelle components to refer to that type even before Isar is present. * Pure/sign/theory: discontinued named name spaces (i.e. classK, typeK, constK, axiomK, oracleK), but provide explicit operations for any of these kinds. For example, Sign.intern typeK is now Sign.intern_type, Theory.hide_space Sign.typeK is now Theory.hide_types. Also note that former Theory.hide_classes/types/consts are now Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions internalize their arguments! INCOMPATIBILITY. * Pure: get_thm interface (of PureThy and ProofContext) expects datatype thmref (with constructors Name and NameSelection) instead of plain string -- INCOMPATIBILITY; * Pure: cases produced by proof methods specify options, where NONE means to remove case bindings -- INCOMPATIBILITY in (RAW_)METHOD_CASES. * Pure: the following operations retrieve axioms or theorems from a theory node or theory hierarchy, respectively: Theory.axioms_of: theory -> (string * term) list Theory.all_axioms_of: theory -> (string * term) list PureThy.thms_of: theory -> (string * thm) list PureThy.all_thms_of: theory -> (string * thm) list * Pure: print_tac now outputs the goal through the trace channel. * Isar toplevel: improved diagnostics, mostly for Poly/ML only. Reference Toplevel.debug (default false) controls detailed printing and tracing of low-level exceptions; Toplevel.profiling (default 0) controls execution profiling -- set to 1 for time and 2 for space (both increase the runtime). * Isar session: The initial use of ROOT.ML is now always timed, i.e. the log will show the actual process times, in contrast to the elapsed wall-clock time that the outer shell wrapper produces. * Simplifier: improved handling of bound variables (nameless representation, avoid allocating new strings). Simprocs that invoke the Simplifier recursively should use Simplifier.inherit_bounds to avoid local name clashes. Failure to do so produces warnings "Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds for further details. * ML functions legacy_bindings and use_legacy_bindings produce ML fact bindings for all theorems stored within a given theory; this may help in porting non-Isar theories to Isar ones, while keeping ML proof scripts for the time being. * ML operator HTML.with_charset specifies the charset begin used for generated HTML files. For example: HTML.with_charset "utf-8" use_thy "Hebrew"; HTML.with_charset "utf-8" use_thy "Chinese"; *** System *** * Allow symlinks to all proper Isabelle executables (Isabelle, isabelle, isatool etc.). * ISABELLE_DOC_FORMAT setting specifies preferred document format (for isatool doc, isatool mkdir, display_drafts etc.). * isatool usedir: option -f allows specification of the ML file to be used by Isabelle; default is ROOT.ML. * New isatool version outputs the version identifier of the Isabelle distribution being used. * HOL: new isatool dimacs2hol converts files in DIMACS CNF format (containing Boolean satisfiability problems) into Isabelle/HOL theories. New in Isabelle2004 (April 2004) -------------------------------- *** General *** * Provers/order.ML: new efficient reasoner for partial and linear orders. Replaces linorder.ML. * Pure: Greek letters (except small lambda, \), as well as Gothic (\...\\...\), calligraphic (\...\), and Euler (\...\), are now considered normal letters, and can therefore be used anywhere where an ASCII letter (a...zA...Z) has until now. COMPATIBILITY: This obviously changes the parsing of some terms, especially where a symbol has been used as a binder, say '\x. ...', which is now a type error since \x will be parsed as an identifier. Fix it by inserting a space around former symbols. Call 'isatool fixgreek' to try to fix parsing errors in existing theory and ML files. * Pure: Macintosh and Windows line-breaks are now allowed in theory files. * Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now allowed in identifiers. Similar to Greek letters \<^isub> is now considered a normal (but invisible) letter. For multiple letter subscripts repeat \<^isub> like this: x\<^isub>1\<^isub>2. * Pure: There are now sub-/superscripts that can span more than one character. Text between \<^bsub> and \<^esub> is set in subscript in ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in superscript. The new control characters are not identifier parts. * Pure: Control-symbols of the form \<^raw:...> will literally print the content of "..." to the latex file instead of \isacntrl... . The "..." may consist of any printable characters excluding the end bracket >. * Pure: Using new Isar command "finalconsts" (or the ML functions Theory.add_finals or Theory.add_finals_i) it is now possible to declare constants "final", which prevents their being given a definition later. It is useful for constants whose behaviour is fixed axiomatically rather than definitionally, such as the meta-logic connectives. * Pure: 'instance' now handles general arities with general sorts (i.e. intersections of classes), * Presentation: generated HTML now uses a CSS style sheet to make layout (somewhat) independent of content. It is copied from lib/html/isabelle.css. It can be changed to alter the colors/layout of generated pages. *** Isar *** * Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac, cut_tac, subgoal_tac and thin_tac: - Now understand static (Isar) contexts. As a consequence, users of Isar locales are no longer forced to write Isar proof scripts. For details see Isar Reference Manual, paragraph 4.3.2: Further tactic emulations. - INCOMPATIBILITY: names of variables to be instantiated may no longer be enclosed in quotes. Instead, precede variable name with `?'. This is consistent with the instantiation attribute "where". * Attributes "where" and "of": - Now take type variables of instantiated theorem into account when reading the instantiation string. This fixes a bug that caused instantiated theorems to have too special types in some circumstances. - "where" permits explicit instantiations of type variables. * Calculation commands "moreover" and "also" no longer interfere with current facts ("this"), admitting arbitrary combinations with "then" and derived forms. * Locales: - Goal statements involving the context element "includes" no longer generate theorems with internal delta predicates (those ending on "_axioms") in the premise. Resolve particular premise with .intro to obtain old form. - Fixed bug in type inference ("unify_frozen") that prevented mix of target specification and "includes" elements in goal statement. - Rule sets .intro and .axioms no longer declared as [intro?] and [elim?] (respectively) by default. - Experimental command for instantiation of locales in proof contexts: instantiate