Isabelle NEWS -- history user-relevant changes ============================================== New in this Isabelle release ---------------------------- *** General *** * Command 'find_theorems': support "*" wildcard in "name:" criterion. *** Document preparation *** * Added antiquotations @{ML_type text} and @{ML_struct text} which check the given source text as ML type/structure, printing verbatim. *** Pure *** * 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. * 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. * Removed obsolete syntactic constant "_K" and its associated parse translation. INCOMPATIBILITY -- use dummy abstraction instead, for example "A -> B" => "Pi A (%_. B)". *** HOL *** * In the context of the assumption "~(s = t)" the Simplifier rewrites "t = s" to False (by simproc "neq_simproc"). For backward compatibility this can be disabled by ML "reset use_neq_simproc". * Tactics 'sat' and 'satx' reimplemented, several improvements: goals no longer need to be stated as " ==> False", equivalences (i.e. "=" on type bool) are handled, variable names of the form "lit_" are no longer reserved, significant speedup. *** ML *** * 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 complements (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. * 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. 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 * 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 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