2 Isabelle NEWS -- history user-relevant changes
3 ==============================================
5 New in Isabelle99-2 (February 2001)
6 -----------------------------------
8 *** Overview of INCOMPATIBILITIES ***
10 * HOL: inductive package no longer splits induction rule aggressively,
11 but only as far as specified by the introductions given; the old
12 format may be recovered via ML function complete_split_rule or attribute
13 'split_rule (complete)';
15 * HOL: induct renamed to lfp_induct, lfp_Tarski to lfp_unfold,
16 gfp_Tarski to gfp_unfold;
18 * HOL: contrapos, contrapos2 renamed to contrapos_nn, contrapos_pp;
20 * HOL: infix "dvd" now has priority 50 rather than 70 (because it is a
21 relation); infix "^^" has been renamed "``"; infix "``" has been
22 renamed "`"; "univalent" has been renamed "single_valued";
24 * HOL/Real: "rinv" and "hrinv" replaced by overloaded "inverse"
27 * HOLCF: infix "`" has been renamed "$"; the symbol syntax is \<cdot>;
29 * Isar: 'obtain' no longer declares "that" fact as simp/intro;
31 * Isar/HOL: method 'induct' now handles non-atomic goals; as a
32 consequence, it is no longer monotonic wrt. the local goal context
33 (which is now passed through the inductive cases);
35 * Document preparation: renamed standard symbols \<ll> to \<lless> and
39 *** Document preparation ***
41 * \isabellestyle{NAME} selects version of Isabelle output (currently
42 available: are "it" for near math-mode best-style output, "sl" for
43 slanted text style, and "tt" for plain type-writer; if no
44 \isabellestyle command is given, output is according to slanted
47 * support sub/super scripts (for single symbols only), input syntax is
48 like this: "A\<^sup>*" or "A\<^sup>\<star>";
50 * some more standard symbols; see Appendix A of the system manual for
51 the complete list of symbols defined in isabellesym.sty;
53 * improved isabelle style files; more abstract symbol implementation
54 (should now use \isamath{...} and \isatext{...} in custom symbol
57 * antiquotation @{goals} and @{subgoals} for output of *dynamic* goals
58 state; Note that presentation of goal states does not conform to
59 actual human-readable proof documents. Please do not include goal
60 states into document output unless you really know what you are doing!
62 * proper indentation of antiquoted output with proportional LaTeX
65 * no_document ML operator temporarily disables LaTeX document
68 * isatool unsymbolize tunes sources for plain ASCII communication;
73 * Pure: Isar now suffers initial goal statements to contain unbound
74 schematic variables (this does not conform to actual readable proof
75 documents, due to unpredictable outcome and non-compositional proof
76 checking); users who know what they are doing may use schematic goals
77 for Prolog-style synthesis of proven results;
79 * Pure: assumption method (an implicit finishing) now handles actual
82 * Pure: improved 'obtain' --- moved to Pure, insert "that" into
83 initial goal, declare "that" only as Pure intro (only for single
84 steps); the "that" rule assumption may now be involved in implicit
85 finishing, thus ".." becomes a feasible for trivial obtains;
87 * Pure: default proof step now includes 'intro_classes'; thus trivial
88 instance proofs may be performed by "..";
90 * Pure: ?thesis / ?this / "..." now work for pure meta-level
93 * Pure: more robust selection of calculational rules;
95 * Pure: the builtin notion of 'finished' goal now includes the ==-refl
96 rule (as well as the assumption rule);
98 * Pure: 'thm_deps' command visualizes dependencies of theorems and
99 lemmas, using the graph browser tool;
101 * Pure: predict failure of "show" in interactive mode;
103 * Pure: 'thms_containing' now takes actual terms as arguments;
105 * HOL: improved method 'induct' --- now handles non-atomic goals
106 (potential INCOMPATIBILITY); tuned error handling;
108 * HOL: cases and induct rules now provide explicit hints about the
109 number of facts to be consumed (0 for "type" and 1 for "set" rules);
110 any remaining facts are inserted into the goal verbatim;
112 * HOL: local contexts (aka cases) may now contain term bindings as
113 well; the 'cases' and 'induct' methods new provide a ?case binding for
114 the result to be shown in each case;
116 * HOL: added 'recdef_tc' command;
118 * isatool convert assists in eliminating legacy ML scripts;
123 * HOL/Library: a collection of generic theories to be used together
124 with main HOL; the theory loader path already includes this directory
125 by default; the following existing theories have been moved here:
126 HOL/Induct/Multiset, HOL/Induct/Acc (as Accessible_Part), HOL/While
127 (as While_Combinator), HOL/Lex/Prefix (as List_Prefix);
129 * HOL/Unix: "Some aspects of Unix file-system security", a typical
130 modelling and verification task performed in Isabelle/HOL +
131 Isabelle/Isar + Isabelle document preparation (by Markus Wenzel).
133 * HOL/Algebra: special summation operator SUM no longer exists, it has
134 been replaced by setsum; infix 'assoc' now has priority 50 (like
135 'dvd'); axiom 'one_not_zero' has been moved from axclass 'ring' to
136 'domain', this makes the theory consistent with mathematical
139 * HOL basics: added overloaded operations "inverse" and "divide"
140 (infix "/"), syntax for generic "abs" operation, generic summation
143 * HOL/typedef: simplified package, provide more useful rules (see also
146 * HOL/datatype: induction rule for arbitrarily branching datatypes is
147 now expressed as a proper nested rule (old-style tactic scripts may
148 require atomize_strip_tac to cope with non-atomic premises);
150 * HOL: renamed theory "Prod" to "Product_Type", renamed "split" rule
151 to "split_conv" (old name still available for compatibility);
153 * HOL: improved concrete syntax for strings (e.g. allows translation
154 rules with string literals);
156 * HOL-Hyperreal: a new target, extending HOL-Real with the hyperreals
157 and Fleuriot's mechanization of analysis;
159 * HOL/Real, HOL/Hyperreal: improved arithmetic simplification;
164 * CTT: x-symbol support for Pi, Sigma, -->, : (membership); note that
165 "lam" is displayed as TWO lambda-symbols
167 * CTT: theory Main now available, containing everything (that is, Bool
173 * Pure: the Simplifier has been implemented properly as a derived rule
174 outside of the actual kernel (at last!); the overall performance
175 penalty in practical applications is about 50%, while reliability of
176 the Isabelle inference kernel has been greatly improved;
178 * print modes "brackets" and "no_brackets" control output of nested =>
179 (types) and ==> (props); the default behaviour is "brackets";
181 * Provers: fast_tac (and friends) now handle actual object-logic rules
182 as assumptions as well;
184 * system: support Poly/ML 4.0;
186 * system: isatool install handles KDE version 1 or 2;
190 New in Isabelle99-1 (October 2000)
191 ----------------------------------
193 *** Overview of INCOMPATIBILITIES ***
195 * HOL: simplification of natural numbers is much changed; to partly
196 recover the old behaviour (e.g. to prevent n+n rewriting to #2*n)
197 issue the following ML commands:
199 Delsimprocs Nat_Numeral_Simprocs.cancel_numerals;
200 Delsimprocs [Nat_Numeral_Simprocs.combine_numerals];
202 * HOL: simplification no longer dives into case-expressions; this is
203 controlled by "t.weak_case_cong" for each datatype t;
205 * HOL: nat_less_induct renamed to less_induct;
207 * HOL: systematic renaming of the SOME (Eps) rules, may use isatool
208 fixsome to patch .thy and .ML sources automatically;
210 select_equality -> some_equality
211 select_eq_Ex -> some_eq_ex
212 selectI2EX -> someI2_ex
215 select1_equality -> some1_equality
216 Eps_sym_eq -> some_sym_eq_trivial
217 Eps_eq -> some_eq_trivial
219 * HOL: exhaust_tac on datatypes superceded by new generic case_tac;
221 * HOL: removed obsolete theorem binding expand_if (refer to split_if
224 * HOL: the recursion equations generated by 'recdef' are now called
225 f.simps instead of f.rules;
227 * HOL: qed_spec_mp now also handles bounded ALL as well;
229 * HOL: 0 is now overloaded, so the type constraint ":: nat" may
232 * HOL: the constant for "f``x" is now "image" rather than "op ``";
234 * HOL: the constant for "f-``x" is now "vimage" rather than "op -``";
236 * HOL: the disjoint sum is now "<+>" instead of "Plus"; the cartesian
237 product is now "<*>" instead of "Times"; the lexicographic product is
238 now "<*lex*>" instead of "**";
240 * HOL: theory Sexp is now in HOL/Induct examples (it used to be part
241 of main HOL, but was unused); better use HOL's datatype package;
243 * HOL: removed "symbols" syntax for constant "override" of theory Map;
244 the old syntax may be recovered as follows:
247 override :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)"
248 (infixl "\\<oplus>" 100)
250 * HOL/Real: "rabs" replaced by overloaded "abs" function;
252 * HOL/ML: even fewer consts are declared as global (see theories Ord,
253 Lfp, Gfp, WF); this only affects ML packages that refer to const names
256 * HOL and ZF: syntax for quotienting wrt an equivalence relation
257 changed from A/r to A//r;
259 * ZF: new treatment of arithmetic (nat & int) may break some old
262 * Isar: renamed some attributes (RS -> THEN, simplify -> simplified,
263 rulify -> rule_format, elimify -> elim_format, ...);
265 * Isar/Provers: intro/elim/dest attributes changed; renamed
266 intro/intro!/intro!! flags to intro!/intro/intro? (in most cases, one
267 should have to change intro!! to intro? only); replaced "delrule" by
270 * Isar/HOL: renamed "intrs" to "intros" in inductive definitions;
272 * Provers: strengthened force_tac by using new first_best_tac;
274 * LaTeX document preparation: several changes of isabelle.sty (see
278 *** Document preparation ***
280 * formal comments (text blocks etc.) in new-style theories may now
281 contain antiquotations of thm/prop/term/typ/text to be presented
282 according to latex print mode; concrete syntax is like this:
283 @{term[show_types] "f(x) = a + x"};
285 * isatool mkdir provides easy setup of Isabelle session directories,
286 including proper document sources;
288 * generated LaTeX sources are now deleted after successful run
289 (isatool document -c); may retain a copy somewhere else via -D option
292 * isatool usedir -D now lets isatool latex -o sty update the Isabelle
293 style files, achieving self-contained LaTeX sources and simplifying
296 * old-style theories now produce (crude) LaTeX output as well;
298 * browser info session directories are now self-contained (may be put
299 on WWW server seperately); improved graphs of nested sessions; removed
300 graph for 'all sessions';
302 * several improvements in isabelle style files; \isabellestyle{it}
303 produces fake math mode output; \isamarkupheader is now \section by
304 default; see lib/texinputs/isabelle.sty etc.;
309 * Isar/Pure: local results and corresponding term bindings are now
310 subject to Hindley-Milner polymorphism (similar to ML); this
311 accommodates incremental type-inference very nicely;
313 * Isar/Pure: new derived language element 'obtain' supports
314 generalized existence reasoning;
316 * Isar/Pure: new calculational elements 'moreover' and 'ultimately'
317 support accumulation of results, without applying any rules yet;
318 useful to collect intermediate results without explicit name
319 references, and for use with transitivity rules with more than 2
322 * Isar/Pure: scalable support for case-analysis type proofs: new
323 'case' language element refers to local contexts symbolically, as
324 produced by certain proof methods; internally, case names are attached
325 to theorems as "tags";
327 * Isar/Pure: theory command 'hide' removes declarations from
328 class/type/const name spaces;
330 * Isar/Pure: theory command 'defs' supports option "(overloaded)" to
331 indicate potential overloading;
333 * Isar/Pure: changed syntax of local blocks from {{ }} to { };
335 * Isar/Pure: syntax of sorts made 'inner', i.e. have to write
336 "{a,b,c}" instead of {a,b,c};
338 * Isar/Pure now provides its own version of intro/elim/dest
339 attributes; useful for building new logics, but beware of confusion
340 with the version in Provers/classical;
342 * Isar/Pure: the local context of (non-atomic) goals is provided via
343 case name 'antecedent';
345 * Isar/Pure: removed obsolete 'transfer' attribute (transfer of thms
346 to the current context is now done automatically);
348 * Isar/Pure: theory command 'method_setup' provides a simple interface
349 for definining proof methods in ML;
351 * Isar/Provers: intro/elim/dest attributes changed; renamed
352 intro/intro!/intro!! flags to intro!/intro/intro? (INCOMPATIBILITY, in
353 most cases, one should have to change intro!! to intro? only);
354 replaced "delrule" by "rule del";
356 * Isar/Provers: new 'hypsubst' method, plain 'subst' method and
357 'symmetric' attribute (the latter supercedes [RS sym]);
359 * Isar/Provers: splitter support (via 'split' attribute and 'simp'
360 method modifier); 'simp' method: 'only:' modifier removes loopers as
361 well (including splits);
363 * Isar/Provers: Simplifier and Classical methods now support all kind
364 of modifiers used in the past, including 'cong', 'iff', etc.
366 * Isar/Provers: added 'fastsimp' and 'clarsimp' methods (combination
367 of Simplifier and Classical reasoner);
369 * Isar/HOL: new proof method 'cases' and improved version of 'induct'
370 now support named cases; major packages (inductive, datatype, primrec,
371 recdef) support case names and properly name parameters;
373 * Isar/HOL: new transitivity rules for substitution in inequalities --
374 monotonicity conditions are extracted to be proven at end of
377 * Isar/HOL: removed 'case_split' thm binding, should use 'cases' proof
380 * Isar/HOL: removed old expand_if = split_if; theorems if_splits =
381 split_if split_if_asm; datatype package provides theorems foo.splits =
382 foo.split foo.split_asm for each datatype;
384 * Isar/HOL: tuned inductive package, rename "intrs" to "intros"
385 (potential INCOMPATIBILITY), emulation of mk_cases feature for proof
386 scripts: new 'inductive_cases' command and 'ind_cases' method; (Note:
387 use "(cases (simplified))" method in proper proof texts);
389 * Isar/HOL: added global 'arith_split' attribute for 'arith' method;
391 * Isar: names of theorems etc. may be natural numbers as well;
393 * Isar: 'pr' command: optional arguments for goals_limit and
394 ProofContext.prems_limit; no longer prints theory contexts, but only
397 * Isar: diagnostic commands 'pr', 'thm', 'prop', 'term', 'typ' admit
398 additional print modes to be specified; e.g. "pr(latex)" will print
399 proof state according to the Isabelle LaTeX style;
401 * Isar: improved support for emulating tactic scripts, including proof
402 methods 'rule_tac' etc., 'cut_tac', 'thin_tac', 'subgoal_tac',
403 'rename_tac', 'rotate_tac', 'tactic', and 'case_tac' / 'induct_tac'
406 * Isar: simplified (more robust) goal selection of proof methods: 1st
407 goal, all goals, or explicit goal specifier (tactic emulation); thus
408 'proof method scripts' have to be in depth-first order;
410 * Isar: tuned 'let' syntax: replaced 'as' keyword by 'and';
412 * Isar: removed 'help' command, which hasn't been too helpful anyway;
413 should instead use individual commands for printing items
414 (print_commands, print_methods etc.);
416 * Isar: added 'nothing' --- the empty list of theorems;
421 * HOL/MicroJava: formalization of a fragment of Java, together with a
422 corresponding virtual machine and a specification of its bytecode
423 verifier and a lightweight bytecode verifier, including proofs of
424 type-safety; by Gerwin Klein, Tobias Nipkow, David von Oheimb, and
425 Cornelia Pusch (see also the homepage of project Bali at
426 http://isabelle.in.tum.de/Bali/);
428 * HOL/Algebra: new theory of rings and univariate polynomials, by
431 * HOL/NumberTheory: fundamental Theorem of Arithmetic, Chinese
432 Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, by Thomas M
435 * HOL/Lattice: fundamental concepts of lattice theory and order
436 structures, including duals, properties of bounds versus algebraic
437 laws, lattice operations versus set-theoretic ones, the Knaster-Tarski
438 Theorem for complete lattices etc.; may also serve as a demonstration
439 for abstract algebraic reasoning using axiomatic type classes, and
440 mathematics-style proof in Isabelle/Isar; by Markus Wenzel;
442 * HOL/Prolog: a (bare-bones) implementation of Lambda-Prolog, by David
445 * HOL/IMPP: extension of IMP with local variables and mutually
446 recursive procedures, by David von Oheimb;
448 * HOL/Lambda: converted into new-style theory and document;
450 * HOL/ex/Multiquote: example of multiple nested quotations and
451 anti-quotations -- basically a generalized version of de-Bruijn
452 representation; very useful in avoiding lifting of operations;
454 * HOL/record: added general record equality rule to simpset; fixed
455 select-update simplification procedure to handle extended records as
456 well; admit "r" as field name;
458 * HOL: 0 is now overloaded over the new sort "zero", allowing its use with
459 other numeric types and also as the identity of groups, rings, etc.;
461 * HOL: new axclass plus_ac0 for addition with the AC-laws and 0 as identity.
462 Types nat and int belong to this axclass;
464 * HOL: greatly improved simplification involving numerals of type nat, int, real:
465 (i + #8 + j) = Suc k simplifies to #7 + (i + j) = k
466 i*j + k + j*#3*i simplifies to #4*(i*j) + k
467 two terms #m*u and #n*u are replaced by #(m+n)*u
468 (where #m, #n and u can implicitly be 1; this is simproc combine_numerals)
469 and the term/formula #m*u+x ~~ #n*u+y simplifies simplifies to #(m-n)+x ~~ y
470 or x ~~ #(n-m)+y, where ~~ is one of = < <= or - (simproc cancel_numerals);
472 * HOL: meson_tac is available (previously in ex/meson.ML); it is a
473 powerful prover for predicate logic but knows nothing of clasets; see
474 ex/mesontest.ML and ex/mesontest2.ML for example applications;
476 * HOL: new version of "case_tac" subsumes both boolean case split and
477 "exhaust_tac" on datatypes; INCOMPATIBILITY: exhaust_tac no longer
478 exists, may define val exhaust_tac = case_tac for ad-hoc portability;
480 * HOL: simplification no longer dives into case-expressions: only the
481 selector expression is simplified, but not the remaining arms; to
482 enable full simplification of case-expressions for datatype t, you may
483 remove t.weak_case_cong from the simpset, either globally (Delcongs
484 [thm"t.weak_case_cong"];) or locally (delcongs [...]).
486 * HOL/recdef: the recursion equations generated by 'recdef' for
487 function 'f' are now called f.simps instead of f.rules; if all
488 termination conditions are proved automatically, these simplification
489 rules are added to the simpset, as in primrec; rules may be named
490 individually as well, resulting in a separate list of theorems for
493 * HOL/While is a new theory that provides a while-combinator. It
494 permits the definition of tail-recursive functions without the
495 provision of a termination measure. The latter is necessary once the
496 invariant proof rule for while is applied.
498 * HOL: new (overloaded) notation for the set of elements below/above
499 some element: {..u}, {..u(}, {l..}, {)l..}. See theory SetInterval.
501 * HOL: theorems impI, allI, ballI bound as "strip";
503 * HOL: new tactic induct_thm_tac: thm -> string -> int -> tactic
504 induct_tac th "x1 ... xn" expects th to have a conclusion of the form
505 P v1 ... vn and abbreviates res_inst_tac [("v1","x1"),...,("vn","xn")] th;
507 * HOL/Real: "rabs" replaced by overloaded "abs" function;
509 * HOL: theory Sexp now in HOL/Induct examples (it used to be part of
510 main HOL, but was unused);
512 * HOL: fewer consts declared as global (e.g. have to refer to
513 "Lfp.lfp" instead of "lfp" internally; affects ML packages only);
515 * HOL: tuned AST representation of nested pairs, avoiding bogus output
516 in case of overlap with user translations (e.g. judgements over
517 tuples); (note that the underlying logical represenation is still
523 * ZF: simplification automatically cancels common terms in arithmetic
524 expressions over nat and int;
526 * ZF: new treatment of nat to minimize type-checking: all operators
527 coerce their operands to a natural number using the function natify,
528 making the algebraic laws unconditional;
530 * ZF: as above, for int: operators coerce their operands to an integer
531 using the function intify;
533 * ZF: the integer library now contains many of the usual laws for the
534 orderings, including $<=, and monotonicity laws for $+ and $*;
536 * ZF: new example ZF/ex/NatSum to demonstrate integer arithmetic
539 * FOL and ZF: AddIffs now available, giving theorems of the form P<->Q
540 to the simplifier and classical reasoner simultaneously;
545 * Provers: blast_tac now handles actual object-logic rules as
546 assumptions; note that auto_tac uses blast_tac internally as well;
548 * Provers: new functions rulify/rulify_no_asm: thm -> thm for turning
549 outer -->/All/Ball into ==>/!!; qed_spec_mp now uses rulify_no_asm;
551 * Provers: delrules now handles destruct rules as well (no longer need
554 * Provers: Blast_tac now warns of and ignores "weak elimination rules" e.g.
555 [| inj ?f; ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
556 use instead the strong form,
557 [| inj ?f; ~ ?W ==> ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
558 in HOL, FOL and ZF the function cla_make_elim will create such rules
561 * Provers: Simplifier.easy_setup provides a fast path to basic
562 Simplifier setup for new object-logics;
564 * Pure: AST translation rules no longer require constant head on LHS;
566 * Pure: improved name spaces: ambiguous output is qualified; support
569 * system: smart setup of canonical ML_HOME, ISABELLE_INTERFACE, and
570 XSYMBOL_HOME; no longer need to do manual configuration in most
573 * system: compression of ML heaps images may now be controlled via -c
574 option of isabelle and isatool usedir (currently only observed by
577 * system: isatool installfonts may handle X-Symbol fonts as well (very
578 useful for remote X11);
580 * system: provide TAGS file for Isabelle sources;
582 * ML: infix 'OF' is a version of 'MRS' with more appropriate argument
585 * ML: renamed flags Syntax.trace_norm_ast to Syntax.trace_ast; global
586 timing flag supersedes proof_timing and Toplevel.trace;
588 * ML: new combinators |>> and |>>> for incremental transformations
589 with secondary results (e.g. certain theory extensions):
591 * ML: PureThy.add_defs gets additional argument to indicate potential
592 overloading (usually false);
594 * ML: PureThy.add_thms/add_axioms/add_defs now return theorems as
599 New in Isabelle99 (October 1999)
600 --------------------------------
602 *** Overview of INCOMPATIBILITIES (see below for more details) ***
604 * HOL: The THEN and ELSE parts of conditional expressions (if P then x else y)
605 are no longer simplified. (This allows the simplifier to unfold recursive
606 functional programs.) To restore the old behaviour, declare
608 Delcongs [if_weak_cong];
610 * HOL: Removed the obsolete syntax "Compl A"; use -A for set
613 * HOL: the predicate "inj" is now defined by translation to "inj_on";
615 * HOL/datatype: mutual_induct_tac no longer exists --
616 use induct_tac "x_1 ... x_n" instead of mutual_induct_tac ["x_1", ..., "x_n"]
618 * HOL/typedef: fixed type inference for representing set; type
619 arguments now have to occur explicitly on the rhs as type constraints;
621 * ZF: The con_defs part of an inductive definition may no longer refer
622 to constants declared in the same theory;
624 * HOL, ZF: the function mk_cases, generated by the inductive
625 definition package, has lost an argument. To simplify its result, it
626 uses the default simpset instead of a supplied list of theorems.
628 * HOL/List: the constructors of type list are now Nil and Cons;
630 * Simplifier: the type of the infix ML functions
631 setSSolver addSSolver setSolver addSolver
632 is now simpset * solver -> simpset where `solver' is a new abstract type
633 for packaging solvers. A solver is created via
634 mk_solver: string -> (thm list -> int -> tactic) -> solver
635 where the string argument is only a comment.
640 * Provers/Arith/fast_lin_arith.ML contains a functor for creating a
641 decision procedure for linear arithmetic. Currently it is used for
642 types `nat', `int', and `real' in HOL (see below); it can, should and
643 will be instantiated for other types and logics as well.
645 * The simplifier now accepts rewrite rules with flexible heads, eg
646 hom ?f ==> ?f(?x+?y) = ?f ?x + ?f ?y
647 They are applied like any rule with a non-pattern lhs, i.e. by first-order
653 * New Isabelle/Isar subsystem provides an alternative to traditional
654 tactical theorem proving; together with the ProofGeneral/isar user
655 interface it offers an interactive environment for developing human
656 readable proof documents (Isar == Intelligible semi-automated
657 reasoning); for further information see isatool doc isar-ref,
658 src/HOL/Isar_examples and http://isabelle.in.tum.de/Isar/
660 * improved and simplified presentation of theories: better HTML markup
661 (including colors), graph views in several sizes; isatool usedir now
662 provides a proper interface for user theories (via -P option); actual
663 document preparation based on (PDF)LaTeX is available as well (for
664 new-style theories only); see isatool doc system for more information;
666 * native support for Proof General, both for classic Isabelle and
669 * ML function thm_deps visualizes dependencies of theorems and lemmas,
670 using the graph browser tool;
672 * Isabelle manuals now also available as PDF;
674 * theory loader rewritten from scratch (may not be fully
675 bug-compatible); old loadpath variable has been replaced by show_path,
676 add_path, del_path, reset_path functions; new operations such as
677 update_thy, touch_thy, remove_thy, use/update_thy_only (see also
680 * improved isatool install: option -k creates KDE application icon,
681 option -p DIR installs standalone binaries;
683 * added ML_PLATFORM setting (useful for cross-platform installations);
684 more robust handling of platform specific ML images for SML/NJ;
686 * the settings environment is now statically scoped, i.e. it is never
687 created again in sub-processes invoked from isabelle, isatool, or
690 * path element specification '~~' refers to '$ISABELLE_HOME';
692 * in locales, the "assumes" and "defines" parts may be omitted if
695 * new print_mode "xsymbols" for extended symbol support (e.g. genuine
698 * new print_mode "HTML";
700 * new flag show_tags controls display of tags of theorems (which are
701 basically just comments that may be attached by some tools);
703 * Isamode 2.6 requires patch to accomodate change of Isabelle font
704 mode and goal output format:
706 diff -r Isamode-2.6/elisp/isa-load.el Isamode/elisp/isa-load.el
708 < (list (isa-getenv "ISABELLE") "-msymbols" logic-name)
710 > (list (isa-getenv "ISABELLE") "-misabelle_font" "-msymbols" logic-name)
711 diff -r Isabelle-2.6/elisp/isa-proofstate.el Isamode/elisp/isa-proofstate.el
713 < (defconst proofstate-proofstart-regexp "^Level [0-9]+$"
715 > (defconst proofstate-proofstart-regexp "^Level [0-9]+"
717 * function bind_thms stores lists of theorems (cf. bind_thm);
719 * new shorthand tactics ftac, eatac, datac, fatac;
721 * qed (and friends) now accept "" as result name; in that case the
722 theorem is not stored, but proper checks and presentation of the
725 * theorem database now also indexes constants "Trueprop", "all",
726 "==>", "=="; thus thms_containing, findI etc. may retrieve more rules;
733 * There are now decision procedures for linear arithmetic over nat and
736 1. arith_tac copes with arbitrary formulae involving `=', `<', `<=',
737 `+', `-', `Suc', `min', `max' and numerical constants; other subterms
738 are treated as atomic; subformulae not involving type `nat' or `int'
739 are ignored; quantified subformulae are ignored unless they are
740 positive universal or negative existential. The tactic has to be
741 invoked by hand and can be a little bit slow. In particular, the
742 running time is exponential in the number of occurrences of `min' and
743 `max', and `-' on `nat'.
745 2. fast_arith_tac is a cut-down version of arith_tac: it only takes
746 (negated) (in)equalities among the premises and the conclusion into
747 account (i.e. no compound formulae) and does not know about `min' and
748 `max', and `-' on `nat'. It is fast and is used automatically by the
751 NB: At the moment, these decision procedures do not cope with mixed
752 nat/int formulae where the two parts interact, such as `m < n ==>
755 * HOL/Numeral provides a generic theory of numerals (encoded
756 efficiently as bit strings); setup for types nat/int/real is in place;
757 INCOMPATIBILITY: since numeral syntax is now polymorphic, rather than
758 int, existing theories and proof scripts may require a few additional
761 * integer division and remainder can now be performed on constant
764 * many properties of integer multiplication, division and remainder
767 * An interface to the Stanford Validity Checker (SVC) is available through the
768 tactic svc_tac. Propositional tautologies and theorems of linear arithmetic
769 are proved automatically. SVC must be installed separately, and its results
770 must be TAKEN ON TRUST (Isabelle does not check the proofs, but tags any
771 invocation of the underlying oracle). For SVC see
772 http://verify.stanford.edu/SVC
774 * IsaMakefile: the HOL-Real target now builds an actual image;
779 * HOL/Real/HahnBanach: the Hahn-Banach theorem for real vector spaces
780 (in Isabelle/Isar) -- by Gertrud Bauer;
782 * HOL/BCV: generic model of bytecode verification, i.e. data-flow
783 analysis for assembly languages with subtypes;
785 * HOL/TLA (Lamport's Temporal Logic of Actions): major reorganization
786 -- avoids syntactic ambiguities and treats state, transition, and
787 temporal levels more uniformly; introduces INCOMPATIBILITIES due to
788 changed syntax and (many) tactics;
790 * HOL/inductive: Now also handles more general introduction rules such
791 as "ALL y. (y, x) : r --> y : acc r ==> x : acc r"; monotonicity
792 theorems are now maintained within the theory (maintained via the
795 * HOL/datatype: Now also handles arbitrarily branching datatypes
796 (using function types) such as
798 datatype 'a tree = Atom 'a | Branch "nat => 'a tree"
800 * HOL/record: record_simproc (part of the default simpset) takes care
801 of selectors applied to updated records; record_split_tac is no longer
802 part of the default claset; update_defs may now be removed from the
803 simpset in many cases; COMPATIBILITY: old behavior achieved by
805 claset_ref () := claset() addSWrapper record_split_wrapper;
806 Delsimprocs [record_simproc]
808 * HOL/typedef: fixed type inference for representing set; type
809 arguments now have to occur explicitly on the rhs as type constraints;
811 * HOL/recdef (TFL): 'congs' syntax now expects comma separated list of theorem
812 names rather than an ML expression;
814 * HOL/defer_recdef (TFL): like recdef but the well-founded relation can be
815 supplied later. Program schemes can be defined, such as
816 "While B C s = (if B s then While B C (C s) else s)"
817 where the well-founded relation can be chosen after B and C have been given.
819 * HOL/List: the constructors of type list are now Nil and Cons;
820 INCOMPATIBILITY: while [] and infix # syntax is still there, of
821 course, ML tools referring to List.list.op # etc. have to be adapted;
823 * HOL_quantifiers flag superseded by "HOL" print mode, which is
824 disabled by default; run isabelle with option -m HOL to get back to
825 the original Gordon/HOL-style output;
827 * HOL/Ord.thy: new bounded quantifier syntax (input only): ALL x<y. P,
828 ALL x<=y. P, EX x<y. P, EX x<=y. P;
830 * HOL basic syntax simplified (more orthogonal): all variants of
831 All/Ex now support plain / symbolic / HOL notation; plain syntax for
832 Eps operator is provided as well: "SOME x. P[x]";
834 * HOL/Sum.thy: sum_case has been moved to HOL/Datatype;
836 * HOL/Univ.thy: infix syntax <*>, <+>, <**>, <+> eliminated and made
837 thus available for user theories;
839 * HOLCF/IOA/Sequents: renamed 'Cons' to 'Consq' to avoid clash with
840 HOL/List; hardly an INCOMPATIBILITY since '>>' syntax is used all the
843 * HOL: new tactic smp_tac: int -> int -> tactic, which applies spec
844 several times and then mp;
849 * the notation <<...>> is now available as a notation for sequences of
852 * the simplifier is now installed
854 * the axiom system has been generalized (thanks to Soren Heilmann)
856 * the classical reasoner now has a default rule database
861 * new primrec section allows primitive recursive functions to be given
862 directly (as in HOL) over datatypes and the natural numbers;
864 * new tactics induct_tac and exhaust_tac for induction (or case
865 analysis) over datatypes and the natural numbers;
867 * the datatype declaration of type T now defines the recursor T_rec;
869 * simplification automatically does freeness reasoning for datatype
872 * automatic type-inference, with AddTCs command to insert new
875 * datatype introduction rules are now added as Safe Introduction rules
878 * the syntax "if P then x else y" is now available in addition to
882 *** Internal programming interfaces ***
884 * tuned simplifier trace output; new flag debug_simp;
886 * structures Vartab / Termtab (instances of TableFun) offer efficient
887 tables indexed by indexname_ord / term_ord (compatible with aconv);
889 * AxClass.axclass_tac lost the theory argument;
891 * tuned current_goals_markers semantics: begin / end goal avoids
892 printing empty lines;
894 * removed prs and prs_fn hook, which was broken because it did not
895 include \n in its semantics, forcing writeln to add one
896 uncoditionally; replaced prs_fn by writeln_fn; consider std_output:
897 string -> unit if you really want to output text without newline;
899 * Symbol.output subject to print mode; INCOMPATIBILITY: defaults to
900 plain output, interface builders may have to enable 'isabelle_font'
901 mode to get Isabelle font glyphs as before;
903 * refined token_translation interface; INCOMPATIBILITY: output length
904 now of type real instead of int;
906 * theory loader actions may be traced via new ThyInfo.add_hook
907 interface (see src/Pure/Thy/thy_info.ML); example application: keep
908 your own database of information attached to *whole* theories -- as
909 opposed to intra-theory data slots offered via TheoryDataFun;
911 * proper handling of dangling sort hypotheses (at last!);
912 Thm.strip_shyps and Drule.strip_shyps_warning take care of removing
913 extra sort hypotheses that can be witnessed from the type signature;
914 the force_strip_shyps flag is gone, any remaining shyps are simply
915 left in the theorem (with a warning issued by strip_shyps_warning);
919 New in Isabelle98-1 (October 1998)
920 ----------------------------------
922 *** Overview of INCOMPATIBILITIES (see below for more details) ***
924 * several changes of automated proof tools;
926 * HOL: major changes to the inductive and datatype packages, including
927 some minor incompatibilities of theory syntax;
929 * HOL: renamed r^-1 to 'converse' from 'inverse'; 'inj_onto' is now
932 * HOL: removed duplicate thms in Arith:
933 less_imp_add_less should be replaced by trans_less_add1
934 le_imp_add_le should be replaced by trans_le_add1
936 * HOL: unary minus is now overloaded (new type constraints may be
939 * HOL and ZF: unary minus for integers is now #- instead of #~. In
940 ZF, expressions such as n#-1 must be changed to n#- 1, since #-1 is
941 now taken as an integer constant.
943 * Pure: ML function 'theory_of' renamed to 'theory';
949 1. Asm_full_simp_tac is now more aggressive.
950 1. It will sometimes reorient premises if that increases their power to
952 2. It does no longer proceed strictly from left to right but may also
953 rotate premises to achieve further simplification.
954 For compatibility reasons there is now Asm_lr_simp_tac which is like the
955 old Asm_full_simp_tac in that it does not rotate premises.
956 2. The simplifier now knows a little bit about nat-arithmetic.
958 * Classical reasoner: wrapper mechanism for the classical reasoner now
959 allows for selected deletion of wrappers, by introduction of names for
960 wrapper functionals. This implies that addbefore, addSbefore,
961 addaltern, and addSaltern now take a pair (name, tactic) as argument,
962 and that adding two tactics with the same name overwrites the first
963 one (emitting a warning).
964 type wrapper = (int -> tactic) -> (int -> tactic)
965 setWrapper, setSWrapper, compWrapper and compSWrapper are replaced by
966 addWrapper, addSWrapper: claset * (string * wrapper) -> claset
967 delWrapper, delSWrapper: claset * string -> claset
968 getWrapper is renamed to appWrappers, getSWrapper to appSWrappers;
970 * Classical reasoner: addbefore/addSbefore now have APPEND/ORELSE
971 semantics; addbefore now affects only the unsafe part of step_tac
972 etc.; this affects addss/auto_tac/force_tac, so EXISTING PROOFS MAY
973 FAIL, but proofs should be fixable easily, e.g. by replacing Auto_tac
976 * Classical reasoner: setwrapper to setWrapper and compwrapper to
977 compWrapper; added safe wrapper (and access functions for it);
979 * HOL/split_all_tac is now much faster and fails if there is nothing
980 to split. Some EXISTING PROOFS MAY REQUIRE ADAPTION because the order
981 and the names of the automatically generated variables have changed.
982 split_all_tac has moved within claset() from unsafe wrappers to safe
983 wrappers, which means that !!-bound variables are split much more
984 aggressively, and safe_tac and clarify_tac now split such variables.
985 If this splitting is not appropriate, use delSWrapper "split_all_tac".
986 Note: the same holds for record_split_tac, which does the job of
987 split_all_tac for record fields.
989 * HOL/Simplifier: Rewrite rules for case distinctions can now be added
990 permanently to the default simpset using Addsplits just like
991 Addsimps. They can be removed via Delsplits just like
992 Delsimps. Lower-case versions are also available.
994 * HOL/Simplifier: The rule split_if is now part of the default
995 simpset. This means that the simplifier will eliminate all occurrences
996 of if-then-else in the conclusion of a goal. To prevent this, you can
997 either remove split_if completely from the default simpset by
998 `Delsplits [split_if]' or remove it in a specific call of the
999 simplifier using `... delsplits [split_if]'. You can also add/delete
1000 other case splitting rules to/from the default simpset: every datatype
1001 generates suitable rules `split_t_case' and `split_t_case_asm' (where
1002 t is the name of the datatype).
1004 * Classical reasoner / Simplifier combination: new force_tac (and
1005 derivatives Force_tac, force) combines rewriting and classical
1006 reasoning (and whatever other tools) similarly to auto_tac, but is
1007 aimed to solve the given subgoal completely.
1012 * new top-level commands `Goal' and `Goalw' that improve upon `goal'
1013 and `goalw': the theory is no longer needed as an explicit argument -
1014 the current theory context is used; assumptions are no longer returned
1015 at the ML-level unless one of them starts with ==> or !!; it is
1016 recommended to convert to these new commands using isatool fixgoal
1017 (backup your sources first!);
1019 * new top-level commands 'thm' and 'thms' for retrieving theorems from
1020 the current theory context, and 'theory' to lookup stored theories;
1022 * new theory section 'locale' for declaring constants, assumptions and
1023 definitions that have local scope;
1025 * new theory section 'nonterminals' for purely syntactic types;
1027 * new theory section 'setup' for generic ML setup functions
1028 (e.g. package initialization);
1030 * the distribution now includes Isabelle icons: see
1031 lib/logo/isabelle-{small,tiny}.xpm;
1033 * isatool install - install binaries with absolute references to
1036 * isatool logo -- create instances of the Isabelle logo (as EPS);
1038 * print mode 'emacs' reserved for Isamode;
1040 * support multiple print (ast) translations per constant name;
1042 * theorems involving oracles are now printed with a suffixed [!];
1047 * there is now a tutorial on Isabelle/HOL (do 'isatool doc tutorial');
1049 * HOL/inductive package reorganized and improved: now supports mutual
1055 oddI "n : EVEN ==> Suc n : ODD"
1056 evenI "n : ODD ==> Suc n : EVEN"
1058 new theorem list "elims" contains an elimination rule for each of the
1059 recursive sets; inductive definitions now handle disjunctive premises
1060 correctly (also ZF);
1062 INCOMPATIBILITIES: requires Inductive as an ancestor; component
1063 "mutual_induct" no longer exists - the induction rule is always
1064 contained in "induct";
1067 * HOL/datatype package re-implemented and greatly improved: now
1068 supports mutually recursive datatypes such as
1071 'a aexp = IF_THEN_ELSE ('a bexp) ('a aexp) ('a aexp)
1072 | SUM ('a aexp) ('a aexp)
1073 | DIFF ('a aexp) ('a aexp)
1076 'a bexp = LESS ('a aexp) ('a aexp)
1077 | AND ('a bexp) ('a bexp)
1078 | OR ('a bexp) ('a bexp)
1080 as well as indirectly recursive datatypes such as
1083 ('a, 'b) term = Var 'a
1084 | App 'b ((('a, 'b) term) list)
1086 The new tactic mutual_induct_tac [<var_1>, ..., <var_n>] i performs
1087 induction on mutually / indirectly recursive datatypes.
1089 Primrec equations are now stored in theory and can be accessed via
1090 <function_name>.simps.
1094 - Theories using datatypes must now have theory Datatype as an
1096 - The specific <typename>.induct_tac no longer exists - use the
1097 generic induct_tac instead.
1098 - natE has been renamed to nat.exhaust - use exhaust_tac
1099 instead of res_inst_tac ... natE. Note that the variable
1100 names in nat.exhaust differ from the names in natE, this
1101 may cause some "fragile" proofs to fail.
1102 - The theorems split_<typename>_case and split_<typename>_case_asm
1103 have been renamed to <typename>.split and <typename>.split_asm.
1104 - Since default sorts of type variables are now handled correctly,
1105 some datatype definitions may have to be annotated with explicit
1107 - Primrec definitions no longer require function name and type
1108 of recursive argument.
1110 Consider using isatool fixdatatype to adapt your theories and proof
1111 scripts to the new package (backup your sources first!).
1114 * HOL/record package: considerably improved implementation; now
1115 includes concrete syntax for record types, terms, updates; theorems
1116 for surjective pairing and splitting !!-bound record variables; proof
1117 support is as follows:
1119 1) standard conversions (selectors or updates applied to record
1120 constructor terms) are part of the standard simpset;
1122 2) inject equations of the form ((x, y) = (x', y')) == x=x' & y=y' are
1123 made part of standard simpset and claset via addIffs;
1125 3) a tactic for record field splitting (record_split_tac) is part of
1126 the standard claset (addSWrapper);
1128 To get a better idea about these rules you may retrieve them via
1129 something like 'thms "foo.simps"' or 'thms "foo.iffs"', where "foo" is
1130 the name of your record type.
1132 The split tactic 3) conceptually simplifies by the following rule:
1134 "(!!x. PROP ?P x) == (!!a b. PROP ?P (a, b))"
1136 Thus any record variable that is bound by meta-all will automatically
1137 blow up into some record constructor term, consequently the
1138 simplifications of 1), 2) apply. Thus force_tac, auto_tac etc. shall
1139 solve record problems automatically.
1142 * reorganized the main HOL image: HOL/Integ and String loaded by
1143 default; theory Main includes everything;
1145 * automatic simplification of integer sums and comparisons, using cancellation;
1147 * added option_map_eq_Some and not_Some_eq to the default simpset and claset;
1149 * added disj_not1 = "(~P | Q) = (P --> Q)" to the default simpset;
1151 * many new identities for unions, intersections, set difference, etc.;
1153 * expand_if, expand_split, expand_sum_case and expand_nat_case are now
1154 called split_if, split_split, split_sum_case and split_nat_case (to go
1155 with add/delsplits);
1157 * HOL/Prod introduces simplification procedure unit_eq_proc rewriting
1158 (?x::unit) = (); this is made part of the default simpset, which COULD
1159 MAKE EXISTING PROOFS FAIL under rare circumstances (consider
1160 'Delsimprocs [unit_eq_proc];' as last resort); also note that
1161 unit_abs_eta_conv is added in order to counter the effect of
1162 unit_eq_proc on (%u::unit. f u), replacing it by f rather than by
1165 * HOL/Fun INCOMPATIBILITY: `inj_onto' is now called `inj_on' (which
1168 * HOL/Set INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
1169 It and 'sym RS equals0D' are now in the default claset, giving automatic
1170 disjointness reasoning but breaking a few old proofs.
1172 * HOL/Relation INCOMPATIBILITY: renamed the relational operator r^-1
1173 to 'converse' from 'inverse' (for compatibility with ZF and some
1176 * HOL/recdef can now declare non-recursive functions, with {} supplied as
1177 the well-founded relation;
1179 * HOL/Set INCOMPATIBILITY: the complement of set A is now written -A instead of
1180 Compl A. The "Compl" syntax remains available as input syntax for this
1183 * HOL/Update: new theory of function updates:
1184 f(a:=b) == %x. if x=a then b else f x
1185 may also be iterated as in f(a:=b,c:=d,...);
1187 * HOL/Vimage: new theory for inverse image of a function, syntax f-``B;
1190 - new function list_update written xs[i:=v] that updates the i-th
1191 list position. May also be iterated as in xs[i:=a,j:=b,...].
1192 - new function `upt' written [i..j(] which generates the list
1193 [i,i+1,...,j-1], i.e. the upper bound is excluded. To include the upper
1194 bound write [i..j], which is a shorthand for [i..j+1(].
1195 - new lexicographic orderings and corresponding wellfoundedness theorems.
1198 - removed 'pred' (predecessor) function;
1199 - generalized some theorems about n-1;
1200 - many new laws about "div" and "mod";
1201 - new laws about greatest common divisors (see theory ex/Primes);
1203 * HOL/Relation: renamed the relational operator r^-1 "converse"
1204 instead of "inverse";
1206 * HOL/Induct/Multiset: a theory of multisets, including the wellfoundedness
1207 of the multiset ordering;
1209 * directory HOL/Real: a construction of the reals using Dedekind cuts
1210 (not included by default);
1212 * directory HOL/UNITY: Chandy and Misra's UNITY formalism;
1214 * directory HOL/Hoare: a new version of Hoare logic which permits many-sorted
1215 programs, i.e. different program variables may have different types.
1217 * calling (stac rew i) now fails if "rew" has no effect on the goal
1218 [previously, this check worked only if the rewrite rule was unconditional]
1219 Now rew can involve either definitions or equalities (either == or =).
1224 * theory Main includes everything; INCOMPATIBILITY: theory ZF.thy contains
1225 only the theorems proved on ZF.ML;
1227 * ZF INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
1228 It and 'sym RS equals0D' are now in the default claset, giving automatic
1229 disjointness reasoning but breaking a few old proofs.
1231 * ZF/Update: new theory of function updates
1232 with default rewrite rule f(x:=y) ` z = if(z=x, y, f`z)
1233 may also be iterated as in f(a:=b,c:=d,...);
1235 * in let x=t in u(x), neither t nor u(x) has to be an FOL term.
1237 * calling (stac rew i) now fails if "rew" has no effect on the goal
1238 [previously, this check worked only if the rewrite rule was unconditional]
1239 Now rew can involve either definitions or equalities (either == or =).
1241 * case_tac provided for compatibility with HOL
1242 (like the old excluded_middle_tac, but with subgoals swapped)
1245 *** Internal programming interfaces ***
1247 * Pure: several new basic modules made available for general use, see
1248 also src/Pure/README;
1250 * improved the theory data mechanism to support encapsulation (data
1251 kind name replaced by private Object.kind, acting as authorization
1252 key); new type-safe user interface via functor TheoryDataFun; generic
1253 print_data function becomes basically useless;
1255 * removed global_names compatibility flag -- all theory declarations
1256 are qualified by default;
1258 * module Pure/Syntax now offers quote / antiquote translation
1259 functions (useful for Hoare logic etc. with implicit dependencies);
1260 see HOL/ex/Antiquote for an example use;
1262 * Simplifier now offers conversions (asm_)(full_)rewrite: simpset ->
1265 * new tactical CHANGED_GOAL for checking that a tactic modifies a
1268 * Display.print_goals function moved to Locale.print_goals;
1270 * standard print function for goals supports current_goals_markers
1271 variable for marking begin of proof, end of proof, start of goal; the
1272 default is ("", "", ""); setting current_goals_markers := ("<proof>",
1273 "</proof>", "<goal>") causes SGML like tagged proof state printing,
1278 New in Isabelle98 (January 1998)
1279 --------------------------------
1281 *** Overview of INCOMPATIBILITIES (see below for more details) ***
1283 * changed lexical syntax of terms / types: dots made part of long
1284 identifiers, e.g. "%x.x" no longer possible, should be "%x. x";
1286 * simpset (and claset) reference variable replaced by functions
1287 simpset / simpset_ref;
1289 * no longer supports theory aliases (via merge) and non-trivial
1290 implicit merge of thms' signatures;
1292 * most internal names of constants changed due to qualified names;
1294 * changed Pure/Sequence interface (see Pure/seq.ML);
1297 *** General Changes ***
1299 * hierachically structured name spaces (for consts, types, axms, thms
1300 etc.); new lexical class 'longid' (e.g. Foo.bar.x) may render much of
1301 old input syntactically incorrect (e.g. "%x.x"); COMPATIBILITY:
1302 isatool fixdots ensures space after dots (e.g. "%x. x"); set
1303 long_names for fully qualified output names; NOTE: ML programs
1304 (special tactics, packages etc.) referring to internal names may have
1305 to be adapted to cope with fully qualified names; in case of severe
1306 backward campatibility problems try setting 'global_names' at compile
1307 time to have enrything declared within a flat name space; one may also
1308 fine tune name declarations in theories via the 'global' and 'local'
1311 * reimplemented the implicit simpset and claset using the new anytype
1312 data filed in signatures; references simpset:simpset ref etc. are
1313 replaced by functions simpset:unit->simpset and
1314 simpset_ref:unit->simpset ref; COMPATIBILITY: use isatool fixclasimp
1315 to patch your ML files accordingly;
1317 * HTML output now includes theory graph data for display with Java
1318 applet or isatool browser; data generated automatically via isatool
1319 usedir (see -i option, ISABELLE_USEDIR_OPTIONS);
1321 * defs may now be conditional; improved rewrite_goals_tac to handle
1322 conditional equations;
1324 * defs now admits additional type arguments, using TYPE('a) syntax;
1326 * theory aliases via merge (e.g. M=A+B+C) no longer supported, always
1327 creates a new theory node; implicit merge of thms' signatures is
1328 restricted to 'trivial' ones; COMPATIBILITY: one may have to use
1329 transfer:theory->thm->thm in (rare) cases;
1331 * improved handling of draft signatures / theories; draft thms (and
1332 ctyps, cterms) are automatically promoted to real ones;
1334 * slightly changed interfaces for oracles: admit many per theory, named
1335 (e.g. oracle foo = mlfun), additional name argument for invoke_oracle;
1337 * print_goals: optional output of const types (set show_consts and
1340 * improved output of warnings (###) and errors (***);
1342 * subgoal_tac displays a warning if the new subgoal has type variables;
1344 * removed old README and Makefiles;
1346 * replaced print_goals_ref hook by print_current_goals_fn and result_error_fn;
1348 * removed obsolete init_pps and init_database;
1350 * deleted the obsolete tactical STATE, which was declared by
1351 fun STATE tacfun st = tacfun st st;
1353 * cd and use now support path variables, e.g. $ISABELLE_HOME, or ~
1354 (which abbreviates $HOME);
1356 * changed Pure/Sequence interface (see Pure/seq.ML); COMPATIBILITY:
1357 use isatool fixseq to adapt your ML programs (this works for fully
1358 qualified references to the Sequence structure only!);
1360 * use_thy no longer requires writable current directory; it always
1361 reloads .ML *and* .thy file, if either one is out of date;
1364 *** Classical Reasoner ***
1366 * Clarify_tac, clarify_tac, clarify_step_tac, Clarify_step_tac: new
1367 tactics that use classical reasoning to simplify a subgoal without
1368 splitting it into several subgoals;
1370 * Safe_tac: like safe_tac but uses the default claset;
1375 * added simplification meta rules:
1376 (asm_)(full_)simplify: simpset -> thm -> thm;
1378 * simplifier.ML no longer part of Pure -- has to be loaded by object
1381 * added prems argument to simplification procedures;
1383 * HOL, FOL, ZF: added infix function `addsplits':
1384 instead of `<simpset> setloop (split_tac <thms>)'
1385 you can simply write `<simpset> addsplits <thms>'
1390 * TYPE('a) syntax for type reflection terms;
1392 * no longer handles consts with name "" -- declare as 'syntax' instead;
1394 * pretty printer: changed order of mixfix annotation preference (again!);
1396 * Pure: fixed idt/idts vs. pttrn/pttrns syntactic categories;
1401 * HOL: there is a new splitter `split_asm_tac' that can be used e.g.
1402 with `addloop' of the simplifier to faciliate case splitting in premises.
1404 * HOL/TLA: Stephan Merz's formalization of Lamport's Temporal Logic of Actions;
1406 * HOL/Auth: new protocol proofs including some for the Internet
1409 * HOL/Map: new theory of `maps' a la VDM;
1411 * HOL/simplifier: simplification procedures nat_cancel_sums for
1412 cancelling out common nat summands from =, <, <= (in)equalities, or
1413 differences; simplification procedures nat_cancel_factor for
1414 cancelling common factor from =, <, <= (in)equalities over natural
1415 sums; nat_cancel contains both kinds of procedures, it is installed by
1416 default in Arith.thy -- this COULD MAKE EXISTING PROOFS FAIL;
1418 * HOL/simplifier: terms of the form
1419 `? x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x)' (or t=x)
1421 `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t)',
1422 and those of the form
1423 `! x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x) --> R(x)' (or t=x)
1425 `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t) --> R(t)',
1428 Each datatype `t' now comes with a theorem `split_t_case' of the form
1430 P(t_case f1 ... fn x) =
1431 ( (!y1 ... ym1. x = C1 y1 ... ym1 --> P(f1 y1 ... ym1)) &
1433 (!y1 ... ymn. x = Cn y1 ... ymn --> P(f1 y1 ... ymn))
1436 and a theorem `split_t_case_asm' of the form
1438 P(t_case f1 ... fn x) =
1439 ~( (? y1 ... ym1. x = C1 y1 ... ym1 & ~P(f1 y1 ... ym1)) |
1441 (? y1 ... ymn. x = Cn y1 ... ymn & ~P(f1 y1 ... ymn))
1443 which can be added to a simpset via `addsplits'. The existing theorems
1444 expand_list_case and expand_option_case have been renamed to
1445 split_list_case and split_option_case.
1448 - `pred n' is automatically converted to `n-1'.
1449 Users are strongly encouraged not to use `pred' any longer,
1450 because it will disappear altogether at some point.
1451 - Users are strongly encouraged to write "0 < n" rather than
1452 "n ~= 0". Theorems and proof tools have been modified towards this
1456 the function "set_of_list" has been renamed "set" (and its theorems too);
1457 the function "nth" now takes its arguments in the reverse order and
1458 has acquired the infix notation "!" as in "xs!n".
1460 * HOL/Set: UNIV is now a constant and is no longer translated to Compl{};
1462 * HOL/Set: The operator (UN x.B x) now abbreviates (UN x:UNIV. B x) and its
1463 specialist theorems (like UN1_I) are gone. Similarly for (INT x.B x);
1465 * HOL/record: extensible records with schematic structural subtyping
1466 (single inheritance); EXPERIMENTAL version demonstrating the encoding,
1467 still lacks various theorems and concrete record syntax;
1472 * removed "axioms" and "generated by" sections;
1474 * replaced "ops" section by extended "consts" section, which is capable of
1475 handling the continuous function space "->" directly;
1478 . proves theorems immediately and stores them in the theory,
1479 . creates hierachical name space,
1480 . now uses normal mixfix annotations (instead of cinfix...),
1481 . minor changes to some names and values (for consistency),
1482 . e.g. cases -> casedist, dists_eq -> dist_eqs, [take_lemma] -> take_lemmas,
1483 . separator between mutual domain defs: changed "," to "and",
1484 . improved handling of sort constraints; now they have to
1485 appear on the left-hand side of the equations only;
1487 * fixed LAM <x,y,zs>.b syntax;
1489 * added extended adm_tac to simplifier in HOLCF -- can now discharge
1490 adm (%x. P (t x)), where P is chainfinite and t continuous;
1495 * FOL: there is a new splitter `split_asm_tac' that can be used e.g.
1496 with `addloop' of the simplifier to faciliate case splitting in premises.
1498 * qed_spec_mp, qed_goal_spec_mp, qed_goalw_spec_mp are available, as
1499 in HOL, they strip ALL and --> from proved theorems;
1503 New in Isabelle94-8 (May 1997)
1504 ------------------------------
1506 *** General Changes ***
1508 * new utilities to build / run / maintain Isabelle etc. (in parts
1509 still somewhat experimental); old Makefiles etc. still functional;
1511 * new 'Isabelle System Manual';
1513 * INSTALL text, together with ./configure and ./build scripts;
1515 * reimplemented type inference for greater efficiency, better error
1516 messages and clean internal interface;
1518 * prlim command for dealing with lots of subgoals (an easier way of
1519 setting goals_limit);
1524 * supports alternative (named) syntax tables (parser and pretty
1525 printer); internal interface is provided by add_modesyntax(_i);
1527 * Pure, FOL, ZF, HOL, HOLCF now support symbolic input and output; to
1528 be used in conjunction with the Isabelle symbol font; uses the
1529 "symbols" syntax table;
1531 * added token_translation interface (may translate name tokens in
1532 arbitrary ways, dependent on their type (free, bound, tfree, ...) and
1533 the current print_mode); IMPORTANT: user print translation functions
1534 are responsible for marking newly introduced bounds
1535 (Syntax.mark_boundT);
1537 * token translations for modes "xterm" and "xterm_color" that display
1538 names in bold, underline etc. or colors (which requires a color
1541 * infixes may now be declared with names independent of their syntax;
1543 * added typed_print_translation (like print_translation, but may
1544 access type of constant);
1547 *** Classical Reasoner ***
1549 Blast_tac: a new tactic! It is often more powerful than fast_tac, but has
1550 some limitations. Blast_tac...
1551 + ignores addss, addbefore, addafter; this restriction is intrinsic
1552 + ignores elimination rules that don't have the correct format
1553 (the conclusion MUST be a formula variable)
1554 + ignores types, which can make HOL proofs fail
1555 + rules must not require higher-order unification, e.g. apply_type in ZF
1556 [message "Function Var's argument not a bound variable" relates to this]
1557 + its proof strategy is more general but can actually be slower
1559 * substitution with equality assumptions no longer permutes other
1562 * minor changes in semantics of addafter (now called addaltern); renamed
1563 setwrapper to setWrapper and compwrapper to compWrapper; added safe wrapper
1564 (and access functions for it);
1566 * improved combination of classical reasoner and simplifier:
1567 + functions for handling clasimpsets
1568 + improvement of addss: now the simplifier is called _after_ the
1570 + safe variant of addss called addSss: uses safe simplifications
1571 _during_ the safe steps. It is more complete as it allows multiple
1572 instantiations of unknowns (e.g. with slow_tac).
1576 * added interface for simplification procedures (functions that
1577 produce *proven* rewrite rules on the fly, depending on current
1580 * ordering on terms as parameter (used for ordered rewriting);
1582 * new functions delcongs, deleqcongs, and Delcongs. richer rep_ss;
1584 * the solver is now split into a safe and an unsafe part.
1585 This should be invisible for the normal user, except that the
1586 functions setsolver and addsolver have been renamed to setSolver and
1587 addSolver; added safe_asm_full_simp_tac;
1592 * a generic induction tactic `induct_tac' which works for all datatypes and
1593 also for type `nat';
1595 * a generic case distinction tactic `exhaust_tac' which works for all
1596 datatypes and also for type `nat';
1598 * each datatype comes with a function `size';
1600 * patterns in case expressions allow tuple patterns as arguments to
1601 constructors, for example `case x of [] => ... | (x,y,z)#ps => ...';
1603 * primrec now also works with type nat;
1605 * recdef: a new declaration form, allows general recursive functions to be
1606 defined in theory files. See HOL/ex/Fib, HOL/ex/Primes, HOL/Subst/Unify.
1608 * the constant for negation has been renamed from "not" to "Not" to
1609 harmonize with FOL, ZF, LK, etc.;
1611 * HOL/ex/LFilter theory of a corecursive "filter" functional for
1614 * HOL/Modelcheck demonstrates invocation of model checker oracle;
1616 * HOL/ex/Ring.thy declares cring_simp, which solves equational
1617 problems in commutative rings, using axiomatic type classes for + and *;
1619 * more examples in HOL/MiniML and HOL/Auth;
1621 * more default rewrite rules for quantifiers, union/intersection;
1623 * a new constant `arbitrary == @x.False';
1625 * HOLCF/IOA replaces old HOL/IOA;
1627 * HOLCF changes: derived all rules and arities
1628 + axiomatic type classes instead of classes
1629 + typedef instead of faking type definitions
1630 + eliminated the internal constants less_fun, less_cfun, UU_fun, UU_cfun etc.
1631 + new axclasses cpo, chfin, flat with flat < chfin < pcpo < cpo < po
1632 + eliminated the types void, one, tr
1633 + use unit lift and bool lift (with translations) instead of one and tr
1634 + eliminated blift from Lift3.thy (use Def instead of blift)
1635 all eliminated rules are derived as theorems --> no visible changes ;
1640 * ZF now has Fast_tac, Simp_tac and Auto_tac. Union_iff is a now a default
1641 rewrite rule; this may affect some proofs. eq_cs is gone but can be put back
1642 as ZF_cs addSIs [equalityI];
1646 New in Isabelle94-7 (November 96)
1647 ---------------------------------
1649 * allowing negative levels (as offsets) in prlev and choplev;
1651 * super-linear speedup for large simplifications;
1653 * FOL, ZF and HOL now use miniscoping: rewriting pushes
1654 quantifications in as far as possible (COULD MAKE EXISTING PROOFS
1655 FAIL); can suppress it using the command Delsimps (ex_simps @
1656 all_simps); De Morgan laws are also now included, by default;
1658 * improved printing of ==> : ~:
1660 * new object-logic "Sequents" adds linear logic, while replacing LK
1661 and Modal (thanks to Sara Kalvala);
1663 * HOL/Auth: correctness proofs for authentication protocols;
1665 * HOL: new auto_tac combines rewriting and classical reasoning (many
1666 examples on HOL/Auth);
1668 * HOL: new command AddIffs for declaring theorems of the form P=Q to
1669 the rewriter and classical reasoner simultaneously;
1671 * function uresult no longer returns theorems in "standard" format;
1672 regain previous version by: val uresult = standard o uresult;
1679 * oracles -- these establish an interface between Isabelle and trusted
1680 external reasoners, which may deliver results as theorems;
1682 * proof objects (in particular record all uses of oracles);
1684 * Simp_tac, Fast_tac, etc. that refer to implicit simpset / claset;
1686 * "constdefs" section in theory files;
1688 * "primrec" section (HOL) no longer requires names;
1690 * internal type "tactic" now simply "thm -> thm Sequence.seq";
1697 * reduced space requirements;
1699 * automatic HTML generation from theories;
1701 * theory files no longer require "..." (quotes) around most types;
1703 * new examples, including two proofs of the Church-Rosser theorem;
1705 * non-curried (1994) version of HOL is no longer distributed;
1712 * greatly reduced space requirements;
1714 * theory files (.thy) no longer require \...\ escapes at line breaks;
1716 * searchable theorem database (see the section "Retrieving theorems" on
1717 page 8 of the Reference Manual);
1719 * new examples, including Grabczewski's monumental case study of the
1722 * The previous version of HOL renamed to Old_HOL;
1724 * The new version of HOL (previously called CHOL) uses a curried syntax
1725 for functions. Application looks like f a b instead of f(a,b);
1727 * Mutually recursive inductive definitions finally work in HOL;
1729 * In ZF, pattern-matching on tuples is now available in all abstractions and
1730 translates to the operator "split";
1737 * new infix operator, addss, allowing the classical reasoner to
1738 perform simplification at each step of its search. Example:
1739 fast_tac (cs addss ss)
1741 * a new logic, CHOL, the same as HOL, but with a curried syntax
1742 for functions. Application looks like f a b instead of f(a,b). Also pairs
1743 look like (a,b) instead of <a,b>;
1745 * PLEASE NOTE: CHOL will eventually replace HOL!
1747 * In CHOL, pattern-matching on tuples is now available in all abstractions.
1748 It translates to the operator "split". A new theory of integers is available;
1750 * In ZF, integer numerals now denote two's-complement binary integers.
1751 Arithmetic operations can be performed by rewriting. See ZF/ex/Bin.ML;
1753 * Many new examples: I/O automata, Church-Rosser theorem, equivalents
1754 of the Axiom of Choice;
1761 * Significantly faster resolution;
1763 * the different sections in a .thy file can now be mixed and repeated
1766 * Database of theorems for FOL, HOL and ZF. New
1767 commands including qed, qed_goal and bind_thm store theorems in the database.
1769 * Simple database queries: return a named theorem (get_thm) or all theorems of
1770 a given theory (thms_of), or find out what theory a theorem was proved in
1773 * Bugs fixed in the inductive definition and datatype packages;
1775 * The classical reasoner provides deepen_tac and depth_tac, making FOL_dup_cs
1776 and HOL_dup_cs obsolete;
1778 * Syntactic ambiguities caused by the new treatment of syntax in Isabelle94-1
1781 * Simpler definition of function space in ZF;
1783 * new results about cardinal and ordinal arithmetic in ZF;
1785 * 'subtype' facility in HOL for introducing new types as subsets of existing