wenzelm@57491: Isabelle NEWS -- history of user-relevant changes wenzelm@57491: ================================================= wenzelm@2553: wenzelm@57695: New in this Isabelle version wenzelm@57695: ---------------------------- wenzelm@57695: wenzelm@57941: *** General *** wenzelm@57941: wenzelm@57941: * Commands 'method_setup' and 'attribute_setup' now work within a wenzelm@57941: local theory context. wenzelm@57941: wenzelm@57946: * Command 'named_theorems' declares a dynamic fact within the context, wenzelm@57946: together with an attribute to maintain the content incrementally. wenzelm@57946: This supersedes functor Named_Thms, but with a subtle change of wenzelm@57946: semantics due to external visual order vs. internal reverse order. wenzelm@57946: wenzelm@58801: * Command 'notepad' requires proper nesting of begin/end and its proof wenzelm@58801: structure in the body: 'oops' is no longer supported here. Minor wenzelm@58801: INCOMPATIBILITY, use 'sorry' instead. wenzelm@58801: wenzelm@58928: * Outer syntax commands are managed authentically within the theory wenzelm@58928: context, without implicit global state. Potential for accidental wenzelm@58928: INCOMPATIBILITY, make sure that required theories are really imported. wenzelm@58928: haftmann@59105: * 'find_theorems': search patterns which are abstractions are haftmann@59105: schematcially expanded before search. Search results match the haftmann@59105: naive expectation more closely, particularly wrt. abbreviations. haftmann@59105: INCOMPATIBILITY. haftmann@59105: wenzelm@57941: wenzelm@58524: *** Prover IDE -- Isabelle/Scala/jEdit *** wenzelm@58524: wenzelm@59308: * Old graph browser (Java/AWT 1.0) is superseded by improved graphview wenzelm@59308: panel, which also includes PDF output. wenzelm@59308: wenzelm@58704: * Improved folding mode "isabelle" based on Isar syntax. wenzelm@58708: Alternatively, the "sidekick" mode may be used for document structure. wenzelm@58704: wenzelm@58758: * Extended bracket matching based on Isar language structure. System wenzelm@58758: option jedit_structure_limit determines maximum number of lines to wenzelm@58758: scan in the buffer. wenzelm@58758: wenzelm@58540: * Support for BibTeX files: context menu, context-sensitive token wenzelm@58540: marker, SideKick parser. wenzelm@58524: wenzelm@58551: * Document antiquotation @{cite} provides formal markup, which is wenzelm@58551: interpreted semi-formally based on .bib files that happen to be opened wenzelm@58592: in the editor (hyperlinks, completion etc.). wenzelm@58551: wenzelm@58785: * Less waste of vertical space via negative line spacing (see Global wenzelm@58785: Options / Text Area). wenzelm@58785: wenzelm@58524: haftmann@58202: *** Pure *** haftmann@58202: haftmann@58202: * Command "class_deps" takes optional sort arguments constraining haftmann@58202: the search space. haftmann@58202: haftmann@58410: * Lexical separation of signed and unsigend numerals: categories "num" haftmann@58410: and "float" are unsigend. INCOMPATIBILITY: subtle change in precedence haftmann@58410: of numeral signs, particulary in expressions involving infix syntax like haftmann@58410: "(- 1) ^ n". haftmann@58410: wenzelm@58421: * Old inner token category "xnum" has been discontinued. Potential wenzelm@58421: INCOMPATIBILITY for exotic syntax: may use mixfix grammar with "num" wenzelm@58421: token category instead. wenzelm@58421: haftmann@58202: blanchet@57737: *** HOL *** blanchet@57737: hoelzl@58775: * Add NO_MATCH-simproc, allows to check for syntactic non-equality hoelzl@58775: haftmann@58649: * Generalized and consolidated some theorems concerning divsibility: haftmann@58649: dvd_reduce ~> dvd_add_triv_right_iff haftmann@58649: dvd_plus_eq_right ~> dvd_add_right_iff haftmann@58649: dvd_plus_eq_left ~> dvd_add_left_iff haftmann@58649: Minor INCOMPATIBILITY. haftmann@58649: haftmann@58770: * "even" and "odd" are mere abbreviations for "2 dvd _" and "~ 2 dvd _" haftmann@58770: and part of HOL-Main. haftmann@58645: even_def ~> even_iff_mod_2_eq_zero haftmann@58740: INCOMPATIBILITY. haftmann@58645: haftmann@58512: * Lemma name consolidation: divide_Numeral1 ~> divide_numeral_1 haftmann@58512: Minor INCOMPATIBILITY. haftmann@58512: hoelzl@58776: * field_simps: Use NO_MATCH-simproc for distribution rules, to avoid hoelzl@58776: non-termination in case of distributing a division. With this change hoelzl@58776: field_simps is in some cases slightly less powerful, if it fails try hoelzl@58776: to add algebra_simps, or use divide_simps. hoelzl@58776: Minor INCOMPATIBILITY. hoelzl@58776: haftmann@58321: * Bootstrap of listsum as special case of abstract product over lists. haftmann@58321: Fact rename: haftmann@58321: listsum_def ~> listsum.eq_foldr haftmann@58321: INCOMPATIBILITY. haftmann@58321: haftmann@58100: * Command and antiquotation "value" provide different evaluation slots (again), haftmann@58100: where the previous strategy (nbe after ML) serves as default. haftmann@58100: Minor INCOMPATIBILITY. haftmann@58100: blanchet@57983: * New (co)datatype package: blanchet@58373: - The 'datatype_new' command has been renamed 'datatype'. The old blanchet@58373: command of that name is now called 'old_datatype' and is provided blanchet@58373: by "~~/src/HOL/Library/Old_Datatype.thy". See blanchet@58373: 'isabelle doc datatypes' for information on porting. blanchet@58373: INCOMPATIBILITY. blanchet@57983: - Renamed theorems: blanchet@57983: disc_corec ~> corec_disc blanchet@57983: disc_corec_iff ~> corec_disc_iff blanchet@57983: disc_exclude ~> distinct_disc blanchet@57983: disc_exhaust ~> exhaust_disc blanchet@57983: disc_map_iff ~> map_disc_iff blanchet@57983: sel_corec ~> corec_sel blanchet@57983: sel_exhaust ~> exhaust_sel blanchet@57983: sel_map ~> map_sel blanchet@57983: sel_set ~> set_sel blanchet@57983: sel_split ~> split_sel blanchet@57983: sel_split_asm ~> split_sel_asm blanchet@57983: strong_coinduct ~> coinduct_strong blanchet@57983: weak_case_cong ~> case_cong_weak blanchet@57983: INCOMPATIBILITY. blanchet@58192: - The "no_code" option to "free_constructors", "datatype_new", and blanchet@58192: "codatatype" has been renamed "plugins del: code". blanchet@58192: INCOMPATIBILITY. blanchet@58044: - The rules "set_empty" have been removed. They are easy blanchet@58044: consequences of other set rules "by auto". blanchet@58044: INCOMPATIBILITY. blanchet@58044: - The rule "set_cases" is now registered with the "[cases set]" blanchet@57990: attribute. This can influence the behavior of the "cases" proof blanchet@57990: method when more than one case rule is applicable (e.g., an blanchet@57990: assumption is of the form "w : set ws" and the method "cases w" blanchet@57990: is invoked). The solution is to specify the case rule explicitly blanchet@57990: (e.g. "cases w rule: widget.exhaust"). blanchet@57990: INCOMPATIBILITY. blanchet@57983: blanchet@57983: * Old datatype package: blanchet@58310: - The old 'datatype' command has been renamed 'old_datatype', and blanchet@58373: 'rep_datatype' has been renamed 'old_rep_datatype'. They are blanchet@58373: provided by "~~/src/HOL/Library/Old_Datatype.thy". See blanchet@58310: 'isabelle doc datatypes' for information on porting. blanchet@58373: INCOMPATIBILITY. blanchet@57983: - Renamed theorems: blanchet@57983: weak_case_cong ~> case_cong_weak blanchet@57983: INCOMPATIBILITY. blanchet@58373: - Renamed theory: blanchet@58373: ~~/src/HOL/Datatype.thy ~> ~~/src/HOL/Library/Old_Datatype.thy blanchet@58373: INCOMPATIBILITY. blanchet@57983: haftmann@58368: * Product over lists via constant "listprod". haftmann@58368: blanchet@59039: * Nitpick: blanchet@59039: - Fixed soundness bug related to the strict and nonstrict subset blanchet@59039: operations. blanchet@59039: blanchet@57737: * Sledgehammer: blanchet@57737: - Minimization is now always enabled by default. blanchet@57737: Removed subcommand: blanchet@57737: min blanchet@59039: - The proof reconstruction, both one-liners and Isar, has been blanchet@59039: dramatically improved. blanchet@59039: - Improved support for CVC4 and veriT. blanchet@57737: blanchet@58062: * Old and new SMT modules: blanchet@58067: - The old 'smt' method has been renamed 'old_smt' and moved to boehmes@59216: 'src/HOL/Library/Old_SMT.thy'. It is provided for compatibility, until blanchet@58067: applications have been ported to use the new 'smt' method. For the blanchet@58067: method to work, an older version of Z3 (e.g. Z3 3.2 or 4.0) must be blanchet@58062: installed, and the environment variable "OLD_Z3_SOLVER" must point to blanchet@58062: it. blanchet@58062: INCOMPATIBILITY. blanchet@58067: - The 'smt2' method has been renamed 'smt'. blanchet@58060: INCOMPATIBILITY. boehmes@59216: - New option 'smt_reconstruction_step_timeout' to limit the reconstruction boehmes@59216: time of Z3 proof steps in the new 'smt' method. boehmes@59216: - New option 'smt_statistics' to display statistics of the new 'smt' boehmes@59216: method, especially runtime statistics of Z3 proof reconstruction. blanchet@58060: nipkow@58247: * List: renamed drop_Suc_conv_tl and nth_drop' to Cons_nth_drop_Suc nipkow@58247: Andreas@58626: * New infrastructure for compiling, running, evaluating and testing Andreas@58626: generated code in target languages in HOL/Library/Code_Test. See Andreas@58626: HOL/Codegenerator_Test/Code_Test* for examples. wenzelm@58008: wenzelm@58630: * Library/Sum_of_Squares: simplified and improved "sos" method. Always wenzelm@58630: use local CSDP executable, which is much faster than the NEOS server. wenzelm@58630: The "sos_cert" functionality is invoked as "sos" with additional wenzelm@58630: argument. Minor INCOMPATIBILITY. wenzelm@58630: immler@58990: * HOL-Decision_Procs: immler@58990: - New counterexample generator quickcheck[approximation] for immler@58990: inequalities of transcendental functions. immler@58990: Uses hardware floating point arithmetic to randomly discover immler@58990: potential counterexamples. Counterexamples are certified with the immler@58990: 'approximation' method. immler@58990: See HOL/Decision_Procs/ex/Approximation_Quickcheck_Ex.thy for immler@58990: examples. immler@58990: hoelzl@59354: * HOL-Probability: Reworked measurability prover hoelzl@59354: - applies destructor rules repeatetly hoelzl@59354: - removed application splitting (replaced by destructor rule) hoelzl@59354: - added congruence rules to rewrite measure spaces under the sets projection wenzelm@58630: wenzelm@58716: *** Document preparation *** wenzelm@58716: wenzelm@58999: * Document markup commands 'chapter', 'section', 'subsection', wenzelm@58999: 'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any wenzelm@58999: context, even before the initial 'theory' command. Obsolete proof wenzelm@58999: commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been wenzelm@58999: discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw' wenzelm@58999: instead. The old 'header' command is still retained for some time, but wenzelm@58999: should be replaced by 'chapter', 'section' etc. (using "isabelle wenzelm@58999: update_header"). Minor INCOMPATIBILITY. wenzelm@58999: wenzelm@58999: * Diagnostic commands and document markup commands within a proof do not wenzelm@58999: affect the command tag for output. Thus commands like 'thm' are subject wenzelm@58999: to proof document structure, and no longer "stick out" accidentally. wenzelm@58999: Commands 'text' and 'txt' merely differ in the LaTeX style, not their wenzelm@58999: tags. Potential INCOMPATIBILITY in exotic situations. wenzelm@58868: wenzelm@58716: * Official support for "tt" style variants, via \isatt{...} or wenzelm@58716: \begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or wenzelm@58716: verbatim environment of LaTeX is no longer used. This allows @{ML} etc. wenzelm@58716: as argument to other macros (such as footnotes). wenzelm@58716: wenzelm@58716: * Document antiquotation @{verbatim} prints ASCII text literally in "tt" wenzelm@58716: style. wenzelm@58716: wenzelm@58716: blanchet@58066: *** ML *** blanchet@58066: wenzelm@59180: * Former combinators NAMED_CRITICAL and CRITICAL for central critical wenzelm@59180: sections have been discontinued, in favour of the more elementary wenzelm@59180: Multithreading.synchronized and its high-level derivative wenzelm@59180: Synchronized.var (which is usually sufficient in applications). Subtle wenzelm@59180: INCOMPATIBILITY: synchronized access needs to be atomic and cannot be wenzelm@59180: nested. wenzelm@59180: wenzelm@59112: * Cartouches within ML sources are turned into values of type wenzelm@59112: Input.source (with formal position information). wenzelm@59112: wenzelm@58963: * Proper context for various elementary tactics: assume_tac, wenzelm@58963: match_tac, compose_tac, Splitter.split_tac etc. Minor wenzelm@58963: INCOMPATIBILITY. wenzelm@58956: blanchet@58066: * Tactical PARALLEL_ALLGOALS is the most common way to refer to blanchet@58066: PARALLEL_GOALS. blanchet@58066: wenzelm@59057: * Basic combinators map, fold, fold_map, split_list, apply are wenzelm@59057: available as parameterized antiquotations, e.g. @{map 4} for lists of wenzelm@59057: quadruples. wenzelm@58634: wenzelm@59058: * Renamed "pairself" to "apply2", in accordance to @{apply 2}. wenzelm@59058: INCOMPATIBILITY. wenzelm@59058: wenzelm@59139: * Synchronized.value (ML) is actually synchronized (as in Scala): wenzelm@59139: subtle change of semantics with minimal potential for INCOMPATIBILITY. wenzelm@59139: blanchet@58066: wenzelm@58610: *** System *** wenzelm@58610: wenzelm@58846: * Support for Proof General and Isar TTY loop has been discontinued. wenzelm@58846: Minor INCOMPATIBILITY. wenzelm@58842: wenzelm@59200: * JVM system property "isabelle.threads" determines size of Scala thread wenzelm@59200: pool, like Isabelle system option "threads" for ML. wenzelm@59200: wenzelm@59201: * JVM system property "isabelle.laf" determines the default Swing wenzelm@59201: look-and-feel, via internal class name or symbolic name as in the jEdit wenzelm@59201: menu Global Options / Appearance. wenzelm@59201: wenzelm@59175: * System option "pretty_margin" is superseded by "thy_output_margin", wenzelm@59175: which is also accessible via document antiquotation option "margin". wenzelm@59175: Only the margin for document output may be changed, but not the global wenzelm@59175: pretty printing: that is 76 for plain console output, and adapted wenzelm@59175: dynamically in GUI front-ends. Implementations of document wenzelm@59175: antiquotations need to observe the margin explicitly according to wenzelm@59175: Thy_Output.string_of_margin. Minor INCOMPATIBILITY. wenzelm@59175: wenzelm@58861: * Historical command-line terminator ";" is no longer accepted. Minor wenzelm@58861: INCOMPATIBILITY, use "isabelle update_semicolons" to remove obsolete wenzelm@58861: semicolons from theory sources. wenzelm@58861: wenzelm@58610: * The Isabelle tool "update_cartouches" changes theory files to use wenzelm@58610: cartouches instead of old-style {* verbatim *} or `alt_string` tokens. wenzelm@58610: wenzelm@58610: wenzelm@57695: wenzelm@57452: New in Isabelle2014 (August 2014) wenzelm@57452: --------------------------------- wenzelm@54055: wenzelm@54702: *** General *** wenzelm@54702: wenzelm@57452: * Support for official Standard ML within the Isabelle context. wenzelm@57452: Command 'SML_file' reads and evaluates the given Standard ML file. wenzelm@57452: Toplevel bindings are stored within the theory context; the initial wenzelm@57452: environment is restricted to the Standard ML implementation of wenzelm@57452: Poly/ML, without the add-ons of Isabelle/ML. Commands 'SML_import' wenzelm@57452: and 'SML_export' allow to exchange toplevel bindings between the two wenzelm@57452: separate environments. See also ~~/src/Tools/SML/Examples.thy for wenzelm@57452: some examples. wenzelm@56499: wenzelm@57504: * Standard tactics and proof methods such as "clarsimp", "auto" and wenzelm@57504: "safe" now preserve equality hypotheses "x = expr" where x is a free wenzelm@57504: variable. Locale assumptions and chained facts containing "x" wenzelm@57504: continue to be useful. The new method "hypsubst_thin" and the wenzelm@57504: configuration option "hypsubst_thin" (within the attribute name space) wenzelm@57504: restore the previous behavior. INCOMPATIBILITY, especially where wenzelm@57504: induction is done after these methods or when the names of free and wenzelm@57504: bound variables clash. As first approximation, old proofs may be wenzelm@57504: repaired by "using [[hypsubst_thin = true]]" in the critical spot. wenzelm@57504: wenzelm@56232: * More static checking of proof methods, which allows the system to wenzelm@56232: form a closure over the concrete syntax. Method arguments should be wenzelm@56232: processed in the original proof context as far as possible, before wenzelm@56232: operating on the goal state. In any case, the standard discipline for wenzelm@56232: subgoal-addressing needs to be observed: no subgoals or a subgoal wenzelm@56232: number that is out of range produces an empty result sequence, not an wenzelm@56232: exception. Potential INCOMPATIBILITY for non-conformant tactical wenzelm@56232: proof tools. wenzelm@56232: wenzelm@57452: * Lexical syntax (inner and outer) supports text cartouches with wenzelm@57452: arbitrary nesting, and without escapes of quotes etc. The Prover IDE wenzelm@57452: supports input via ` (backquote). wenzelm@57452: wenzelm@57452: * The outer syntax categories "text" (for formal comments and document wenzelm@57452: markup commands) and "altstring" (for literal fact references) allow wenzelm@57452: cartouches as well, in addition to the traditional mix of quotations. wenzelm@57452: wenzelm@57452: * Syntax of document antiquotation @{rail} now uses \ instead wenzelm@57452: of "\\", to avoid the optical illusion of escaped backslash within wenzelm@57491: string token. General renovation of its syntax using text cartouches. wenzelm@57452: Minor INCOMPATIBILITY. wenzelm@57452: wenzelm@57452: * Discontinued legacy_isub_isup, which was a temporary workaround for wenzelm@57452: Isabelle/ML in Isabelle2013-1. The prover process no longer accepts wenzelm@57452: old identifier syntax with \<^isub> or \<^isup>. Potential wenzelm@57452: INCOMPATIBILITY. wenzelm@57452: wenzelm@57452: * Document antiquotation @{url} produces markup for the given URL, wenzelm@57452: which results in an active hyperlink within the text. wenzelm@57452: wenzelm@57452: * Document antiquotation @{file_unchecked} is like @{file}, but does wenzelm@57452: not check existence within the file-system. wenzelm@57452: wenzelm@57452: * Updated and extended manuals: codegen, datatypes, implementation, wenzelm@57452: isar-ref, jedit, system. wenzelm@57423: wenzelm@54702: wenzelm@54533: *** Prover IDE -- Isabelle/Scala/jEdit *** wenzelm@54533: wenzelm@57650: * Improved Document panel: simplified interaction where every single wenzelm@57452: mouse click (re)opens document via desktop environment or as jEdit wenzelm@57452: buffer. wenzelm@57452: wenzelm@57452: * Support for Navigator plugin (with toolbar buttons), with connection wenzelm@57452: to PIDE hyperlinks. wenzelm@57452: wenzelm@57452: * Auxiliary files ('ML_file' etc.) are managed by the Prover IDE. wenzelm@57452: Open text buffers take precedence over copies within the file-system. wenzelm@57452: wenzelm@57452: * Improved support for Isabelle/ML, with jEdit mode "isabelle-ml" for wenzelm@57452: auxiliary ML files. wenzelm@57423: wenzelm@57423: * Improved syntactic and semantic completion mechanism, with simple wenzelm@57423: templates, completion language context, name-space completion, wenzelm@57423: file-name completion, spell-checker completion. wenzelm@57423: wenzelm@57423: * Refined GUI popup for completion: more robust key/mouse event wenzelm@57423: handling and propagation to enclosing text area -- avoid loosing wenzelm@57423: keystrokes with slow / remote graphics displays. wenzelm@57423: wenzelm@57833: * Completion popup supports both ENTER and TAB (default) to select an wenzelm@57833: item, depending on Isabelle options. wenzelm@57833: wenzelm@57423: * Refined insertion of completion items wrt. jEdit text: multiple wenzelm@57423: selections, rectangular selections, rectangular selection as "tall wenzelm@57423: caret". wenzelm@56342: wenzelm@56580: * Integrated spell-checker for document text, comments etc. with wenzelm@57423: completion popup and context-menu. wenzelm@56554: wenzelm@56879: * More general "Query" panel supersedes "Find" panel, with GUI access wenzelm@56879: to commands 'find_theorems' and 'find_consts', as well as print wenzelm@56879: operations for the context. Minor incompatibility in keyboard wenzelm@56879: shortcuts etc.: replace action isabelle-find by isabelle-query. wenzelm@56761: wenzelm@56901: * Search field for all output panels ("Output", "Query", "Info" etc.) wenzelm@56901: to highlight text via regular expression. wenzelm@56901: wenzelm@54881: * Option "jedit_print_mode" (see also "Plugin Options / Isabelle / wenzelm@54881: General") allows to specify additional print modes for the prover wenzelm@54881: process, without requiring old-fashioned command-line invocation of wenzelm@54881: "isabelle jedit -m MODE". wenzelm@54881: wenzelm@56505: * More support for remote files (e.g. http) using standard Java wenzelm@56505: networking operations instead of jEdit virtual file-systems. wenzelm@56505: wenzelm@57822: * Empty editors buffers that are no longer required (e.g.\ via theory wenzelm@57822: imports) are automatically removed from the document model. wenzelm@57822: wenzelm@57869: * Improved monitor panel. wenzelm@57869: wenzelm@56838: * Improved Console/Scala plugin: more uniform scala.Console output, wenzelm@56838: more robust treatment of threads and interrupts. wenzelm@56838: wenzelm@56939: * Improved management of dockable windows: clarified keyboard focus wenzelm@56939: and window placement wrt. main editor view; optional menu item to wenzelm@56939: "Detach" a copy where this makes sense. wenzelm@56939: wenzelm@57452: * New Simplifier Trace panel provides an interactive view of the wenzelm@57591: simplification process, enabled by the "simp_trace_new" attribute wenzelm@57452: within the context. wenzelm@57452: wenzelm@57452: wenzelm@55001: *** Pure *** wenzelm@55001: wenzelm@57504: * Low-level type-class commands 'classes', 'classrel', 'arities' have wenzelm@57504: been discontinued to avoid the danger of non-trivial axiomatization wenzelm@57504: that is not immediately visible. INCOMPATIBILITY, use regular wenzelm@57504: 'instance' command with proof. The required OFCLASS(...) theorem wenzelm@57504: might be postulated via 'axiomatization' beforehand, or the proof wenzelm@57504: finished trivially if the underlying class definition is made vacuous wenzelm@57504: (without any assumptions). See also Isabelle/ML operations wenzelm@57504: Axclass.class_axiomatization, Axclass.classrel_axiomatization, wenzelm@57504: Axclass.arity_axiomatization. wenzelm@57504: wenzelm@56245: * Basic constants of Pure use more conventional names and are always wenzelm@56245: qualified. Rare INCOMPATIBILITY, but with potentially serious wenzelm@56245: consequences, notably for tools in Isabelle/ML. The following wenzelm@56245: renaming needs to be applied: wenzelm@56245: wenzelm@56245: == ~> Pure.eq wenzelm@56245: ==> ~> Pure.imp wenzelm@56245: all ~> Pure.all wenzelm@56245: TYPE ~> Pure.type wenzelm@56245: dummy_pattern ~> Pure.dummy_pattern wenzelm@56245: wenzelm@56245: Systematic porting works by using the following theory setup on a wenzelm@56245: *previous* Isabelle version to introduce the new name accesses for the wenzelm@56245: old constants: wenzelm@56245: wenzelm@56245: setup {* wenzelm@56245: fn thy => thy wenzelm@56245: |> Sign.root_path wenzelm@56245: |> Sign.const_alias (Binding.qualify true "Pure" @{binding eq}) "==" wenzelm@56245: |> Sign.const_alias (Binding.qualify true "Pure" @{binding imp}) "==>" wenzelm@56245: |> Sign.const_alias (Binding.qualify true "Pure" @{binding all}) "all" wenzelm@56245: |> Sign.restore_naming thy wenzelm@56245: *} wenzelm@56245: wenzelm@56245: Thus ML antiquotations like @{const_name Pure.eq} may be used already. wenzelm@56245: Later the application is moved to the current Isabelle version, and wenzelm@56245: the auxiliary aliases are deleted. wenzelm@56245: wenzelm@55143: * Attributes "where" and "of" allow an optional context of local wenzelm@55143: variables ('for' declaration): these variables become schematic in the wenzelm@55143: instantiated theorem. wenzelm@55143: wenzelm@55152: * Obsolete attribute "standard" has been discontinued (legacy since wenzelm@55152: Isabelle2012). Potential INCOMPATIBILITY, use explicit 'for' context wenzelm@55152: where instantiations with schematic variables are intended (for wenzelm@55152: declaration commands like 'lemmas' or attributes like "of"). The wenzelm@55152: following temporary definition may help to port old applications: wenzelm@55152: wenzelm@55152: attribute_setup standard = wenzelm@55152: "Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))" wenzelm@55152: wenzelm@55001: * More thorough check of proof context for goal statements and wenzelm@55006: attributed fact expressions (concerning background theory, declared wenzelm@55006: hyps). Potential INCOMPATIBILITY, tools need to observe standard wenzelm@55006: context discipline. See also Assumption.add_assumes and the more wenzelm@55006: primitive Thm.assume_hyps. wenzelm@55001: wenzelm@55108: * Inner syntax token language allows regular quoted strings "..." wenzelm@55108: (only makes sense in practice, if outer syntax is delimited wenzelm@57452: differently, e.g. via cartouches). wenzelm@57452: wenzelm@57504: * Command 'print_term_bindings' supersedes 'print_binds' for clarity, wenzelm@57504: but the latter is retained some time as Proof General legacy. wenzelm@57504: wenzelm@57452: * Code generator preprocessor: explicit control of simp tracing on a wenzelm@57452: per-constant basis. See attribute "code_preproc". haftmann@57430: wenzelm@55001: haftmann@54227: *** HOL *** haftmann@54227: wenzelm@57504: * Code generator: enforce case of identifiers only for strict target wenzelm@57504: language requirements. INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * Code generator: explicit proof contexts in many ML interfaces. wenzelm@57504: INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * Code generator: minimize exported identifiers by default. Minor wenzelm@57504: INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * Code generation for SML and OCaml: dropped arcane "no_signatures" wenzelm@57504: option. Minor INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * "declare [[code abort: ...]]" replaces "code_abort ...". wenzelm@57504: INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * "declare [[code drop: ...]]" drops all code equations associated wenzelm@57504: with the given constants. wenzelm@57504: wenzelm@57504: * Code generations are provided for make, fields, extend and truncate wenzelm@57504: operations on records. haftmann@57437: wenzelm@57452: * Command and antiquotation "value" are now hardcoded against nbe and wenzelm@57452: ML. Minor INCOMPATIBILITY. wenzelm@57452: wenzelm@57504: * Renamed command 'enriched_type' to 'functor'. INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * The symbol "\" may be used within char or string literals wenzelm@57504: to represent (Char Nibble0 NibbleA), i.e. ASCII newline. wenzelm@57504: wenzelm@57504: * Qualified String.implode and String.explode. INCOMPATIBILITY. haftmann@56923: wenzelm@57452: * Simplifier: Enhanced solver of preconditions of rewrite rules can wenzelm@57452: now deal with conjunctions. For help with converting proofs, the old wenzelm@57452: behaviour of the simplifier can be restored like this: declare/using wenzelm@57452: [[simp_legacy_precond]]. This configuration option will disappear wenzelm@57452: again in the future. INCOMPATIBILITY. nipkow@56073: wenzelm@55139: * Simproc "finite_Collect" is no longer enabled by default, due to wenzelm@55139: spurious crashes and other surprises. Potential INCOMPATIBILITY. wenzelm@55139: wenzelm@57452: * Moved new (co)datatype package and its dependencies from session wenzelm@57452: "HOL-BNF" to "HOL". The commands 'bnf', 'wrap_free_constructors', wenzelm@57452: 'datatype_new', 'codatatype', 'primcorec', 'primcorecursive' are now wenzelm@57452: part of theory "Main". wenzelm@57452: blanchet@55098: Theory renamings: blanchet@55098: FunDef.thy ~> Fun_Def.thy (and Fun_Def_Base.thy) blanchet@55098: Library/Wfrec.thy ~> Wfrec.thy blanchet@55098: Library/Zorn.thy ~> Zorn.thy blanchet@55098: Cardinals/Order_Relation.thy ~> Order_Relation.thy blanchet@55098: Library/Order_Union.thy ~> Cardinals/Order_Union.thy blanchet@55098: Cardinals/Cardinal_Arithmetic_Base.thy ~> BNF_Cardinal_Arithmetic.thy blanchet@55098: Cardinals/Cardinal_Order_Relation_Base.thy ~> BNF_Cardinal_Order_Relation.thy blanchet@55098: Cardinals/Constructions_on_Wellorders_Base.thy ~> BNF_Constructions_on_Wellorders.thy blanchet@55098: Cardinals/Wellorder_Embedding_Base.thy ~> BNF_Wellorder_Embedding.thy blanchet@55098: Cardinals/Wellorder_Relation_Base.thy ~> BNF_Wellorder_Relation.thy blanchet@55098: BNF/Ctr_Sugar.thy ~> Ctr_Sugar.thy blanchet@55098: BNF/Basic_BNFs.thy ~> Basic_BNFs.thy blanchet@55098: BNF/BNF_Comp.thy ~> BNF_Comp.thy blanchet@55098: BNF/BNF_Def.thy ~> BNF_Def.thy blanchet@55098: BNF/BNF_FP_Base.thy ~> BNF_FP_Base.thy blanchet@55098: BNF/BNF_GFP.thy ~> BNF_GFP.thy blanchet@55098: BNF/BNF_LFP.thy ~> BNF_LFP.thy blanchet@55098: BNF/BNF_Util.thy ~> BNF_Util.thy blanchet@55098: BNF/Coinduction.thy ~> Coinduction.thy blanchet@55098: BNF/More_BNFs.thy ~> Library/More_BNFs.thy blanchet@55098: BNF/Countable_Type.thy ~> Library/Countable_Set_Type.thy blanchet@55098: BNF/Examples/* ~> BNF_Examples/* wenzelm@57452: blanchet@55098: New theories: blanchet@55098: Wellorder_Extension.thy (split from Zorn.thy) blanchet@55098: Library/Cardinal_Notations.thy traytel@56942: Library/BNF_Axomatization.thy blanchet@55098: BNF_Examples/Misc_Primcorec.thy blanchet@55098: BNF_Examples/Stream_Processor.thy wenzelm@57452: blanchet@55519: Discontinued theories: blanchet@55098: BNF/BNF.thy blanchet@55098: BNF/Equiv_Relations_More.thy wenzelm@57452: wenzelm@57452: INCOMPATIBILITY. blanchet@55098: blanchet@56118: * New (co)datatype package: wenzelm@57452: - Command 'primcorec' is fully implemented. wenzelm@57452: - Command 'datatype_new' generates size functions ("size_xxx" and wenzelm@57452: "size") as required by 'fun'. wenzelm@57452: - BNFs are integrated with the Lifting tool and new-style wenzelm@57452: (co)datatypes with Transfer. wenzelm@57452: - Renamed commands: blanchet@55875: datatype_new_compat ~> datatype_compat blanchet@55875: primrec_new ~> primrec blanchet@55875: wrap_free_constructors ~> free_constructors blanchet@55875: INCOMPATIBILITY. wenzelm@57452: - The generated constants "xxx_case" and "xxx_rec" have been renamed blanchet@55875: "case_xxx" and "rec_xxx" (e.g., "prod_case" ~> "case_prod"). blanchet@55875: INCOMPATIBILITY. wenzelm@57452: - The constant "xxx_(un)fold" and related theorems are no longer wenzelm@57452: generated. Use "xxx_(co)rec" or define "xxx_(un)fold" manually wenzelm@57452: using "prim(co)rec". blanchet@55875: INCOMPATIBILITY. wenzelm@57452: - No discriminators are generated for nullary constructors by wenzelm@57452: default, eliminating the need for the odd "=:" syntax. blanchet@57091: INCOMPATIBILITY. wenzelm@57452: - No discriminators or selectors are generated by default by blanchet@57094: "datatype_new", unless custom names are specified or the new blanchet@57094: "discs_sels" option is passed. blanchet@57094: INCOMPATIBILITY. blanchet@55875: blanchet@55643: * Old datatype package: wenzelm@57452: - The generated theorems "xxx.cases" and "xxx.recs" have been wenzelm@57452: renamed "xxx.case" and "xxx.rec" (e.g., "sum.cases" -> wenzelm@57452: "sum.case"). INCOMPATIBILITY. wenzelm@57452: - The generated constants "xxx_case", "xxx_rec", and "xxx_size" have wenzelm@57452: been renamed "case_xxx", "rec_xxx", and "size_xxx" (e.g., wenzelm@57452: "prod_case" ~> "case_prod"). INCOMPATIBILITY. wenzelm@57452: wenzelm@57452: * The types "'a list" and "'a option", their set and map functions, wenzelm@57452: their relators, and their selectors are now produced using the new wenzelm@57452: BNF-based datatype package. wenzelm@57452: blanchet@55519: Renamed constants: blanchet@55519: Option.set ~> set_option blanchet@55519: Option.map ~> map_option blanchet@55525: option_rel ~> rel_option wenzelm@57452: blanchet@55519: Renamed theorems: blanchet@55585: set_def ~> set_rec[abs_def] blanchet@55519: map_def ~> map_rec[abs_def] blanchet@55519: Option.map_def ~> map_option_case[abs_def] (with "case_option" instead of "rec_option") blanchet@56652: option.recs ~> option.rec blanchet@55524: list_all2_def ~> list_all2_iff blanchet@55585: set.simps ~> set_simps (or the slightly different "list.set") blanchet@55519: map.simps ~> list.map blanchet@55519: hd.simps ~> list.sel(1) blanchet@55519: tl.simps ~> list.sel(2-3) blanchet@55519: the.simps ~> option.sel wenzelm@57452: wenzelm@57452: INCOMPATIBILITY. blanchet@55519: blanchet@55933: * The following map functions and relators have been renamed: blanchet@55939: sum_map ~> map_sum blanchet@55939: map_pair ~> map_prod blanchet@55944: prod_rel ~> rel_prod blanchet@55943: sum_rel ~> rel_sum blanchet@55945: fun_rel ~> rel_fun blanchet@55942: set_rel ~> rel_set blanchet@55942: filter_rel ~> rel_filter wenzelm@57452: fset_rel ~> rel_fset (in "src/HOL/Library/FSet.thy") wenzelm@57452: cset_rel ~> rel_cset (in "src/HOL/Library/Countable_Set_Type.thy") wenzelm@57452: vset ~> rel_vset (in "src/HOL/Library/Quotient_Set.thy") wenzelm@57452: wenzelm@57452: INCOMPATIBILITY. wenzelm@57452: kuncar@57826: * Lifting and Transfer: kuncar@57826: - a type variable as a raw type is supported kuncar@57826: - stronger reflexivity prover kuncar@57826: - rep_eq is always generated by lift_definition wenzelm@57856: - setup for Lifting/Transfer is now automated for BNFs kuncar@57826: + holds for BNFs that do not contain a dead variable wenzelm@57856: + relator_eq, relator_mono, relator_distr, relator_domain, kuncar@57826: relator_eq_onp, quot_map, transfer rules for bi_unique, bi_total, kuncar@57826: right_unique, right_total, left_unique, left_total are proved kuncar@57826: automatically kuncar@57826: + definition of a predicator is generated automatically kuncar@57826: + simplification rules for a predicator definition are proved kuncar@57826: automatically for datatypes kuncar@57826: - consolidation of the setup of Lifting/Transfer wenzelm@57856: + property that a relator preservers reflexivity is not needed any kuncar@57826: more kuncar@57826: Minor INCOMPATIBILITY. wenzelm@57856: + left_total and left_unique rules are now transfer rules kuncar@57826: (reflexivity_rule attribute not needed anymore) kuncar@57826: INCOMPATIBILITY. wenzelm@57856: + Domainp does not have to be a separate assumption in kuncar@57826: relator_domain theorems (=> more natural statement) kuncar@57826: INCOMPATIBILITY. kuncar@57826: - registration of code equations is more robust kuncar@57826: Potential INCOMPATIBILITY. kuncar@57826: - respectfulness proof obligation is preprocessed to a more readable kuncar@57826: form kuncar@57826: Potential INCOMPATIBILITY. kuncar@57826: - eq_onp is always unfolded in respectfulness proof obligation kuncar@57826: Potential INCOMPATIBILITY. wenzelm@57856: - unregister lifting setup for Code_Numeral.integer and kuncar@57826: Code_Numeral.natural kuncar@57826: Potential INCOMPATIBILITY. kuncar@57826: - Lifting.invariant -> eq_onp kuncar@57826: INCOMPATIBILITY. wenzelm@57856: wenzelm@57508: * New internal SAT solver "cdclite" that produces models and proof wenzelm@57508: traces. This solver replaces the internal SAT solvers "enumerate" and wenzelm@57508: "dpll". Applications that explicitly used one of these two SAT wenzelm@57508: solvers should use "cdclite" instead. In addition, "cdclite" is now wenzelm@57508: the default SAT solver for the "sat" and "satx" proof methods and wenzelm@57508: corresponding tactics; the old default can be restored using "declare wenzelm@57508: [[sat_solver = zchaff_with_proofs]]". Minor INCOMPATIBILITY. wenzelm@57508: wenzelm@57508: * SMT module: A new version of the SMT module, temporarily called wenzelm@57508: "SMT2", uses SMT-LIB 2 and supports recent versions of Z3 (e.g., wenzelm@57508: 4.3). The new proof method is called "smt2". CVC3 and CVC4 are also wenzelm@57508: supported as oracles. Yices is no longer supported, because no version wenzelm@57508: of the solver can handle both SMT-LIB 2 and quantifiers. wenzelm@57508: wenzelm@57508: * Activation of Z3 now works via "z3_non_commercial" system option wenzelm@57508: (without requiring restart), instead of former settings variable wenzelm@57508: "Z3_NON_COMMERCIAL". The option can be edited in Isabelle/jEdit menu wenzelm@57508: Plugin Options / Isabelle / General. wenzelm@57508: wenzelm@57508: * Sledgehammer: wenzelm@57508: - Z3 can now produce Isar proofs. wenzelm@57508: - MaSh overhaul: blanchet@57532: . New SML-based learning algorithms eliminate the dependency on wenzelm@57508: Python and increase performance and reliability. wenzelm@57508: . MaSh and MeSh are now used by default together with the wenzelm@57508: traditional MePo (Meng-Paulson) relevance filter. To disable wenzelm@57508: MaSh, set the "MaSh" system option in Isabelle/jEdit Plugin wenzelm@57508: Options / Isabelle / General to "none". wenzelm@57508: - New option: wenzelm@57508: smt_proofs wenzelm@57508: - Renamed options: wenzelm@57508: isar_compress ~> compress wenzelm@57508: isar_try0 ~> try0 wenzelm@57508: wenzelm@57508: INCOMPATIBILITY. wenzelm@57508: wenzelm@57508: * Removed solvers remote_cvc3 and remote_z3. Use cvc3 and z3 instead. wenzelm@57508: wenzelm@57508: * Nitpick: wenzelm@57508: - Fixed soundness bug whereby mutually recursive datatypes could wenzelm@57508: take infinite values. wenzelm@57508: - Fixed soundness bug with low-level number functions such as wenzelm@57508: "Abs_Integ" and "Rep_Integ". wenzelm@57508: - Removed "std" option. wenzelm@57508: - Renamed "show_datatypes" to "show_types" and "hide_datatypes" to wenzelm@57508: "hide_types". wenzelm@57508: wenzelm@57508: * Metis: Removed legacy proof method 'metisFT'. Use 'metis wenzelm@57508: (full_types)' instead. INCOMPATIBILITY. wenzelm@57508: wenzelm@57508: * Try0: Added 'algebra' and 'meson' to the set of proof methods. wenzelm@57508: wenzelm@57508: * Adjustion of INF and SUP operations: wenzelm@57508: - Elongated constants INFI and SUPR to INFIMUM and SUPREMUM. wenzelm@57508: - Consolidated theorem names containing INFI and SUPR: have INF and wenzelm@57508: SUP instead uniformly. wenzelm@57508: - More aggressive normalization of expressions involving INF and Inf wenzelm@57508: or SUP and Sup. wenzelm@57508: - INF_image and SUP_image do not unfold composition. wenzelm@57508: - Dropped facts INF_comp, SUP_comp. wenzelm@57508: - Default congruence rules strong_INF_cong and strong_SUP_cong, with wenzelm@57508: simplifier implication in premises. Generalize and replace former wenzelm@57508: INT_cong, SUP_cong wenzelm@57508: wenzelm@57508: INCOMPATIBILITY. wenzelm@57508: wenzelm@57508: * SUP and INF generalized to conditionally_complete_lattice. wenzelm@57508: wenzelm@57508: * Swapped orientation of facts image_comp and vimage_comp: wenzelm@57508: wenzelm@57508: image_compose ~> image_comp [symmetric] wenzelm@57508: image_comp ~> image_comp [symmetric] wenzelm@57508: vimage_compose ~> vimage_comp [symmetric] wenzelm@57508: vimage_comp ~> vimage_comp [symmetric] wenzelm@57508: wenzelm@57508: INCOMPATIBILITY. wenzelm@57508: wenzelm@57504: * Theory reorganization: split of Big_Operators.thy into wenzelm@57504: Groups_Big.thy and Lattices_Big.thy. blanchet@55098: haftmann@57418: * Consolidated some facts about big group operators: haftmann@57418: haftmann@57418: setsum_0' ~> setsum.neutral haftmann@57418: setsum_0 ~> setsum.neutral_const haftmann@57418: setsum_addf ~> setsum.distrib haftmann@57418: setsum_cartesian_product ~> setsum.cartesian_product haftmann@57418: setsum_cases ~> setsum.If_cases haftmann@57418: setsum_commute ~> setsum.commute haftmann@57418: setsum_cong ~> setsum.cong haftmann@57418: setsum_delta ~> setsum.delta haftmann@57418: setsum_delta' ~> setsum.delta' haftmann@57418: setsum_diff1' ~> setsum.remove haftmann@57418: setsum_empty ~> setsum.empty haftmann@57418: setsum_infinite ~> setsum.infinite haftmann@57418: setsum_insert ~> setsum.insert haftmann@57418: setsum_inter_restrict'' ~> setsum.inter_filter haftmann@57418: setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left haftmann@57418: setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right haftmann@57418: setsum_mono_zero_left ~> setsum.mono_neutral_left haftmann@57418: setsum_mono_zero_right ~> setsum.mono_neutral_right haftmann@57418: setsum_reindex ~> setsum.reindex haftmann@57418: setsum_reindex_cong ~> setsum.reindex_cong haftmann@57418: setsum_reindex_nonzero ~> setsum.reindex_nontrivial haftmann@57418: setsum_restrict_set ~> setsum.inter_restrict haftmann@57418: setsum_Plus ~> setsum.Plus haftmann@57418: setsum_setsum_restrict ~> setsum.commute_restrict haftmann@57418: setsum_Sigma ~> setsum.Sigma haftmann@57418: setsum_subset_diff ~> setsum.subset_diff haftmann@57418: setsum_Un_disjoint ~> setsum.union_disjoint haftmann@57418: setsum_UN_disjoint ~> setsum.UNION_disjoint haftmann@57418: setsum_Un_Int ~> setsum.union_inter haftmann@57418: setsum_Union_disjoint ~> setsum.Union_disjoint haftmann@57418: setsum_UNION_zero ~> setsum.Union_comp haftmann@57418: setsum_Un_zero ~> setsum.union_inter_neutral haftmann@57418: strong_setprod_cong ~> setprod.strong_cong haftmann@57418: strong_setsum_cong ~> setsum.strong_cong haftmann@57418: setprod_1' ~> setprod.neutral haftmann@57418: setprod_1 ~> setprod.neutral_const haftmann@57418: setprod_cartesian_product ~> setprod.cartesian_product haftmann@57418: setprod_cong ~> setprod.cong haftmann@57418: setprod_delta ~> setprod.delta haftmann@57418: setprod_delta' ~> setprod.delta' haftmann@57418: setprod_empty ~> setprod.empty haftmann@57418: setprod_infinite ~> setprod.infinite haftmann@57418: setprod_insert ~> setprod.insert haftmann@57418: setprod_mono_one_cong_left ~> setprod.mono_neutral_cong_left haftmann@57418: setprod_mono_one_cong_right ~> setprod.mono_neutral_cong_right haftmann@57418: setprod_mono_one_left ~> setprod.mono_neutral_left haftmann@57418: setprod_mono_one_right ~> setprod.mono_neutral_right haftmann@57418: setprod_reindex ~> setprod.reindex haftmann@57418: setprod_reindex_cong ~> setprod.reindex_cong haftmann@57418: setprod_reindex_nonzero ~> setprod.reindex_nontrivial haftmann@57418: setprod_Sigma ~> setprod.Sigma haftmann@57418: setprod_subset_diff ~> setprod.subset_diff haftmann@57418: setprod_timesf ~> setprod.distrib haftmann@57418: setprod_Un2 ~> setprod.union_diff2 haftmann@57418: setprod_Un_disjoint ~> setprod.union_disjoint haftmann@57418: setprod_UN_disjoint ~> setprod.UNION_disjoint haftmann@57418: setprod_Un_Int ~> setprod.union_inter haftmann@57418: setprod_Union_disjoint ~> setprod.Union_disjoint haftmann@57418: setprod_Un_one ~> setprod.union_inter_neutral haftmann@57418: haftmann@57418: Dropped setsum_cong2 (simple variant of setsum.cong). haftmann@57418: Dropped setsum_inter_restrict' (simple variant of setsum.inter_restrict) haftmann@57418: Dropped setsum_reindex_id, setprod_reindex_id haftmann@57418: (simple variants of setsum.reindex [symmetric], setprod.reindex [symmetric]). haftmann@57418: wenzelm@57452: INCOMPATIBILITY. wenzelm@57452: haftmann@54864: * Abolished slightly odd global lattice interpretation for min/max. haftmann@54864: wenzelm@57452: Fact consolidations: haftmann@54864: min_max.inf_assoc ~> min.assoc haftmann@54864: min_max.inf_commute ~> min.commute haftmann@54864: min_max.inf_left_commute ~> min.left_commute haftmann@54864: min_max.inf_idem ~> min.idem haftmann@54864: min_max.inf_left_idem ~> min.left_idem haftmann@54864: min_max.inf_right_idem ~> min.right_idem haftmann@54864: min_max.sup_assoc ~> max.assoc haftmann@54864: min_max.sup_commute ~> max.commute haftmann@54864: min_max.sup_left_commute ~> max.left_commute haftmann@54864: min_max.sup_idem ~> max.idem haftmann@54864: min_max.sup_left_idem ~> max.left_idem haftmann@54864: min_max.sup_inf_distrib1 ~> max_min_distrib2 haftmann@54864: min_max.sup_inf_distrib2 ~> max_min_distrib1 haftmann@54864: min_max.inf_sup_distrib1 ~> min_max_distrib2 haftmann@54864: min_max.inf_sup_distrib2 ~> min_max_distrib1 haftmann@54864: min_max.distrib ~> min_max_distribs haftmann@54864: min_max.inf_absorb1 ~> min.absorb1 haftmann@54864: min_max.inf_absorb2 ~> min.absorb2 haftmann@54864: min_max.sup_absorb1 ~> max.absorb1 haftmann@54864: min_max.sup_absorb2 ~> max.absorb2 haftmann@54864: min_max.le_iff_inf ~> min.absorb_iff1 haftmann@54864: min_max.le_iff_sup ~> max.absorb_iff2 haftmann@54864: min_max.inf_le1 ~> min.cobounded1 haftmann@54864: min_max.inf_le2 ~> min.cobounded2 haftmann@54864: le_maxI1, min_max.sup_ge1 ~> max.cobounded1 haftmann@54864: le_maxI2, min_max.sup_ge2 ~> max.cobounded2 haftmann@54864: min_max.le_infI1 ~> min.coboundedI1 haftmann@54864: min_max.le_infI2 ~> min.coboundedI2 haftmann@54864: min_max.le_supI1 ~> max.coboundedI1 haftmann@54864: min_max.le_supI2 ~> max.coboundedI2 haftmann@54864: min_max.less_infI1 ~> min.strict_coboundedI1 haftmann@54864: min_max.less_infI2 ~> min.strict_coboundedI2 haftmann@54864: min_max.less_supI1 ~> max.strict_coboundedI1 haftmann@54864: min_max.less_supI2 ~> max.strict_coboundedI2 haftmann@54864: min_max.inf_mono ~> min.mono haftmann@54864: min_max.sup_mono ~> max.mono haftmann@54864: min_max.le_infI, min_max.inf_greatest ~> min.boundedI haftmann@54864: min_max.le_supI, min_max.sup_least ~> max.boundedI haftmann@54864: min_max.le_inf_iff ~> min.bounded_iff haftmann@54864: min_max.le_sup_iff ~> max.bounded_iff haftmann@54864: haftmann@54864: For min_max.inf_sup_aci, prefer (one of) min.commute, min.assoc, haftmann@54864: min.left_commute, min.left_idem, max.commute, max.assoc, haftmann@54864: max.left_commute, max.left_idem directly. haftmann@54864: wenzelm@57452: For min_max.inf_sup_ord, prefer (one of) min.cobounded1, wenzelm@57452: min.cobounded2, max.cobounded1m max.cobounded2 directly. haftmann@54864: haftmann@56807: For min_ac or max_ac, prefer more general collection ac_simps. haftmann@54864: wenzelm@58604: INCOMPATIBILITY. haftmann@54864: wenzelm@57452: * Theorem disambiguation Inf_le_Sup (on finite sets) ~> wenzelm@57452: Inf_fin_le_Sup_fin. INCOMPATIBILITY. haftmann@54745: haftmann@54295: * Qualified constant names Wellfounded.acc, Wellfounded.accp. haftmann@54295: INCOMPATIBILITY. haftmann@54295: haftmann@54228: * Fact generalization and consolidation: haftmann@54228: neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1 wenzelm@57452: wenzelm@57452: INCOMPATIBILITY. wenzelm@57452: wenzelm@57452: * Purely algebraic definition of even. Fact generalization and wenzelm@57452: consolidation: haftmann@54228: nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd haftmann@54228: even_zero_(nat|int) ~> even_zero wenzelm@57452: haftmann@54228: INCOMPATIBILITY. wenzelm@54055: haftmann@54489: * Abolished neg_numeral. wenzelm@57452: - Canonical representation for minus one is "- 1". wenzelm@57452: - Canonical representation for other negative numbers is "- (numeral _)". wenzelm@57452: - When devising rule sets for number calculation, consider the haftmann@54587: following canonical cases: 0, 1, numeral _, - 1, - numeral _. wenzelm@57452: - HOLogic.dest_number also recognizes numerals in non-canonical forms wenzelm@54893: like "numeral One", "- numeral One", "- 0" and even "- ... - _". wenzelm@57452: - Syntax for negative numerals is mere input syntax. wenzelm@57452: haftmann@56964: INCOMPATIBILITY. haftmann@54489: wenzelm@57517: * Reduced name variants for rules on associativity and commutativity: wenzelm@57517: wenzelm@57517: add_assoc ~> add.assoc wenzelm@57517: add_commute ~> add.commute wenzelm@57517: add_left_commute ~> add.left_commute wenzelm@57517: mult_assoc ~> mult.assoc wenzelm@57517: mult_commute ~> mult.commute wenzelm@57517: mult_left_commute ~> mult.left_commute wenzelm@57517: nat_add_assoc ~> add.assoc wenzelm@57517: nat_add_commute ~> add.commute wenzelm@57517: nat_add_left_commute ~> add.left_commute wenzelm@57517: nat_mult_assoc ~> mult.assoc wenzelm@57517: nat_mult_commute ~> mult.commute wenzelm@57517: eq_assoc ~> iff_assoc wenzelm@57517: eq_left_commute ~> iff_left_commute wenzelm@57517: wenzelm@57517: INCOMPATIBILITY. wenzelm@57517: wenzelm@57650: * Fact collections add_ac and mult_ac are considered old-fashioned. wenzelm@57637: Prefer ac_simps instead, or specify rules wenzelm@57637: (add|mult).(assoc|commute|left_commute) individually. wenzelm@57517: haftmann@54230: * Elimination of fact duplicates: haftmann@54230: equals_zero_I ~> minus_unique haftmann@54230: diff_eq_0_iff_eq ~> right_minus_eq traytel@54588: nat_infinite ~> infinite_UNIV_nat traytel@54588: int_infinite ~> infinite_UNIV_int wenzelm@57452: haftmann@54230: INCOMPATIBILITY. haftmann@54230: haftmann@54230: * Fact name consolidation: haftmann@54230: diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus haftmann@54250: minus_le_self_iff ~> neg_less_eq_nonneg haftmann@54250: le_minus_self_iff ~> less_eq_neg_nonpos haftmann@54250: neg_less_nonneg ~> neg_less_pos haftmann@54250: less_minus_self_iff ~> less_neg_neg [simp] wenzelm@57452: haftmann@54230: INCOMPATIBILITY. haftmann@54230: haftmann@54230: * More simplification rules on unary and binary minus: haftmann@54230: add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1, haftmann@54230: add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2, haftmann@54230: add_minus_cancel, diff_add_cancel, le_add_same_cancel1, haftmann@54230: le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2, wenzelm@57452: minus_add_cancel, uminus_add_conv_diff. These correspondingly have wenzelm@57452: been taken away from fact collections algebra_simps and field_simps. wenzelm@57452: INCOMPATIBILITY. haftmann@54230: haftmann@54230: To restore proofs, the following patterns are helpful: haftmann@54230: haftmann@54230: a) Arbitrary failing proof not involving "diff_def": haftmann@54230: Consider simplification with algebra_simps or field_simps. haftmann@54230: haftmann@54230: b) Lifting rules from addition to subtraction: wenzelm@54893: Try with "using of [... "- _" ...]" by simp". haftmann@54230: haftmann@54230: c) Simplification with "diff_def": just drop "diff_def". haftmann@54230: Consider simplification with algebra_simps or field_simps; haftmann@54230: or the brute way with haftmann@54230: "simp add: diff_conv_add_uminus del: add_uminus_conv_diff". haftmann@54230: wenzelm@57452: * Introduce bdd_above and bdd_below in theory wenzelm@57452: Conditionally_Complete_Lattices, use them instead of explicitly wenzelm@57452: stating boundedness of sets. wenzelm@57452: wenzelm@57452: * ccpo.admissible quantifies only over non-empty chains to allow more wenzelm@57452: syntax-directed proof rules; the case of the empty chain shows up as wenzelm@57452: additional case in fixpoint induction proofs. INCOMPATIBILITY. hoelzl@54264: hoelzl@56214: * Removed and renamed theorems in Series: hoelzl@56214: summable_le ~> suminf_le hoelzl@56214: suminf_le ~> suminf_le_const hoelzl@56214: series_pos_le ~> setsum_le_suminf hoelzl@56214: series_pos_less ~> setsum_less_suminf hoelzl@56214: suminf_ge_zero ~> suminf_nonneg hoelzl@56214: suminf_gt_zero ~> suminf_pos hoelzl@56214: suminf_gt_zero_iff ~> suminf_pos_iff hoelzl@56214: summable_sumr_LIMSEQ_suminf ~> summable_LIMSEQ hoelzl@56214: suminf_0_le ~> suminf_nonneg [rotate] hoelzl@56214: pos_summable ~> summableI_nonneg_bounded hoelzl@56214: ratio_test ~> summable_ratio_test hoelzl@56214: hoelzl@56214: removed series_zero, replaced by sums_finite hoelzl@56214: hoelzl@56214: removed auxiliary lemmas: wenzelm@57452: hoelzl@56214: sumr_offset, sumr_offset2, sumr_offset3, sumr_offset4, sumr_group, wenzelm@57452: half, le_Suc_ex_iff, lemma_realpow_diff_sumr, wenzelm@57452: real_setsum_nat_ivl_bounded, summable_le2, ratio_test_lemma2, wenzelm@57452: sumr_minus_one_realpow_zerom, sumr_one_lb_realpow_zero, wenzelm@57452: summable_convergent_sumr_iff, sumr_diff_mult_const wenzelm@57452: hoelzl@56214: INCOMPATIBILITY. hoelzl@56214: hoelzl@56214: * Replace (F)DERIV syntax by has_derivative: hoelzl@56214: - "(f has_derivative f') (at x within s)" replaces "FDERIV f x : s : f'" hoelzl@56214: hoelzl@56214: - "(f has_field_derivative f') (at x within s)" replaces "DERIV f x : s : f'" hoelzl@56214: hoelzl@56214: - "f differentiable at x within s" replaces "_ differentiable _ in _" syntax hoelzl@56214: hoelzl@56214: - removed constant isDiff hoelzl@56214: wenzelm@57452: - "DERIV f x : f'" and "FDERIV f x : f'" syntax is only available as wenzelm@57452: input syntax. wenzelm@57452: wenzelm@57452: - "DERIV f x : s : f'" and "FDERIV f x : s : f'" syntax removed. hoelzl@56214: hoelzl@56214: - Renamed FDERIV_... lemmas to has_derivative_... hoelzl@56214: hoelzl@56381: - renamed deriv (the syntax constant used for "DERIV _ _ :> _") to DERIV hoelzl@56381: hoelzl@56381: - removed DERIV_intros, has_derivative_eq_intros hoelzl@56381: wenzelm@57452: - introduced derivative_intros and deriative_eq_intros which wenzelm@57452: includes now rules for DERIV, has_derivative and wenzelm@57452: has_vector_derivative. hoelzl@56381: hoelzl@56214: - Other renamings: hoelzl@56214: differentiable_def ~> real_differentiable_def hoelzl@56214: differentiableE ~> real_differentiableE hoelzl@56214: fderiv_def ~> has_derivative_at hoelzl@56214: field_fderiv_def ~> field_has_derivative_at hoelzl@56214: isDiff_der ~> differentiable_def hoelzl@56214: deriv_fderiv ~> has_field_derivative_def hoelzl@56381: deriv_def ~> DERIV_def wenzelm@57452: wenzelm@57452: INCOMPATIBILITY. wenzelm@57452: wenzelm@57452: * Include more theorems in continuous_intros. Remove the wenzelm@57452: continuous_on_intros, isCont_intros collections, these facts are now wenzelm@57452: in continuous_intros. wenzelm@57452: wenzelm@57452: * Theorems about complex numbers are now stated only using Re and Im, wenzelm@57452: the Complex constructor is not used anymore. It is possible to use wenzelm@57452: primcorec to defined the behaviour of a complex-valued function. wenzelm@57452: wenzelm@57452: Removed theorems about the Complex constructor from the simpset, they wenzelm@57452: are available as the lemma collection legacy_Complex_simps. This wenzelm@57452: especially removes wenzelm@57452: hoelzl@56889: i_complex_of_real: "ii * complex_of_real r = Complex 0 r". hoelzl@56889: wenzelm@57452: Instead the reverse direction is supported with hoelzl@56889: Complex_eq: "Complex a b = a + \ * b" hoelzl@56889: wenzelm@57452: Moved csqrt from Fundamental_Algebra_Theorem to Complex. hoelzl@56889: hoelzl@56889: Renamings: hoelzl@56889: Re/Im ~> complex.sel hoelzl@56889: complex_Re/Im_zero ~> zero_complex.sel hoelzl@56889: complex_Re/Im_add ~> plus_complex.sel hoelzl@56889: complex_Re/Im_minus ~> uminus_complex.sel hoelzl@56889: complex_Re/Im_diff ~> minus_complex.sel hoelzl@56889: complex_Re/Im_one ~> one_complex.sel hoelzl@56889: complex_Re/Im_mult ~> times_complex.sel hoelzl@56889: complex_Re/Im_inverse ~> inverse_complex.sel hoelzl@56889: complex_Re/Im_scaleR ~> scaleR_complex.sel hoelzl@56889: complex_Re/Im_i ~> ii.sel hoelzl@56889: complex_Re/Im_cnj ~> cnj.sel hoelzl@56889: Re/Im_cis ~> cis.sel hoelzl@56889: hoelzl@56889: complex_divide_def ~> divide_complex_def hoelzl@56889: complex_norm_def ~> norm_complex_def hoelzl@56889: cmod_def ~> norm_complex_de hoelzl@56889: hoelzl@56889: Removed theorems: hoelzl@56889: complex_zero_def hoelzl@56889: complex_add_def hoelzl@56889: complex_minus_def hoelzl@56889: complex_diff_def hoelzl@56889: complex_one_def hoelzl@56889: complex_mult_def hoelzl@56889: complex_inverse_def hoelzl@56889: complex_scaleR_def hoelzl@56889: wenzelm@57452: INCOMPATIBILITY. wenzelm@57452: wenzelm@57504: * Theory Lubs moved HOL image to HOL-Library. It is replaced by wenzelm@57504: Conditionally_Complete_Lattices. INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * HOL-Library: new theory src/HOL/Library/Tree.thy. wenzelm@57504: wenzelm@57504: * HOL-Library: removed theory src/HOL/Library/Kleene_Algebra.thy; it wenzelm@57504: is subsumed by session Kleene_Algebra in AFP. wenzelm@57504: wenzelm@57856: * HOL-Library / theory RBT: various constants and facts are hidden; wenzelm@57856: lifting setup is unregistered. INCOMPATIBILITY. kuncar@57826: wenzelm@57504: * HOL-Cardinals: new theory src/HOL/Cardinals/Ordinal_Arithmetic.thy. wenzelm@57504: wenzelm@57504: * HOL-Word: bit representations prefer type bool over type bit. wenzelm@57504: INCOMPATIBILITY. wenzelm@57504: wenzelm@57504: * HOL-Word: wenzelm@57504: - Abandoned fact collection "word_arith_alts", which is a duplicate wenzelm@57504: of "word_arith_wis". wenzelm@57504: - Dropped first (duplicated) element in fact collections wenzelm@57504: "sint_word_ariths", "word_arith_alts", "uint_word_ariths", wenzelm@57504: "uint_word_arith_bintrs". wenzelm@57504: wenzelm@57504: * HOL-Number_Theory: wenzelm@57504: - consolidated the proofs of the binomial theorem wenzelm@57504: - the function fib is again of type nat => nat and not overloaded wenzelm@57504: - no more references to Old_Number_Theory in the HOL libraries wenzelm@57504: (except the AFP) wenzelm@57504: wenzelm@57504: INCOMPATIBILITY. wenzelm@57504: immler@54787: * HOL-Multivariate_Analysis: wenzelm@57452: - Type class ordered_real_vector for ordered vector spaces. wenzelm@57452: - New theory Complex_Basic_Analysis defining complex derivatives, lp15@57253: holomorphic functions, etc., ported from HOL Light's canal.ml. wenzelm@57452: - Changed order of ordered_euclidean_space to be compatible with immler@54787: pointwise ordering on products. Therefore instance of immler@54787: conditionally_complete_lattice and ordered_real_vector. immler@54787: INCOMPATIBILITY: use box instead of greaterThanLessThan or wenzelm@57452: explicit set-comprehensions with eucl_less for other (half-)open immler@54787: intervals. immler@57476: - removed dependencies on type class ordered_euclidean_space with immler@57476: introduction of "cbox" on euclidean_space immler@57476: - renamed theorems: immler@57476: interval ~> box immler@57476: mem_interval ~> mem_box immler@57476: interval_eq_empty ~> box_eq_empty immler@57476: interval_ne_empty ~> box_ne_empty immler@57476: interval_sing(1) ~> cbox_sing immler@57476: interval_sing(2) ~> box_sing immler@57476: subset_interval_imp ~> subset_box_imp immler@57476: subset_interval ~> subset_box immler@57476: open_interval ~> open_box immler@57476: closed_interval ~> closed_cbox immler@57476: interior_closed_interval ~> interior_cbox immler@57476: bounded_closed_interval ~> bounded_cbox immler@57476: compact_interval ~> compact_cbox immler@57476: bounded_subset_closed_interval_symmetric ~> bounded_subset_cbox_symmetric immler@57476: bounded_subset_closed_interval ~> bounded_subset_cbox immler@57476: mem_interval_componentwiseI ~> mem_box_componentwiseI immler@57476: convex_box ~> convex_prod immler@57476: rel_interior_real_interval ~> rel_interior_real_box immler@57476: convex_interval ~> convex_box immler@57476: convex_hull_eq_real_interval ~> convex_hull_eq_real_cbox immler@57476: frechet_derivative_within_closed_interval ~> frechet_derivative_within_cbox immler@57476: content_closed_interval' ~> content_cbox' immler@57476: elementary_subset_interval ~> elementary_subset_box immler@57476: diameter_closed_interval ~> diameter_cbox immler@57476: frontier_closed_interval ~> frontier_cbox immler@57476: frontier_open_interval ~> frontier_box immler@57476: bounded_subset_open_interval_symmetric ~> bounded_subset_box_symmetric immler@57476: closure_open_interval ~> closure_box immler@57476: open_closed_interval_convex ~> open_cbox_convex immler@57476: open_interval_midpoint ~> box_midpoint immler@57476: content_image_affinity_interval ~> content_image_affinity_cbox immler@57476: is_interval_interval ~> is_interval_cbox + is_interval_box + is_interval_closed_interval immler@57476: bounded_interval ~> bounded_closed_interval + bounded_boxes immler@57476: immler@57476: - respective theorems for intervals over the reals: immler@57476: content_closed_interval + content_cbox immler@57476: has_integral + has_integral_real immler@57476: fine_division_exists + fine_division_exists_real immler@57476: has_integral_null + has_integral_null_real immler@57476: tagged_division_union_interval + tagged_division_union_interval_real immler@57476: has_integral_const + has_integral_const_real immler@57476: integral_const + integral_const_real immler@57476: has_integral_bound + has_integral_bound_real immler@57476: integrable_continuous + integrable_continuous_real immler@57476: integrable_subinterval + integrable_subinterval_real immler@57476: has_integral_reflect_lemma + has_integral_reflect_lemma_real immler@57476: integrable_reflect + integrable_reflect_real immler@57476: integral_reflect + integral_reflect_real immler@57476: image_affinity_interval + image_affinity_cbox immler@57476: image_smult_interval + image_smult_cbox immler@57476: integrable_const + integrable_const_ivl immler@57476: integrable_on_subinterval + integrable_on_subcbox immler@57476: hoelzl@56369: - renamed theorems: hoelzl@56369: derivative_linear ~> has_derivative_bounded_linear hoelzl@56369: derivative_is_linear ~> has_derivative_linear hoelzl@56369: bounded_linear_imp_linear ~> bounded_linear.linear hoelzl@56369: hoelzl@56993: * HOL-Probability: hoelzl@57825: - Renamed positive_integral to nn_integral: hoelzl@57825: hoelzl@57825: . Renamed all lemmas "*positive_integral*" to *nn_integral*" hoelzl@57825: positive_integral_positive ~> nn_integral_nonneg hoelzl@57825: hoelzl@57825: . Renamed abbreviation integral\<^sup>P to integral\<^sup>N. hoelzl@57825: wenzelm@57452: - replaced the Lebesgue integral on real numbers by the more general wenzelm@57452: Bochner integral for functions into a real-normed vector space. hoelzl@56993: hoelzl@56993: integral_zero ~> integral_zero / integrable_zero hoelzl@56993: integral_minus ~> integral_minus / integrable_minus hoelzl@56993: integral_add ~> integral_add / integrable_add hoelzl@56993: integral_diff ~> integral_diff / integrable_diff hoelzl@56993: integral_setsum ~> integral_setsum / integrable_setsum hoelzl@56993: integral_multc ~> integral_mult_left / integrable_mult_left hoelzl@56993: integral_cmult ~> integral_mult_right / integrable_mult_right hoelzl@56993: integral_triangle_inequality~> integral_norm_bound hoelzl@56993: integrable_nonneg ~> integrableI_nonneg hoelzl@56993: integral_positive ~> integral_nonneg_AE hoelzl@56993: integrable_abs_iff ~> integrable_abs_cancel hoelzl@57825: positive_integral_lim_INF ~> nn_integral_liminf hoelzl@56993: lebesgue_real_affine ~> lborel_real_affine hoelzl@56993: borel_integral_has_integral ~> has_integral_lebesgue_integral hoelzl@56993: integral_indicator ~> hoelzl@56993: integral_real_indicator / integrable_real_indicator hoelzl@57825: positive_integral_fst ~> nn_integral_fst' hoelzl@57825: positive_integral_fst_measurable ~> nn_integral_fst hoelzl@57825: positive_integral_snd_measurable ~> nn_integral_snd hoelzl@56993: hoelzl@56993: integrable_fst_measurable ~> hoelzl@56993: integral_fst / integrable_fst / AE_integrable_fst hoelzl@56993: hoelzl@56993: integrable_snd_measurable ~> hoelzl@56993: integral_snd / integrable_snd / AE_integrable_snd hoelzl@56993: hoelzl@56993: integral_monotone_convergence ~> hoelzl@56993: integral_monotone_convergence / integrable_monotone_convergence hoelzl@56993: hoelzl@56993: integral_monotone_convergence_at_top ~> hoelzl@56993: integral_monotone_convergence_at_top / hoelzl@56993: integrable_monotone_convergence_at_top hoelzl@56993: hoelzl@56993: has_integral_iff_positive_integral_lebesgue ~> hoelzl@56993: has_integral_iff_has_bochner_integral_lebesgue_nonneg hoelzl@56993: hoelzl@56993: lebesgue_integral_has_integral ~> hoelzl@56993: has_integral_integrable_lebesgue_nonneg hoelzl@56993: hoelzl@56993: positive_integral_lebesgue_has_integral ~> hoelzl@56993: integral_has_integral_lebesgue_nonneg / hoelzl@56993: integrable_has_integral_lebesgue_nonneg hoelzl@56993: hoelzl@56993: lebesgue_integral_real_affine ~> hoelzl@57825: nn_integral_real_affine hoelzl@56993: hoelzl@56993: has_integral_iff_positive_integral_lborel ~> hoelzl@56993: integral_has_integral_nonneg / integrable_has_integral_nonneg hoelzl@56993: hoelzl@56993: The following theorems where removed: hoelzl@56993: hoelzl@56993: lebesgue_integral_nonneg hoelzl@56993: lebesgue_integral_uminus hoelzl@56993: lebesgue_integral_cmult hoelzl@56993: lebesgue_integral_multc hoelzl@56993: lebesgue_integral_cmult_nonneg hoelzl@56993: integral_cmul_indicator hoelzl@56993: integral_real wenzelm@54672: wenzelm@57452: - Formalized properties about exponentially, Erlang, and normal wenzelm@57452: distributed random variables. wenzelm@57452: wenzelm@57504: * HOL-Decision_Procs: Separate command 'approximate' for approximative wenzelm@57504: computation in src/HOL/Decision_Procs/Approximation. Minor wenzelm@57504: INCOMPATIBILITY. wenzelm@57452: nipkow@57112: wenzelm@55622: *** Scala *** wenzelm@55622: wenzelm@55622: * The signature and semantics of Document.Snapshot.cumulate_markup / wenzelm@55622: select_markup have been clarified. Markup is now traversed in the wenzelm@55622: order of reports given by the prover: later markup is usually more wenzelm@55622: specific and may override results accumulated so far. The elements wenzelm@55622: guard is mandatory and checked precisely. Subtle INCOMPATIBILITY. wenzelm@55622: wenzelm@57452: * Substantial reworking of internal PIDE protocol communication wenzelm@57452: channels. INCOMPATIBILITY. wenzelm@57452: wenzelm@55622: wenzelm@54449: *** ML *** wenzelm@54449: wenzelm@57504: * Subtle change of semantics of Thm.eq_thm: theory stamps are not wenzelm@57504: compared (according to Thm.thm_ord), but assumed to be covered by the wenzelm@57504: current background theory. Thus equivalent data produced in different wenzelm@57504: branches of the theory graph usually coincides (e.g. relevant for wenzelm@57504: theory merge). Note that the softer Thm.eq_thm_prop is often more wenzelm@57504: appropriate than Thm.eq_thm. wenzelm@57504: wenzelm@57504: * Proper context for basic Simplifier operations: rewrite_rule, wenzelm@57504: rewrite_goals_rule, rewrite_goals_tac etc. INCOMPATIBILITY, need to wenzelm@57504: pass runtime Proof.context (and ensure that the simplified entity wenzelm@57504: actually belongs to it). wenzelm@57504: wenzelm@57504: * Proper context discipline for read_instantiate and instantiate_tac: wenzelm@57504: variables that are meant to become schematic need to be given as wenzelm@57504: fixed, and are generalized by the explicit context of local variables. wenzelm@57504: This corresponds to Isar attributes "where" and "of" with 'for' wenzelm@57504: declaration. INCOMPATIBILITY, also due to potential change of indices wenzelm@57504: of schematic variables. wenzelm@57504: wenzelm@56303: * Moved ML_Compiler.exn_trace and other operations on exceptions to wenzelm@56303: structure Runtime. Minor INCOMPATIBILITY. wenzelm@56303: wenzelm@56279: * Discontinued old Toplevel.debug in favour of system option wenzelm@57452: "ML_exception_trace", which may be also declared within the context wenzelm@57452: via "declare [[ML_exception_trace = true]]". Minor INCOMPATIBILITY. wenzelm@56279: wenzelm@56281: * Renamed configuration option "ML_trace" to "ML_source_trace". Minor wenzelm@56281: INCOMPATIBILITY. wenzelm@56281: wenzelm@56281: * Configuration option "ML_print_depth" controls the pretty-printing wenzelm@56281: depth of the ML compiler within the context. The old print_depth in wenzelm@56285: ML is still available as default_print_depth, but rarely used. Minor wenzelm@56285: INCOMPATIBILITY. wenzelm@56279: wenzelm@54449: * Toplevel function "use" refers to raw ML bootstrap environment, wenzelm@54449: without Isar context nor antiquotations. Potential INCOMPATIBILITY. wenzelm@54449: Note that 'ML_file' is the canonical command to load ML files into the wenzelm@54449: formal context. wenzelm@54449: wenzelm@56205: * Simplified programming interface to define ML antiquotations, see wenzelm@56205: structure ML_Antiquotation. Minor INCOMPATIBILITY. wenzelm@56069: wenzelm@56071: * ML antiquotation @{here} refers to its source position, which is wenzelm@56071: occasionally useful for experimentation and diagnostic purposes. wenzelm@56071: wenzelm@56135: * ML antiquotation @{path} produces a Path.T value, similarly to wenzelm@56135: Path.explode, but with compile-time check against the file-system and wenzelm@56135: some PIDE markup. Note that unlike theory source, ML does not have a wenzelm@56135: well-defined master directory, so an absolute symbolic path wenzelm@56135: specification is usually required, e.g. "~~/src/HOL". wenzelm@56135: wenzelm@56399: * ML antiquotation @{print} inlines a function to print an arbitrary wenzelm@56399: ML value, which is occasionally useful for diagnostic or demonstration wenzelm@56399: purposes. wenzelm@56399: wenzelm@54449: wenzelm@54683: *** System *** wenzelm@54683: wenzelm@57443: * Proof General with its traditional helper scripts is now an optional wenzelm@57504: Isabelle component, e.g. see ProofGeneral-4.2-2 from the Isabelle wenzelm@57504: component repository http://isabelle.in.tum.de/components/. Note that wenzelm@57504: the "system" manual provides general explanations about add-on wenzelm@57504: components, especially those that are not bundled with the release. wenzelm@57443: wenzelm@56439: * The raw Isabelle process executable has been renamed from wenzelm@56439: "isabelle-process" to "isabelle_process", which conforms to common wenzelm@56439: shell naming conventions, and allows to define a shell function within wenzelm@56439: the Isabelle environment to avoid dynamic path lookup. Rare wenzelm@57504: incompatibility for old tools that do not use the ISABELLE_PROCESS wenzelm@57504: settings variable. wenzelm@56439: wenzelm@57439: * Former "isabelle tty" has been superseded by "isabelle console", wenzelm@57439: with implicit build like "isabelle jedit", and without the mostly wenzelm@57439: obsolete Isar TTY loop. wenzelm@57439: wenzelm@57452: * Simplified "isabelle display" tool. Settings variables DVI_VIEWER wenzelm@57452: and PDF_VIEWER now refer to the actual programs, not shell wenzelm@57452: command-lines. Discontinued option -c: invocation may be asynchronous wenzelm@57452: via desktop environment, without any special precautions. Potential wenzelm@57452: INCOMPATIBILITY with ambitious private settings. wenzelm@57452: wenzelm@57413: * Removed obsolete "isabelle unsymbolize". Note that the usual format wenzelm@57413: for email communication is the Unicode rendering of Isabelle symbols, wenzelm@57413: as produced by Isabelle/jEdit, for example. wenzelm@57413: wenzelm@57452: * Removed obsolete tool "wwwfind". Similar functionality may be wenzelm@57452: integrated into Isabelle/jEdit eventually. wenzelm@57452: wenzelm@57452: * Improved 'display_drafts' concerning desktop integration and wenzelm@57452: repeated invocation in PIDE front-end: re-use single file wenzelm@57452: $ISABELLE_HOME_USER/tmp/drafts.pdf and corresponding views. wenzelm@57084: wenzelm@57452: * Session ROOT specifications require explicit 'document_files' for wenzelm@57452: robust dependencies on LaTeX sources. Only these explicitly given wenzelm@57452: files are copied to the document output directory, before document wenzelm@57452: processing is started. wenzelm@57452: wenzelm@57504: * Windows: support for regular TeX installation (e.g. MiKTeX) instead wenzelm@57504: of TeX Live from Cygwin. wenzelm@57504: wenzelm@57084: wenzelm@57693: wenzelm@54639: New in Isabelle2013-2 (December 2013) wenzelm@54639: ------------------------------------- wenzelm@54639: wenzelm@54639: *** Prover IDE -- Isabelle/Scala/jEdit *** wenzelm@54639: wenzelm@54639: * More robust editing of running commands with internal forks, wenzelm@54639: e.g. non-terminating 'by' steps. wenzelm@54639: wenzelm@54641: * More relaxed Sledgehammer panel: avoid repeated application of query wenzelm@54641: after edits surrounding the command location. wenzelm@54641: wenzelm@54648: * More status information about commands that are interrupted wenzelm@54648: accidentally (via physical event or Poly/ML runtime system signal, wenzelm@54648: e.g. out-of-memory). wenzelm@54648: wenzelm@54653: wenzelm@54653: *** System *** wenzelm@54653: wenzelm@54653: * More robust termination of external processes managed by wenzelm@54664: Isabelle/ML: support cancellation of tasks within the range of wenzelm@54664: milliseconds, as required for PIDE document editing with automatically wenzelm@54664: tried tools (e.g. Sledgehammer). wenzelm@54653: wenzelm@54648: * Reactivated Isabelle/Scala kill command for external processes on wenzelm@54648: Mac OS X, which was accidentally broken in Isabelle2013-1 due to a wenzelm@54648: workaround for some Debian/Ubuntu Linux versions from 2013. wenzelm@54648: wenzelm@54639: wenzelm@54639: wenzelm@53971: New in Isabelle2013-1 (November 2013) wenzelm@53971: ------------------------------------- wenzelm@50994: wenzelm@51293: *** General *** wenzelm@51293: wenzelm@53971: * Discontinued obsolete 'uses' within theory header. Note that wenzelm@53971: commands like 'ML_file' work without separate declaration of file wenzelm@53971: dependencies. Minor INCOMPATIBILITY. wenzelm@53971: wenzelm@53971: * Discontinued redundant 'use' command, which was superseded by wenzelm@53971: 'ML_file' in Isabelle2013. Minor INCOMPATIBILITY. wenzelm@53971: wenzelm@53016: * Simplified subscripts within identifiers, using plain \<^sub> wenzelm@53016: instead of the second copy \<^isub> and \<^isup>. Superscripts are wenzelm@53016: only for literal tokens within notation; explicit mixfix annotations wenzelm@53016: for consts or fixed variables may be used as fall-back for unusual wenzelm@53016: names. Obsolete \ has been expanded to \<^sup>2 in wenzelm@53016: Isabelle/HOL. INCOMPATIBILITY, use "isabelle update_sub_sup" to wenzelm@53016: standardize symbols as a starting point for further manual cleanup. wenzelm@53016: The ML reference variable "legacy_isub_isup" may be set as temporary wenzelm@53016: workaround, to make the prover accept a subset of the old identifier wenzelm@53016: syntax. wenzelm@53016: wenzelm@53021: * Document antiquotations: term style "isub" has been renamed to wenzelm@53021: "sub". Minor INCOMPATIBILITY. wenzelm@53021: wenzelm@52487: * Uniform management of "quick_and_dirty" as system option (see also wenzelm@52487: "isabelle options"), configuration option within the context (see also wenzelm@52487: Config.get in Isabelle/ML), and attribute in Isabelle/Isar. Minor wenzelm@52487: INCOMPATIBILITY, need to use more official Isabelle means to access wenzelm@52487: quick_and_dirty, instead of historical poking into mutable reference. wenzelm@52059: wenzelm@52060: * Renamed command 'print_configs' to 'print_options'. Minor wenzelm@52060: INCOMPATIBILITY. wenzelm@52060: wenzelm@52430: * Proper diagnostic command 'print_state'. Old 'pr' (with its wenzelm@52430: implicit change of some global references) is retained for now as wenzelm@52430: control command, e.g. for ProofGeneral 3.7.x. wenzelm@52430: wenzelm@52549: * Discontinued 'print_drafts' command with its old-fashioned PS output wenzelm@52549: and Unix command-line print spooling. Minor INCOMPATIBILITY: use wenzelm@52549: 'display_drafts' instead and print via the regular document viewer. wenzelm@52549: wenzelm@53971: * Updated and extended "isar-ref" and "implementation" manual, wenzelm@53971: eliminated old "ref" manual. wenzelm@53971: wenzelm@51293: wenzelm@51533: *** Prover IDE -- Isabelle/Scala/jEdit *** wenzelm@51533: wenzelm@53971: * New manual "jedit" for Isabelle/jEdit, see isabelle doc or wenzelm@53852: Documentation panel. wenzelm@53852: wenzelm@53971: * Dockable window "Documentation" provides access to Isabelle wenzelm@53971: documentation. wenzelm@52646: wenzelm@52949: * Dockable window "Find" provides query operations for formal entities wenzelm@52949: (GUI front-end to 'find_theorems' command). wenzelm@52949: wenzelm@53050: * Dockable window "Sledgehammer" manages asynchronous / parallel wenzelm@53050: sledgehammer runs over existing document sources, independently of wenzelm@53050: normal editing and checking process. wenzelm@53050: wenzelm@51533: * Dockable window "Timing" provides an overview of relevant command wenzelm@54332: timing information, depending on option jedit_timing_threshold. The wenzelm@54332: same timing information is shown in the extended tooltip of the wenzelm@54332: command keyword, when hovering the mouse over it while the CONTROL or wenzelm@54332: COMMAND modifier is pressed. wenzelm@51533: wenzelm@53971: * Improved dockable window "Theories": Continuous checking of proof wenzelm@53971: document (visible and required parts) may be controlled explicitly, wenzelm@53971: using check box or shortcut "C+e ENTER". Individual theory nodes may wenzelm@53971: be marked explicitly as required and checked in full, using check box wenzelm@53971: or shortcut "C+e SPACE". wenzelm@53971: wenzelm@54305: * Improved completion mechanism, which is now managed by the wenzelm@54305: Isabelle/jEdit plugin instead of SideKick. Refined table of Isabelle wenzelm@54305: symbol abbreviations (see $ISABELLE_HOME/etc/symbols). wenzelm@54305: wenzelm@54319: * Standard jEdit keyboard shortcut C+b complete-word is remapped to wenzelm@54319: isabelle.complete for explicit completion in Isabelle sources. wenzelm@54319: INCOMPATIBILITY wrt. jEdit defaults, may have to invent new shortcuts wenzelm@54319: to resolve conflict. wenzelm@54319: wenzelm@54305: * Improved support of various "minor modes" for Isabelle NEWS, wenzelm@54305: options, session ROOT etc., with completion and SideKick tree view. wenzelm@54305: wenzelm@53971: * Strictly monotonic document update, without premature cancellation of wenzelm@53971: running transactions that are still needed: avoid reset/restart of wenzelm@53971: such command executions while editing. wenzelm@53271: wenzelm@53971: * Support for asynchronous print functions, as overlay to existing wenzelm@53971: document content. wenzelm@53971: wenzelm@53971: * Support for automatic tools in HOL, which try to prove or disprove wenzelm@53971: toplevel theorem statements. wenzelm@53971: wenzelm@53971: * Action isabelle.reset-font-size resets main text area font size wenzelm@54365: according to Isabelle/Scala plugin option "jedit_font_reset_size" (see wenzelm@54365: also "Plugin Options / Isabelle / General"). It can be bound to some wenzelm@54365: keyboard shortcut by the user (e.g. C+0 and/or C+NUMPAD0). wenzelm@53971: wenzelm@53971: * File specifications in jEdit (e.g. file browser) may refer to wenzelm@54351: $ISABELLE_HOME and $ISABELLE_HOME_USER on all platforms. Discontinued wenzelm@54351: obsolete $ISABELLE_HOME_WINDOWS variable. wenzelm@53971: wenzelm@53971: * Improved support for Linux look-and-feel "GTK+", see also "Utilities wenzelm@53971: / Global Options / Appearance". wenzelm@53971: wenzelm@53971: * Improved support of native Mac OS X functionality via "MacOSX" wenzelm@53971: plugin, which is now enabled by default. wenzelm@53971: wenzelm@51533: wenzelm@51313: *** Pure *** wenzelm@51313: ballarin@54049: * Commands 'interpretation' and 'sublocale' are now target-sensitive. ballarin@54049: In particular, 'interpretation' allows for non-persistent ballarin@54049: interpretation within "context ... begin ... end" blocks offering a ballarin@54049: light-weight alternative to 'sublocale'. See "isar-ref" manual for ballarin@54049: details. haftmann@51747: ballarin@51565: * Improved locales diagnostic command 'print_dependencies'. ballarin@51565: wenzelm@51313: * Discontinued obsolete 'axioms' command, which has been marked as wenzelm@51313: legacy since Isabelle2009-2. INCOMPATIBILITY, use 'axiomatization' wenzelm@51313: instead, while observing its uniform scope for polymorphism. wenzelm@51313: wenzelm@51316: * Discontinued empty name bindings in 'axiomatization'. wenzelm@51316: INCOMPATIBILITY. wenzelm@51316: wenzelm@53971: * System option "proofs" has been discontinued. Instead the global wenzelm@53971: state of Proofterm.proofs is persistently compiled into logic images wenzelm@53971: as required, notably HOL-Proofs. Users no longer need to change wenzelm@53971: Proofterm.proofs dynamically. Minor INCOMPATIBILITY. wenzelm@53971: wenzelm@53971: * Syntax translation functions (print_translation etc.) always depend wenzelm@53971: on Proof.context. Discontinued former "(advanced)" option -- this is wenzelm@53971: now the default. Minor INCOMPATIBILITY. wenzelm@53971: wenzelm@53971: * Former global reference trace_unify_fail is now available as wenzelm@53971: configuration option "unify_trace_failure" (global context only). wenzelm@53971: wenzelm@52463: * SELECT_GOAL now retains the syntactic context of the overall goal wenzelm@52463: state (schematic variables etc.). Potential INCOMPATIBILITY in rare wenzelm@52463: situations. wenzelm@52463: wenzelm@51313: hoelzl@51002: *** HOL *** hoelzl@51002: wenzelm@54032: * Stronger precedence of syntax for big intersection and union on wenzelm@54032: sets, in accordance with corresponding lattice operations. wenzelm@54032: INCOMPATIBILITY. wenzelm@54032: wenzelm@54032: * Notation "{p:A. P}" now allows tuple patterns as well. wenzelm@54032: wenzelm@54032: * Nested case expressions are now translated in a separate check phase wenzelm@54032: rather than during parsing. The data for case combinators is separated wenzelm@54032: from the datatype package. The declaration attribute wenzelm@54032: "case_translation" can be used to register new case combinators: wenzelm@54032: wenzelm@54032: declare [[case_translation case_combinator constructor1 ... constructorN]] haftmann@52637: haftmann@52435: * Code generator: wenzelm@53160: - 'code_printing' unifies 'code_const' / 'code_type' / 'code_class' / wenzelm@53160: 'code_instance'. wenzelm@53160: - 'code_identifier' declares name hints for arbitrary identifiers in wenzelm@53160: generated code, subsuming 'code_modulename'. wenzelm@53983: wenzelm@53983: See the isar-ref manual for syntax diagrams, and the HOL theories for wenzelm@53983: examples. haftmann@52435: wenzelm@54032: * Attibute 'code': 'code' now declares concrete and abstract code wenzelm@54032: equations uniformly. Use explicit 'code equation' and 'code abstract' wenzelm@54032: to distinguish both when desired. wenzelm@54032: wenzelm@54032: * Discontinued theories Code_Integer and Efficient_Nat by a more wenzelm@54032: fine-grain stack of theories Code_Target_Int, Code_Binary_Nat, wenzelm@54032: Code_Target_Nat and Code_Target_Numeral. See the tutorial on code wenzelm@54032: generation for details. INCOMPATIBILITY. wenzelm@54032: wenzelm@54032: * Numeric types are mapped by default to target language numerals: wenzelm@54032: natural (replaces former code_numeral) and integer (replaces former wenzelm@54032: code_int). Conversions are available as integer_of_natural / wenzelm@54032: natural_of_integer / integer_of_nat / nat_of_integer (in HOL) and wenzelm@54032: Code_Numeral.integer_of_natural / Code_Numeral.natural_of_integer (in wenzelm@54032: ML). INCOMPATIBILITY. wenzelm@54032: wenzelm@54032: * Function package: For mutually recursive functions f and g, separate wenzelm@54032: cases rules f.cases and g.cases are generated instead of unusable wenzelm@54032: f_g.cases which exposed internal sum types. Potential INCOMPATIBILITY, wenzelm@54032: in the case that the unusable rule was used nevertheless. wenzelm@54032: wenzelm@54032: * Function package: For each function f, new rules f.elims are wenzelm@54032: generated, which eliminate equalities of the form "f x = t". wenzelm@54032: wenzelm@54032: * New command 'fun_cases' derives ad-hoc elimination rules for wenzelm@54032: function equations as simplified instances of f.elims, analogous to wenzelm@54032: inductive_cases. See ~~/src/HOL/ex/Fundefs.thy for some examples. blanchet@53307: kuncar@54021: * Lifting: kuncar@54021: - parametrized correspondence relations are now supported: wenzelm@54378: + parametricity theorems for the raw term can be specified in kuncar@54021: the command lift_definition, which allow us to generate stronger kuncar@54021: transfer rules kuncar@54021: + setup_lifting generates stronger transfer rules if parametric kuncar@54021: correspondence relation can be generated kuncar@54021: + various new properties of the relator must be specified to support kuncar@54021: parametricity kuncar@54021: + parametricity theorem for the Quotient relation can be specified kuncar@54021: - setup_lifting generates domain rules for the Transfer package kuncar@54021: - stronger reflexivity prover of respectfulness theorems for type kuncar@54021: copies kuncar@54021: - ===> and --> are now local. The symbols can be introduced kuncar@54021: by interpreting the locale lifting_syntax (typically in an kuncar@54021: anonymous context) wenzelm@54378: - Lifting/Transfer relevant parts of Library/Quotient_* are now in kuncar@54021: Main. Potential INCOMPATIBILITY kuncar@54021: - new commands for restoring and deleting Lifting/Transfer context: kuncar@54021: lifting_forget, lifting_update wenzelm@54378: - the command print_quotmaps was renamed to print_quot_maps. kuncar@54021: INCOMPATIBILITY kuncar@54021: kuncar@54021: * Transfer: wenzelm@54378: - better support for domains in Transfer: replace Domainp T kuncar@54021: by the actual invariant in a transferred goal kuncar@54021: - transfer rules can have as assumptions other transfer rules kuncar@54021: - Experimental support for transferring from the raw level to the kuncar@54021: abstract level: Transfer.transferred attribute kuncar@54021: - Attribute version of the transfer method: untransferred attribute kuncar@54021: haftmann@52286: * Reification and reflection: wenzelm@53160: - Reification is now directly available in HOL-Main in structure wenzelm@53160: "Reification". wenzelm@53160: - Reflection now handles multiple lists with variables also. wenzelm@53160: - The whole reflection stack has been decomposed into conversions. haftmann@52286: INCOMPATIBILITY. haftmann@52286: haftmann@51489: * Revised devices for recursive definitions over finite sets: haftmann@51489: - Only one fundamental fold combinator on finite set remains: haftmann@51489: Finite_Set.fold :: ('a => 'b => 'b) => 'b => 'a set => 'b haftmann@51489: This is now identity on infinite sets. wenzelm@52745: - Locales ("mini packages") for fundamental definitions with haftmann@51489: Finite_Set.fold: folding, folding_idem. haftmann@51489: - Locales comm_monoid_set, semilattice_order_set and haftmann@51489: semilattice_neutr_order_set for big operators on sets. haftmann@51489: See theory Big_Operators for canonical examples. haftmann@51489: Note that foundational constants comm_monoid_set.F and haftmann@51489: semilattice_set.F correspond to former combinators fold_image haftmann@51489: and fold1 respectively. These are now gone. You may use haftmann@51490: those foundational constants as substitutes, but it is wenzelm@53983: preferable to interpret the above locales accordingly. haftmann@51489: - Dropped class ab_semigroup_idem_mult (special case of lattice, haftmann@51489: no longer needed in connection with Finite_Set.fold etc.) haftmann@51489: - Fact renames: haftmann@51489: card.union_inter ~> card_Un_Int [symmetric] haftmann@51489: card.union_disjoint ~> card_Un_disjoint haftmann@51489: INCOMPATIBILITY. haftmann@51489: haftmann@51487: * Locale hierarchy for abstract orderings and (semi)lattices. haftmann@51487: wenzelm@53526: * Complete_Partial_Order.admissible is defined outside the type class wenzelm@53526: ccpo, but with mandatory prefix ccpo. Admissibility theorems lose the wenzelm@53526: class predicate assumption or sort constraint when possible. Andreas@53362: INCOMPATIBILITY. Andreas@53362: wenzelm@53160: * Introduce type class "conditionally_complete_lattice": Like a wenzelm@53160: complete lattice but does not assume the existence of the top and wenzelm@53160: bottom elements. Allows to generalize some lemmas about reals and wenzelm@53160: extended reals. Removed SupInf and replaced it by the instantiation wenzelm@53160: of conditionally_complete_lattice for real. Renamed lemmas about wenzelm@53160: conditionally-complete lattice from Sup_... to cSup_... and from wenzelm@53160: Inf_... to cInf_... to avoid hidding of similar complete lattice wenzelm@53160: lemmas. wenzelm@53160: wenzelm@53160: * Introduce type class linear_continuum as combination of wenzelm@53160: conditionally-complete lattices and inner dense linorders which have wenzelm@53160: more than one element. INCOMPATIBILITY. wenzelm@53160: wenzelm@53983: * Introduced type classes order_top and order_bot. The old classes top wenzelm@53983: and bot only contain the syntax without assumptions. INCOMPATIBILITY: wenzelm@53983: Rename bot -> order_bot, top -> order_top lammich@53683: wenzelm@53160: * Introduce type classes "no_top" and "no_bot" for orderings without wenzelm@53160: top and bottom elements. hoelzl@51732: hoelzl@51732: * Split dense_linorder into inner_dense_order and no_top, no_bot. hoelzl@51732: hoelzl@51732: * Complex_Main: Unify and move various concepts from wenzelm@53160: HOL-Multivariate_Analysis to HOL-Complex_Main. hoelzl@51732: wenzelm@53983: - Introduce type class (lin)order_topology and wenzelm@53983: linear_continuum_topology. Allows to generalize theorems about wenzelm@53983: limits and order. Instances are reals and extended reals. hoelzl@51732: hoelzl@51732: - continuous and continuos_on from Multivariate_Analysis: wenzelm@53983: "continuous" is the continuity of a function at a filter. "isCont" wenzelm@53983: is now an abbrevitation: "isCont x f == continuous (at _) f". wenzelm@53983: wenzelm@53983: Generalized continuity lemmas from isCont to continuous on an wenzelm@53983: arbitrary filter. wenzelm@53983: wenzelm@53983: - compact from Multivariate_Analysis. Use Bolzano's lemma to prove wenzelm@53983: compactness of closed intervals on reals. Continuous functions wenzelm@53983: attain infimum and supremum on compact sets. The inverse of a wenzelm@53983: continuous function is continuous, when the function is continuous wenzelm@53983: on a compact set. hoelzl@51732: hoelzl@51732: - connected from Multivariate_Analysis. Use it to prove the hoelzl@51775: intermediate value theorem. Show connectedness of intervals on hoelzl@51775: linear_continuum_topology). hoelzl@51732: hoelzl@51732: - first_countable_topology from Multivariate_Analysis. Is used to wenzelm@53983: show equivalence of properties on the neighbourhood filter of x and wenzelm@53983: on all sequences converging to x. wenzelm@53983: wenzelm@53983: - FDERIV: Definition of has_derivative moved to Deriv.thy. Moved wenzelm@53983: theorems from Library/FDERIV.thy to Deriv.thy and base the wenzelm@53983: definition of DERIV on FDERIV. Add variants of DERIV and FDERIV wenzelm@53983: which are restricted to sets, i.e. to represent derivatives from wenzelm@53983: left or right. hoelzl@51732: hoelzl@51732: - Removed the within-filter. It is replaced by the principal filter: hoelzl@51732: hoelzl@51732: F within X = inf F (principal X) hoelzl@51732: hoelzl@51732: - Introduce "at x within U" as a single constant, "at x" is now an hoelzl@51732: abbreviation for "at x within UNIV" hoelzl@51732: wenzelm@53983: - Introduce named theorem collections tendsto_intros, wenzelm@53983: continuous_intros, continuous_on_intros and FDERIV_intros. Theorems wenzelm@53983: in tendsto_intros (or FDERIV_intros) are also available as wenzelm@53983: tendsto_eq_intros (or FDERIV_eq_intros) where the right-hand side wenzelm@53983: is replaced by a congruence rule. This allows to apply them as wenzelm@53983: intro rules and then proving equivalence by the simplifier. hoelzl@51732: hoelzl@51732: - Restructured theories in HOL-Complex_Main: hoelzl@51732: hoelzl@51732: + Moved RealDef and RComplete into Real hoelzl@51732: hoelzl@51732: + Introduced Topological_Spaces and moved theorems about hoelzl@51732: topological spaces, filters, limits and continuity to it hoelzl@51732: hoelzl@51732: + Renamed RealVector to Real_Vector_Spaces hoelzl@51732: wenzelm@53983: + Split Lim, SEQ, Series into Topological_Spaces, wenzelm@53983: Real_Vector_Spaces, and Limits hoelzl@51732: hoelzl@51732: + Moved Ln and Log to Transcendental hoelzl@51732: hoelzl@51732: + Moved theorems about continuity from Deriv to Topological_Spaces hoelzl@51732: hoelzl@51732: - Remove various auxiliary lemmas. hoelzl@51732: hoelzl@51732: INCOMPATIBILITY. hoelzl@51002: blanchet@53738: * Nitpick: blanchet@55889: - Added option "spy". blanchet@55889: - Reduce incidence of "too high arity" errors. blanchet@53738: blanchet@51137: * Sledgehammer: blanchet@51137: - Renamed option: blanchet@51137: isar_shrink ~> isar_compress blanchet@53738: INCOMPATIBILITY. blanchet@55889: - Added options "isar_try0", "spy". blanchet@55889: - Better support for "isar_proofs". blanchet@55889: - MaSh has been fined-tuned and now runs as a local server. blanchet@51137: wenzelm@54032: * Improved support for ad hoc overloading of constants (see also wenzelm@54032: isar-ref manual and ~~/src/HOL/ex/Adhoc_Overloading_Examples.thy). wenzelm@54032: wenzelm@54032: * Library/Polynomial.thy: wenzelm@54032: - Use lifting for primitive definitions. wenzelm@54032: - Explicit conversions from and to lists of coefficients, used for wenzelm@54032: generated code. wenzelm@54032: - Replaced recursion operator poly_rec by fold_coeffs. wenzelm@54032: - Prefer pre-existing gcd operation for gcd. wenzelm@54032: - Fact renames: wenzelm@54032: poly_eq_iff ~> poly_eq_poly_eq_iff wenzelm@54032: poly_ext ~> poly_eqI wenzelm@54032: expand_poly_eq ~> poly_eq_iff wenzelm@54032: IMCOMPATIBILITY. wenzelm@54032: wenzelm@54032: * New Library/Simps_Case_Conv.thy: Provides commands simps_of_case and wenzelm@54032: case_of_simps to convert function definitions between a list of wenzelm@54032: equations with patterns on the lhs and a single equation with case wenzelm@54032: expressions on the rhs. See also Ex/Simps_Case_Conv_Examples.thy. wenzelm@54032: wenzelm@54032: * New Library/FSet.thy: type of finite sets defined as a subtype of wenzelm@54032: sets defined by Lifting/Transfer. wenzelm@54032: wenzelm@54032: * Discontinued theory src/HOL/Library/Eval_Witness. INCOMPATIBILITY. wenzelm@54032: wenzelm@54032: * Consolidation of library theories on product orders: wenzelm@54032: wenzelm@54032: Product_Lattice ~> Product_Order -- pointwise order on products wenzelm@54032: Product_ord ~> Product_Lexorder -- lexicographic order on products wenzelm@54032: wenzelm@54032: INCOMPATIBILITY. wenzelm@54032: wenzelm@53160: * Imperative-HOL: The MREC combinator is considered legacy and no wenzelm@53160: longer included by default. INCOMPATIBILITY, use partial_function wenzelm@53160: instead, or import theory Legacy_Mrec as a fallback. wenzelm@53160: wenzelm@53983: * HOL-Algebra: Discontinued theories ~~/src/HOL/Algebra/abstract and wenzelm@53983: ~~/src/HOL/Algebra/poly. Existing theories should be based on wenzelm@53983: ~~/src/HOL/Library/Polynomial instead. The latter provides wenzelm@53983: integration with HOL's type classes for rings. INCOMPATIBILITY. ballarin@51517: wenzelm@54033: * HOL-BNF: wenzelm@54032: - Various improvements to BNF-based (co)datatype package, including wenzelm@54032: new commands "primrec_new", "primcorec", and wenzelm@54032: "datatype_new_compat", as well as documentation. See wenzelm@54032: "datatypes.pdf" for details. wenzelm@54032: - New "coinduction" method to avoid some boilerplate (compared to wenzelm@54032: coinduct). wenzelm@54032: - Renamed keywords: wenzelm@54032: data ~> datatype_new wenzelm@54032: codata ~> codatatype wenzelm@54032: bnf_def ~> bnf wenzelm@54032: - Renamed many generated theorems, including wenzelm@54032: discs ~> disc wenzelm@54032: map_comp' ~> map_comp wenzelm@54032: map_id' ~> map_id wenzelm@54032: sels ~> sel wenzelm@54032: set_map' ~> set_map wenzelm@54032: sets ~> set wenzelm@54032: IMCOMPATIBILITY. wenzelm@54032: ballarin@51517: wenzelm@51551: *** ML *** wenzelm@51551: wenzelm@53971: * Spec_Check is a Quickcheck tool for Isabelle/ML. The ML function wenzelm@53971: "check_property" allows to check specifications of the form "ALL x y wenzelm@53971: z. prop x y z". See also ~~/src/Tools/Spec_Check/ with its wenzelm@53971: Examples.thy in particular. wenzelm@53971: wenzelm@53709: * Improved printing of exception trace in Poly/ML 5.5.1, with regular wenzelm@53709: tracing output in the command transaction context instead of physical wenzelm@53709: stdout. See also Toplevel.debug, Toplevel.debugging and wenzelm@53709: ML_Compiler.exn_trace. wenzelm@53709: wenzelm@53971: * ML type "theory" is now immutable, without any special treatment of wenzelm@53971: drafts or linear updates (which could lead to "stale theory" errors in wenzelm@53971: the past). Discontinued obsolete operations like Theory.copy, wenzelm@53971: Theory.checkpoint, and the auxiliary type theory_ref. Minor wenzelm@53971: INCOMPATIBILITY. wenzelm@53164: wenzelm@51551: * More uniform naming of goal functions for skipped proofs: wenzelm@51551: wenzelm@51551: Skip_Proof.prove ~> Goal.prove_sorry wenzelm@51551: Skip_Proof.prove_global ~> Goal.prove_sorry_global wenzelm@51551: wenzelm@53971: Minor INCOMPATIBILITY. wenzelm@51703: wenzelm@51717: * Simplifier tactics and tools use proper Proof.context instead of wenzelm@51717: historic type simpset. Old-style declarations like addsimps, wenzelm@51717: addsimprocs etc. operate directly on Proof.context. Raw type simpset wenzelm@51717: retains its use as snapshot of the main Simplifier context, using wenzelm@51717: simpset_of and put_simpset on Proof.context. INCOMPATIBILITY -- port wenzelm@51717: old tools by making them depend on (ctxt : Proof.context) instead of wenzelm@51717: (ss : simpset), then turn (simpset_of ctxt) into ctxt. wenzelm@51717: wenzelm@53971: * Modifiers for classical wrappers (e.g. addWrapper, delWrapper) wenzelm@53971: operate on Proof.context instead of claset, for uniformity with addIs, wenzelm@53971: addEs, addDs etc. Note that claset_of and put_claset allow to manage wenzelm@53971: clasets separately from the context. wenzelm@53971: wenzelm@51717: * Discontinued obsolete ML antiquotations @{claset} and @{simpset}. wenzelm@51717: INCOMPATIBILITY, use @{context} instead. wenzelm@51717: wenzelm@53971: * Antiquotation @{theory_context A} is similar to @{theory A}, but wenzelm@53971: presents the result as initial Proof.context. wenzelm@53971: wenzelm@51551: wenzelm@51398: *** System *** wenzelm@51398: wenzelm@52052: * Discontinued obsolete isabelle usedir, mkdir, make -- superseded by wenzelm@52052: "isabelle build" in Isabelle2013. INCOMPATIBILITY. wenzelm@51398: wenzelm@52054: * Discontinued obsolete isabelle-process options -f and -u (former wenzelm@52054: administrative aliases of option -e). Minor INCOMPATIBILITY. wenzelm@52054: wenzelm@52550: * Discontinued obsolete isabelle print tool, and PRINT_COMMAND wenzelm@52550: settings variable. wenzelm@52550: wenzelm@52746: * Discontinued ISABELLE_DOC_FORMAT settings variable and historic wenzelm@52746: document formats: dvi.gz, ps, ps.gz -- the default document format is wenzelm@52746: always pdf. wenzelm@52743: wenzelm@52053: * Isabelle settings variable ISABELLE_BUILD_JAVA_OPTIONS allows to wenzelm@52053: specify global resources of the JVM process run by isabelle build. wenzelm@52053: wenzelm@52116: * Toplevel executable $ISABELLE_HOME/bin/isabelle_scala_script allows wenzelm@52116: to run Isabelle/Scala source files as standalone programs. wenzelm@52116: wenzelm@52439: * Improved "isabelle keywords" tool (for old-style ProofGeneral wenzelm@52439: keyword tables): use Isabelle/Scala operations, which inspect outer wenzelm@52439: syntax without requiring to build sessions first. wenzelm@52439: wenzelm@53971: * Sessions may be organized via 'chapter' specifications in the ROOT wenzelm@53971: file, which determines a two-level hierarchy of browser info. The old wenzelm@53971: tree-like organization via implicit sub-session relation (with its wenzelm@53971: tendency towards erratic fluctuation of URLs) has been discontinued. wenzelm@53971: The default chapter is called "Unsorted". Potential INCOMPATIBILITY wenzelm@53971: for HTML presentation of theories. wenzelm@53971: wenzelm@51398: wenzelm@51398: wenzelm@50993: New in Isabelle2013 (February 2013) wenzelm@50993: ----------------------------------- wenzelm@47887: wenzelm@47967: *** General *** wenzelm@47967: wenzelm@50126: * Theorem status about oracles and unfinished/failed future proofs is wenzelm@50126: no longer printed by default, since it is incompatible with wenzelm@50126: incremental / parallel checking of the persistent document model. ML wenzelm@50126: function Thm.peek_status may be used to inspect a snapshot of the wenzelm@50126: ongoing evaluation process. Note that in batch mode --- notably wenzelm@50126: isabelle build --- the system ensures that future proofs of all wenzelm@50126: accessible theorems in the theory context are finished (as before). wenzelm@50126: wenzelm@49699: * Configuration option show_markup controls direct inlining of markup wenzelm@49699: into the printed representation of formal entities --- notably type wenzelm@49699: and sort constraints. This enables Prover IDE users to retrieve that wenzelm@49699: information via tooltips in the output window, for example. wenzelm@49699: wenzelm@48890: * Command 'ML_file' evaluates ML text from a file directly within the wenzelm@48890: theory, without any predeclaration via 'uses' in the theory header. wenzelm@48890: wenzelm@49243: * Old command 'use' command and corresponding keyword 'uses' in the wenzelm@49243: theory header are legacy features and will be discontinued soon. wenzelm@49243: Tools that load their additional source files may imitate the wenzelm@49243: 'ML_file' implementation, such that the system can take care of wenzelm@49243: dependencies properly. wenzelm@49243: wenzelm@47967: * Discontinued obsolete method fastsimp / tactic fast_simp_tac, which wenzelm@47967: is called fastforce / fast_force_tac already since Isabelle2011-1. wenzelm@47967: wenzelm@50110: * Updated and extended "isar-ref" and "implementation" manual, reduced wenzelm@50110: remaining material in old "ref" manual. wenzelm@48120: wenzelm@51050: * Improved support for auxiliary contexts that indicate block structure wenzelm@51050: for specifications. Nesting of "context fixes ... context assumes ..." wenzelm@49841: and "class ... context ...". wenzelm@49841: wenzelm@50772: * Attribute "consumes" allows a negative value as well, which is wenzelm@50778: interpreted relatively to the total number of premises of the rule in wenzelm@50772: the target context. This form of declaration is stable when exported wenzelm@50772: from a nested 'context' with additional assumptions. It is the wenzelm@50772: preferred form for definitional packages, notably cases/rules produced wenzelm@50772: in HOL/inductive and HOL/function. wenzelm@50772: wenzelm@49869: * More informative error messages for Isar proof commands involving wenzelm@49869: lazy enumerations (method applications etc.). wenzelm@49869: wenzelm@50213: * Refined 'help' command to retrieve outer syntax commands according wenzelm@50213: to name patterns (with clickable results). wenzelm@50213: wenzelm@47967: wenzelm@49968: *** Prover IDE -- Isabelle/Scala/jEdit *** wenzelm@49968: wenzelm@49968: * Parallel terminal proofs ('by') are enabled by default, likewise wenzelm@49968: proofs that are built into packages like 'datatype', 'function'. This wenzelm@49968: allows to "run ahead" checking the theory specifications on the wenzelm@49968: surface, while the prover is still crunching on internal wenzelm@49968: justifications. Unfinished / cancelled proofs are restarted as wenzelm@49968: required to complete full proof checking eventually. wenzelm@49968: wenzelm@49968: * Improved output panel with tooltips, hyperlinks etc. based on the wenzelm@49968: same Rich_Text_Area as regular Isabelle/jEdit buffers. Activation of wenzelm@49968: tooltips leads to some window that supports the same recursively, wenzelm@49968: which can lead to stacks of tooltips as the semantic document content wenzelm@49968: is explored. ESCAPE closes the whole stack, individual windows may be wenzelm@49968: closed separately, or detached to become independent jEdit dockables. wenzelm@49968: wenzelm@50717: * Improved support for commands that produce graph output: the text wenzelm@50717: message contains a clickable area to open a new instance of the graph wenzelm@50717: browser on demand. wenzelm@50717: wenzelm@49968: * More robust incremental parsing of outer syntax (partial comments, wenzelm@49968: malformed symbols). Changing the balance of open/close quotes and wenzelm@49968: comment delimiters works more conveniently with unfinished situations wenzelm@49968: that frequently occur in user interaction. wenzelm@49968: wenzelm@49968: * More efficient painting and improved reactivity when editing large wenzelm@49968: files. More scalable management of formal document content. wenzelm@49968: wenzelm@50505: * Smarter handling of tracing messages: prover process pauses after wenzelm@50505: certain number of messages per command transaction, with some user wenzelm@50505: dialog to stop or continue. This avoids swamping the front-end with wenzelm@50119: potentially infinite message streams. wenzelm@49968: wenzelm@49968: * More plugin options and preferences, based on Isabelle/Scala. The wenzelm@49968: jEdit plugin option panel provides access to some Isabelle/Scala wenzelm@49968: options, including tuning parameters for editor reactivity and color wenzelm@49968: schemes. wenzelm@49968: wenzelm@50184: * Dockable window "Symbols" provides some editing support for Isabelle wenzelm@50184: symbols. wenzelm@50184: wenzelm@51082: * Dockable window "Monitor" shows ML runtime statistics. Note that wenzelm@51082: continuous display of the chart slows down the system. wenzelm@50701: wenzelm@50183: * Improved editing support for control styles: subscript, superscript, wenzelm@50183: bold, reset of style -- operating on single symbols or text wenzelm@50198: selections. Cf. keyboard shortcuts C+e DOWN/UP/RIGHT/LEFT. wenzelm@50198: wenzelm@50198: * Actions isabelle.increase-font-size and isabelle.decrease-font-size wenzelm@50198: adjust the main text area font size, and its derivatives for output, wenzelm@50836: tooltips etc. Cf. keyboard shortcuts C-PLUS and C-MINUS, which often wenzelm@50836: need to be adapted to local keyboard layouts. wenzelm@50183: wenzelm@50730: * More reactive completion popup by default: use \t (TAB) instead of wenzelm@50730: \n (NEWLINE) to minimize intrusion into regular flow of editing. See wenzelm@50730: also "Plugin Options / SideKick / General / Code Completion Options". wenzelm@50730: wenzelm@50406: * Implicit check and build dialog of the specified logic session wenzelm@50406: image. For example, HOL, HOLCF, HOL-Nominal can be produced on wenzelm@50406: demand, without bundling big platform-dependent heap images in the wenzelm@50406: Isabelle distribution. wenzelm@50406: wenzelm@49968: * Uniform Java 7 platform on Linux, Mac OS X, Windows: recent updates wenzelm@49968: from Oracle provide better multi-platform experience. This version is wenzelm@49968: now bundled exclusively with Isabelle. wenzelm@49968: wenzelm@49968: wenzelm@48205: *** Pure *** wenzelm@48205: haftmann@48431: * Code generation for Haskell: restrict unqualified imports from haftmann@48431: Haskell Prelude to a small set of fundamental operations. haftmann@48431: wenzelm@50646: * Command 'export_code': relative file names are interpreted wenzelm@50646: relatively to master directory of current theory rather than the wenzelm@50646: rather arbitrary current working directory. INCOMPATIBILITY. haftmann@48371: wenzelm@48205: * Discontinued obsolete attribute "COMP". Potential INCOMPATIBILITY, wenzelm@48205: use regular rule composition via "OF" / "THEN", or explicit proof wenzelm@48205: structure instead. Note that Isabelle/ML provides a variety of wenzelm@48205: operators like COMP, INCR_COMP, COMP_INCR, which need to be applied wenzelm@48205: with some care where this is really required. wenzelm@48205: wenzelm@48792: * Command 'typ' supports an additional variant with explicit sort wenzelm@48792: constraint, to infer and check the most general type conforming to a wenzelm@51063: given sort. Example (in HOL): wenzelm@48792: wenzelm@48792: typ "_ * _ * bool * unit" :: finite wenzelm@48792: wenzelm@50716: * Command 'locale_deps' visualizes all locales and their relations as wenzelm@50716: a Hasse diagram. wenzelm@50716: wenzelm@48205: bulwahn@48013: *** HOL *** bulwahn@48013: wenzelm@50646: * Sledgehammer: wenzelm@50646: wenzelm@50646: - Added MaSh relevance filter based on machine-learning; see the wenzelm@50646: Sledgehammer manual for details. wenzelm@50646: - Polished Isar proofs generated with "isar_proofs" option. wenzelm@50646: - Rationalized type encodings ("type_enc" option). blanchet@50720: - Renamed "kill_provers" subcommand to "kill_all". wenzelm@50646: - Renamed options: wenzelm@50646: isar_proof ~> isar_proofs wenzelm@50646: isar_shrink_factor ~> isar_shrink wenzelm@50646: max_relevant ~> max_facts wenzelm@50646: relevance_thresholds ~> fact_thresholds wenzelm@50646: wenzelm@50646: * Quickcheck: added an optimisation for equality premises. It is wenzelm@50646: switched on by default, and can be switched off by setting the wenzelm@50646: configuration quickcheck_optimise_equality to false. wenzelm@50646: kuncar@50878: * Quotient: only one quotient can be defined by quotient_type kuncar@50878: INCOMPATIBILITY. kuncar@50878: kuncar@50878: * Lifting: kuncar@50878: - generation of an abstraction function equation in lift_definition kuncar@50878: - quot_del attribute kuncar@50878: - renamed no_abs_code -> no_code (INCOMPATIBILITY.) kuncar@50878: wenzelm@50646: * Simproc "finite_Collect" rewrites set comprehensions into pointfree wenzelm@50646: expressions. wenzelm@50646: wenzelm@50646: * Preprocessing of the code generator rewrites set comprehensions into wenzelm@50646: pointfree expressions. wenzelm@50646: wenzelm@50646: * The SMT solver Z3 has now by default a restricted set of directly wenzelm@50646: supported features. For the full set of features (div/mod, nonlinear wenzelm@50646: arithmetic, datatypes/records) with potential proof reconstruction wenzelm@50646: failures, enable the configuration option "z3_with_extensions". Minor wenzelm@50646: INCOMPATIBILITY. haftmann@49948: wenzelm@49836: * Simplified 'typedef' specifications: historical options for implicit wenzelm@49836: set definition and alternative name have been discontinued. The wenzelm@49836: former behavior of "typedef (open) t = A" is now the default, but wenzelm@49836: written just "typedef t = A". INCOMPATIBILITY, need to adapt theories wenzelm@49836: accordingly. wenzelm@49836: wenzelm@50646: * Removed constant "chars"; prefer "Enum.enum" on type "char" wenzelm@50646: directly. INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Moved operation product, sublists and n_lists from theory Enum to wenzelm@50646: List. INCOMPATIBILITY. haftmann@49822: haftmann@49739: * Theorem UN_o generalized to SUP_comp. INCOMPATIBILITY. haftmann@49739: haftmann@49738: * Class "comm_monoid_diff" formalises properties of bounded haftmann@49388: subtraction, with natural numbers and multisets as typical instances. haftmann@49388: wenzelm@50646: * Added combinator "Option.these" with type "'a option set => 'a set". wenzelm@50646: wenzelm@50646: * Theory "Transitive_Closure": renamed lemmas wenzelm@50646: wenzelm@50646: reflcl_tranclp -> reflclp_tranclp wenzelm@50646: rtranclp_reflcl -> rtranclp_reflclp wenzelm@50646: wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Theory "Rings": renamed lemmas (in class semiring) wenzelm@50646: wenzelm@50646: left_distrib ~> distrib_right wenzelm@50646: right_distrib ~> distrib_left wenzelm@50646: wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Generalized the definition of limits: wenzelm@50646: wenzelm@50646: - Introduced the predicate filterlim (LIM x F. f x :> G) which wenzelm@50646: expresses that when the input values x converge to F then the wenzelm@50646: output f x converges to G. wenzelm@50646: wenzelm@50646: - Added filters for convergence to positive (at_top) and negative wenzelm@50646: infinity (at_bot). wenzelm@50646: wenzelm@50646: - Moved infinity in the norm (at_infinity) from wenzelm@50646: Multivariate_Analysis to Complex_Main. wenzelm@50646: wenzelm@50646: - Removed real_tendsto_inf, it is superseded by "LIM x F. f x :> wenzelm@50646: at_top". wenzelm@50646: wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Theory "Library/Option_ord" provides instantiation of option type to wenzelm@50646: lattice type classes. wenzelm@50646: wenzelm@50646: * Theory "Library/Multiset": renamed wenzelm@50646: wenzelm@50646: constant fold_mset ~> Multiset.fold wenzelm@50646: fact fold_mset_commute ~> fold_mset_comm wenzelm@50646: wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Renamed theory Library/List_Prefix to Library/Sublist, with related wenzelm@50646: changes as follows. wenzelm@50646: wenzelm@50646: - Renamed constants (and related lemmas) Christian@49145: Christian@49145: prefix ~> prefixeq Christian@49145: strict_prefix ~> prefix Christian@49145: wenzelm@50646: - Replaced constant "postfix" by "suffixeq" with swapped argument wenzelm@50646: order (i.e., "postfix xs ys" is now "suffixeq ys xs") and dropped wenzelm@50646: old infix syntax "xs >>= ys"; use "suffixeq ys xs" instead. wenzelm@50646: Renamed lemmas accordingly. wenzelm@50646: wenzelm@50646: - Added constant "list_hembeq" for homeomorphic embedding on wenzelm@50646: lists. Added abbreviation "sublisteq" for special case wenzelm@50646: "list_hembeq (op =)". wenzelm@50646: wenzelm@50646: - Theory Library/Sublist no longer provides "order" and "bot" type wenzelm@50646: class instances for the prefix order (merely corresponding locale wenzelm@50646: interpretations). The type class instances are now in theory wenzelm@50646: Library/Prefix_Order. wenzelm@50646: wenzelm@50646: - The sublist relation of theory Library/Sublist_Order is now based wenzelm@50646: on "Sublist.sublisteq". Renamed lemmas accordingly: Christian@50516: Christian@50516: le_list_append_le_same_iff ~> Sublist.sublisteq_append_le_same_iff Christian@50516: le_list_append_mono ~> Sublist.list_hembeq_append_mono Christian@50516: le_list_below_empty ~> Sublist.list_hembeq_Nil, Sublist.list_hembeq_Nil2 Christian@50516: le_list_Cons_EX ~> Sublist.list_hembeq_ConsD Christian@50516: le_list_drop_Cons2 ~> Sublist.sublisteq_Cons2' Christian@50516: le_list_drop_Cons_neq ~> Sublist.sublisteq_Cons2_neq Christian@50516: le_list_drop_Cons ~> Sublist.sublisteq_Cons' Christian@50516: le_list_drop_many ~> Sublist.sublisteq_drop_many Christian@50516: le_list_filter_left ~> Sublist.sublisteq_filter_left Christian@50516: le_list_rev_drop_many ~> Sublist.sublisteq_rev_drop_many Christian@50516: le_list_rev_take_iff ~> Sublist.sublisteq_append Christian@50516: le_list_same_length ~> Sublist.sublisteq_same_length Christian@50516: le_list_take_many_iff ~> Sublist.sublisteq_append' Christian@49145: less_eq_list.drop ~> less_eq_list_drop Christian@49145: less_eq_list.induct ~> less_eq_list_induct Christian@50516: not_le_list_length ~> Sublist.not_sublisteq_length Christian@49145: wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * New theory Library/Countable_Set. wenzelm@50646: wenzelm@50646: * Theory Library/Debug and Library/Parallel provide debugging and wenzelm@50646: parallel execution for code generated towards Isabelle/ML. wenzelm@50646: wenzelm@50646: * Theory Library/FuncSet: Extended support for Pi and extensional and wenzelm@50646: introduce the extensional dependent function space "PiE". Replaced wenzelm@50646: extensional_funcset by an abbreviation, and renamed lemmas from wenzelm@50646: extensional_funcset to PiE as follows: wenzelm@50646: wenzelm@50646: extensional_empty ~> PiE_empty wenzelm@50646: extensional_funcset_empty_domain ~> PiE_empty_domain wenzelm@50646: extensional_funcset_empty_range ~> PiE_empty_range wenzelm@50646: extensional_funcset_arb ~> PiE_arb wenzelm@50646: extensional_funcset_mem ~> PiE_mem wenzelm@50646: extensional_funcset_extend_domainI ~> PiE_fun_upd wenzelm@50646: extensional_funcset_restrict_domain ~> fun_upd_in_PiE wenzelm@50646: extensional_funcset_extend_domain_eq ~> PiE_insert_eq wenzelm@50646: card_extensional_funcset ~> card_PiE wenzelm@50646: finite_extensional_funcset ~> finite_PiE wenzelm@50646: wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Theory Library/FinFun: theory of almost everywhere constant wenzelm@50646: functions (supersedes the AFP entry "Code Generation for Functions as wenzelm@50646: Data"). wenzelm@50646: wenzelm@50646: * Theory Library/Phantom: generic phantom type to make a type wenzelm@50646: parameter appear in a constant's type. This alternative to adding wenzelm@50646: TYPE('a) as another parameter avoids unnecessary closures in generated wenzelm@50646: code. wenzelm@50646: wenzelm@50646: * Theory Library/RBT_Impl: efficient construction of red-black trees wenzelm@50646: from sorted associative lists. Merging two trees with rbt_union may wenzelm@50646: return a structurally different tree than before. Potential wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@50646: * Theory Library/IArray: immutable arrays with code generation. wenzelm@50646: wenzelm@50646: * Theory Library/Finite_Lattice: theory of finite lattices. wenzelm@50646: wenzelm@50646: * HOL/Multivariate_Analysis: replaced wenzelm@50646: wenzelm@50646: "basis :: 'a::euclidean_space => nat => real" wenzelm@50646: "\\ :: (nat => real) => 'a::euclidean_space" wenzelm@50646: wenzelm@50646: on euclidean spaces by using the inner product "_ \ _" with wenzelm@50646: vectors from the Basis set: "\\ i. f i" is superseded by wenzelm@50646: "SUM i : Basis. f i * r i". wenzelm@50646: wenzelm@50646: With this change the following constants are also changed or removed: wenzelm@50646: wenzelm@50646: DIM('a) :: nat ~> card (Basis :: 'a set) (is an abbreviation) wenzelm@50646: a $$ i ~> inner a i (where i : Basis) wenzelm@50646: cart_base i removed wenzelm@50646: \, \' removed hoelzl@50526: hoelzl@50526: Theorems about these constants where removed. hoelzl@50526: hoelzl@50526: Renamed lemmas: hoelzl@50526: wenzelm@50646: component_le_norm ~> Basis_le_norm wenzelm@50646: euclidean_eq ~> euclidean_eq_iff wenzelm@50646: differential_zero_maxmin_component ~> differential_zero_maxmin_cart wenzelm@50646: euclidean_simps ~> inner_simps wenzelm@50646: independent_basis ~> independent_Basis wenzelm@50646: span_basis ~> span_Basis wenzelm@50646: in_span_basis ~> in_span_Basis wenzelm@50646: norm_bound_component_le ~> norm_boound_Basis_le wenzelm@50646: norm_bound_component_lt ~> norm_boound_Basis_lt wenzelm@50646: component_le_infnorm ~> Basis_le_infnorm wenzelm@50646: wenzelm@50646: INCOMPATIBILITY. hoelzl@50526: hoelzl@50141: * HOL/Probability: wenzelm@50646: wenzelm@50646: - Added simproc "measurable" to automatically prove measurability. wenzelm@50646: wenzelm@50646: - Added induction rules for sigma sets with disjoint union wenzelm@50646: (sigma_sets_induct_disjoint) and for Borel-measurable functions wenzelm@50646: (borel_measurable_induct). wenzelm@50646: wenzelm@50646: - Added the Daniell-Kolmogorov theorem (the existence the limit of a wenzelm@50646: projective family). wenzelm@50646: wenzelm@50646: * HOL/Cardinals: Theories of ordinals and cardinals (supersedes the wenzelm@50646: AFP entry "Ordinals_and_Cardinals"). wenzelm@50646: wenzelm@50646: * HOL/BNF: New (co)datatype package based on bounded natural functors wenzelm@50646: with support for mixed, nested recursion and interesting non-free wenzelm@50646: datatypes. blanchet@48094: wenzelm@50991: * HOL/Finite_Set and Relation: added new set and relation operations kuncar@50878: expressed by Finite_Set.fold. kuncar@50878: kuncar@50878: * New theory HOL/Library/RBT_Set: implementation of sets by red-black kuncar@50878: trees for the code generator. kuncar@50878: kuncar@50878: * HOL/Library/RBT and HOL/Library/Mapping have been converted to kuncar@50878: Lifting/Transfer. kuncar@50878: possible INCOMPATIBILITY. kuncar@50878: kuncar@50878: * HOL/Set: renamed Set.project -> Set.filter kuncar@50878: INCOMPATIBILITY. kuncar@50878: wenzelm@48120: wenzelm@48206: *** Document preparation *** wenzelm@48206: wenzelm@50646: * Dropped legacy antiquotations "term_style" and "thm_style", since wenzelm@50646: styles may be given as arguments to "term" and "thm" already. wenzelm@50646: Discontinued legacy styles "prem1" .. "prem19". wenzelm@50646: wenzelm@50646: * Default LaTeX rendering for \ is now based on eurosym package, wenzelm@50646: instead of slightly exotic babel/greek. wenzelm@48206: wenzelm@48616: * Document variant NAME may use different LaTeX entry point wenzelm@48616: document/root_NAME.tex if that file exists, instead of the common wenzelm@48616: document/root.tex. wenzelm@48616: wenzelm@48657: * Simplified custom document/build script, instead of old-style wenzelm@48657: document/IsaMakefile. Minor INCOMPATIBILITY. wenzelm@48657: wenzelm@48206: wenzelm@48992: *** ML *** wenzelm@48992: wenzelm@50646: * The default limit for maximum number of worker threads is now 8, wenzelm@50646: instead of 4, in correspondence to capabilities of contemporary wenzelm@50646: hardware and Poly/ML runtime system. wenzelm@50646: wenzelm@49869: * Type Seq.results and related operations support embedded error wenzelm@49869: messages within lazy enumerations, and thus allow to provide wenzelm@49869: informative errors in the absence of any usable results. wenzelm@49869: wenzelm@48992: * Renamed Position.str_of to Position.here to emphasize that this is a wenzelm@48992: formal device to inline positions into message text, but not wenzelm@48992: necessarily printing visible text. wenzelm@48992: wenzelm@48992: wenzelm@48206: *** System *** wenzelm@48206: wenzelm@48585: * Advanced support for Isabelle sessions and build management, see wenzelm@48585: "system" manual for the chapter of that name, especially the "isabelle wenzelm@51056: build" tool and its examples. The "isabelle mkroot" tool prepares wenzelm@51056: session root directories for use with "isabelle build", similar to wenzelm@51056: former "isabelle mkdir" for "isabelle usedir". Note that this affects wenzelm@51056: document preparation as well. INCOMPATIBILITY, isabelle usedir / wenzelm@48736: mkdir / make are rendered obsolete. wenzelm@48736: wenzelm@51056: * Discontinued obsolete Isabelle/build script, it is superseded by the wenzelm@51056: regular isabelle build tool. For example: wenzelm@51056: wenzelm@51056: isabelle build -s -b HOL wenzelm@51056: wenzelm@48736: * Discontinued obsolete "isabelle makeall". wenzelm@48585: wenzelm@48722: * Discontinued obsolete IsaMakefile and ROOT.ML files from the wenzelm@48722: Isabelle distribution, except for rudimentary src/HOL/IsaMakefile that wenzelm@48722: provides some traditional targets that invoke "isabelle build". Note wenzelm@48722: that this is inefficient! Applications of Isabelle/HOL involving wenzelm@48722: "isabelle make" should be upgraded to use "isabelle build" directly. wenzelm@48722: wenzelm@48693: * The "isabelle options" tool prints Isabelle system options, as wenzelm@48693: required for "isabelle build", for example. wenzelm@48693: wenzelm@50646: * The "isabelle logo" tool produces EPS and PDF format simultaneously. wenzelm@50646: Minor INCOMPATIBILITY in command-line options. wenzelm@50646: wenzelm@50646: * The "isabelle install" tool has now a simpler command-line. Minor wenzelm@50646: INCOMPATIBILITY. wenzelm@50646: wenzelm@48844: * The "isabelle components" tool helps to resolve add-on components wenzelm@48844: that are not bundled, or referenced from a bare-bones repository wenzelm@48844: version of Isabelle. wenzelm@48844: wenzelm@50646: * Settings variable ISABELLE_PLATFORM_FAMILY refers to the general wenzelm@50646: platform family: "linux", "macos", "windows". wenzelm@50646: wenzelm@50646: * The ML system is configured as regular component, and no longer wenzelm@50646: picked up from some surrounding directory. Potential INCOMPATIBILITY wenzelm@50646: for home-made settings. wenzelm@50132: wenzelm@50701: * Improved ML runtime statistics (heap, threads, future tasks etc.). wenzelm@50701: wenzelm@48206: * Discontinued support for Poly/ML 5.2.1, which was the last version wenzelm@48206: without exception positions and advanced ML compiler/toplevel wenzelm@48206: configuration. wenzelm@48206: wenzelm@48574: * Discontinued special treatment of Proof General -- no longer guess wenzelm@48574: PROOFGENERAL_HOME based on accidental file-system layout. Minor wenzelm@48574: INCOMPATIBILITY: provide PROOFGENERAL_HOME and PROOFGENERAL_OPTIONS wenzelm@48574: settings manually, or use a Proof General version that has been wenzelm@48574: bundled as Isabelle component. wenzelm@48206: wenzelm@50182: wenzelm@48120: wenzelm@47462: New in Isabelle2012 (May 2012) wenzelm@47462: ------------------------------ wenzelm@45109: wenzelm@45593: *** General *** wenzelm@45593: wenzelm@45614: * Prover IDE (PIDE) improvements: wenzelm@45614: wenzelm@47585: - more robust Sledgehammer integration (as before the sledgehammer wenzelm@47806: command-line needs to be typed into the source buffer) wenzelm@45614: - markup for bound variables wenzelm@47806: - markup for types of term variables (displayed as tooltips) wenzelm@46956: - support for user-defined Isar commands within the running session wenzelm@47158: - improved support for Unicode outside original 16bit range wenzelm@47158: e.g. glyph for \ (thanks to jEdit 4.5.1) wenzelm@45614: wenzelm@47806: * Forward declaration of outer syntax keywords within the theory wenzelm@47806: header -- minor INCOMPATIBILITY for user-defined commands. Allow new wenzelm@47806: commands to be used in the same theory where defined. wenzelm@46485: wenzelm@47482: * Auxiliary contexts indicate block structure for specifications with wenzelm@47482: additional parameters and assumptions. Such unnamed contexts may be wenzelm@47482: nested within other targets, like 'theory', 'locale', 'class', wenzelm@47482: 'instantiation' etc. Results from the local context are generalized wenzelm@47482: accordingly and applied to the enclosing target context. Example: wenzelm@47482: wenzelm@47482: context wenzelm@47482: fixes x y z :: 'a wenzelm@47482: assumes xy: "x = y" and yz: "y = z" wenzelm@47482: begin wenzelm@47482: wenzelm@47482: lemma my_trans: "x = z" using xy yz by simp wenzelm@47482: wenzelm@47482: end wenzelm@47482: wenzelm@47482: thm my_trans wenzelm@47482: wenzelm@47482: The most basic application is to factor-out context elements of wenzelm@47482: several fixes/assumes/shows theorem statements, e.g. see wenzelm@47482: ~~/src/HOL/Isar_Examples/Group_Context.thy wenzelm@47482: wenzelm@47482: Any other local theory specification element works within the "context wenzelm@47482: ... begin ... end" block as well. wenzelm@47482: wenzelm@47484: * Bundled declarations associate attributed fact expressions with a wenzelm@47484: given name in the context. These may be later included in other wenzelm@47484: contexts. This allows to manage context extensions casually, without wenzelm@47855: the logical dependencies of locales and locale interpretation. See wenzelm@47855: commands 'bundle', 'include', 'including' etc. in the isar-ref manual. wenzelm@47484: wenzelm@47829: * Commands 'lemmas' and 'theorems' allow local variables using 'for' wenzelm@47829: declaration, and results are standardized before being stored. Thus wenzelm@47829: old-style "standard" after instantiation or composition of facts wenzelm@47829: becomes obsolete. Minor INCOMPATIBILITY, due to potential change of wenzelm@47829: indices of schematic variables. wenzelm@47829: wenzelm@47829: * Rule attributes in local theory declarations (e.g. locale or class) wenzelm@47829: are now statically evaluated: the resulting theorem is stored instead wenzelm@47829: of the original expression. INCOMPATIBILITY in rare situations, where wenzelm@47829: the historic accident of dynamic re-evaluation in interpretations wenzelm@47829: etc. was exploited. wenzelm@47829: wenzelm@47829: * New tutorial "Programming and Proving in Isabelle/HOL" wenzelm@47829: ("prog-prove"). It completely supersedes "A Tutorial Introduction to wenzelm@47829: Structured Isar Proofs" ("isar-overview"), which has been removed. It wenzelm@47829: also supersedes "Isabelle/HOL, A Proof Assistant for Higher-Order wenzelm@47829: Logic" as the recommended beginners tutorial, but does not cover all wenzelm@47829: of the material of that old tutorial. wenzelm@47829: wenzelm@47829: * Updated and extended reference manuals: "isar-ref", wenzelm@47829: "implementation", "system"; reduced remaining material in old "ref" wenzelm@47829: manual. wenzelm@47829: wenzelm@47829: wenzelm@47829: *** Pure *** wenzelm@47829: wenzelm@46976: * Command 'definition' no longer exports the foundational "raw_def" wenzelm@46976: into the user context. Minor INCOMPATIBILITY, may use the regular wenzelm@46976: "def" result with attribute "abs_def" to imitate the old version. wenzelm@46976: wenzelm@47855: * Attribute "abs_def" turns an equation of the form "f x y == t" into wenzelm@47855: "f == %x y. t", which ensures that "simp" or "unfold" steps always wenzelm@47855: expand it. This also works for object-logic equality. (Formerly wenzelm@47855: undocumented feature.) wenzelm@47855: wenzelm@47856: * Sort constraints are now propagated in simultaneous statements, just wenzelm@47856: like type constraints. INCOMPATIBILITY in rare situations, where wenzelm@47856: distinct sorts used to be assigned accidentally. For example: wenzelm@47856: wenzelm@47856: lemma "P (x::'a::foo)" and "Q (y::'a::bar)" -- "now illegal" wenzelm@47856: wenzelm@47856: lemma "P (x::'a)" and "Q (y::'a::bar)" wenzelm@47856: -- "now uniform 'a::bar instead of default sort for first occurrence (!)" wenzelm@47856: wenzelm@47856: * Rule composition via attribute "OF" (or ML functions OF/MRS) is more wenzelm@47856: tolerant against multiple unifiers, as long as the final result is wenzelm@47856: unique. (As before, rules are composed in canonical right-to-left wenzelm@47856: order to accommodate newly introduced premises.) wenzelm@47856: wenzelm@47806: * Renamed some inner syntax categories: wenzelm@47806: wenzelm@47806: num ~> num_token wenzelm@47806: xnum ~> xnum_token wenzelm@47806: xstr ~> str_token wenzelm@47806: wenzelm@47806: Minor INCOMPATIBILITY. Note that in practice "num_const" or wenzelm@47806: "num_position" etc. are mainly used instead (which also include wenzelm@47806: position information via constraints). wenzelm@47806: wenzelm@47829: * Simplified configuration options for syntax ambiguity: see wenzelm@47829: "syntax_ambiguity_warning" and "syntax_ambiguity_limit" in isar-ref wenzelm@47829: manual. Minor INCOMPATIBILITY. wenzelm@47829: wenzelm@47856: * Discontinued configuration option "syntax_positions": atomic terms wenzelm@47856: in parse trees are always annotated by position constraints. wenzelm@45134: wenzelm@47464: * Old code generator for SML and its commands 'code_module', wenzelm@45383: 'code_library', 'consts_code', 'types_code' have been discontinued. haftmann@46028: Use commands of the generic code generator instead. INCOMPATIBILITY. wenzelm@45383: wenzelm@47464: * Redundant attribute "code_inline" has been discontinued. Use wenzelm@47464: "code_unfold" instead. INCOMPATIBILITY. wenzelm@47464: wenzelm@47464: * Dropped attribute "code_unfold_post" in favor of the its dual wenzelm@47464: "code_abbrev", which yields a common pattern in definitions like haftmann@46028: haftmann@46028: definition [code_abbrev]: "f = t" haftmann@46028: haftmann@46028: INCOMPATIBILITY. wenzelm@45383: wenzelm@47856: * Obsolete 'types' command has been discontinued. Use 'type_synonym' wenzelm@47856: instead. INCOMPATIBILITY. wenzelm@47856: wenzelm@47856: * Discontinued old "prems" fact, which used to refer to the accidental wenzelm@47856: collection of foundational premises in the context (already marked as wenzelm@47856: legacy since Isabelle2011). wenzelm@47855: wenzelm@45427: huffman@45122: *** HOL *** huffman@45122: wenzelm@47464: * Type 'a set is now a proper type constructor (just as before wenzelm@47464: Isabelle2008). Definitions mem_def and Collect_def have disappeared. wenzelm@47464: Non-trivial INCOMPATIBILITY. For developments keeping predicates and wenzelm@47855: sets separate, it is often sufficient to rephrase some set S that has wenzelm@47855: been accidentally used as predicates by "%x. x : S", and some wenzelm@47855: predicate P that has been accidentally used as set by "{x. P x}". wenzelm@47855: Corresponding proofs in a first step should be pruned from any wenzelm@47855: tinkering with former theorems mem_def and Collect_def as far as wenzelm@47855: possible. wenzelm@47855: wenzelm@47855: For developments which deliberately mix predicates and sets, a wenzelm@47464: planning step is necessary to determine what should become a predicate wenzelm@47464: and what a set. It can be helpful to carry out that step in wenzelm@47464: Isabelle2011-1 before jumping right into the current release. wenzelm@47464: wenzelm@47855: * Code generation by default implements sets as container type rather wenzelm@47855: than predicates. INCOMPATIBILITY. wenzelm@47855: wenzelm@47855: * New type synonym 'a rel = ('a * 'a) set wenzelm@47855: wenzelm@47464: * The representation of numerals has changed. Datatype "num" wenzelm@47464: represents strictly positive binary numerals, along with functions wenzelm@47464: "numeral :: num => 'a" and "neg_numeral :: num => 'a" to represent wenzelm@47855: positive and negated numeric literals, respectively. See also wenzelm@47855: definitions in ~~/src/HOL/Num.thy. Potential INCOMPATIBILITY, some wenzelm@47855: user theories may require adaptations as follows: huffman@47108: huffman@47108: - Theorems with number_ring or number_semiring constraints: These huffman@47108: classes are gone; use comm_ring_1 or comm_semiring_1 instead. huffman@47108: huffman@47108: - Theories defining numeric types: Remove number, number_semiring, huffman@47108: and number_ring instances. Defer all theorems about numerals until huffman@47108: after classes one and semigroup_add have been instantiated. huffman@47108: huffman@47108: - Numeral-only simp rules: Replace each rule having a "number_of v" huffman@47108: pattern with two copies, one for numeral and one for neg_numeral. huffman@47108: huffman@47108: - Theorems about subclasses of semiring_1 or ring_1: These classes huffman@47108: automatically support numerals now, so more simp rules and huffman@47108: simprocs may now apply within the proof. huffman@47108: huffman@47108: - Definitions and theorems using old constructors Pls/Min/Bit0/Bit1: huffman@47108: Redefine using other integer operations. huffman@47108: wenzelm@47855: * Transfer: New package intended to generalize the existing wenzelm@47855: "descending" method and related theorem attributes from the Quotient wenzelm@47855: package. (Not all functionality is implemented yet, but future wenzelm@47855: development will focus on Transfer as an eventual replacement for the wenzelm@47855: corresponding parts of the Quotient package.) wenzelm@47809: wenzelm@47809: - transfer_rule attribute: Maintains a collection of transfer rules, wenzelm@47809: which relate constants at two different types. Transfer rules may wenzelm@47809: relate different type instances of the same polymorphic constant, wenzelm@47809: or they may relate an operation on a raw type to a corresponding wenzelm@47809: operation on an abstract type (quotient or subtype). For example: wenzelm@47809: wenzelm@47809: ((A ===> B) ===> list_all2 A ===> list_all2 B) map map wenzelm@47809: (cr_int ===> cr_int ===> cr_int) (%(x,y) (u,v). (x+u, y+v)) plus_int wenzelm@47809: wenzelm@47809: - transfer method: Replaces a subgoal on abstract types with an wenzelm@47809: equivalent subgoal on the corresponding raw types. Constants are wenzelm@47809: replaced with corresponding ones according to the transfer rules. wenzelm@47809: Goals are generalized over all free variables by default; this is huffman@47851: necessary for variables whose types change, but can be overridden wenzelm@47855: for specific variables with e.g. "transfer fixing: x y z". The wenzelm@47809: variant transfer' method allows replacing a subgoal with one that wenzelm@47809: is logically stronger (rather than equivalent). wenzelm@47809: wenzelm@47809: - relator_eq attribute: Collects identity laws for relators of wenzelm@47809: various type constructors, e.g. "list_all2 (op =) = (op =)". The wenzelm@47809: transfer method uses these lemmas to infer transfer rules for wenzelm@47809: non-polymorphic constants on the fly. wenzelm@47809: wenzelm@47809: - transfer_prover method: Assists with proving a transfer rule for a wenzelm@47809: new constant, provided the constant is defined in terms of other wenzelm@47809: constants that already have transfer rules. It should be applied wenzelm@47809: after unfolding the constant definitions. wenzelm@47809: wenzelm@47809: - HOL/ex/Transfer_Int_Nat.thy: Example theory demonstrating transfer wenzelm@47809: from type nat to type int. wenzelm@47809: huffman@47851: * Lifting: New package intended to generalize the quotient_definition huffman@47851: facility of the Quotient package; designed to work with Transfer. wenzelm@47809: wenzelm@47809: - lift_definition command: Defines operations on an abstract type in wenzelm@47809: terms of a corresponding operation on a representation wenzelm@47809: type. Example syntax: wenzelm@47809: wenzelm@47809: lift_definition dlist_insert :: "'a => 'a dlist => 'a dlist" wenzelm@47809: is List.insert wenzelm@47809: wenzelm@47809: Users must discharge a respectfulness proof obligation when each wenzelm@47809: constant is defined. (For a type copy, i.e. a typedef with UNIV, wenzelm@47809: the proof is discharged automatically.) The obligation is wenzelm@47809: presented in a user-friendly, readable form; a respectfulness wenzelm@47809: theorem in the standard format and a transfer rule are generated wenzelm@47809: by the package. wenzelm@47809: wenzelm@47809: - Integration with code_abstype: For typedefs (e.g. subtypes wenzelm@47809: corresponding to a datatype invariant, such as dlist), wenzelm@47809: lift_definition generates a code certificate theorem and sets up wenzelm@47809: code generation for each constant. wenzelm@47809: wenzelm@47809: - setup_lifting command: Sets up the Lifting package to work with a wenzelm@47809: user-defined type. The user must provide either a quotient theorem wenzelm@47809: or a type_definition theorem. The package configures transfer wenzelm@47809: rules for equality and quantifiers on the type, and sets up the wenzelm@47809: lift_definition command to work with the type. wenzelm@47809: wenzelm@47809: - Usage examples: See Quotient_Examples/Lift_DList.thy, huffman@47851: Quotient_Examples/Lift_RBT.thy, Quotient_Examples/Lift_FSet.thy, huffman@47851: Word/Word.thy and Library/Float.thy. wenzelm@47809: wenzelm@47809: * Quotient package: wenzelm@47809: wenzelm@47809: - The 'quotient_type' command now supports a 'morphisms' option with wenzelm@47809: rep and abs functions, similar to typedef. wenzelm@47809: wenzelm@47809: - 'quotient_type' sets up new types to work with the Lifting and wenzelm@47809: Transfer packages, as with 'setup_lifting'. wenzelm@47809: wenzelm@47809: - The 'quotient_definition' command now requires the user to prove a wenzelm@47809: respectfulness property at the point where the constant is wenzelm@47809: defined, similar to lift_definition; INCOMPATIBILITY. wenzelm@47809: wenzelm@47809: - Renamed predicate 'Quotient' to 'Quotient3', and renamed theorems wenzelm@47809: accordingly, INCOMPATIBILITY. wenzelm@47809: wenzelm@47809: * New diagnostic command 'find_unused_assms' to find potentially wenzelm@47809: superfluous assumptions in theorems using Quickcheck. wenzelm@47809: wenzelm@47809: * Quickcheck: wenzelm@47809: wenzelm@47809: - Quickcheck returns variable assignments as counterexamples, which wenzelm@47809: allows to reveal the underspecification of functions under test. wenzelm@47809: For example, refuting "hd xs = x", it presents the variable wenzelm@47809: assignment xs = [] and x = a1 as a counterexample, assuming that wenzelm@47809: any property is false whenever "hd []" occurs in it. wenzelm@47809: wenzelm@47809: These counterexample are marked as potentially spurious, as wenzelm@47809: Quickcheck also returns "xs = []" as a counterexample to the wenzelm@47809: obvious theorem "hd xs = hd xs". wenzelm@47809: wenzelm@47809: After finding a potentially spurious counterexample, Quickcheck wenzelm@47809: continues searching for genuine ones. wenzelm@47809: wenzelm@47809: By default, Quickcheck shows potentially spurious and genuine wenzelm@47809: counterexamples. The option "genuine_only" sets quickcheck to only wenzelm@47809: show genuine counterexamples. wenzelm@47809: wenzelm@47809: - The command 'quickcheck_generator' creates random and exhaustive wenzelm@47809: value generators for a given type and operations. wenzelm@47809: wenzelm@47809: It generates values by using the operations as if they were wenzelm@47809: constructors of that type. wenzelm@47809: wenzelm@47809: - Support for multisets. wenzelm@47809: wenzelm@47809: - Added "use_subtype" options. wenzelm@47809: wenzelm@47809: - Added "quickcheck_locale" configuration to specify how to process wenzelm@47809: conjectures in a locale context. wenzelm@47809: wenzelm@47855: * Nitpick: Fixed infinite loop caused by the 'peephole_optim' option wenzelm@47855: and affecting 'rat' and 'real'. wenzelm@47809: wenzelm@47809: * Sledgehammer: wenzelm@47809: - Integrated more tightly with SPASS, as described in the ITP 2012 wenzelm@47809: paper "More SPASS with Isabelle". wenzelm@47809: - Made it try "smt" as a fallback if "metis" fails or times out. wenzelm@47809: - Added support for the following provers: Alt-Ergo (via Why3 and wenzelm@47809: TFF1), iProver, iProver-Eq. wenzelm@47809: - Sped up the minimizer. wenzelm@47809: - Added "lam_trans", "uncurry_aliases", and "minimize" options. wenzelm@47809: - Renamed "slicing" ("no_slicing") option to "slice" ("dont_slice"). wenzelm@47809: - Renamed "sound" option to "strict". wenzelm@47809: wenzelm@47855: * Metis: Added possibility to specify lambda translations scheme as a wenzelm@47855: parenthesized argument (e.g., "by (metis (lifting) ...)"). wenzelm@47855: wenzelm@47855: * SMT: Renamed "smt_fixed" option to "smt_read_only_certificates". wenzelm@47855: wenzelm@47855: * Command 'try0': Renamed from 'try_methods'. INCOMPATIBILITY. wenzelm@47809: wenzelm@47856: * New "case_product" attribute to generate a case rule doing multiple wenzelm@47856: case distinctions at the same time. E.g. wenzelm@47856: wenzelm@47856: list.exhaust [case_product nat.exhaust] wenzelm@47856: wenzelm@47856: produces a rule which can be used to perform case distinction on both wenzelm@47856: a list and a nat. wenzelm@47856: wenzelm@47809: * New "eventually_elim" method as a generalized variant of the wenzelm@47855: eventually_elim* rules. Supports structured proofs. wenzelm@47855: wenzelm@47702: * Typedef with implicit set definition is considered legacy. Use wenzelm@47702: "typedef (open)" form instead, which will eventually become the wenzelm@47702: default. wenzelm@47702: wenzelm@47856: * Record: code generation can be switched off manually with wenzelm@47856: wenzelm@47856: declare [[record_coden = false]] -- "default true" wenzelm@47856: wenzelm@47856: * Datatype: type parameters allow explicit sort constraints. wenzelm@47856: wenzelm@47855: * Concrete syntax for case expressions includes constraints for source wenzelm@47855: positions, and thus produces Prover IDE markup for its bindings. wenzelm@47855: INCOMPATIBILITY for old-style syntax translations that augment the wenzelm@47855: pattern notation; e.g. see src/HOL/HOLCF/One.thy for translations of wenzelm@47855: one_case. wenzelm@47855: wenzelm@47855: * Clarified attribute "mono_set": pure declaration without modifying wenzelm@47855: the result of the fact expression. wenzelm@47855: haftmann@46752: * More default pred/set conversions on a couple of relation operations wenzelm@47464: and predicates. Added powers of predicate relations. Consolidation wenzelm@47464: of some relation theorems: haftmann@46752: haftmann@46752: converse_def ~> converse_unfold haftmann@47549: rel_comp_def ~> relcomp_unfold haftmann@47820: symp_def ~> (modified, use symp_def and sym_def instead) haftmann@46752: transp_def ~> transp_trans haftmann@46752: Domain_def ~> Domain_unfold haftmann@46752: Range_def ~> Domain_converse [symmetric] haftmann@46752: haftmann@46981: Generalized theorems INF_INT_eq, INF_INT_eq2, SUP_UN_eq, SUP_UN_eq2. haftmann@46981: wenzelm@47464: See theory "Relation" for examples for making use of pred/set wenzelm@47464: conversions by means of attributes "to_set" and "to_pred". haftmann@47086: haftmann@46752: INCOMPATIBILITY. haftmann@46752: bulwahn@46363: * Renamed facts about the power operation on relations, i.e., relpow wenzelm@47464: to match the constant's name: wenzelm@47463: wenzelm@46458: rel_pow_1 ~> relpow_1 bulwahn@46363: rel_pow_0_I ~> relpow_0_I bulwahn@46363: rel_pow_Suc_I ~> relpow_Suc_I bulwahn@46363: rel_pow_Suc_I2 ~> relpow_Suc_I2 bulwahn@46363: rel_pow_0_E ~> relpow_0_E bulwahn@46363: rel_pow_Suc_E ~> relpow_Suc_E bulwahn@46363: rel_pow_E ~> relpow_E wenzelm@46458: rel_pow_Suc_D2 ~> relpow_Suc_D2 wenzelm@47463: rel_pow_Suc_E2 ~> relpow_Suc_E2 bulwahn@46363: rel_pow_Suc_D2' ~> relpow_Suc_D2' bulwahn@46363: rel_pow_E2 ~> relpow_E2 bulwahn@46363: rel_pow_add ~> relpow_add bulwahn@46363: rel_pow_commute ~> relpow bulwahn@46363: rel_pow_empty ~> relpow_empty: bulwahn@46363: rtrancl_imp_UN_rel_pow ~> rtrancl_imp_UN_relpow bulwahn@46363: rel_pow_imp_rtrancl ~> relpow_imp_rtrancl bulwahn@46363: rtrancl_is_UN_rel_pow ~> rtrancl_is_UN_relpow bulwahn@46363: rtrancl_imp_rel_pow ~> rtrancl_imp_relpow bulwahn@46363: rel_pow_fun_conv ~> relpow_fun_conv bulwahn@46363: rel_pow_finite_bounded1 ~> relpow_finite_bounded1 bulwahn@46363: rel_pow_finite_bounded ~> relpow_finite_bounded bulwahn@46363: rtrancl_finite_eq_rel_pow ~> rtrancl_finite_eq_relpow bulwahn@46363: trancl_finite_eq_rel_pow ~> trancl_finite_eq_relpow bulwahn@46363: single_valued_rel_pow ~> single_valued_relpow wenzelm@47463: bulwahn@46363: INCOMPATIBILITY. bulwahn@46363: bulwahn@47448: * Theory Relation: Consolidated constant name for relation composition wenzelm@47464: and corresponding theorem names: wenzelm@47464: haftmann@47549: - Renamed constant rel_comp to relcomp. wenzelm@47464: bulwahn@47448: - Dropped abbreviation pred_comp. Use relcompp instead. wenzelm@47464: bulwahn@47448: - Renamed theorems: wenzelm@47464: bulwahn@47448: rel_compI ~> relcompI bulwahn@47448: rel_compEpair ~> relcompEpair bulwahn@47448: rel_compE ~> relcompE bulwahn@47448: pred_comp_rel_comp_eq ~> relcompp_relcomp_eq bulwahn@47448: rel_comp_empty1 ~> relcomp_empty1 bulwahn@47448: rel_comp_mono ~> relcomp_mono bulwahn@47448: rel_comp_subset_Sigma ~> relcomp_subset_Sigma bulwahn@47448: rel_comp_distrib ~> relcomp_distrib bulwahn@47448: rel_comp_distrib2 ~> relcomp_distrib2 bulwahn@47448: rel_comp_UNION_distrib ~> relcomp_UNION_distrib bulwahn@47448: rel_comp_UNION_distrib2 ~> relcomp_UNION_distrib2 bulwahn@47448: single_valued_rel_comp ~> single_valued_relcomp haftmann@47549: rel_comp_def ~> relcomp_unfold bulwahn@47448: converse_rel_comp ~> converse_relcomp bulwahn@47448: pred_compI ~> relcomppI bulwahn@47448: pred_compE ~> relcomppE bulwahn@47448: pred_comp_bot1 ~> relcompp_bot1 bulwahn@47448: pred_comp_bot2 ~> relcompp_bot2 bulwahn@47448: transp_pred_comp_less_eq ~> transp_relcompp_less_eq bulwahn@47448: pred_comp_mono ~> relcompp_mono bulwahn@47448: pred_comp_distrib ~> relcompp_distrib bulwahn@47448: pred_comp_distrib2 ~> relcompp_distrib2 bulwahn@47448: converse_pred_comp ~> converse_relcompp wenzelm@47464: bulwahn@47448: finite_rel_comp ~> finite_relcomp wenzelm@47464: bulwahn@47448: set_rel_comp ~> set_relcomp bulwahn@47448: bulwahn@47448: INCOMPATIBILITY. bulwahn@47448: haftmann@47550: * Theory Divides: Discontinued redundant theorems about div and mod. haftmann@47550: INCOMPATIBILITY, use the corresponding generic theorems instead. haftmann@47550: haftmann@47550: DIVISION_BY_ZERO ~> div_by_0, mod_by_0 haftmann@47550: zdiv_self ~> div_self haftmann@47550: zmod_self ~> mod_self haftmann@47550: zdiv_zero ~> div_0 haftmann@47550: zmod_zero ~> mod_0 haftmann@47550: zdiv_zmod_equality ~> div_mod_equality2 haftmann@47550: zdiv_zmod_equality2 ~> div_mod_equality haftmann@47550: zmod_zdiv_trivial ~> mod_div_trivial haftmann@47550: zdiv_zminus_zminus ~> div_minus_minus haftmann@47550: zmod_zminus_zminus ~> mod_minus_minus haftmann@47550: zdiv_zminus2 ~> div_minus_right haftmann@47550: zmod_zminus2 ~> mod_minus_right haftmann@47550: zdiv_minus1_right ~> div_minus1_right haftmann@47550: zmod_minus1_right ~> mod_minus1_right haftmann@47550: zdvd_mult_div_cancel ~> dvd_mult_div_cancel haftmann@47550: zmod_zmult1_eq ~> mod_mult_right_eq haftmann@47550: zpower_zmod ~> power_mod haftmann@47550: zdvd_zmod ~> dvd_mod haftmann@47550: zdvd_zmod_imp_zdvd ~> dvd_mod_imp_dvd haftmann@47550: mod_mult_distrib ~> mult_mod_left haftmann@47550: mod_mult_distrib2 ~> mult_mod_right haftmann@47550: haftmann@47550: * Removed redundant theorems nat_mult_2 and nat_mult_2_right; use haftmann@47550: generic mult_2 and mult_2_right instead. INCOMPATIBILITY. haftmann@47550: haftmann@47551: * Finite_Set.fold now qualified. INCOMPATIBILITY. haftmann@47551: haftmann@47552: * Consolidated theorem names concerning fold combinators: haftmann@47550: haftmann@47550: inf_INFI_fold_inf ~> inf_INF_fold_inf haftmann@47550: sup_SUPR_fold_sup ~> sup_SUP_fold_sup haftmann@47550: INFI_fold_inf ~> INF_fold_inf haftmann@47550: SUPR_fold_sup ~> SUP_fold_sup haftmann@47550: union_set ~> union_set_fold haftmann@47550: minus_set ~> minus_set_fold haftmann@47550: INFI_set_fold ~> INF_set_fold haftmann@47550: SUPR_set_fold ~> SUP_set_fold haftmann@47550: INF_code ~> INF_set_foldr haftmann@47550: SUP_code ~> SUP_set_foldr haftmann@47550: foldr.simps ~> foldr.simps (in point-free formulation) haftmann@47550: foldr_fold_rev ~> foldr_conv_fold haftmann@47550: foldl_fold ~> foldl_conv_fold haftmann@47550: foldr_foldr ~> foldr_conv_foldl haftmann@47550: foldl_foldr ~> foldl_conv_foldr haftmann@47552: fold_set_remdups ~> fold_set_fold_remdups haftmann@47552: fold_set ~> fold_set_fold haftmann@47552: fold1_set ~> fold1_set_fold haftmann@47550: haftmann@47550: INCOMPATIBILITY. haftmann@47550: haftmann@47550: * Dropped rarely useful theorems concerning fold combinators: haftmann@47550: foldl_apply, foldl_fun_comm, foldl_rev, fold_weak_invariant, haftmann@47550: rev_foldl_cons, fold_set_remdups, fold_set, fold_set1, haftmann@47550: concat_conv_foldl, foldl_weak_invariant, foldl_invariant, haftmann@47550: foldr_invariant, foldl_absorb0, foldl_foldr1_lemma, foldl_foldr1, haftmann@47550: listsum_conv_fold, listsum_foldl, sort_foldl_insort, foldl_assoc, haftmann@47550: foldr_conv_foldl, start_le_sum, elem_le_sum, sum_eq_0_conv. haftmann@47550: INCOMPATIBILITY. For the common phrases "%xs. List.foldr plus xs 0" haftmann@47550: and "List.foldl plus 0", prefer "List.listsum". Otherwise it can be haftmann@47550: useful to boil down "List.foldr" and "List.foldl" to "List.fold" by haftmann@47550: unfolding "foldr_conv_fold" and "foldl_conv_fold". haftmann@47550: haftmann@47550: * Dropped lemmas minus_set_foldr, union_set_foldr, union_coset_foldr, haftmann@47550: inter_coset_foldr, Inf_fin_set_foldr, Sup_fin_set_foldr, haftmann@47550: Min_fin_set_foldr, Max_fin_set_foldr, Inf_set_foldr, Sup_set_foldr, haftmann@47550: INF_set_foldr, SUP_set_foldr. INCOMPATIBILITY. Prefer corresponding haftmann@47550: lemmas over fold rather than foldr, or make use of lemmas haftmann@47550: fold_conv_foldr and fold_rev. haftmann@47550: haftmann@47550: * Congruence rules Option.map_cong and Option.bind_cong for recursion haftmann@47550: through option types. haftmann@47550: wenzelm@47856: * "Transitive_Closure.ntrancl": bounded transitive closure on wenzelm@47856: relations. wenzelm@47856: wenzelm@47856: * Constant "Set.not_member" now qualified. INCOMPATIBILITY. wenzelm@47856: wenzelm@47856: * Theory Int: Discontinued many legacy theorems specific to type int. wenzelm@47856: INCOMPATIBILITY, use the corresponding generic theorems instead. wenzelm@47856: wenzelm@47856: zminus_zminus ~> minus_minus wenzelm@47856: zminus_0 ~> minus_zero wenzelm@47856: zminus_zadd_distrib ~> minus_add_distrib wenzelm@47856: zadd_commute ~> add_commute wenzelm@47856: zadd_assoc ~> add_assoc wenzelm@47856: zadd_left_commute ~> add_left_commute wenzelm@47856: zadd_ac ~> add_ac wenzelm@47856: zmult_ac ~> mult_ac wenzelm@47856: zadd_0 ~> add_0_left wenzelm@47856: zadd_0_right ~> add_0_right wenzelm@47856: zadd_zminus_inverse2 ~> left_minus wenzelm@47856: zmult_zminus ~> mult_minus_left wenzelm@47856: zmult_commute ~> mult_commute wenzelm@47856: zmult_assoc ~> mult_assoc wenzelm@47856: zadd_zmult_distrib ~> left_distrib wenzelm@47856: zadd_zmult_distrib2 ~> right_distrib wenzelm@47856: zdiff_zmult_distrib ~> left_diff_distrib wenzelm@47856: zdiff_zmult_distrib2 ~> right_diff_distrib wenzelm@47856: zmult_1 ~> mult_1_left wenzelm@47856: zmult_1_right ~> mult_1_right wenzelm@47856: zle_refl ~> order_refl wenzelm@47856: zle_trans ~> order_trans wenzelm@47856: zle_antisym ~> order_antisym wenzelm@47856: zle_linear ~> linorder_linear wenzelm@47856: zless_linear ~> linorder_less_linear wenzelm@47856: zadd_left_mono ~> add_left_mono wenzelm@47856: zadd_strict_right_mono ~> add_strict_right_mono wenzelm@47856: zadd_zless_mono ~> add_less_le_mono wenzelm@47856: int_0_less_1 ~> zero_less_one wenzelm@47856: int_0_neq_1 ~> zero_neq_one wenzelm@47856: zless_le ~> less_le wenzelm@47856: zpower_zadd_distrib ~> power_add wenzelm@47856: zero_less_zpower_abs_iff ~> zero_less_power_abs_iff wenzelm@47856: zero_le_zpower_abs ~> zero_le_power_abs wenzelm@47856: wenzelm@47856: * Theory Deriv: Renamed wenzelm@47856: wenzelm@47856: DERIV_nonneg_imp_nonincreasing ~> DERIV_nonneg_imp_nondecreasing wenzelm@47856: wenzelm@47856: * Theory Library/Multiset: Improved code generation of multisets. wenzelm@47856: wenzelm@47855: * Theory HOL/Library/Set_Algebras: Addition and multiplication on sets krauss@47703: are expressed via type classes again. The special syntax krauss@47703: \/\ has been replaced by plain +/*. Removed constant krauss@47703: setsum_set, which is now subsumed by Big_Operators.setsum. krauss@47703: INCOMPATIBILITY. krauss@47703: wenzelm@46160: * Theory HOL/Library/Diagonalize has been removed. INCOMPATIBILITY, wenzelm@46160: use theory HOL/Library/Nat_Bijection instead. wenzelm@46160: wenzelm@47464: * Theory HOL/Library/RBT_Impl: Backing implementation of red-black wenzelm@47464: trees is now inside a type class context. Names of affected wenzelm@47464: operations and lemmas have been prefixed by rbt_. INCOMPATIBILITY for wenzelm@47464: theories working directly with raw red-black trees, adapt the names as wenzelm@47464: follows: Andreas@47452: Andreas@47452: Operations: Andreas@47452: bulkload -> rbt_bulkload Andreas@47452: del_from_left -> rbt_del_from_left Andreas@47452: del_from_right -> rbt_del_from_right Andreas@47452: del -> rbt_del Andreas@47452: delete -> rbt_delete Andreas@47452: ins -> rbt_ins Andreas@47452: insert -> rbt_insert Andreas@47452: insertw -> rbt_insert_with Andreas@47452: insert_with_key -> rbt_insert_with_key Andreas@47452: map_entry -> rbt_map_entry Andreas@47452: lookup -> rbt_lookup Andreas@47452: sorted -> rbt_sorted Andreas@47452: tree_greater -> rbt_greater Andreas@47452: tree_less -> rbt_less Andreas@47452: tree_less_symbol -> rbt_less_symbol Andreas@47452: union -> rbt_union Andreas@47452: union_with -> rbt_union_with Andreas@47452: union_with_key -> rbt_union_with_key Andreas@47452: Andreas@47452: Lemmas: Andreas@47452: balance_left_sorted -> balance_left_rbt_sorted Andreas@47452: balance_left_tree_greater -> balance_left_rbt_greater Andreas@47452: balance_left_tree_less -> balance_left_rbt_less Andreas@47452: balance_right_sorted -> balance_right_rbt_sorted Andreas@47452: balance_right_tree_greater -> balance_right_rbt_greater Andreas@47452: balance_right_tree_less -> balance_right_rbt_less Andreas@47452: balance_sorted -> balance_rbt_sorted Andreas@47452: balance_tree_greater -> balance_rbt_greater Andreas@47452: balance_tree_less -> balance_rbt_less Andreas@47452: bulkload_is_rbt -> rbt_bulkload_is_rbt Andreas@47452: combine_sorted -> combine_rbt_sorted Andreas@47452: combine_tree_greater -> combine_rbt_greater Andreas@47452: combine_tree_less -> combine_rbt_less Andreas@47452: delete_in_tree -> rbt_delete_in_tree Andreas@47452: delete_is_rbt -> rbt_delete_is_rbt Andreas@47452: del_from_left_tree_greater -> rbt_del_from_left_rbt_greater Andreas@47452: del_from_left_tree_less -> rbt_del_from_left_rbt_less Andreas@47452: del_from_right_tree_greater -> rbt_del_from_right_rbt_greater Andreas@47452: del_from_right_tree_less -> rbt_del_from_right_rbt_less Andreas@47452: del_in_tree -> rbt_del_in_tree Andreas@47452: del_inv1_inv2 -> rbt_del_inv1_inv2 Andreas@47452: del_sorted -> rbt_del_rbt_sorted Andreas@47452: del_tree_greater -> rbt_del_rbt_greater Andreas@47452: del_tree_less -> rbt_del_rbt_less Andreas@47452: dom_lookup_Branch -> dom_rbt_lookup_Branch Andreas@47452: entries_lookup -> entries_rbt_lookup Andreas@47452: finite_dom_lookup -> finite_dom_rbt_lookup Andreas@47452: insert_sorted -> rbt_insert_rbt_sorted Andreas@47452: insertw_is_rbt -> rbt_insertw_is_rbt Andreas@47452: insertwk_is_rbt -> rbt_insertwk_is_rbt Andreas@47452: insertwk_sorted -> rbt_insertwk_rbt_sorted Andreas@47452: insertw_sorted -> rbt_insertw_rbt_sorted Andreas@47452: ins_sorted -> ins_rbt_sorted Andreas@47452: ins_tree_greater -> ins_rbt_greater Andreas@47452: ins_tree_less -> ins_rbt_less Andreas@47452: is_rbt_sorted -> is_rbt_rbt_sorted Andreas@47452: lookup_balance -> rbt_lookup_balance Andreas@47452: lookup_bulkload -> rbt_lookup_rbt_bulkload Andreas@47452: lookup_delete -> rbt_lookup_rbt_delete Andreas@47452: lookup_Empty -> rbt_lookup_Empty Andreas@47452: lookup_from_in_tree -> rbt_lookup_from_in_tree Andreas@47452: lookup_in_tree -> rbt_lookup_in_tree Andreas@47452: lookup_ins -> rbt_lookup_ins Andreas@47452: lookup_insert -> rbt_lookup_rbt_insert Andreas@47452: lookup_insertw -> rbt_lookup_rbt_insertw Andreas@47452: lookup_insertwk -> rbt_lookup_rbt_insertwk Andreas@47452: lookup_keys -> rbt_lookup_keys Andreas@47452: lookup_map -> rbt_lookup_map Andreas@47452: lookup_map_entry -> rbt_lookup_rbt_map_entry Andreas@47452: lookup_tree_greater -> rbt_lookup_rbt_greater Andreas@47452: lookup_tree_less -> rbt_lookup_rbt_less Andreas@47452: lookup_union -> rbt_lookup_rbt_union Andreas@47452: map_entry_color_of -> rbt_map_entry_color_of Andreas@47452: map_entry_inv1 -> rbt_map_entry_inv1 Andreas@47452: map_entry_inv2 -> rbt_map_entry_inv2 Andreas@47452: map_entry_is_rbt -> rbt_map_entry_is_rbt Andreas@47452: map_entry_sorted -> rbt_map_entry_rbt_sorted Andreas@47452: map_entry_tree_greater -> rbt_map_entry_rbt_greater Andreas@47452: map_entry_tree_less -> rbt_map_entry_rbt_less Andreas@47452: map_tree_greater -> map_rbt_greater Andreas@47452: map_tree_less -> map_rbt_less Andreas@47452: map_sorted -> map_rbt_sorted Andreas@47452: paint_sorted -> paint_rbt_sorted Andreas@47452: paint_lookup -> paint_rbt_lookup Andreas@47452: paint_tree_greater -> paint_rbt_greater Andreas@47452: paint_tree_less -> paint_rbt_less Andreas@47452: sorted_entries -> rbt_sorted_entries Andreas@47452: tree_greater_eq_trans -> rbt_greater_eq_trans Andreas@47452: tree_greater_nit -> rbt_greater_nit Andreas@47452: tree_greater_prop -> rbt_greater_prop Andreas@47452: tree_greater_simps -> rbt_greater_simps Andreas@47452: tree_greater_trans -> rbt_greater_trans Andreas@47452: tree_less_eq_trans -> rbt_less_eq_trans Andreas@47452: tree_less_nit -> rbt_less_nit Andreas@47452: tree_less_prop -> rbt_less_prop Andreas@47452: tree_less_simps -> rbt_less_simps Andreas@47452: tree_less_trans -> rbt_less_trans Andreas@47452: tree_ord_props -> rbt_ord_props Andreas@47452: union_Branch -> rbt_union_Branch Andreas@47452: union_is_rbt -> rbt_union_is_rbt Andreas@47452: unionw_is_rbt -> rbt_unionw_is_rbt Andreas@47452: unionwk_is_rbt -> rbt_unionwk_is_rbt Andreas@47452: unionwk_sorted -> rbt_unionwk_rbt_sorted Andreas@47452: wenzelm@47807: * Theory HOL/Library/Float: Floating point numbers are now defined as wenzelm@47807: a subset of the real numbers. All operations are defined using the wenzelm@47807: lifing-framework and proofs use the transfer method. INCOMPATIBILITY. hoelzl@47616: hoelzl@47616: Changed Operations: hoelzl@47622: float_abs -> abs hoelzl@47622: float_nprt -> nprt hoelzl@47622: float_pprt -> pprt hoelzl@47622: pow2 -> use powr hoelzl@47622: round_down -> float_round_down hoelzl@47622: round_up -> float_round_up hoelzl@47622: scale -> exponent hoelzl@47622: hoelzl@47622: Removed Operations: hoelzl@47622: ceiling_fl, lb_mult, lb_mod, ub_mult, ub_mod hoelzl@47622: hoelzl@47622: Renamed Lemmas: hoelzl@47622: abs_float_def -> Float.compute_float_abs hoelzl@47622: bitlen_ge0 -> bitlen_nonneg hoelzl@47622: bitlen.simps -> Float.compute_bitlen hoelzl@47622: float_components -> Float_mantissa_exponent hoelzl@47622: float_divl.simps -> Float.compute_float_divl hoelzl@47622: float_divr.simps -> Float.compute_float_divr hoelzl@47622: float_eq_odd -> mult_powr_eq_mult_powr_iff hoelzl@47622: float_power -> real_of_float_power hoelzl@47622: lapprox_posrat_def -> Float.compute_lapprox_posrat hoelzl@47622: lapprox_rat.simps -> Float.compute_lapprox_rat hoelzl@47622: le_float_def' -> Float.compute_float_le hoelzl@47622: le_float_def -> less_eq_float.rep_eq hoelzl@47622: less_float_def' -> Float.compute_float_less hoelzl@47622: less_float_def -> less_float.rep_eq hoelzl@47622: normfloat_def -> Float.compute_normfloat hoelzl@47622: normfloat_imp_odd_or_zero -> mantissa_not_dvd and mantissa_noteq_0 hoelzl@47622: normfloat -> normfloat_def hoelzl@47622: normfloat_unique -> use normfloat_def hoelzl@47622: number_of_float_Float -> Float.compute_float_numeral, Float.compute_float_neg_numeral hoelzl@47622: one_float_def -> Float.compute_float_one hoelzl@47622: plus_float_def -> Float.compute_float_plus hoelzl@47622: rapprox_posrat_def -> Float.compute_rapprox_posrat hoelzl@47622: rapprox_rat.simps -> Float.compute_rapprox_rat hoelzl@47622: real_of_float_0 -> zero_float.rep_eq hoelzl@47622: real_of_float_1 -> one_float.rep_eq hoelzl@47622: real_of_float_abs -> abs_float.rep_eq hoelzl@47622: real_of_float_add -> plus_float.rep_eq hoelzl@47622: real_of_float_minus -> uminus_float.rep_eq hoelzl@47622: real_of_float_mult -> times_float.rep_eq hoelzl@47622: real_of_float_simp -> Float.rep_eq hoelzl@47622: real_of_float_sub -> minus_float.rep_eq hoelzl@47622: round_down.simps -> Float.compute_float_round_down hoelzl@47622: round_up.simps -> Float.compute_float_round_up hoelzl@47622: times_float_def -> Float.compute_float_times hoelzl@47622: uminus_float_def -> Float.compute_float_uminus hoelzl@47622: zero_float_def -> Float.compute_float_zero hoelzl@47622: hoelzl@47622: Lemmas not necessary anymore, use the transfer method: hoelzl@47622: bitlen_B0, bitlen_B1, bitlen_ge1, bitlen_Min, bitlen_Pls, float_divl, hoelzl@47622: float_divr, float_le_simp, float_less1_mantissa_bound, hoelzl@47622: float_less_simp, float_less_zero, float_le_zero, hoelzl@47622: float_pos_less1_e_neg, float_pos_m_pos, float_split, float_split2, hoelzl@47622: floor_pos_exp, lapprox_posrat, lapprox_posrat_bottom, lapprox_rat, hoelzl@47622: lapprox_rat_bottom, normalized_float, rapprox_posrat, hoelzl@47622: rapprox_posrat_le1, rapprox_rat, real_of_float_ge0_exp, hoelzl@47622: real_of_float_neg_exp, real_of_float_nge0_exp, round_down floor_fl, hoelzl@47622: round_up, zero_le_float, zero_less_float hoelzl@47616: wenzelm@47856: * New theory HOL/Library/DAList provides an abstract type for wenzelm@47856: association lists with distinct keys. noschinl@45791: wenzelm@47866: * Session HOL/IMP: Added new theory of abstract interpretation of wenzelm@47866: annotated commands. wenzelm@47866: wenzelm@47855: * Session HOL-Import: Re-implementation from scratch is faster, wenzelm@47855: simpler, and more scalable. Requires a proof bundle, which is wenzelm@47855: available as an external component. Discontinued old (and mostly wenzelm@47855: dead) Importer for HOL4 and HOL Light. INCOMPATIBILITY. wenzelm@47855: wenzelm@47855: * Session HOL-Word: Discontinued many redundant theorems specific to wenzelm@47855: type 'a word. INCOMPATIBILITY, use the corresponding generic theorems wenzelm@47855: instead. wenzelm@47855: wenzelm@47855: word_sub_alt ~> word_sub_wi wenzelm@47855: word_add_alt ~> word_add_def wenzelm@47855: word_mult_alt ~> word_mult_def wenzelm@47855: word_minus_alt ~> word_minus_def wenzelm@47855: word_0_alt ~> word_0_wi wenzelm@47855: word_1_alt ~> word_1_wi wenzelm@47855: word_add_0 ~> add_0_left wenzelm@47855: word_add_0_right ~> add_0_right wenzelm@47855: word_mult_1 ~> mult_1_left wenzelm@47855: word_mult_1_right ~> mult_1_right wenzelm@47855: word_add_commute ~> add_commute wenzelm@47855: word_add_assoc ~> add_assoc wenzelm@47855: word_add_left_commute ~> add_left_commute wenzelm@47855: word_mult_commute ~> mult_commute wenzelm@47855: word_mult_assoc ~> mult_assoc wenzelm@47855: word_mult_left_commute ~> mult_left_commute wenzelm@47855: word_left_distrib ~> left_distrib wenzelm@47855: word_right_distrib ~> right_distrib wenzelm@47855: word_left_minus ~> left_minus wenzelm@47855: word_diff_0_right ~> diff_0_right wenzelm@47855: word_diff_self ~> diff_self wenzelm@47855: word_sub_def ~> diff_minus wenzelm@47855: word_diff_minus ~> diff_minus wenzelm@47855: word_add_ac ~> add_ac wenzelm@47855: word_mult_ac ~> mult_ac wenzelm@47855: word_plus_ac0 ~> add_0_left add_0_right add_ac wenzelm@47855: word_times_ac1 ~> mult_1_left mult_1_right mult_ac wenzelm@47855: word_order_trans ~> order_trans wenzelm@47855: word_order_refl ~> order_refl wenzelm@47855: word_order_antisym ~> order_antisym wenzelm@47855: word_order_linear ~> linorder_linear wenzelm@47855: lenw1_zero_neq_one ~> zero_neq_one wenzelm@47855: word_number_of_eq ~> number_of_eq wenzelm@47855: word_of_int_add_hom ~> wi_hom_add wenzelm@47855: word_of_int_sub_hom ~> wi_hom_sub wenzelm@47855: word_of_int_mult_hom ~> wi_hom_mult wenzelm@47855: word_of_int_minus_hom ~> wi_hom_neg wenzelm@47855: word_of_int_succ_hom ~> wi_hom_succ wenzelm@47855: word_of_int_pred_hom ~> wi_hom_pred wenzelm@47855: word_of_int_0_hom ~> word_0_wi wenzelm@47855: word_of_int_1_hom ~> word_1_wi wenzelm@47855: wenzelm@47809: * Session HOL-Word: New proof method "word_bitwise" for splitting wenzelm@47809: machine word equalities and inequalities into logical circuits, wenzelm@47809: defined in HOL/Word/WordBitwise.thy. Supports addition, subtraction, wenzelm@47809: multiplication, shifting by constants, bitwise operators and numeric wenzelm@47809: constants. Requires fixed-length word types, not 'a word. Solves wenzelm@47854: many standard word identities outright and converts more into first wenzelm@47809: order problems amenable to blast or similar. See also examples in wenzelm@47809: HOL/Word/Examples/WordExamples.thy. wenzelm@47809: wenzelm@47807: * Session HOL-Probability: Introduced the type "'a measure" to wenzelm@47807: represent measures, this replaces the records 'a algebra and 'a wenzelm@47807: measure_space. The locales based on subset_class now have two wenzelm@47856: locale-parameters the space \ and the set of measurable sets M. wenzelm@47856: The product of probability spaces uses now the same constant as the wenzelm@47856: finite product of sigma-finite measure spaces "PiM :: ('i => 'a) wenzelm@47807: measure". Most constants are defined now outside of locales and gain wenzelm@47807: an additional parameter, like null_sets, almost_eventually or \'. wenzelm@47807: Measure space constructions for distributions and densities now got wenzelm@47807: their own constants distr and density. Instead of using locales to wenzelm@47807: describe measure spaces with a finite space, the measure count_space wenzelm@47807: and point_measure is introduced. INCOMPATIBILITY. hoelzl@47694: hoelzl@47694: Renamed constants: hoelzl@47694: measure -> emeasure hoelzl@47694: finite_measure.\' -> measure hoelzl@47694: product_algebra_generator -> prod_algebra hoelzl@47694: product_prob_space.emb -> prod_emb hoelzl@47694: product_prob_space.infprod_algebra -> PiM hoelzl@47694: hoelzl@47694: Removed locales: hoelzl@47694: completeable_measure_space hoelzl@47694: finite_measure_space hoelzl@47694: finite_prob_space hoelzl@47694: finite_product_finite_prob_space hoelzl@47694: finite_product_sigma_algebra hoelzl@47694: finite_sigma_algebra hoelzl@47694: measure_space hoelzl@47694: pair_finite_prob_space hoelzl@47694: pair_finite_sigma_algebra hoelzl@47694: pair_finite_space hoelzl@47694: pair_sigma_algebra hoelzl@47694: product_sigma_algebra hoelzl@47694: hoelzl@47694: Removed constants: hoelzl@47751: conditional_space hoelzl@47694: distribution -> use distr measure, or distributed predicate hoelzl@47751: image_space hoelzl@47694: joint_distribution -> use distr measure, or distributed predicate hoelzl@47751: pair_measure_generator hoelzl@47694: product_prob_space.infprod_algebra -> use PiM hoelzl@47694: subvimage hoelzl@47694: hoelzl@47694: Replacement theorems: hoelzl@47751: finite_additivity_sufficient -> ring_of_sets.countably_additiveI_finite hoelzl@47751: finite_measure.empty_measure -> measure_empty hoelzl@47751: finite_measure.finite_continuity_from_above -> finite_measure.finite_Lim_measure_decseq hoelzl@47751: finite_measure.finite_continuity_from_below -> finite_measure.finite_Lim_measure_incseq hoelzl@47751: finite_measure.finite_measure_countably_subadditive -> finite_measure.finite_measure_subadditive_countably hoelzl@47751: finite_measure.finite_measure_eq -> finite_measure.emeasure_eq_measure hoelzl@47751: finite_measure.finite_measure -> finite_measure.emeasure_finite hoelzl@47751: finite_measure.finite_measure_finite_singleton -> finite_measure.finite_measure_eq_setsum_singleton hoelzl@47751: finite_measure.positive_measure' -> measure_nonneg hoelzl@47751: finite_measure.real_measure -> finite_measure.emeasure_real hoelzl@47751: finite_product_prob_space.finite_measure_times -> finite_product_prob_space.finite_measure_PiM_emb hoelzl@47751: finite_product_sigma_algebra.in_P -> sets_PiM_I_finite hoelzl@47751: finite_product_sigma_algebra.P_empty -> space_PiM_empty, sets_PiM_empty hoelzl@47751: information_space.conditional_entropy_eq -> information_space.conditional_entropy_simple_distributed hoelzl@47751: information_space.conditional_entropy_positive -> information_space.conditional_entropy_nonneg_simple hoelzl@47751: information_space.conditional_mutual_information_eq_mutual_information -> information_space.conditional_mutual_information_eq_mutual_information_simple hoelzl@47751: information_space.conditional_mutual_information_generic_positive -> information_space.conditional_mutual_information_nonneg_simple hoelzl@47751: information_space.conditional_mutual_information_positive -> information_space.conditional_mutual_information_nonneg_simple hoelzl@47751: information_space.entropy_commute -> information_space.entropy_commute_simple hoelzl@47751: information_space.entropy_eq -> information_space.entropy_simple_distributed hoelzl@47751: information_space.entropy_generic_eq -> information_space.entropy_simple_distributed hoelzl@47751: information_space.entropy_positive -> information_space.entropy_nonneg_simple hoelzl@47751: information_space.entropy_uniform_max -> information_space.entropy_uniform hoelzl@47751: information_space.KL_eq_0_imp -> information_space.KL_eq_0_iff_eq hoelzl@47751: information_space.KL_eq_0 -> information_space.KL_same_eq_0 hoelzl@47751: information_space.KL_ge_0 -> information_space.KL_nonneg hoelzl@47751: information_space.mutual_information_eq -> information_space.mutual_information_simple_distributed hoelzl@47751: information_space.mutual_information_positive -> information_space.mutual_information_nonneg_simple hoelzl@47751: Int_stable_cuboids -> Int_stable_atLeastAtMost hoelzl@47751: Int_stable_product_algebra_generator -> positive_integral hoelzl@47751: measure_preserving -> equality "distr M N f = N" "f : measurable M N" hoelzl@47694: measure_space.additive -> emeasure_additive hoelzl@47751: measure_space.AE_iff_null_set -> AE_iff_null hoelzl@47751: measure_space.almost_everywhere_def -> eventually_ae_filter hoelzl@47751: measure_space.almost_everywhere_vimage -> AE_distrD hoelzl@47751: measure_space.continuity_from_above -> INF_emeasure_decseq hoelzl@47751: measure_space.continuity_from_above_Lim -> Lim_emeasure_decseq hoelzl@47751: measure_space.continuity_from_below_Lim -> Lim_emeasure_incseq hoelzl@47694: measure_space.continuity_from_below -> SUP_emeasure_incseq hoelzl@47751: measure_space_density -> emeasure_density hoelzl@47751: measure_space.density_is_absolutely_continuous -> absolutely_continuousI_density hoelzl@47751: measure_space.integrable_vimage -> integrable_distr hoelzl@47751: measure_space.integral_translated_density -> integral_density hoelzl@47751: measure_space.integral_vimage -> integral_distr hoelzl@47751: measure_space.measure_additive -> plus_emeasure hoelzl@47751: measure_space.measure_compl -> emeasure_compl hoelzl@47751: measure_space.measure_countable_increasing -> emeasure_countable_increasing hoelzl@47751: measure_space.measure_countably_subadditive -> emeasure_subadditive_countably hoelzl@47694: measure_space.measure_decseq -> decseq_emeasure hoelzl@47751: measure_space.measure_Diff -> emeasure_Diff hoelzl@47751: measure_space.measure_Diff_null_set -> emeasure_Diff_null_set hoelzl@47694: measure_space.measure_eq_0 -> emeasure_eq_0 hoelzl@47694: measure_space.measure_finitely_subadditive -> emeasure_subadditive_finite hoelzl@47751: measure_space.measure_finite_singleton -> emeasure_eq_setsum_singleton hoelzl@47751: measure_space.measure_incseq -> incseq_emeasure hoelzl@47751: measure_space.measure_insert -> emeasure_insert hoelzl@47751: measure_space.measure_mono -> emeasure_mono hoelzl@47751: measure_space.measure_not_negative -> emeasure_not_MInf hoelzl@47751: measure_space.measure_preserving_Int_stable -> measure_eqI_generator_eq hoelzl@47751: measure_space.measure_setsum -> setsum_emeasure hoelzl@47751: measure_space.measure_setsum_split -> setsum_emeasure_cover hoelzl@47694: measure_space.measure_space_vimage -> emeasure_distr hoelzl@47751: measure_space.measure_subadditive_finite -> emeasure_subadditive_finite hoelzl@47751: measure_space.measure_subadditive -> subadditive hoelzl@47751: measure_space.measure_top -> emeasure_space hoelzl@47751: measure_space.measure_UN_eq_0 -> emeasure_UN_eq_0 hoelzl@47751: measure_space.measure_Un_null_set -> emeasure_Un_null_set hoelzl@47751: measure_space.positive_integral_translated_density -> positive_integral_density hoelzl@47751: measure_space.positive_integral_vimage -> positive_integral_distr hoelzl@47694: measure_space.real_continuity_from_above -> Lim_measure_decseq hoelzl@47751: measure_space.real_continuity_from_below -> Lim_measure_incseq hoelzl@47694: measure_space.real_measure_countably_subadditive -> measure_subadditive_countably hoelzl@47751: measure_space.real_measure_Diff -> measure_Diff hoelzl@47751: measure_space.real_measure_finite_Union -> measure_finite_Union hoelzl@47751: measure_space.real_measure_setsum_singleton -> measure_eq_setsum_singleton hoelzl@47751: measure_space.real_measure_subadditive -> measure_subadditive hoelzl@47751: measure_space.real_measure_Union -> measure_Union hoelzl@47751: measure_space.real_measure_UNION -> measure_UNION hoelzl@47694: measure_space.simple_function_vimage -> simple_function_comp hoelzl@47694: measure_space.simple_integral_vimage -> simple_integral_distr hoelzl@47751: measure_space.simple_integral_vimage -> simple_integral_distr hoelzl@47751: measure_unique_Int_stable -> measure_eqI_generator_eq hoelzl@47751: measure_unique_Int_stable_vimage -> measure_eqI_generator_eq hoelzl@47694: pair_sigma_algebra.measurable_cut_fst -> sets_Pair1 hoelzl@47694: pair_sigma_algebra.measurable_cut_snd -> sets_Pair2 hoelzl@47694: pair_sigma_algebra.measurable_pair_image_fst -> measurable_Pair1 hoelzl@47694: pair_sigma_algebra.measurable_pair_image_snd -> measurable_Pair2 hoelzl@47694: pair_sigma_algebra.measurable_product_swap -> measurable_pair_swap_iff hoelzl@47694: pair_sigma_algebra.pair_sigma_algebra_measurable -> measurable_pair_swap hoelzl@47694: pair_sigma_algebra.pair_sigma_algebra_swap_measurable -> measurable_pair_swap' hoelzl@47694: pair_sigma_algebra.sets_swap -> sets_pair_swap hoelzl@47751: pair_sigma_finite.measure_cut_measurable_fst -> pair_sigma_finite.measurable_emeasure_Pair1 hoelzl@47751: pair_sigma_finite.measure_cut_measurable_snd -> pair_sigma_finite.measurable_emeasure_Pair2 hoelzl@47751: pair_sigma_finite.measure_preserving_swap -> pair_sigma_finite.distr_pair_swap hoelzl@47751: pair_sigma_finite.pair_measure_alt2 -> pair_sigma_finite.emeasure_pair_measure_alt2 hoelzl@47751: pair_sigma_finite.pair_measure_alt -> pair_sigma_finite.emeasure_pair_measure_alt hoelzl@47751: pair_sigma_finite.pair_measure_times -> pair_sigma_finite.emeasure_pair_measure_Times hoelzl@47751: prob_space.indep_distribution_eq_measure -> prob_space.indep_vars_iff_distr_eq_PiM hoelzl@47751: prob_space.indep_var_distributionD -> prob_space.indep_var_distribution_eq hoelzl@47694: prob_space.measure_space_1 -> prob_space.emeasure_space_1 hoelzl@47694: prob_space.prob_space_vimage -> prob_space_distr hoelzl@47694: prob_space.random_variable_restrict -> measurable_restrict hoelzl@47751: prob_space_unique_Int_stable -> measure_eqI_prob_space hoelzl@47751: product_algebraE -> prod_algebraE_all hoelzl@47751: product_algebra_generator_der -> prod_algebra_eq_finite hoelzl@47751: product_algebra_generator_into_space -> prod_algebra_sets_into_space hoelzl@47751: product_algebraI -> sets_PiM_I_finite hoelzl@47751: product_measure_exists -> product_sigma_finite.sigma_finite hoelzl@47694: product_prob_space.finite_index_eq_finite_product -> product_prob_space.sets_PiM_generator hoelzl@47694: product_prob_space.finite_measure_infprod_emb_Pi -> product_prob_space.measure_PiM_emb hoelzl@47694: product_prob_space.infprod_spec -> product_prob_space.emeasure_PiM_emb_not_empty hoelzl@47694: product_prob_space.measurable_component -> measurable_component_singleton hoelzl@47694: product_prob_space.measurable_emb -> measurable_prod_emb hoelzl@47694: product_prob_space.measurable_into_infprod_algebra -> measurable_PiM_single hoelzl@47694: product_prob_space.measurable_singleton_infprod -> measurable_component_singleton hoelzl@47694: product_prob_space.measure_emb -> emeasure_prod_emb hoelzl@47751: product_prob_space.measure_preserving_restrict -> product_prob_space.distr_restrict hoelzl@47751: product_sigma_algebra.product_algebra_into_space -> space_closed hoelzl@47751: product_sigma_finite.measure_fold -> product_sigma_finite.distr_merge hoelzl@47751: product_sigma_finite.measure_preserving_component_singelton -> product_sigma_finite.distr_singleton hoelzl@47751: product_sigma_finite.measure_preserving_merge -> product_sigma_finite.distr_merge hoelzl@47694: sequence_space.measure_infprod -> sequence_space.measure_PiM_countable hoelzl@47751: sets_product_algebra -> sets_PiM hoelzl@47751: sigma_algebra.measurable_sigma -> measurable_measure_of hoelzl@47751: sigma_finite_measure.disjoint_sigma_finite -> sigma_finite_disjoint hoelzl@47751: sigma_finite_measure.RN_deriv_vimage -> sigma_finite_measure.RN_deriv_distr hoelzl@47751: sigma_product_algebra_sigma_eq -> sigma_prod_algebra_sigma_eq hoelzl@47751: space_product_algebra -> space_PiM hoelzl@47694: wenzelm@47855: * Session HOL-TPTP: support to parse and import TPTP problems (all wenzelm@47855: languages) into Isabelle/HOL. wenzelm@47413: blanchet@45398: noschinl@45160: *** FOL *** noschinl@45160: wenzelm@45383: * New "case_product" attribute (see HOL). noschinl@45160: wenzelm@45109: wenzelm@47463: *** ZF *** wenzelm@47463: wenzelm@47463: * Greater support for structured proofs involving induction or case wenzelm@47463: analysis. wenzelm@47463: wenzelm@47463: * Much greater use of mathematical symbols. wenzelm@47463: wenzelm@47463: * Removal of many ML theorem bindings. INCOMPATIBILITY. wenzelm@47463: wenzelm@47463: wenzelm@45128: *** ML *** wenzelm@45128: wenzelm@46948: * Antiquotation @{keyword "name"} produces a parser for outer syntax wenzelm@46948: from a minor keyword introduced via theory header declaration. wenzelm@46948: wenzelm@46961: * Antiquotation @{command_spec "name"} produces the wenzelm@46961: Outer_Syntax.command_spec from a major keyword introduced via theory wenzelm@46961: header declaration; it can be passed to Outer_Syntax.command etc. wenzelm@46961: wenzelm@46916: * Local_Theory.define no longer hard-wires default theorem name wenzelm@46992: "foo_def", but retains the binding as given. If that is Binding.empty wenzelm@46992: / Attrib.empty_binding, the result is not registered as user-level wenzelm@46992: fact. The Local_Theory.define_internal variant allows to specify a wenzelm@46992: non-empty name (used for the foundation in the background theory), wenzelm@46992: while omitting the fact binding in the user-context. Potential wenzelm@46992: INCOMPATIBILITY for derived definitional packages: need to specify wenzelm@46992: naming policy for primitive definitions more explicitly. wenzelm@46916: wenzelm@46497: * Renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in wenzelm@46497: conformance with similar operations in structure Term and Logic. wenzelm@46497: wenzelm@45592: * Antiquotation @{attributes [...]} embeds attribute source wenzelm@45592: representation into the ML text, which is particularly useful with wenzelm@45592: declarations like Local_Theory.note. wenzelm@45592: wenzelm@45128: * Structure Proof_Context follows standard naming scheme. Old wenzelm@45128: ProofContext has been discontinued. INCOMPATIBILITY. wenzelm@45128: wenzelm@45293: * Refined Local_Theory.declaration {syntax, pervasive}, with subtle wenzelm@45298: change of semantics: update is applied to auxiliary local theory wenzelm@45293: context as well. wenzelm@45293: wenzelm@45620: * Modernized some old-style infix operations: wenzelm@45620: wenzelm@45620: addeqcongs ~> Simplifier.add_eqcong wenzelm@45620: deleqcongs ~> Simplifier.del_eqcong wenzelm@45620: addcongs ~> Simplifier.add_cong wenzelm@45620: delcongs ~> Simplifier.del_cong wenzelm@45625: setmksimps ~> Simplifier.set_mksimps wenzelm@45625: setmkcong ~> Simplifier.set_mkcong wenzelm@45625: setmksym ~> Simplifier.set_mksym wenzelm@45625: setmkeqTrue ~> Simplifier.set_mkeqTrue wenzelm@45625: settermless ~> Simplifier.set_termless wenzelm@45625: setsubgoaler ~> Simplifier.set_subgoaler wenzelm@45620: addsplits ~> Splitter.add_split wenzelm@45620: delsplits ~> Splitter.del_split wenzelm@45620: wenzelm@45128: wenzelm@47461: *** System *** wenzelm@47461: wenzelm@47661: * USER_HOME settings variable points to cross-platform user home wenzelm@47661: directory, which coincides with HOME on POSIX systems only. Likewise, wenzelm@47661: the Isabelle path specification "~" now expands to $USER_HOME, instead wenzelm@47661: of former $HOME. A different default for USER_HOME may be set wenzelm@47661: explicitly in shell environment, before Isabelle settings are wenzelm@47661: evaluated. Minor INCOMPATIBILITY: need to adapt Isabelle path where wenzelm@47661: the generic user home was intended. wenzelm@47661: wenzelm@47807: * ISABELLE_HOME_WINDOWS refers to ISABELLE_HOME in windows file name wenzelm@47807: notation, which is useful for the jEdit file browser, for example. wenzelm@47807: wenzelm@47464: * ISABELLE_JDK_HOME settings variable points to JDK with javac and jar wenzelm@47464: (not just JRE). wenzelm@47464: wenzelm@47461: wenzelm@45109: wenzelm@44801: New in Isabelle2011-1 (October 2011) wenzelm@44801: ------------------------------------ wenzelm@41651: wenzelm@41703: *** General *** wenzelm@41703: wenzelm@44760: * Improved Isabelle/jEdit Prover IDE (PIDE), which can be invoked as wenzelm@44968: "isabelle jedit" or "ISABELLE_HOME/Isabelle" on the command line. wenzelm@44760: wenzelm@44777: - Management of multiple theory files directly from the editor wenzelm@44760: buffer store -- bypassing the file-system (no requirement to save wenzelm@44760: files for checking). wenzelm@44760: wenzelm@44777: - Markup of formal entities within the text buffer, with semantic wenzelm@44760: highlighting, tooltips and hyperlinks to jump to defining source wenzelm@44760: positions. wenzelm@44760: wenzelm@44777: - Improved text rendering, with sub/superscripts in the source wenzelm@44777: buffer (including support for copy/paste wrt. output panel, HTML wenzelm@44777: theory output and other non-Isabelle text boxes). wenzelm@44777: wenzelm@44777: - Refined scheduling of proof checking and printing of results, wenzelm@44760: based on interactive editor view. (Note: jEdit folding and wenzelm@44760: narrowing allows to restrict buffer perspectives explicitly.) wenzelm@44760: wenzelm@44777: - Reduced CPU performance requirements, usable on machines with few wenzelm@44760: cores. wenzelm@44760: wenzelm@44777: - Reduced memory requirements due to pruning of unused document wenzelm@44760: versions (garbage collection). wenzelm@44760: wenzelm@44760: See also ~~/src/Tools/jEdit/README.html for further information, wenzelm@44760: including some remaining limitations. wenzelm@44760: wenzelm@44800: * Theory loader: source files are exclusively located via the master wenzelm@44800: directory of each theory node (where the .thy file itself resides). wenzelm@44800: The global load path (such as src/HOL/Library) has been discontinued. wenzelm@44800: Note that the path element ~~ may be used to reference theories in the wenzelm@44800: Isabelle home folder -- for instance, "~~/src/HOL/Library/FuncSet". wenzelm@44800: INCOMPATIBILITY. wenzelm@44800: wenzelm@41955: * Theory loader: source files are identified by content via SHA1 wenzelm@41955: digests. Discontinued former path/modtime identification and optional wenzelm@41955: ISABELLE_FILE_IDENT plugin scripts. wenzelm@41955: wenzelm@41703: * Parallelization of nested Isar proofs is subject to wenzelm@41703: Goal.parallel_proofs_threshold (default 100). See also isabelle wenzelm@41703: usedir option -Q. wenzelm@41703: wenzelm@42669: * Name space: former unsynchronized references are now proper wenzelm@42669: configuration options, with more conventional names: wenzelm@42669: wenzelm@42669: long_names ~> names_long wenzelm@42669: short_names ~> names_short wenzelm@42669: unique_names ~> names_unique wenzelm@42669: wenzelm@42669: Minor INCOMPATIBILITY, need to declare options in context like this: wenzelm@42669: wenzelm@42669: declare [[names_unique = false]] wenzelm@42358: wenzelm@42502: * Literal facts `prop` may contain dummy patterns, e.g. `_ = _`. Note wenzelm@42502: that the result needs to be unique, which means fact specifications wenzelm@42502: may have to be refined after enriching a proof context. wenzelm@42502: wenzelm@44800: * Attribute "case_names" has been refined: the assumptions in each case wenzelm@44800: can be named now by following the case name with [name1 name2 ...]. wenzelm@44800: wenzelm@44968: * Isabelle/Isar reference manual has been updated and extended: wenzelm@44968: - "Synopsis" provides a catalog of main Isar language concepts. wenzelm@44968: - Formal references in syntax diagrams, via @{rail} antiquotation. wenzelm@44968: - Updated material from classic "ref" manual, notably about wenzelm@44968: "Classical Reasoner". wenzelm@42633: wenzelm@41703: blanchet@41727: *** HOL *** blanchet@41727: wenzelm@44968: * Class bot and top require underlying partial order rather than wenzelm@44800: preorder: uniqueness of bot and top is guaranteed. INCOMPATIBILITY. haftmann@43815: haftmann@43940: * Class complete_lattice: generalized a couple of lemmas from sets; wenzelm@44800: generalized theorems INF_cong and SUP_cong. New type classes for wenzelm@44800: complete boolean algebras and complete linear orders. Lemmas wenzelm@44800: Inf_less_iff, less_Sup_iff, INF_less_iff, less_SUP_iff now reside in wenzelm@44800: class complete_linorder. wenzelm@44800: wenzelm@44800: Changed proposition of lemmas Inf_bool_def, Sup_bool_def, Inf_fun_def, wenzelm@44800: Sup_fun_def, Inf_apply, Sup_apply. wenzelm@44800: wenzelm@45088: Removed redundant lemmas (the right hand side gives hints how to wenzelm@45088: replace them for (metis ...), or (simp only: ...) proofs): hoelzl@45041: hoelzl@45041: Inf_singleton ~> Inf_insert [where A="{}", unfolded Inf_empty inf_top_right] hoelzl@45041: Sup_singleton ~> Sup_insert [where A="{}", unfolded Sup_empty sup_bot_right] hoelzl@45041: Inf_binary ~> Inf_insert, Inf_empty, and inf_top_right hoelzl@45041: Sup_binary ~> Sup_insert, Sup_empty, and sup_bot_right hoelzl@45041: Int_eq_Inter ~> Inf_insert, Inf_empty, and inf_top_right hoelzl@45041: Un_eq_Union ~> Sup_insert, Sup_empty, and sup_bot_right hoelzl@45041: Inter_def ~> INF_def, image_def hoelzl@45041: Union_def ~> SUP_def, image_def hoelzl@45041: INT_eq ~> INF_def, and image_def hoelzl@45041: UN_eq ~> SUP_def, and image_def hoelzl@45041: INF_subset ~> INF_superset_mono [OF _ order_refl] wenzelm@44800: wenzelm@44800: More consistent and comprehensive names: wenzelm@44800: hoelzl@45041: INTER_eq_Inter_image ~> INF_def hoelzl@45041: UNION_eq_Union_image ~> SUP_def haftmann@43872: INFI_def ~> INF_def haftmann@43872: SUPR_def ~> SUP_def haftmann@44103: INF_leI ~> INF_lower haftmann@44103: INF_leI2 ~> INF_lower2 haftmann@44103: le_INFI ~> INF_greatest haftmann@44103: le_SUPI ~> SUP_upper haftmann@44103: le_SUPI2 ~> SUP_upper2 haftmann@44103: SUP_leI ~> SUP_least haftmann@43873: INFI_bool_eq ~> INF_bool_eq haftmann@43873: SUPR_bool_eq ~> SUP_bool_eq haftmann@43873: INFI_apply ~> INF_apply haftmann@43873: SUPR_apply ~> SUP_apply haftmann@44103: INTER_def ~> INTER_eq haftmann@44103: UNION_def ~> UNION_eq haftmann@44103: haftmann@43865: INCOMPATIBILITY. haftmann@43865: wenzelm@44973: * Renamed theory Complete_Lattice to Complete_Lattices. wenzelm@44973: INCOMPATIBILITY. wenzelm@44973: wenzelm@44973: * Theory Complete_Lattices: lemmas Inf_eq_top_iff, INF_eq_top_iff, wenzelm@44973: INF_image, Inf_insert, INF_top, Inf_top_conv, INF_top_conv, SUP_bot, wenzelm@44973: Sup_bot_conv, SUP_bot_conv, Sup_eq_top_iff, SUP_eq_top_iff, SUP_image, wenzelm@44973: Sup_insert are now declared as [simp]. INCOMPATIBILITY. wenzelm@44973: wenzelm@44973: * Theory Lattice: lemmas compl_inf_bot, compl_le_comp_iff, wenzelm@44973: compl_sup_top, inf_idem, inf_left_idem, inf_sup_absorb, sup_idem, wenzelm@44973: sup_inf_absob, sup_left_idem are now declared as [simp]. Minor wenzelm@44973: INCOMPATIBILITY. wenzelm@44973: krauss@44845: * Added syntactic classes "inf" and "sup" for the respective krauss@44845: constants. INCOMPATIBILITY: Changes in the argument order of the krauss@44845: (mostly internal) locale predicates for some derived classes. krauss@44845: wenzelm@44800: * Theorem collections ball_simps and bex_simps do not contain theorems wenzelm@44800: referring to UNION any longer; these have been moved to collection wenzelm@44800: UN_ball_bex_simps. INCOMPATIBILITY. wenzelm@44800: wenzelm@44800: * Theory Archimedean_Field: floor now is defined as parameter of a wenzelm@44800: separate type class floor_ceiling. wenzelm@44800: wenzelm@44800: * Theory Finite_Set: more coherent development of fold_set locales: haftmann@42874: haftmann@42874: locale fun_left_comm ~> locale comp_fun_commute haftmann@42874: locale fun_left_comm_idem ~> locale comp_fun_idem wenzelm@44800: wenzelm@44800: Both use point-free characterization; interpretation proofs may need wenzelm@44800: adjustment. INCOMPATIBILITY. haftmann@42874: wenzelm@44800: * Theory Limits: Type "'a net" has been renamed to "'a filter", in huffman@44081: accordance with standard mathematical terminology. INCOMPATIBILITY. huffman@44081: wenzelm@44800: * Theory Complex_Main: The locale interpretations for the wenzelm@44800: bounded_linear and bounded_bilinear locales have been removed, in wenzelm@44800: order to reduce the number of duplicate lemmas. Users must use the wenzelm@44800: original names for distributivity theorems, potential INCOMPATIBILITY. huffman@44282: huffman@44282: divide.add ~> add_divide_distrib huffman@44282: divide.diff ~> diff_divide_distrib huffman@44282: divide.setsum ~> setsum_divide_distrib huffman@44282: mult.add_right ~> right_distrib huffman@44282: mult.diff_right ~> right_diff_distrib huffman@44282: mult_right.setsum ~> setsum_right_distrib huffman@44282: mult_left.diff ~> left_diff_distrib huffman@44282: wenzelm@44800: * Theory Complex_Main: Several redundant theorems have been removed or huffman@44568: replaced by more general versions. INCOMPATIBILITY. huffman@44522: huffman@45051: real_diff_def ~> minus_real_def huffman@45051: real_divide_def ~> divide_real_def huffman@45051: real_less_def ~> less_le huffman@45051: real_abs_def ~> abs_real_def huffman@45051: real_sgn_def ~> sgn_real_def huffman@45051: real_mult_commute ~> mult_commute huffman@45051: real_mult_assoc ~> mult_assoc huffman@45051: real_mult_1 ~> mult_1_left huffman@45051: real_add_mult_distrib ~> left_distrib huffman@45051: real_zero_not_eq_one ~> zero_neq_one huffman@45051: real_mult_inverse_left ~> left_inverse huffman@45051: INVERSE_ZERO ~> inverse_zero huffman@45051: real_le_refl ~> order_refl huffman@45051: real_le_antisym ~> order_antisym huffman@45051: real_le_trans ~> order_trans huffman@45051: real_le_linear ~> linear huffman@45051: real_le_eq_diff ~> le_iff_diff_le_0 huffman@45051: real_add_left_mono ~> add_left_mono huffman@45051: real_mult_order ~> mult_pos_pos huffman@45051: real_mult_less_mono2 ~> mult_strict_left_mono huffman@44822: real_of_int_real_of_nat ~> real_of_int_of_nat_eq huffman@44522: real_0_le_divide_iff ~> zero_le_divide_iff huffman@44522: realpow_two_disj ~> power2_eq_iff huffman@44522: real_squared_diff_one_factored ~> square_diff_one_factored huffman@44522: realpow_two_diff ~> square_diff_square_factored huffman@44669: reals_complete2 ~> complete_real huffman@44749: real_sum_squared_expand ~> power2_sum huffman@44522: exp_ln_eq ~> ln_unique huffman@44711: expi_add ~> exp_add huffman@44711: expi_zero ~> exp_zero huffman@44522: lemma_DERIV_subst ~> DERIV_cong huffman@44568: LIMSEQ_Zfun_iff ~> tendsto_Zfun_iff huffman@44568: LIMSEQ_const ~> tendsto_const huffman@44568: LIMSEQ_norm ~> tendsto_norm huffman@44568: LIMSEQ_add ~> tendsto_add huffman@44568: LIMSEQ_minus ~> tendsto_minus huffman@44568: LIMSEQ_minus_cancel ~> tendsto_minus_cancel huffman@44568: LIMSEQ_diff ~> tendsto_diff huffman@44568: bounded_linear.LIMSEQ ~> bounded_linear.tendsto huffman@44568: bounded_bilinear.LIMSEQ ~> bounded_bilinear.tendsto huffman@44568: LIMSEQ_mult ~> tendsto_mult huffman@44568: LIMSEQ_inverse ~> tendsto_inverse huffman@44568: LIMSEQ_divide ~> tendsto_divide huffman@44568: LIMSEQ_pow ~> tendsto_power huffman@44568: LIMSEQ_setsum ~> tendsto_setsum huffman@44568: LIMSEQ_setprod ~> tendsto_setprod huffman@44568: LIMSEQ_norm_zero ~> tendsto_norm_zero_iff huffman@44568: LIMSEQ_rabs_zero ~> tendsto_rabs_zero_iff huffman@44568: LIMSEQ_imp_rabs ~> tendsto_rabs huffman@44710: LIMSEQ_add_minus ~> tendsto_add [OF _ tendsto_minus] huffman@44710: LIMSEQ_add_const ~> tendsto_add [OF _ tendsto_const] huffman@44710: LIMSEQ_diff_const ~> tendsto_diff [OF _ tendsto_const] huffman@44748: LIMSEQ_Complex ~> tendsto_Complex huffman@44568: LIM_ident ~> tendsto_ident_at huffman@44568: LIM_const ~> tendsto_const huffman@44568: LIM_add ~> tendsto_add huffman@44568: LIM_add_zero ~> tendsto_add_zero huffman@44568: LIM_minus ~> tendsto_minus huffman@44568: LIM_diff ~> tendsto_diff huffman@44568: LIM_norm ~> tendsto_norm huffman@44568: LIM_norm_zero ~> tendsto_norm_zero huffman@44568: LIM_norm_zero_cancel ~> tendsto_norm_zero_cancel huffman@44568: LIM_norm_zero_iff ~> tendsto_norm_zero_iff huffman@44568: LIM_rabs ~> tendsto_rabs huffman@44568: LIM_rabs_zero ~> tendsto_rabs_zero huffman@44568: LIM_rabs_zero_cancel ~> tendsto_rabs_zero_cancel huffman@44568: LIM_rabs_zero_iff ~> tendsto_rabs_zero_iff huffman@44568: LIM_compose ~> tendsto_compose huffman@44568: LIM_mult ~> tendsto_mult huffman@44568: LIM_scaleR ~> tendsto_scaleR huffman@44568: LIM_of_real ~> tendsto_of_real huffman@44568: LIM_power ~> tendsto_power huffman@44568: LIM_inverse ~> tendsto_inverse huffman@44568: LIM_sgn ~> tendsto_sgn huffman@44568: isCont_LIM_compose ~> isCont_tendsto_compose huffman@44568: bounded_linear.LIM ~> bounded_linear.tendsto huffman@44568: bounded_linear.LIM_zero ~> bounded_linear.tendsto_zero huffman@44568: bounded_bilinear.LIM ~> bounded_bilinear.tendsto huffman@44568: bounded_bilinear.LIM_prod_zero ~> bounded_bilinear.tendsto_zero huffman@44568: bounded_bilinear.LIM_left_zero ~> bounded_bilinear.tendsto_left_zero huffman@44568: bounded_bilinear.LIM_right_zero ~> bounded_bilinear.tendsto_right_zero huffman@44568: LIM_inverse_fun ~> tendsto_inverse [OF tendsto_ident_at] huffman@44522: wenzelm@44967: * Theory Complex_Main: The definition of infinite series was wenzelm@44967: generalized. Now it is defined on the type class {topological_space, wenzelm@44967: comm_monoid_add}. Hence it is useable also for extended real numbers. wenzelm@42484: huffman@44908: * Theory Complex_Main: The complex exponential function "expi" is now huffman@44908: a type-constrained abbreviation for "exp :: complex => complex"; thus huffman@44908: several polymorphic lemmas about "exp" are now applicable to "expi". huffman@44908: wenzelm@44968: * Code generation: wenzelm@44968: wenzelm@44968: - Theory Library/Code_Char_ord provides native ordering of wenzelm@44968: characters in the target language. wenzelm@44968: wenzelm@44968: - Commands code_module and code_library are legacy, use export_code wenzelm@44968: instead. wenzelm@44968: wenzelm@44968: - Method "evaluation" is legacy, use method "eval" instead. wenzelm@44968: wenzelm@44968: - Legacy evaluator "SML" is deactivated by default. May be wenzelm@44968: reactivated by the following theory command: wenzelm@44968: wenzelm@44968: setup {* Value.add_evaluator ("SML", Codegen.eval_term) *} wenzelm@44968: wenzelm@44968: * Declare ext [intro] by default. Rare INCOMPATIBILITY. wenzelm@44968: wenzelm@45088: * New proof method "induction" that gives induction hypotheses the wenzelm@45088: name "IH", thus distinguishing them from further hypotheses that come wenzelm@45088: from rule induction. The latter are still called "hyps". Method wenzelm@45088: "induction" is a thin wrapper around "induct" and follows the same wenzelm@45088: syntax. wenzelm@45088: wenzelm@44968: * Method "fastsimp" has been renamed to "fastforce", but "fastsimp" is wenzelm@44968: still available as a legacy feature for some time. wenzelm@44968: wenzelm@44968: * Nitpick: wenzelm@44968: - Added "need" and "total_consts" options. wenzelm@44968: - Reintroduced "show_skolems" option by popular demand. wenzelm@44968: - Renamed attribute: nitpick_def ~> nitpick_unfold. wenzelm@44968: INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: * Sledgehammer: wenzelm@44968: - Use quasi-sound (and efficient) translations by default. wenzelm@44968: - Added support for the following provers: E-ToFoF, LEO-II, wenzelm@44968: Satallax, SNARK, Waldmeister, and Z3 with TPTP syntax. wenzelm@44968: - Automatically preplay and minimize proofs before showing them if wenzelm@44968: this can be done within reasonable time. wenzelm@44968: - sledgehammer available_provers ~> sledgehammer supported_provers. wenzelm@44968: INCOMPATIBILITY. wenzelm@44968: - Added "preplay_timeout", "slicing", "type_enc", "sound", wenzelm@44968: "max_mono_iters", and "max_new_mono_instances" options. wenzelm@44968: - Removed "explicit_apply" and "full_types" options as well as "Full wenzelm@44968: Types" Proof General menu item. INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: * Metis: wenzelm@44968: - Removed "metisF" -- use "metis" instead. INCOMPATIBILITY. wenzelm@44968: - Obsoleted "metisFT" -- use "metis (full_types)" instead. wenzelm@44968: INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: * Command 'try': wenzelm@44968: - Renamed 'try_methods' and added "simp:", "intro:", "dest:", and wenzelm@44968: "elim:" options. INCOMPATIBILITY. wenzelm@44968: - Introduced 'try' that not only runs 'try_methods' but also wenzelm@44968: 'solve_direct', 'sledgehammer', 'quickcheck', and 'nitpick'. wenzelm@44968: wenzelm@44968: * Quickcheck: wenzelm@44968: - Added "eval" option to evaluate terms for the found counterexample wenzelm@44968: (currently only supported by the default (exhaustive) tester). wenzelm@44968: - Added post-processing of terms to obtain readable counterexamples wenzelm@44968: (currently only supported by the default (exhaustive) tester). wenzelm@44968: - New counterexample generator quickcheck[narrowing] enables wenzelm@44968: narrowing-based testing. Requires the Glasgow Haskell compiler wenzelm@44968: with its installation location defined in the Isabelle settings wenzelm@44968: environment as ISABELLE_GHC. wenzelm@44968: - Removed quickcheck tester "SML" based on the SML code generator wenzelm@44968: (formly in HOL/Library). wenzelm@44968: wenzelm@44968: * Function package: discontinued option "tailrec". INCOMPATIBILITY, wenzelm@44968: use 'partial_function' instead. wenzelm@44968: wenzelm@44968: * Theory Library/Extended_Reals replaces now the positive extended wenzelm@44968: reals found in probability theory. This file is extended by wenzelm@44968: Multivariate_Analysis/Extended_Real_Limits. wenzelm@44968: wenzelm@44974: * Theory Library/Old_Recdef: old 'recdef' package has been moved here, wenzelm@44974: from where it must be imported explicitly if it is really required. wenzelm@44974: INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: * Theory Library/Wfrec: well-founded recursion combinator "wfrec" has wenzelm@44968: been moved here. INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: * Theory Library/Saturated provides type of numbers with saturated wenzelm@44968: arithmetic. wenzelm@44968: wenzelm@44968: * Theory Library/Product_Lattice defines a pointwise ordering for the wenzelm@44968: product type 'a * 'b, and provides instance proofs for various order wenzelm@44968: and lattice type classes. wenzelm@44968: wenzelm@44968: * Theory Library/Countable now provides the "countable_datatype" proof wenzelm@44968: method for proving "countable" class instances for datatypes. wenzelm@44968: wenzelm@44968: * Theory Library/Cset_Monad allows do notation for computable sets wenzelm@44968: (cset) via the generic monad ad-hoc overloading facility. wenzelm@44968: wenzelm@44968: * Library: Theories of common data structures are split into theories wenzelm@44968: for implementation, an invariant-ensuring type, and connection to an wenzelm@44968: abstract type. INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: - RBT is split into RBT and RBT_Mapping. wenzelm@44968: - AssocList is split and renamed into AList and AList_Mapping. wenzelm@44968: - DList is split into DList_Impl, DList, and DList_Cset. wenzelm@44968: - Cset is split into Cset and List_Cset. wenzelm@44968: wenzelm@44968: * Theory Library/Nat_Infinity has been renamed to wenzelm@44968: Library/Extended_Nat, with name changes of the following types and wenzelm@44968: constants: wenzelm@44968: wenzelm@44968: type inat ~> type enat wenzelm@44968: Fin ~> enat wenzelm@44968: Infty ~> infinity (overloaded) wenzelm@44968: iSuc ~> eSuc wenzelm@44968: the_Fin ~> the_enat wenzelm@44968: wenzelm@44968: Every theorem name containing "inat", "Fin", "Infty", or "iSuc" has wenzelm@44968: been renamed accordingly. INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: * Session Multivariate_Analysis: The euclidean_space type class now wenzelm@44968: fixes a constant "Basis :: 'a set" consisting of the standard wenzelm@44968: orthonormal basis for the type. Users now have the option of wenzelm@44968: quantifying over this set instead of using the "basis" function, e.g. wenzelm@44968: "ALL x:Basis. P x" vs "ALL i vec_eq_iff wenzelm@44968: dist_nth_le_cart ~> dist_vec_nth_le wenzelm@44968: tendsto_vector ~> vec_tendstoI wenzelm@44968: Cauchy_vector ~> vec_CauchyI wenzelm@44968: wenzelm@44968: * Session Multivariate_Analysis: Several duplicate theorems have been wenzelm@44968: removed, and other theorems have been renamed or replaced with more wenzelm@44968: general versions. INCOMPATIBILITY. wenzelm@44968: wenzelm@44968: finite_choice ~> finite_set_choice wenzelm@44968: eventually_conjI ~> eventually_conj wenzelm@44968: eventually_and ~> eventually_conj_iff wenzelm@44968: eventually_false ~> eventually_False wenzelm@44968: setsum_norm ~> norm_setsum wenzelm@44968: Lim_sequentially ~> LIMSEQ_def wenzelm@44968: Lim_ident_at ~> LIM_ident wenzelm@44968: Lim_const ~> tendsto_const wenzelm@44968: Lim_cmul ~> tendsto_scaleR [OF tendsto_const] wenzelm@44968: Lim_neg ~> tendsto_minus wenzelm@44968: Lim_add ~> tendsto_add wenzelm@44968: Lim_sub ~> tendsto_diff wenzelm@44968: Lim_mul ~> tendsto_scaleR wenzelm@44968: Lim_vmul ~> tendsto_scaleR [OF _ tendsto_const] wenzelm@44968: Lim_null_norm ~> tendsto_norm_zero_iff [symmetric] wenzelm@44968: Lim_linear ~> bounded_linear.tendsto wenzelm@44968: Lim_component ~> tendsto_euclidean_component wenzelm@44968: Lim_component_cart ~> tendsto_vec_nth wenzelm@44968: Lim_inner ~> tendsto_inner [OF tendsto_const] wenzelm@44968: dot_lsum ~> inner_setsum_left wenzelm@44968: dot_rsum ~> inner_setsum_right wenzelm@44968: continuous_cmul ~> continuous_scaleR [OF continuous_const] wenzelm@44968: continuous_neg ~> continuous_minus wenzelm@44968: continuous_sub ~> continuous_diff wenzelm@44968: continuous_vmul ~> continuous_scaleR [OF _ continuous_const] wenzelm@44968: continuous_mul ~> continuous_scaleR wenzelm@44968: continuous_inv ~> continuous_inverse wenzelm@44968: continuous_at_within_inv ~> continuous_at_within_inverse wenzelm@44968: continuous_at_inv ~> continuous_at_inverse wenzelm@44968: continuous_at_norm ~> continuous_norm [OF continuous_at_id] wenzelm@44968: continuous_at_infnorm ~> continuous_infnorm [OF continuous_at_id] wenzelm@44968: continuous_at_component ~> continuous_component [OF continuous_at_id] wenzelm@44968: continuous_on_neg ~> continuous_on_minus wenzelm@44968: continuous_on_sub ~> continuous_on_diff wenzelm@44968: continuous_on_cmul ~> continuous_on_scaleR [OF continuous_on_const] wenzelm@44968: continuous_on_vmul ~> continuous_on_scaleR [OF _ continuous_on_const] wenzelm@44968: continuous_on_mul ~> continuous_on_scaleR wenzelm@44968: continuous_on_mul_real ~> continuous_on_mult wenzelm@44968: continuous_on_inner ~> continuous_on_inner [OF continuous_on_const] wenzelm@44968: continuous_on_norm ~> continuous_on_norm [OF continuous_on_id] wenzelm@44968: continuous_on_inverse ~> continuous_on_inv wenzelm@44968: uniformly_continuous_on_neg ~> uniformly_continuous_on_minus wenzelm@44968: uniformly_continuous_on_sub ~> uniformly_continuous_on_diff wenzelm@44968: subset_interior ~> interior_mono wenzelm@44968: subset_closure ~> closure_mono wenzelm@44968: closure_univ ~> closure_UNIV wenzelm@44968: real_arch_lt ~> reals_Archimedean2 wenzelm@44968: real_arch ~> reals_Archimedean3 wenzelm@44968: real_abs_norm ~> abs_norm_cancel wenzelm@44968: real_abs_sub_norm ~> norm_triangle_ineq3 wenzelm@44968: norm_cauchy_schwarz_abs ~> Cauchy_Schwarz_ineq2 wenzelm@44968: wenzelm@44968: * Session HOL-Probability: wenzelm@44968: - Caratheodory's extension lemma is now proved for ring_of_sets. wenzelm@44968: - Infinite products of probability measures are now available. wenzelm@44968: - Sigma closure is independent, if the generator is independent wenzelm@44968: - Use extended reals instead of positive extended wenzelm@44968: reals. INCOMPATIBILITY. wenzelm@44968: huffman@45049: * Session HOLCF: Discontinued legacy theorem names, INCOMPATIBILITY. huffman@45049: huffman@45049: expand_fun_below ~> fun_below_iff huffman@45049: below_fun_ext ~> fun_belowI huffman@45049: expand_cfun_eq ~> cfun_eq_iff huffman@45049: ext_cfun ~> cfun_eqI huffman@45049: expand_cfun_below ~> cfun_below_iff huffman@45049: below_cfun_ext ~> cfun_belowI huffman@45049: monofun_fun_fun ~> fun_belowD huffman@45049: monofun_fun_arg ~> monofunE huffman@45049: monofun_lub_fun ~> adm_monofun [THEN admD] huffman@45049: cont_lub_fun ~> adm_cont [THEN admD] huffman@45049: cont2cont_Rep_CFun ~> cont2cont_APP huffman@45049: cont_Rep_CFun_app ~> cont_APP_app huffman@45049: cont_Rep_CFun_app_app ~> cont_APP_app_app huffman@45049: cont_cfun_fun ~> cont_Rep_cfun1 [THEN contE] huffman@45049: cont_cfun_arg ~> cont_Rep_cfun2 [THEN contE] huffman@45049: contlub_cfun ~> lub_APP [symmetric] huffman@45049: contlub_LAM ~> lub_LAM [symmetric] huffman@45049: thelubI ~> lub_eqI huffman@45049: UU_I ~> bottomI huffman@45049: lift_distinct1 ~> lift.distinct(1) huffman@45049: lift_distinct2 ~> lift.distinct(2) huffman@45049: Def_not_UU ~> lift.distinct(2) huffman@45049: Def_inject ~> lift.inject huffman@45049: below_UU_iff ~> below_bottom_iff huffman@45049: eq_UU_iff ~> eq_bottom_iff huffman@45049: huffman@44903: krauss@41685: *** Document preparation *** krauss@41685: wenzelm@44800: * Antiquotation @{rail} layouts railroad syntax diagrams, see also wenzelm@44800: isar-ref manual, both for description and actual application of the wenzelm@44800: same. wenzelm@44800: wenzelm@44800: * Antiquotation @{value} evaluates the given term and presents its wenzelm@44800: result. wenzelm@44800: wenzelm@44800: * Antiquotations: term style "isub" provides ad-hoc conversion of wenzelm@44800: variables x1, y23 into subscripted form x\<^isub>1, wenzelm@44800: y\<^isub>2\<^isub>3. wenzelm@41651: wenzelm@42484: * Predefined LaTeX macros for Isabelle symbols \ and \ wenzelm@42484: (e.g. see ~~/src/HOL/Library/Monad_Syntax.thy). wenzelm@42484: wenzelm@44967: * Localized \isabellestyle switch can be used within blocks or groups wenzelm@44967: like this: wenzelm@44967: wenzelm@44967: \isabellestyle{it} %preferred default wenzelm@44967: {\isabellestylett @{text "typewriter stuff"}} wenzelm@44967: wenzelm@44967: * Discontinued special treatment of hard tabulators. Implicit wenzelm@44967: tab-width is now defined as 1. Potential INCOMPATIBILITY for visual wenzelm@44967: layouts. wenzelm@44800: wenzelm@41651: wenzelm@41944: *** ML *** wenzelm@41944: wenzelm@43731: * The inner syntax of sort/type/term/prop supports inlined YXML wenzelm@43731: representations within quoted string tokens. By encoding logical wenzelm@43731: entities via Term_XML (in ML or Scala) concrete syntax can be wenzelm@43731: bypassed, which is particularly useful for producing bits of text wenzelm@43731: under external program control. wenzelm@43731: wenzelm@43565: * Antiquotations for ML and document preparation are managed as theory wenzelm@43565: data, which requires explicit setup. wenzelm@43565: wenzelm@42897: * Isabelle_Process.is_active allows tools to check if the official wenzelm@42897: process wrapper is running (Isabelle/Scala/jEdit) or the old TTY loop wenzelm@42897: (better known as Proof General). wenzelm@42897: wenzelm@42360: * Structure Proof_Context follows standard naming scheme. Old wenzelm@42360: ProofContext is still available for some time as legacy alias. wenzelm@42360: wenzelm@42015: * Structure Timing provides various operations for timing; supersedes wenzelm@42015: former start_timing/end_timing etc. wenzelm@42015: wenzelm@41944: * Path.print is the official way to show file-system paths to users wenzelm@41944: (including quotes etc.). wenzelm@41944: wenzelm@42056: * Inner syntax: identifiers in parse trees of generic categories wenzelm@42056: "logic", "aprop", "idt" etc. carry position information (disguised as wenzelm@42056: type constraints). Occasional INCOMPATIBILITY with non-compliant wenzelm@42057: translations that choke on unexpected type constraints. Positions can wenzelm@42057: be stripped in ML translations via Syntax.strip_positions / wenzelm@42057: Syntax.strip_positions_ast, or via the syntax constant wenzelm@42057: "_strip_positions" within parse trees. As last resort, positions can wenzelm@42057: be disabled via the configuration option Syntax.positions, which is wenzelm@42057: called "syntax_positions" in Isar attribute syntax. wenzelm@42056: wenzelm@42290: * Discontinued special status of various ML structures that contribute wenzelm@42290: to structure Syntax (Ast, Lexicon, Mixfix, Parser, Printer etc.): less wenzelm@42290: pervasive content, no inclusion in structure Syntax. INCOMPATIBILITY, wenzelm@42290: refer directly to Ast.Constant, Lexicon.is_identifier, wenzelm@42290: Syntax_Trans.mk_binder_tr etc. wenzelm@42224: wenzelm@42247: * Typed print translation: discontinued show_sorts argument, which is wenzelm@42247: already available via context of "advanced" translation. wenzelm@42247: wenzelm@42370: * Refined PARALLEL_GOALS tactical: degrades gracefully for schematic wenzelm@42370: goal states; body tactic needs to address all subgoals uniformly. wenzelm@42370: wenzelm@42403: * Slightly more special eq_list/eq_set, with shortcut involving wenzelm@42403: pointer equality (assumes that eq relation is reflexive). wenzelm@42403: wenzelm@42793: * Classical tactics use proper Proof.context instead of historic types wenzelm@42793: claset/clasimpset. Old-style declarations like addIs, addEs, addDs wenzelm@42793: operate directly on Proof.context. Raw type claset retains its use as wenzelm@42793: snapshot of the classical context, which can be recovered via wenzelm@42793: (put_claset HOL_cs) etc. Type clasimpset has been discontinued. wenzelm@42793: INCOMPATIBILITY, classical tactics and derived proof methods require wenzelm@42793: proper Proof.context. wenzelm@42793: wenzelm@44803: wenzelm@44803: *** System *** wenzelm@44803: wenzelm@44968: * Discontinued support for Poly/ML 5.2, which was the last version wenzelm@44968: without proper multithreading and TimeLimit implementation. wenzelm@44968: wenzelm@44968: * Discontinued old lib/scripts/polyml-platform, which has been wenzelm@44968: obsolete since Isabelle2009-2. wenzelm@44968: wenzelm@44967: * Various optional external tools are referenced more robustly and wenzelm@44967: uniformly by explicit Isabelle settings as follows: wenzelm@44967: wenzelm@44967: ISABELLE_CSDP (formerly CSDP_EXE) wenzelm@44967: ISABELLE_GHC (formerly EXEC_GHC or GHC_PATH) wenzelm@44967: ISABELLE_OCAML (formerly EXEC_OCAML) wenzelm@44967: ISABELLE_SWIPL (formerly EXEC_SWIPL) wenzelm@44967: ISABELLE_YAP (formerly EXEC_YAP) wenzelm@44967: wenzelm@44967: Note that automated detection from the file-system or search path has wenzelm@44967: been discontinued. INCOMPATIBILITY. wenzelm@44967: wenzelm@43752: * Scala layer provides JVM method invocation service for static wenzelm@44800: methods of type (String)String, see Invoke_Scala.method in ML. For wenzelm@44800: example: wenzelm@43752: wenzelm@43752: Invoke_Scala.method "java.lang.System.getProperty" "java.home" wenzelm@43752: wenzelm@44967: Together with YXML.string_of_body/parse_body and XML.Encode/Decode wenzelm@44967: this allows to pass structured values between ML and Scala. wenzelm@44800: wenzelm@44803: * The IsabelleText fonts includes some further glyphs to support the wenzelm@44803: Prover IDE. Potential INCOMPATIBILITY: users who happen to have wenzelm@44803: installed a local copy (which is normally *not* required) need to wenzelm@44803: delete or update it from ~~/lib/fonts/. wenzelm@41944: wenzelm@41703: wenzelm@45089: wenzelm@41512: New in Isabelle2011 (January 2011) wenzelm@41512: ---------------------------------- wenzelm@37383: wenzelm@37536: *** General *** wenzelm@37536: wenzelm@41573: * Experimental Prover IDE based on Isabelle/Scala and jEdit (see wenzelm@41612: src/Tools/jEdit). This also serves as IDE for Isabelle/ML, with wenzelm@41612: useful tooltips and hyperlinks produced from its static analysis. The wenzelm@41612: bundled component provides an executable Isabelle tool that can be run wenzelm@41612: like this: wenzelm@41612: wenzelm@41612: Isabelle2011/bin/isabelle jedit wenzelm@41573: wenzelm@40948: * Significantly improved Isabelle/Isar implementation manual. wenzelm@40948: wenzelm@41594: * System settings: ISABELLE_HOME_USER now includes ISABELLE_IDENTIFIER wenzelm@41595: (and thus refers to something like $HOME/.isabelle/Isabelle2011), wenzelm@41594: while the default heap location within that directory lacks that extra wenzelm@41594: suffix. This isolates multiple Isabelle installations from each wenzelm@41594: other, avoiding problems with old settings in new versions. wenzelm@41594: INCOMPATIBILITY, need to copy/upgrade old user settings manually. wenzelm@41594: wenzelm@40947: * Source files are always encoded as UTF-8, instead of old-fashioned wenzelm@40947: ISO-Latin-1. INCOMPATIBILITY. Isabelle LaTeX documents might require wenzelm@40948: the following package declarations: wenzelm@40947: wenzelm@40947: \usepackage[utf8]{inputenc} wenzelm@40947: \usepackage{textcomp} wenzelm@40947: krauss@41440: * Explicit treatment of UTF-8 sequences as Isabelle symbols, such that wenzelm@37536: a Unicode character is treated as a single symbol, not a sequence of wenzelm@37536: non-ASCII bytes as before. Since Isabelle/ML string literals may wenzelm@37536: contain symbols without further backslash escapes, Unicode can now be wenzelm@37536: used here as well. Recall that Symbol.explode in ML provides a wenzelm@37536: consistent view on symbols, while raw explode (or String.explode) wenzelm@37536: merely give a byte-oriented representation. wenzelm@37536: wenzelm@41594: * Theory loader: source files are primarily located via the master wenzelm@41594: directory of each theory node (where the .thy file itself resides). wenzelm@41594: The global load path is still partially available as legacy feature. wenzelm@41594: Minor INCOMPATIBILITY due to subtle change in file lookup: use wenzelm@41594: explicit paths, relatively to the theory. wenzelm@38135: wenzelm@37939: * Special treatment of ML file names has been discontinued. wenzelm@37939: Historically, optional extensions .ML or .sml were added on demand -- wenzelm@37939: at the cost of clarity of file dependencies. Recall that Isabelle/ML wenzelm@58604: files exclusively use the .ML extension. Minor INCOMPATIBILITY. wenzelm@37939: wenzelm@38980: * Various options that affect pretty printing etc. are now properly wenzelm@38767: handled within the context via configuration options, instead of wenzelm@40879: unsynchronized references or print modes. There are both ML Config.T wenzelm@40879: entities and Isar declaration attributes to access these. wenzelm@38767: wenzelm@39125: ML (Config.T) Isar (attribute) wenzelm@39125: wenzelm@39128: eta_contract eta_contract wenzelm@39137: show_brackets show_brackets wenzelm@39134: show_sorts show_sorts wenzelm@39134: show_types show_types wenzelm@39126: show_question_marks show_question_marks wenzelm@39126: show_consts show_consts wenzelm@40879: show_abbrevs show_abbrevs wenzelm@39126: wenzelm@41379: Syntax.ast_trace syntax_ast_trace wenzelm@41379: Syntax.ast_stat syntax_ast_stat wenzelm@39126: Syntax.ambiguity_level syntax_ambiguity_level wenzelm@39126: wenzelm@39126: Goal_Display.goals_limit goals_limit wenzelm@39126: Goal_Display.show_main_goal show_main_goal wenzelm@39126: wenzelm@41379: Method.rule_trace rule_trace wenzelm@41379: wenzelm@39125: Thy_Output.display thy_output_display wenzelm@39125: Thy_Output.quotes thy_output_quotes wenzelm@39125: Thy_Output.indent thy_output_indent wenzelm@39125: Thy_Output.source thy_output_source wenzelm@39125: Thy_Output.break thy_output_break wenzelm@39125: krauss@41440: Note that corresponding "..._default" references in ML may only be wenzelm@38767: changed globally at the ROOT session setup, but *not* within a theory. wenzelm@40879: The option "show_abbrevs" supersedes the former print mode wenzelm@40879: "no_abbrevs" with inverted meaning. wenzelm@38767: wenzelm@40878: * More systematic naming of some configuration options. huffman@41294: INCOMPATIBILITY. wenzelm@40878: wenzelm@40878: trace_simp ~> simp_trace wenzelm@40878: debug_simp ~> simp_debug wenzelm@40878: wenzelm@40291: * Support for real valued configuration options, using simplistic wenzelm@40291: floating-point notation that coincides with the inner syntax for wenzelm@40291: float_token. wenzelm@40291: wenzelm@41594: * Support for real valued preferences (with approximative PGIP type): wenzelm@41594: front-ends need to accept "pgint" values in float notation. wenzelm@41594: INCOMPATIBILITY. wenzelm@41573: wenzelm@41573: * The IsabelleText font now includes Cyrillic, Hebrew, Arabic from wenzelm@41573: DejaVu Sans. wenzelm@41573: wenzelm@41594: * Discontinued support for Poly/ML 5.0 and 5.1 versions. wenzelm@41594: wenzelm@40948: wenzelm@40948: *** Pure *** wenzelm@40948: wenzelm@41249: * Command 'type_synonym' (with single argument) replaces somewhat wenzelm@41249: outdated 'types', which is still available as legacy feature for some wenzelm@41249: time. wenzelm@41249: wenzelm@41249: * Command 'nonterminal' (with 'and' separated list of arguments) wenzelm@41249: replaces somewhat outdated 'nonterminals'. INCOMPATIBILITY. wenzelm@41229: wenzelm@40965: * Command 'notepad' replaces former 'example_proof' for wenzelm@41020: experimentation in Isar without any result. INCOMPATIBILITY. wenzelm@40965: ballarin@41435: * Locale interpretation commands 'interpret' and 'sublocale' accept ballarin@41435: lists of equations to map definitions in a locale to appropriate ballarin@41435: entities in the context of the interpretation. The 'interpretation' ballarin@41435: command already provided this functionality. ballarin@41435: wenzelm@41594: * Diagnostic command 'print_dependencies' prints the locale instances wenzelm@41594: that would be activated if the specified expression was interpreted in wenzelm@41594: the current context. Variant "print_dependencies!" assumes a context wenzelm@41594: without interpretations. ballarin@38110: ballarin@38110: * Diagnostic command 'print_interps' prints interpretations in proofs ballarin@38110: in addition to interpretations in theories. ballarin@38110: wenzelm@38708: * Discontinued obsolete 'global' and 'local' commands to manipulate wenzelm@38708: the theory name space. Rare INCOMPATIBILITY. The ML functions wenzelm@38708: Sign.root_path and Sign.local_path may be applied directly where this wenzelm@38708: feature is still required for historical reasons. wenzelm@38708: wenzelm@40948: * Discontinued obsolete 'constdefs' command. INCOMPATIBILITY, use haftmann@39215: 'definition' instead. haftmann@39215: wenzelm@41574: * The "prems" fact, which refers to the accidental collection of wenzelm@41574: foundational premises in the context, is now explicitly marked as wenzelm@41594: legacy feature and will be discontinued soon. Consider using "assms" wenzelm@41594: of the head statement or reference facts by explicit names. wenzelm@41574: wenzelm@40801: * Document antiquotations @{class} and @{type} print classes and type wenzelm@40801: constructors. wenzelm@40801: wenzelm@40801: * Document antiquotation @{file} checks file/directory entries within wenzelm@40801: the local file system. haftmann@39305: ballarin@38110: haftmann@37387: *** HOL *** haftmann@37387: wenzelm@41594: * Coercive subtyping: functions can be declared as coercions and type wenzelm@41594: inference will add them as necessary upon input of a term. Theory wenzelm@41594: Complex_Main declares real :: nat => real and real :: int => real as wenzelm@41594: coercions. A coercion function f is declared like this: wenzelm@40939: wenzelm@40939: declare [[coercion f]] nipkow@40866: wenzelm@41571: To lift coercions through type constructors (e.g. from nat => real to nipkow@40866: nat list => real list), map functions can be declared, e.g. nipkow@40866: wenzelm@40939: declare [[coercion_map map]] wenzelm@40939: wenzelm@40939: Currently coercion inference is activated only in theories including wenzelm@40939: real numbers, i.e. descendants of Complex_Main. This is controlled by wenzelm@41020: the configuration option "coercion_enabled", e.g. it can be enabled in wenzelm@40939: other theories like this: wenzelm@40939: wenzelm@40939: declare [[coercion_enabled]] nipkow@40866: wenzelm@41571: * Command 'partial_function' provides basic support for recursive wenzelm@41571: function definitions over complete partial orders. Concrete instances krauss@40183: are provided for i) the option type, ii) tail recursion on arbitrary wenzelm@41571: types, and iii) the heap monad of Imperative_HOL. See wenzelm@41571: src/HOL/ex/Fundefs.thy and src/HOL/Imperative_HOL/ex/Linked_Lists.thy wenzelm@41571: for examples. krauss@40183: wenzelm@41571: * Function package: f.psimps rules are no longer implicitly declared wenzelm@41571: as [simp]. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Datatype package: theorems generated for executable equality (class wenzelm@41571: "eq") carry proper names and are treated as default code equations. wenzelm@41571: wenzelm@41594: * Inductive package: now offers command 'inductive_simps' to wenzelm@41594: automatically derive instantiated and simplified equations for wenzelm@41594: inductive predicates, similar to 'inductive_cases'. wenzelm@41594: wenzelm@41571: * Command 'enriched_type' allows to register properties of the wenzelm@41571: functorial structure of types. haftmann@39771: haftmann@39644: * Improved infrastructure for term evaluation using code generator haftmann@39644: techniques, in particular static evaluation conversions. haftmann@39644: wenzelm@41594: * Code generator: Scala (2.8 or higher) has been added to the target wenzelm@41594: languages. wenzelm@41594: haftmann@41398: * Code generator: globbing constant expressions "*" and "Theory.*" haftmann@41398: have been replaced by the more idiomatic "_" and "Theory._". haftmann@41398: INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Code generator: export_code without explicit file declaration prints haftmann@41398: to standard output. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Code generator: do not print function definitions for case haftmann@41398: combinators any longer. haftmann@41398: wenzelm@41594: * Code generator: simplification with rules determined with wenzelm@41571: src/Tools/Code/code_simp.ML and method "code_simp". wenzelm@41571: wenzelm@41594: * Code generator for records: more idiomatic representation of record wenzelm@40948: types. Warning: records are not covered by ancient SML code wenzelm@40948: generation any longer. INCOMPATIBILITY. In cases of need, a suitable wenzelm@40948: rep_datatype declaration helps to succeed then: haftmann@38537: haftmann@38537: record 'a foo = ... haftmann@38537: ... haftmann@38537: rep_datatype foo_ext ... haftmann@38535: wenzelm@41594: * Records: logical foundation type for records does not carry a wenzelm@41594: '_type' suffix any longer (obsolete due to authentic syntax). wenzelm@41594: INCOMPATIBILITY. wenzelm@41594: haftmann@41398: * Quickcheck now by default uses exhaustive testing instead of random wenzelm@41571: testing. Random testing can be invoked by "quickcheck [random]", wenzelm@41571: exhaustive testing by "quickcheck [exhaustive]". haftmann@41398: haftmann@41398: * Quickcheck instantiates polymorphic types with small finite haftmann@41398: datatypes by default. This enables a simple execution mechanism to haftmann@41398: handle quantifiers and function equality over the finite datatypes. haftmann@41398: wenzelm@41571: * Quickcheck random generator has been renamed from "code" to wenzelm@41571: "random". INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Quickcheck now has a configurable time limit which is set to 30 haftmann@41398: seconds by default. This can be changed by adding [timeout = n] to the haftmann@41398: quickcheck command. The time limit for Auto Quickcheck is still set haftmann@41398: independently. haftmann@38461: haftmann@38461: * Quickcheck in locales considers interpretations of that locale for haftmann@38461: counter example search. haftmann@38461: blanchet@40059: * Sledgehammer: wenzelm@41571: - Added "smt" and "remote_smt" provers based on the "smt" proof wenzelm@41571: method. See the Sledgehammer manual for details ("isabelle doc wenzelm@41571: sledgehammer"). blanchet@40059: - Renamed commands: blanchet@40059: sledgehammer atp_info ~> sledgehammer running_provers blanchet@40059: sledgehammer atp_kill ~> sledgehammer kill_provers blanchet@40059: sledgehammer available_atps ~> sledgehammer available_provers blanchet@40059: INCOMPATIBILITY. blanchet@40059: - Renamed options: blanchet@40059: sledgehammer [atps = ...] ~> sledgehammer [provers = ...] blanchet@40062: sledgehammer [atp = ...] ~> sledgehammer [prover = ...] blanchet@40341: sledgehammer [timeout = 77 s] ~> sledgehammer [timeout = 77] blanchet@40341: (and "ms" and "min" are no longer supported) blanchet@40341: INCOMPATIBILITY. blanchet@40341: blanchet@40341: * Nitpick: blanchet@40341: - Renamed options: blanchet@40341: nitpick [timeout = 77 s] ~> nitpick [timeout = 77] blanchet@40341: nitpick [tac_timeout = 777 ms] ~> nitpick [tac_timeout = 0.777] blanchet@40059: INCOMPATIBILITY. blanchet@40725: - Added support for partial quotient types. blanchet@40725: - Added local versions of the "Nitpick.register_xxx" functions. blanchet@40725: - Added "whack" option. blanchet@40725: - Allow registration of quotient types as codatatypes. blanchet@40725: - Improved "merge_type_vars" option to merge more types. blanchet@40725: - Removed unsound "fast_descrs" option. blanchet@40725: - Added custom symmetry breaking for datatypes, making it possible to reach blanchet@40725: higher cardinalities. blanchet@40725: - Prevent the expansion of too large definitions. blanchet@39957: wenzelm@41571: * Proof methods "metis" and "meson" now have configuration options wenzelm@41571: "meson_trace", "metis_trace", and "metis_verbose" that can be enabled wenzelm@41571: to diagnose these tools. E.g. wenzelm@41571: wenzelm@41571: using [[metis_trace = true]] wenzelm@41571: haftmann@41398: * Auto Solve: Renamed "Auto Solve Direct". The tool is now available haftmann@41398: manually as command 'solve_direct'. haftmann@41398: boehmes@41601: * The default SMT solver Z3 must be enabled explicitly (due to boehmes@41601: licensing issues) by setting the environment variable wenzelm@41603: Z3_NON_COMMERCIAL in etc/settings of the component, for example. For wenzelm@41603: commercial applications, the SMT solver CVC3 is provided as fall-back; wenzelm@41603: changing the SMT solver is done via the configuration option wenzelm@41603: "smt_solver". boehmes@41432: boehmes@41432: * Remote SMT solvers need to be referred to by the "remote_" prefix, wenzelm@41571: i.e. "remote_cvc3" and "remote_z3". wenzelm@41571: wenzelm@41571: * Added basic SMT support for datatypes, records, and typedefs using wenzelm@41571: the oracle mode (no proofs). Direct support of pairs has been dropped wenzelm@41571: in exchange (pass theorems fst_conv snd_conv pair_collapse to the SMT wenzelm@41571: support for a similar behavior). Minor INCOMPATIBILITY. boehmes@41432: boehmes@40162: * Changed SMT configuration options: boehmes@40162: - Renamed: boehmes@41432: z3_proofs ~> smt_oracle (with inverted meaning) boehmes@40162: z3_trace_assms ~> smt_trace_used_facts boehmes@40162: INCOMPATIBILITY. boehmes@40162: - Added: boehmes@40424: smt_verbose boehmes@41432: smt_random_seed boehmes@40424: smt_datatypes boehmes@41432: smt_infer_triggers boehmes@41432: smt_monomorph_limit boehmes@40162: cvc3_options boehmes@41432: remote_cvc3_options boehmes@41432: remote_z3_options boehmes@40162: yices_options blanchet@39957: wenzelm@40948: * Boogie output files (.b2i files) need to be declared in the theory wenzelm@40948: header. boehmes@40580: wenzelm@41594: * Simplification procedure "list_to_set_comprehension" rewrites list wenzelm@41594: comprehensions applied to List.set to set comprehensions. Occasional wenzelm@41594: INCOMPATIBILITY, may be deactivated like this: wenzelm@41594: wenzelm@41594: declare [[simproc del: list_to_set_comprehension]] wenzelm@41594: wenzelm@41573: * Removed old version of primrec package. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Removed simplifier congruence rule of "prod_case", as has for long haftmann@41398: been the case with "split". INCOMPATIBILITY. haftmann@41398: haftmann@41398: * String.literal is a type, but not a datatype. INCOMPATIBILITY. haftmann@41398: krauss@40388: * Removed [split_format ... and ... and ...] version of krauss@40388: [split_format]. Potential INCOMPATIBILITY. krauss@40388: wenzelm@41571: * Predicate "sorted" now defined inductively, with nice induction wenzelm@41571: rules. INCOMPATIBILITY: former sorted.simps now named sorted_simps. haftmann@41398: haftmann@41398: * Constant "contents" renamed to "the_elem", to free the generic name haftmann@41398: contents for other uses. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Renamed class eq and constant eq (for code generation) to class haftmann@41398: equal and constant equal, plus renaming of related facts and various haftmann@41398: tuning. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Dropped type classes mult_mono and mult_mono1. INCOMPATIBILITY. haftmann@41398: wenzelm@41571: * Removed output syntax "'a ~=> 'b" for "'a => 'b option". wenzelm@41571: INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Renamed theory Fset to Cset, type Fset.fset to Cset.set, in order to haftmann@41398: avoid confusion with finite sets. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Abandoned locales equiv, congruent and congruent2 for equivalence haftmann@41398: relations. INCOMPATIBILITY: use equivI rather than equiv_intro (same haftmann@41398: for congruent(2)). haftmann@41398: haftmann@41398: * Some previously unqualified names have been qualified: haftmann@41398: haftmann@41398: types haftmann@41398: bool ~> HOL.bool haftmann@41398: nat ~> Nat.nat haftmann@41398: haftmann@41398: constants haftmann@41398: Trueprop ~> HOL.Trueprop haftmann@41398: True ~> HOL.True haftmann@41398: False ~> HOL.False haftmann@41398: op & ~> HOL.conj haftmann@41398: op | ~> HOL.disj haftmann@41398: op --> ~> HOL.implies haftmann@41398: op = ~> HOL.eq haftmann@41398: Not ~> HOL.Not haftmann@41398: The ~> HOL.The haftmann@41398: All ~> HOL.All haftmann@41398: Ex ~> HOL.Ex haftmann@41398: Ex1 ~> HOL.Ex1 haftmann@41398: Let ~> HOL.Let haftmann@41398: If ~> HOL.If haftmann@41398: Ball ~> Set.Ball haftmann@41398: Bex ~> Set.Bex haftmann@41398: Suc ~> Nat.Suc haftmann@41398: Pair ~> Product_Type.Pair haftmann@41398: fst ~> Product_Type.fst haftmann@41398: snd ~> Product_Type.snd haftmann@41398: curry ~> Product_Type.curry haftmann@41398: op : ~> Set.member haftmann@41398: Collect ~> Set.Collect haftmann@41398: haftmann@41398: INCOMPATIBILITY. haftmann@41398: haftmann@41398: * More canonical naming convention for some fundamental definitions: haftmann@41398: haftmann@41398: bot_bool_eq ~> bot_bool_def haftmann@41398: top_bool_eq ~> top_bool_def haftmann@41398: inf_bool_eq ~> inf_bool_def haftmann@41398: sup_bool_eq ~> sup_bool_def haftmann@41398: bot_fun_eq ~> bot_fun_def haftmann@41398: top_fun_eq ~> top_fun_def haftmann@41398: inf_fun_eq ~> inf_fun_def haftmann@41398: sup_fun_eq ~> sup_fun_def haftmann@41398: haftmann@41398: INCOMPATIBILITY. haftmann@41398: haftmann@41398: * More stylized fact names: haftmann@41398: haftmann@41398: expand_fun_eq ~> fun_eq_iff haftmann@41398: expand_set_eq ~> set_eq_iff haftmann@41398: set_ext ~> set_eqI haftmann@41398: nat_number ~> eval_nat_numeral haftmann@41398: haftmann@41398: INCOMPATIBILITY. haftmann@41398: wenzelm@41571: * Refactoring of code-generation specific operations in theory List: haftmann@41398: haftmann@41398: constants haftmann@41398: null ~> List.null haftmann@41398: haftmann@41398: facts haftmann@41398: mem_iff ~> member_def haftmann@41398: null_empty ~> null_def haftmann@41398: haftmann@41398: INCOMPATIBILITY. Note that these were not supposed to be used haftmann@41398: regularly unless for striking reasons; their main purpose was code haftmann@41398: generation. haftmann@41398: haftmann@41398: Various operations from the Haskell prelude are used for generating haftmann@41398: Haskell code. haftmann@41398: wenzelm@41571: * Term "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". Term wenzelm@41571: "surj f" is now an abbreviation of "range f = UNIV". The theorems wenzelm@41571: bij_def and surj_def are unchanged. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Abolished some non-alphabetic type names: "prod" and "sum" replace haftmann@41398: "*" and "+" respectively. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Name "Plus" of disjoint sum operator "<+>" is now hidden. Write wenzelm@41571: "Sum_Type.Plus" instead. haftmann@41398: haftmann@41398: * Constant "split" has been merged with constant "prod_case"; names of haftmann@41398: ML functions, facts etc. involving split have been retained so far, haftmann@41398: though. INCOMPATIBILITY. haftmann@41398: haftmann@41398: * Dropped old infix syntax "_ mem _" for List.member; use "_ : set _" haftmann@41398: instead. INCOMPATIBILITY. haftmann@41398: wenzelm@41571: * Removed lemma "Option.is_none_none" which duplicates "is_none_def". haftmann@41398: INCOMPATIBILITY. haftmann@41398: wenzelm@41594: * Former theory Library/Enum is now part of the HOL-Main image. wenzelm@41594: INCOMPATIBILITY: all constants of the Enum theory now have to be wenzelm@41594: referred to by its qualified name. wenzelm@41594: wenzelm@41594: enum ~> Enum.enum wenzelm@41594: nlists ~> Enum.nlists wenzelm@41594: product ~> Enum.product wenzelm@41594: wenzelm@41594: * Theory Library/Monad_Syntax provides do-syntax for monad types. wenzelm@41594: Syntax in Library/State_Monad has been changed to avoid ambiguities. wenzelm@41594: INCOMPATIBILITY. wenzelm@41594: wenzelm@41594: * Theory Library/SetsAndFunctions has been split into wenzelm@41594: Library/Function_Algebras and Library/Set_Algebras; canonical names wenzelm@41594: for instance definitions for functions; various improvements. wenzelm@41594: INCOMPATIBILITY. wenzelm@41594: wenzelm@41594: * Theory Library/Multiset provides stable quicksort implementation of wenzelm@41594: sort_key. wenzelm@41594: wenzelm@41594: * Theory Library/Multiset: renamed empty_idemp ~> empty_neutral. wenzelm@41594: INCOMPATIBILITY. wenzelm@41594: wenzelm@41594: * Session Multivariate_Analysis: introduced a type class for euclidean wenzelm@41594: space. Most theorems are now stated in terms of euclidean spaces wenzelm@41594: instead of finite cartesian products. wenzelm@41594: wenzelm@41594: types wenzelm@41594: real ^ 'n ~> 'a::real_vector wenzelm@41594: ~> 'a::euclidean_space wenzelm@41594: ~> 'a::ordered_euclidean_space wenzelm@41594: (depends on your needs) wenzelm@41594: wenzelm@41594: constants wenzelm@41594: _ $ _ ~> _ $$ _ wenzelm@41594: \ x. _ ~> \\ x. _ wenzelm@41594: CARD('n) ~> DIM('a) wenzelm@41594: wenzelm@41594: Also note that the indices are now natural numbers and not from some wenzelm@41594: finite type. Finite cartesian products of euclidean spaces, products wenzelm@41594: of euclidean spaces the real and complex numbers are instantiated to wenzelm@41594: be euclidean_spaces. INCOMPATIBILITY. wenzelm@41594: wenzelm@41594: * Session Probability: introduced pextreal as positive extended real wenzelm@41594: numbers. Use pextreal as value for measures. Introduce the wenzelm@41594: Radon-Nikodym derivative, product spaces and Fubini's theorem for wenzelm@41594: arbitrary sigma finite measures. Introduces Lebesgue measure based on wenzelm@41594: the integral in Multivariate Analysis. INCOMPATIBILITY. wenzelm@41594: wenzelm@41594: * Session Imperative_HOL: revamped, corrected dozens of inadequacies. wenzelm@41594: INCOMPATIBILITY. wenzelm@41594: wenzelm@41594: * Session SPARK (with image HOL-SPARK) provides commands to load and wenzelm@41594: prove verification conditions generated by the SPARK Ada program wenzelm@41594: verifier. See also src/HOL/SPARK and src/HOL/SPARK/Examples. berghofe@41567: huffman@40621: ballarin@41433: *** HOL-Algebra *** ballarin@41433: ballarin@41433: * Theorems for additive ring operations (locale abelian_monoid and ballarin@41433: descendants) are generated by interpretation from their multiplicative ballarin@41434: counterparts. Names (in particular theorem names) have the mandatory ballarin@41434: qualifier 'add'. Previous theorem names are redeclared for ballarin@41434: compatibility. ballarin@41434: wenzelm@41571: * Structure "int_ring" is now an abbreviation (previously a ballarin@41434: definition). This fits more natural with advanced interpretations. ballarin@41433: ballarin@41433: huffman@40621: *** HOLCF *** huffman@40621: huffman@40621: * The domain package now runs in definitional mode by default: The wenzelm@41571: former command 'new_domain' is now called 'domain'. To use the domain huffman@40621: package in its original axiomatic mode, use 'domain (unsafe)'. huffman@40621: INCOMPATIBILITY. huffman@40621: wenzelm@41571: * The new class "domain" is now the default sort. Class "predomain" wenzelm@41571: is an unpointed version of "domain". Theories can be updated by wenzelm@41571: replacing sort annotations as shown below. INCOMPATIBILITY. huffman@40621: huffman@40621: 'a::type ~> 'a::countable huffman@40621: 'a::cpo ~> 'a::predomain huffman@40621: 'a::pcpo ~> 'a::domain huffman@40621: wenzelm@41571: * The old type class "rep" has been superseded by class "domain". huffman@40621: Accordingly, users of the definitional package must remove any wenzelm@41571: "default_sort rep" declarations. INCOMPATIBILITY. huffman@40621: huffman@41401: * The domain package (definitional mode) now supports unpointed huffman@41401: predomain argument types, as long as they are marked 'lazy'. (Strict wenzelm@41571: arguments must be in class "domain".) For example, the following huffman@41401: domain definition now works: huffman@41401: huffman@41401: domain natlist = nil | cons (lazy "nat discr") (lazy "natlist") huffman@41401: huffman@41401: * Theory HOLCF/Library/HOL_Cpo provides cpo and predomain class wenzelm@41571: instances for types from main HOL: bool, nat, int, char, 'a + 'b, wenzelm@41571: 'a option, and 'a list. Additionally, it configures fixrec and the wenzelm@41571: domain package to work with these types. For example: huffman@41401: huffman@41401: fixrec isInl :: "('a + 'b) u -> tr" huffman@41401: where "isInl$(up$(Inl x)) = TT" | "isInl$(up$(Inr y)) = FF" huffman@41401: huffman@41401: domain V = VFun (lazy "V -> V") | VCon (lazy "nat") (lazy "V list") huffman@41401: wenzelm@41571: * The "(permissive)" option of fixrec has been replaced with a wenzelm@41571: per-equation "(unchecked)" option. See wenzelm@41571: src/HOL/HOLCF/Tutorial/Fixrec_ex.thy for examples. INCOMPATIBILITY. wenzelm@41571: wenzelm@41571: * The "bifinite" class no longer fixes a constant "approx"; the class wenzelm@41571: now just asserts that such a function exists. INCOMPATIBILITY. wenzelm@41571: wenzelm@41571: * Former type "alg_defl" has been renamed to "defl". HOLCF no longer huffman@41287: defines an embedding of type 'a defl into udom by default; instances wenzelm@41571: of "bifinite" and "domain" classes are available in wenzelm@41571: src/HOL/HOLCF/Library/Defl_Bifinite.thy. wenzelm@41571: wenzelm@41571: * The syntax "REP('a)" has been replaced with "DEFL('a)". wenzelm@41571: wenzelm@41571: * The predicate "directed" has been removed. INCOMPATIBILITY. wenzelm@41571: wenzelm@41571: * The type class "finite_po" has been removed. INCOMPATIBILITY. wenzelm@41571: wenzelm@41571: * The function "cprod_map" has been renamed to "prod_map". huffman@41401: INCOMPATIBILITY. huffman@41401: huffman@41401: * The monadic bind operator on each powerdomain has new binder syntax wenzelm@41571: similar to sets, e.g. "\\x\xs. t" represents wenzelm@41571: "upper_bind\xs\(\ x. t)". huffman@41401: huffman@41401: * The infix syntax for binary union on each powerdomain has changed wenzelm@41571: from e.g. "+\" to "\\", for consistency with set wenzelm@41571: syntax. INCOMPATIBILITY. wenzelm@41571: wenzelm@41571: * The constant "UU" has been renamed to "bottom". The syntax "UU" is huffman@41429: still supported as an input translation. huffman@41429: huffman@40621: * Renamed some theorems (the original names are also still available). wenzelm@41571: huffman@40621: expand_fun_below ~> fun_below_iff huffman@40621: below_fun_ext ~> fun_belowI huffman@40621: expand_cfun_eq ~> cfun_eq_iff huffman@40621: ext_cfun ~> cfun_eqI huffman@40621: expand_cfun_below ~> cfun_below_iff huffman@40621: below_cfun_ext ~> cfun_belowI huffman@40621: cont2cont_Rep_CFun ~> cont2cont_APP huffman@40621: huffman@40621: * The Abs and Rep functions for various types have changed names. wenzelm@40948: Related theorem names have also changed to match. INCOMPATIBILITY. wenzelm@41571: huffman@40621: Rep_CFun ~> Rep_cfun huffman@40621: Abs_CFun ~> Abs_cfun huffman@40621: Rep_Sprod ~> Rep_sprod huffman@40621: Abs_Sprod ~> Abs_sprod huffman@40621: Rep_Ssum ~> Rep_ssum huffman@40621: Abs_Ssum ~> Abs_ssum huffman@40621: huffman@40621: * Lemmas with names of the form *_defined_iff or *_strict_iff have wenzelm@41571: been renamed to *_bottom_iff. INCOMPATIBILITY. huffman@40621: huffman@40621: * Various changes to bisimulation/coinduction with domain package: wenzelm@41571: wenzelm@41571: - Definitions of "bisim" constants no longer mention definedness. wenzelm@41571: - With mutual recursion, "bisim" predicate is now curried. huffman@40621: - With mutual recursion, each type gets a separate coind theorem. huffman@40621: - Variable names in bisim_def and coinduct rules have changed. wenzelm@41571: huffman@40621: INCOMPATIBILITY. huffman@40621: wenzelm@41571: * Case combinators generated by the domain package for type "foo" are wenzelm@41571: now named "foo_case" instead of "foo_when". INCOMPATIBILITY. huffman@40621: huffman@40771: * Several theorems have been renamed to more accurately reflect the wenzelm@41571: names of constants and types involved. INCOMPATIBILITY. wenzelm@41571: huffman@40771: thelub_const ~> lub_const huffman@40771: lub_const ~> is_lub_const huffman@40771: thelubI ~> lub_eqI huffman@40771: is_lub_lub ~> is_lubD2 huffman@40771: lubI ~> is_lub_lub huffman@40771: unique_lub ~> is_lub_unique huffman@40771: is_ub_lub ~> is_lub_rangeD1 huffman@40771: lub_bin_chain ~> is_lub_bin_chain huffman@41030: lub_fun ~> is_lub_fun huffman@41030: thelub_fun ~> lub_fun huffman@41031: thelub_cfun ~> lub_cfun huffman@40771: thelub_Pair ~> lub_Pair huffman@40771: lub_cprod ~> is_lub_prod huffman@40771: thelub_cprod ~> lub_prod huffman@40771: minimal_cprod ~> minimal_prod huffman@40771: inst_cprod_pcpo ~> inst_prod_pcpo huffman@41430: UU_I ~> bottomI huffman@41430: compact_UU ~> compact_bottom huffman@41430: deflation_UU ~> deflation_bottom huffman@41430: finite_deflation_UU ~> finite_deflation_bottom huffman@40771: wenzelm@41571: * Many legacy theorem names have been discontinued. INCOMPATIBILITY. wenzelm@41571: huffman@40621: sq_ord_less_eq_trans ~> below_eq_trans huffman@40621: sq_ord_eq_less_trans ~> eq_below_trans huffman@40621: refl_less ~> below_refl huffman@40621: trans_less ~> below_trans huffman@40621: antisym_less ~> below_antisym huffman@40621: antisym_less_inverse ~> po_eq_conv [THEN iffD1] huffman@40621: box_less ~> box_below huffman@40621: rev_trans_less ~> rev_below_trans huffman@40621: not_less2not_eq ~> not_below2not_eq huffman@40621: less_UU_iff ~> below_UU_iff huffman@40621: flat_less_iff ~> flat_below_iff huffman@40621: adm_less ~> adm_below huffman@40621: adm_not_less ~> adm_not_below huffman@40621: adm_compact_not_less ~> adm_compact_not_below huffman@40621: less_fun_def ~> below_fun_def huffman@40621: expand_fun_less ~> fun_below_iff huffman@40621: less_fun_ext ~> fun_belowI huffman@40621: less_discr_def ~> below_discr_def huffman@40621: discr_less_eq ~> discr_below_eq huffman@40621: less_unit_def ~> below_unit_def huffman@40621: less_cprod_def ~> below_prod_def huffman@40621: prod_lessI ~> prod_belowI huffman@40621: Pair_less_iff ~> Pair_below_iff huffman@40621: fst_less_iff ~> fst_below_iff huffman@40621: snd_less_iff ~> snd_below_iff huffman@40621: expand_cfun_less ~> cfun_below_iff huffman@40621: less_cfun_ext ~> cfun_belowI huffman@40621: injection_less ~> injection_below huffman@40621: less_up_def ~> below_up_def huffman@40621: not_Iup_less ~> not_Iup_below huffman@40621: Iup_less ~> Iup_below huffman@40621: up_less ~> up_below huffman@40621: Def_inject_less_eq ~> Def_below_Def huffman@40621: Def_less_is_eq ~> Def_below_iff huffman@40621: spair_less_iff ~> spair_below_iff huffman@40621: less_sprod ~> below_sprod huffman@40621: spair_less ~> spair_below huffman@40621: sfst_less_iff ~> sfst_below_iff huffman@40621: ssnd_less_iff ~> ssnd_below_iff huffman@40621: fix_least_less ~> fix_least_below huffman@40621: dist_less_one ~> dist_below_one huffman@40621: less_ONE ~> below_ONE huffman@40621: ONE_less_iff ~> ONE_below_iff huffman@40621: less_sinlD ~> below_sinlD huffman@40621: less_sinrD ~> below_sinrD huffman@40621: huffman@40621: wenzelm@40948: *** FOL and ZF *** haftmann@38522: wenzelm@41310: * All constant names are now qualified internally and use proper wenzelm@41310: identifiers, e.g. "IFOL.eq" instead of "op =". INCOMPATIBILITY. wenzelm@41310: haftmann@38522: wenzelm@37868: *** ML *** wenzelm@37868: wenzelm@41594: * Antiquotation @{assert} inlines a function bool -> unit that raises wenzelm@41594: Fail if the argument is false. Due to inlining the source position of wenzelm@41594: failed assertions is included in the error output. wenzelm@41594: wenzelm@41594: * Discontinued antiquotation @{theory_ref}, which is obsolete since ML wenzelm@41594: text is in practice always evaluated with a stable theory checkpoint. wenzelm@41594: Minor INCOMPATIBILITY, use (Theory.check_thy @{theory}) instead. wenzelm@41594: wenzelm@41594: * Antiquotation @{theory A} refers to theory A from the ancestry of wenzelm@41594: the current context, not any accidental theory loader state as before. wenzelm@41594: Potential INCOMPATIBILITY, subtle change in semantics. wenzelm@41228: wenzelm@40956: * Syntax.pretty_priority (default 0) configures the required priority wenzelm@40956: of pretty-printed output and thus affects insertion of parentheses. wenzelm@40956: wenzelm@40959: * Syntax.default_root (default "any") configures the inner syntax wenzelm@40959: category (nonterminal symbol) for parsing of terms. wenzelm@40959: wenzelm@40722: * Former exception Library.UnequalLengths now coincides with wenzelm@40722: ListPair.UnequalLengths. wenzelm@40722: wenzelm@41594: * Renamed structure MetaSimplifier to Raw_Simplifier. Note that the wenzelm@41594: main functionality is provided by structure Simplifier. wenzelm@41594: wenzelm@40627: * Renamed raw "explode" function to "raw_explode" to emphasize its wenzelm@40627: meaning. Note that internally to Isabelle, Symbol.explode is used in wenzelm@40627: almost all situations. wenzelm@40627: wenzelm@40318: * Discontinued obsolete function sys_error and exception SYS_ERROR. wenzelm@40318: See implementation manual for further details on exceptions in wenzelm@40318: Isabelle/ML. wenzelm@40318: wenzelm@39616: * Renamed setmp_noncritical to Unsynchronized.setmp to emphasize its wenzelm@39616: meaning. wenzelm@39616: wenzelm@39557: * Renamed structure PureThy to Pure_Thy and moved most of its wenzelm@39557: operations to structure Global_Theory, to emphasize that this is wenzelm@39557: rarely-used global-only stuff. wenzelm@39557: wenzelm@39513: * Discontinued Output.debug. Minor INCOMPATIBILITY, use plain writeln wenzelm@39513: instead (or tracing for high-volume output). wenzelm@39513: wenzelm@38980: * Configuration option show_question_marks only affects regular pretty wenzelm@38980: printing of types and terms, not raw Term.string_of_vname. wenzelm@38980: wenzelm@39164: * ML_Context.thm and ML_Context.thms are no longer pervasive. Rare wenzelm@39164: INCOMPATIBILITY, superseded by static antiquotations @{thm} and wenzelm@39164: @{thms} for most purposes. wenzelm@39164: wenzelm@41594: * ML structure Unsynchronized is never opened, not even in Isar wenzelm@38980: interaction mode as before. Old Unsynchronized.set etc. have been wenzelm@38980: discontinued -- use plain := instead. This should be *rare* anyway, wenzelm@38980: since modern tools always work via official context data, notably wenzelm@38980: configuration options. wenzelm@38980: wenzelm@39239: * Parallel and asynchronous execution requires special care concerning wenzelm@39239: interrupts. Structure Exn provides some convenience functions that wenzelm@39239: avoid working directly with raw Interrupt. User code must not absorb wenzelm@39239: interrupts -- intermediate handling (for cleanup etc.) needs to be wenzelm@39239: followed by re-raising of the original exception. Another common wenzelm@39239: source of mistakes are "handle _" patterns, which make the meaning of wenzelm@39239: the program subject to physical effects of the environment. wenzelm@39239: wenzelm@37868: wenzelm@37868: wenzelm@37144: New in Isabelle2009-2 (June 2010) wenzelm@37144: --------------------------------- haftmann@33993: wenzelm@35260: *** General *** wenzelm@35260: wenzelm@35436: * Authentic syntax for *all* logical entities (type classes, type wenzelm@35436: constructors, term constants): provides simple and robust wenzelm@35436: correspondence between formal entities and concrete syntax. Within wenzelm@35436: the parse tree / AST representations, "constants" are decorated by wenzelm@35436: their category (class, type, const) and spelled out explicitly with wenzelm@35436: their full internal name. wenzelm@35436: wenzelm@35436: Substantial INCOMPATIBILITY concerning low-level syntax declarations wenzelm@35436: and translations (translation rules and translation functions in ML). wenzelm@35436: Some hints on upgrading: wenzelm@35260: wenzelm@35260: - Many existing uses of 'syntax' and 'translations' can be replaced wenzelm@35436: by more modern 'type_notation', 'notation' and 'abbreviation', wenzelm@35436: which are independent of this issue. wenzelm@35260: wenzelm@35260: - 'translations' require markup within the AST; the term syntax wenzelm@35260: provides the following special forms: wenzelm@35260: wenzelm@35260: CONST c -- produces syntax version of constant c from context wenzelm@35261: XCONST c -- literally c, checked as constant from context wenzelm@35261: c -- literally c, if declared by 'syntax' wenzelm@35261: wenzelm@35261: Plain identifiers are treated as AST variables -- occasionally the wenzelm@35261: system indicates accidental variables via the error "rhs contains wenzelm@35261: extra variables". wenzelm@35260: wenzelm@35436: Type classes and type constructors are marked according to their wenzelm@35436: concrete syntax. Some old translations rules need to be written wenzelm@35436: for the "type" category, using type constructor application wenzelm@35436: instead of pseudo-term application of the default category wenzelm@35436: "logic". wenzelm@35436: wenzelm@35260: - 'parse_translation' etc. in ML may use the following wenzelm@35260: antiquotations: wenzelm@35260: wenzelm@35436: @{class_syntax c} -- type class c within parse tree / AST wenzelm@35436: @{term_syntax c} -- type constructor c within parse tree / AST wenzelm@35260: @{const_syntax c} -- ML version of "CONST c" above wenzelm@35260: @{syntax_const c} -- literally c (checked wrt. 'syntax' declarations) wenzelm@35260: wenzelm@35436: - Literal types within 'typed_print_translations', i.e. those *not* wenzelm@35436: represented as pseudo-terms are represented verbatim. Use @{class wenzelm@35436: c} or @{type_name c} here instead of the above syntax wenzelm@35436: antiquotations. wenzelm@35436: wenzelm@35260: Note that old non-authentic syntax was based on unqualified base wenzelm@35436: names, so all of the above "constant" names would coincide. Recall wenzelm@35436: that 'print_syntax' and ML_command "set Syntax.trace_ast" help to wenzelm@35436: diagnose syntax problems. wenzelm@35260: wenzelm@35351: * Type constructors admit general mixfix syntax, not just infix. wenzelm@35351: wenzelm@36508: * Concrete syntax may be attached to local entities without a proof wenzelm@36508: body, too. This works via regular mixfix annotations for 'fix', wenzelm@36508: 'def', 'obtain' etc. or via the explicit 'write' command, which is wenzelm@36508: similar to the 'notation' command in theory specifications. wenzelm@36508: wenzelm@37351: * Discontinued unnamed infix syntax (legacy feature for many years) -- wenzelm@37351: need to specify constant name and syntax separately. Internal ML wenzelm@37351: datatype constructors have been renamed from InfixName to Infix etc. wenzelm@37351: Minor INCOMPATIBILITY. wenzelm@37351: wenzelm@37351: * Schematic theorem statements need to be explicitly markup as such, wenzelm@37351: via commands 'schematic_lemma', 'schematic_theorem', wenzelm@37351: 'schematic_corollary'. Thus the relevance of the proof is made wenzelm@37351: syntactically clear, which impacts performance in a parallel or wenzelm@37351: asynchronous interactive environment. Minor INCOMPATIBILITY. wenzelm@37351: wenzelm@35613: * Use of cumulative prems via "!" in some proof methods has been wenzelm@37351: discontinued (old legacy feature). wenzelm@35613: boehmes@35979: * References 'trace_simp' and 'debug_simp' have been replaced by wenzelm@36857: configuration options stored in the context. Enabling tracing (the wenzelm@36857: case of debugging is similar) in proofs works via wenzelm@36857: wenzelm@36857: using [[trace_simp = true]] wenzelm@36857: wenzelm@36857: Tracing is then active for all invocations of the simplifier in wenzelm@36857: subsequent goal refinement steps. Tracing may also still be enabled or wenzelm@40780: disabled via the ProofGeneral settings menu. boehmes@35979: wenzelm@36177: * Separate commands 'hide_class', 'hide_type', 'hide_const', wenzelm@36177: 'hide_fact' replace the former 'hide' KIND command. Minor wenzelm@36177: INCOMPATIBILITY. wenzelm@36177: wenzelm@37298: * Improved parallelism of proof term normalization: usedir -p2 -q0 is wenzelm@37298: more efficient than combinations with -q1 or -q2. wenzelm@37298: wenzelm@35260: haftmann@34170: *** Pure *** haftmann@34170: wenzelm@37351: * Proofterms record type-class reasoning explicitly, using the wenzelm@37351: "unconstrain" operation internally. This eliminates all sort wenzelm@37351: constraints from a theorem and proof, introducing explicit wenzelm@37351: OFCLASS-premises. On the proof term level, this operation is wenzelm@37351: automatically applied at theorem boundaries, such that closed proofs wenzelm@37351: are always free of sort constraints. INCOMPATIBILITY for tools that wenzelm@37351: inspect proof terms. haftmann@36147: wenzelm@35765: * Local theory specifications may depend on extra type variables that wenzelm@35765: are not present in the result type -- arguments TYPE('a) :: 'a itself wenzelm@35765: are added internally. For example: wenzelm@35765: wenzelm@35765: definition unitary :: bool where "unitary = (ALL (x::'a) y. x = y)" wenzelm@35765: wenzelm@37351: * Predicates of locales introduced by classes carry a mandatory wenzelm@37351: "class" prefix. INCOMPATIBILITY. wenzelm@37351: wenzelm@37351: * Vacuous class specifications observe default sort. INCOMPATIBILITY. wenzelm@37351: wenzelm@37351: * Old 'axclass' command has been discontinued. INCOMPATIBILITY, use wenzelm@37351: 'class' instead. wenzelm@37351: wenzelm@37351: * Command 'code_reflect' allows to incorporate generated ML code into wenzelm@37351: runtime environment; replaces immature code_datatype antiquotation. wenzelm@37351: INCOMPATIBILITY. wenzelm@37351: wenzelm@37351: * Code generator: simple concept for abstract datatypes obeying wenzelm@37351: invariants. wenzelm@37351: wenzelm@36857: * Code generator: details of internal data cache have no impact on the wenzelm@36857: user space functionality any longer. wenzelm@36857: wenzelm@37351: * Methods "unfold_locales" and "intro_locales" ignore non-locale wenzelm@37351: subgoals. This is more appropriate for interpretations with 'where'. wenzelm@36857: INCOMPATIBILITY. haftmann@34170: wenzelm@36356: * Command 'example_proof' opens an empty proof body. This allows to wenzelm@36356: experiment with Isar, without producing any persistent result. wenzelm@36356: wenzelm@35413: * Commands 'type_notation' and 'no_type_notation' declare type syntax wenzelm@35413: within a local theory context, with explicit checking of the wenzelm@35413: constructors involved (in contrast to the raw 'syntax' versions). wenzelm@35413: wenzelm@36178: * Commands 'types' and 'typedecl' now work within a local theory wenzelm@36178: context -- without introducing dependencies on parameters or wenzelm@36178: assumptions, which is not possible in Isabelle/Pure. wenzelm@35681: wenzelm@36857: * Command 'defaultsort' has been renamed to 'default_sort', it works wenzelm@36857: within a local theory context. Minor INCOMPATIBILITY. wenzelm@36454: haftmann@34170: haftmann@33993: *** HOL *** haftmann@33993: wenzelm@37351: * Command 'typedef' now works within a local theory context -- without wenzelm@37351: introducing dependencies on parameters or assumptions, which is not wenzelm@37351: possible in Isabelle/Pure/HOL. Note that the logical environment may wenzelm@37351: contain multiple interpretations of local typedefs (with different wenzelm@37351: non-emptiness proofs), even in a global theory context. wenzelm@37351: wenzelm@37351: * New package for quotient types. Commands 'quotient_type' and wenzelm@37351: 'quotient_definition' may be used for defining types and constants by wenzelm@37351: quotient constructions. An example is the type of integers created by wenzelm@37351: quotienting pairs of natural numbers: wenzelm@37380: wenzelm@37351: fun wenzelm@37380: intrel :: "(nat * nat) => (nat * nat) => bool" wenzelm@37351: where wenzelm@37351: "intrel (x, y) (u, v) = (x + v = u + y)" wenzelm@37351: wenzelm@37380: quotient_type int = "nat * nat" / intrel wenzelm@37351: by (auto simp add: equivp_def expand_fun_eq) wenzelm@37380: wenzelm@37351: quotient_definition wenzelm@37351: "0::int" is "(0::nat, 0::nat)" wenzelm@37351: wenzelm@37351: The method "lifting" can be used to lift of theorems from the wenzelm@37351: underlying "raw" type to the quotient type. The example wenzelm@37351: src/HOL/Quotient_Examples/FSet.thy includes such a quotient wenzelm@37351: construction and provides a reasoning infrastructure for finite sets. wenzelm@37351: wenzelm@37351: * Renamed Library/Quotient.thy to Library/Quotient_Type.thy to avoid wenzelm@37351: clash with new theory Quotient in Main HOL. wenzelm@37351: wenzelm@37351: * Moved the SMT binding into the main HOL session, eliminating wenzelm@37351: separate HOL-SMT session. wenzelm@37351: haftmann@37020: * List membership infix mem operation is only an input abbreviation. haftmann@37020: INCOMPATIBILITY. haftmann@37020: wenzelm@37144: * Theory Library/Word.thy has been removed. Use library Word/Word.thy wenzelm@37144: for future developements; former Library/Word.thy is still present in wenzelm@37144: the AFP entry RSAPPS. haftmann@36963: wenzelm@36857: * Theorem Int.int_induct renamed to Int.int_of_nat_induct and is no wenzelm@36857: longer shadowed. INCOMPATIBILITY. haftmann@36808: huffman@36836: * Dropped theorem duplicate comp_arith; use semiring_norm instead. huffman@36836: INCOMPATIBILITY. huffman@36836: huffman@36836: * Dropped theorem RealPow.real_sq_order; use power2_le_imp_le instead. huffman@36836: INCOMPATIBILITY. haftmann@36714: wenzelm@36857: * Dropped normalizing_semiring etc; use the facts in semiring classes wenzelm@36857: instead. INCOMPATIBILITY. wenzelm@36857: huffman@36979: * Dropped several real-specific versions of lemmas about floor and wenzelm@37351: ceiling; use the generic lemmas from theory "Archimedean_Field" wenzelm@37351: instead. INCOMPATIBILITY. huffman@36979: huffman@36979: floor_number_of_eq ~> floor_number_of huffman@36979: le_floor_eq_number_of ~> number_of_le_floor huffman@36979: le_floor_eq_zero ~> zero_le_floor huffman@36979: le_floor_eq_one ~> one_le_floor huffman@36979: floor_less_eq_number_of ~> floor_less_number_of huffman@36979: floor_less_eq_zero ~> floor_less_zero huffman@36979: floor_less_eq_one ~> floor_less_one huffman@36979: less_floor_eq_number_of ~> number_of_less_floor huffman@36979: less_floor_eq_zero ~> zero_less_floor huffman@36979: less_floor_eq_one ~> one_less_floor huffman@36979: floor_le_eq_number_of ~> floor_le_number_of huffman@36979: floor_le_eq_zero ~> floor_le_zero huffman@36979: floor_le_eq_one ~> floor_le_one huffman@36979: floor_subtract_number_of ~> floor_diff_number_of huffman@36979: floor_subtract_one ~> floor_diff_one huffman@36979: ceiling_number_of_eq ~> ceiling_number_of huffman@36979: ceiling_le_eq_number_of ~> ceiling_le_number_of huffman@36979: ceiling_le_zero_eq ~> ceiling_le_zero huffman@36979: ceiling_le_eq_one ~> ceiling_le_one huffman@36979: less_ceiling_eq_number_of ~> number_of_less_ceiling huffman@36979: less_ceiling_eq_zero ~> zero_less_ceiling huffman@36979: less_ceiling_eq_one ~> one_less_ceiling huffman@36979: ceiling_less_eq_number_of ~> ceiling_less_number_of huffman@36979: ceiling_less_eq_zero ~> ceiling_less_zero huffman@36979: ceiling_less_eq_one ~> ceiling_less_one huffman@36979: le_ceiling_eq_number_of ~> number_of_le_ceiling huffman@36979: le_ceiling_eq_zero ~> zero_le_ceiling huffman@36979: le_ceiling_eq_one ~> one_le_ceiling huffman@36979: ceiling_subtract_number_of ~> ceiling_diff_number_of huffman@36979: ceiling_subtract_one ~> ceiling_diff_one huffman@36979: wenzelm@37144: * Theory "Finite_Set": various folding_XXX locales facilitate the wenzelm@36857: application of the various fold combinators on finite sets. wenzelm@36857: wenzelm@36857: * Library theory "RBT" renamed to "RBT_Impl"; new library theory "RBT" wenzelm@36857: provides abstract red-black tree type which is backed by "RBT_Impl" as wenzelm@58604: implementation. INCOMPATIBILITY. haftmann@36147: huffman@36830: * Theory Library/Coinductive_List has been removed -- superseded by wenzelm@35763: AFP/thys/Coinductive. wenzelm@35763: huffman@36829: * Theory PReal, including the type "preal" and related operations, has huffman@36829: been removed. INCOMPATIBILITY. huffman@36829: wenzelm@37380: * Real: new development using Cauchy Sequences. wenzelm@37380: wenzelm@37351: * Split off theory "Big_Operators" containing setsum, setprod, wenzelm@37351: Inf_fin, Sup_fin, Min, Max from theory Finite_Set. INCOMPATIBILITY. wenzelm@36857: wenzelm@36857: * Theory "Rational" renamed to "Rat", for consistency with "Nat", wenzelm@36857: "Int" etc. INCOMPATIBILITY. wenzelm@36857: wenzelm@37351: * Constant Rat.normalize needs to be qualified. INCOMPATIBILITY. wenzelm@37143: wenzelm@36857: * New set of rules "ac_simps" provides combined assoc / commute wenzelm@36857: rewrites for all interpretations of the appropriate generic locales. wenzelm@36857: wenzelm@36857: * Renamed theory "OrderedGroup" to "Groups" and split theory wenzelm@36857: "Ring_and_Field" into theories "Rings" and "Fields"; for more wenzelm@36857: appropriate and more consistent names suitable for name prefixes wenzelm@36857: within the HOL theories. INCOMPATIBILITY. haftmann@35050: haftmann@35084: * Some generic constants have been put to appropriate theories: wenzelm@36857: - less_eq, less: Orderings wenzelm@36857: - zero, one, plus, minus, uminus, times, abs, sgn: Groups wenzelm@36857: - inverse, divide: Rings haftmann@35084: INCOMPATIBILITY. haftmann@35084: wenzelm@36857: * More consistent naming of type classes involving orderings (and wenzelm@36857: lattices): haftmann@35027: haftmann@35027: lower_semilattice ~> semilattice_inf haftmann@35027: upper_semilattice ~> semilattice_sup haftmann@35027: haftmann@35027: dense_linear_order ~> dense_linorder haftmann@35027: haftmann@35027: pordered_ab_group_add ~> ordered_ab_group_add haftmann@35027: pordered_ab_group_add_abs ~> ordered_ab_group_add_abs haftmann@35027: pordered_ab_semigroup_add ~> ordered_ab_semigroup_add haftmann@35027: pordered_ab_semigroup_add_imp_le ~> ordered_ab_semigroup_add_imp_le haftmann@35027: pordered_cancel_ab_semigroup_add ~> ordered_cancel_ab_semigroup_add haftmann@35027: pordered_cancel_comm_semiring ~> ordered_cancel_comm_semiring haftmann@35027: pordered_cancel_semiring ~> ordered_cancel_semiring haftmann@35027: pordered_comm_monoid_add ~> ordered_comm_monoid_add haftmann@35027: pordered_comm_ring ~> ordered_comm_ring haftmann@35027: pordered_comm_semiring ~> ordered_comm_semiring haftmann@35027: pordered_ring ~> ordered_ring haftmann@35027: pordered_ring_abs ~> ordered_ring_abs haftmann@35027: pordered_semiring ~> ordered_semiring haftmann@35027: haftmann@35027: ordered_ab_group_add ~> linordered_ab_group_add haftmann@35027: ordered_ab_semigroup_add ~> linordered_ab_semigroup_add haftmann@35027: ordered_cancel_ab_semigroup_add ~> linordered_cancel_ab_semigroup_add haftmann@35027: ordered_comm_semiring_strict ~> linordered_comm_semiring_strict haftmann@35027: ordered_field ~> linordered_field haftmann@35027: ordered_field_no_lb ~> linordered_field_no_lb haftmann@35027: ordered_field_no_ub ~> linordered_field_no_ub haftmann@35027: ordered_field_dense_linear_order ~> dense_linordered_field haftmann@35027: ordered_idom ~> linordered_idom haftmann@35027: ordered_ring ~> linordered_ring haftmann@35027: ordered_ring_le_cancel_factor ~> linordered_ring_le_cancel_factor haftmann@35027: ordered_ring_less_cancel_factor ~> linordered_ring_less_cancel_factor haftmann@35027: ordered_ring_strict ~> linordered_ring_strict haftmann@35027: ordered_semidom ~> linordered_semidom haftmann@35027: ordered_semiring ~> linordered_semiring haftmann@35027: ordered_semiring_1 ~> linordered_semiring_1 haftmann@35027: ordered_semiring_1_strict ~> linordered_semiring_1_strict haftmann@35027: ordered_semiring_strict ~> linordered_semiring_strict haftmann@35027: wenzelm@36857: The following slightly odd type classes have been moved to a wenzelm@37351: separate theory Library/Lattice_Algebras: haftmann@35032: haftmann@35032: lordered_ab_group_add ~> lattice_ab_group_add haftmann@35032: lordered_ab_group_add_abs ~> lattice_ab_group_add_abs haftmann@35032: lordered_ab_group_add_meet ~> semilattice_inf_ab_group_add haftmann@35032: lordered_ab_group_add_join ~> semilattice_sup_ab_group_add haftmann@35032: lordered_ring ~> lattice_ring haftmann@35032: haftmann@35027: INCOMPATIBILITY. haftmann@35027: haftmann@36416: * Refined field classes: wenzelm@36857: - classes division_ring_inverse_zero, field_inverse_zero, wenzelm@36857: linordered_field_inverse_zero include rule inverse 0 = 0 -- wenzelm@36857: subsumes former division_by_zero class; wenzelm@36857: - numerous lemmas have been ported from field to division_ring. wenzelm@36857: INCOMPATIBILITY. haftmann@36416: haftmann@36416: * Refined algebra theorem collections: wenzelm@36857: - dropped theorem group group_simps, use algebra_simps instead; wenzelm@36857: - dropped theorem group ring_simps, use field_simps instead; wenzelm@36857: - proper theorem collection field_simps subsumes former theorem wenzelm@36857: groups field_eq_simps and field_simps; wenzelm@36857: - dropped lemma eq_minus_self_iff which is a duplicate for wenzelm@36857: equal_neg_zero. wenzelm@36857: INCOMPATIBILITY. wenzelm@35009: wenzelm@35009: * Theory Finite_Set and List: some lemmas have been generalized from wenzelm@34076: sets to lattices: wenzelm@34076: haftmann@34007: fun_left_comm_idem_inter ~> fun_left_comm_idem_inf haftmann@34007: fun_left_comm_idem_union ~> fun_left_comm_idem_sup haftmann@34007: inter_Inter_fold_inter ~> inf_Inf_fold_inf haftmann@34007: union_Union_fold_union ~> sup_Sup_fold_sup haftmann@34007: Inter_fold_inter ~> Inf_fold_inf haftmann@34007: Union_fold_union ~> Sup_fold_sup haftmann@34007: inter_INTER_fold_inter ~> inf_INFI_fold_inf haftmann@34007: union_UNION_fold_union ~> sup_SUPR_fold_sup haftmann@34007: INTER_fold_inter ~> INFI_fold_inf haftmann@34007: UNION_fold_union ~> SUPR_fold_sup haftmann@34007: wenzelm@44973: * Theory "Complete_Lattice": lemmas top_def and bot_def have been wenzelm@44973: replaced by the more convenient lemmas Inf_empty and Sup_empty. wenzelm@44973: Dropped lemmas Inf_insert_simp and Sup_insert_simp, which are subsumed wenzelm@44973: by Inf_insert and Sup_insert. Lemmas Inf_UNIV and Sup_UNIV replace wenzelm@44973: former Inf_Univ and Sup_Univ. Lemmas inf_top_right and sup_bot_right wenzelm@44973: subsume inf_top and sup_bot respectively. INCOMPATIBILITY. haftmann@36416: wenzelm@36857: * Reorganized theory Multiset: swapped notation of pointwise and wenzelm@36857: multiset order: wenzelm@37351: wenzelm@36857: - pointwise ordering is instance of class order with standard syntax wenzelm@36857: <= and <; wenzelm@36857: - multiset ordering has syntax <=# and <#; partial order properties wenzelm@36857: are provided by means of interpretation with prefix wenzelm@36857: multiset_order; wenzelm@36857: - less duplication, less historical organization of sections, wenzelm@36857: conversion from associations lists to multisets, rudimentary code wenzelm@36857: generation; wenzelm@36857: - use insert_DiffM2 [symmetric] instead of elem_imp_eq_diff_union, wenzelm@36857: if needed. wenzelm@37351: nipkow@36903: Renamed: wenzelm@37351: wenzelm@37351: multiset_eq_conv_count_eq ~> multiset_ext_iff wenzelm@37351: multi_count_ext ~> multiset_ext wenzelm@37351: diff_union_inverse2 ~> diff_union_cancelR wenzelm@37351: wenzelm@36857: INCOMPATIBILITY. haftmann@36416: nipkow@36903: * Theory Permutation: replaced local "remove" by List.remove1. nipkow@36903: haftmann@36416: * Code generation: ML and OCaml code is decorated with signatures. haftmann@36416: wenzelm@35009: * Theory List: added transpose. wenzelm@35009: huffman@35810: * Library/Nat_Bijection.thy is a collection of bijective functions huffman@35810: between nat and other types, which supersedes the older libraries huffman@35810: Library/Nat_Int_Bij.thy and HOLCF/NatIso.thy. INCOMPATIBILITY. huffman@35810: huffman@35810: Constants: huffman@35810: Nat_Int_Bij.nat2_to_nat ~> prod_encode huffman@35810: Nat_Int_Bij.nat_to_nat2 ~> prod_decode huffman@35810: Nat_Int_Bij.int_to_nat_bij ~> int_encode huffman@35810: Nat_Int_Bij.nat_to_int_bij ~> int_decode huffman@35810: Countable.pair_encode ~> prod_encode huffman@35810: NatIso.prod2nat ~> prod_encode huffman@35810: NatIso.nat2prod ~> prod_decode huffman@35810: NatIso.sum2nat ~> sum_encode huffman@35810: NatIso.nat2sum ~> sum_decode huffman@35810: NatIso.list2nat ~> list_encode huffman@35810: NatIso.nat2list ~> list_decode huffman@35810: NatIso.set2nat ~> set_encode huffman@35810: NatIso.nat2set ~> set_decode huffman@35810: huffman@35810: Lemmas: huffman@35810: Nat_Int_Bij.bij_nat_to_int_bij ~> bij_int_decode huffman@35810: Nat_Int_Bij.nat2_to_nat_inj ~> inj_prod_encode huffman@35810: Nat_Int_Bij.nat2_to_nat_surj ~> surj_prod_encode huffman@35810: Nat_Int_Bij.nat_to_nat2_inj ~> inj_prod_decode huffman@35810: Nat_Int_Bij.nat_to_nat2_surj ~> surj_prod_decode huffman@35810: Nat_Int_Bij.i2n_n2i_id ~> int_encode_inverse huffman@35810: Nat_Int_Bij.n2i_i2n_id ~> int_decode_inverse huffman@35810: Nat_Int_Bij.surj_nat_to_int_bij ~> surj_int_encode huffman@35810: Nat_Int_Bij.surj_int_to_nat_bij ~> surj_int_decode huffman@35810: Nat_Int_Bij.inj_nat_to_int_bij ~> inj_int_encode huffman@35810: Nat_Int_Bij.inj_int_to_nat_bij ~> inj_int_decode huffman@35810: Nat_Int_Bij.bij_nat_to_int_bij ~> bij_int_encode huffman@35810: Nat_Int_Bij.bij_int_to_nat_bij ~> bij_int_decode huffman@35810: blanchet@36929: * Sledgehammer: blanchet@36929: - Renamed ATP commands: blanchet@36929: atp_info ~> sledgehammer running_atps blanchet@36929: atp_kill ~> sledgehammer kill_atps blanchet@36929: atp_messages ~> sledgehammer messages blanchet@36929: atp_minimize ~> sledgehammer minimize blanchet@36929: print_atps ~> sledgehammer available_atps blanchet@36929: INCOMPATIBILITY. blanchet@36929: - Added user's manual ("isabelle doc sledgehammer"). blanchet@36929: - Added option syntax and "sledgehammer_params" to customize blanchet@36929: Sledgehammer's behavior. See the manual for details. blanchet@36929: - Modified the Isar proof reconstruction code so that it produces blanchet@36929: direct proofs rather than proofs by contradiction. (This feature blanchet@36929: is still experimental.) blanchet@36929: - Made Isar proof reconstruction work for SPASS, remote ATPs, and in blanchet@36929: full-typed mode. blanchet@36929: - Added support for TPTP syntax for SPASS via the "spass_tptp" ATP. blanchet@36929: blanchet@36928: * Nitpick: blanchet@36928: - Added and implemented "binary_ints" and "bits" options. blanchet@36928: - Added "std" option and implemented support for nonstandard models. blanchet@36928: - Added and implemented "finitize" option to improve the precision blanchet@36928: of infinite datatypes based on a monotonicity analysis. blanchet@36928: - Added support for quotient types. blanchet@36928: - Added support for "specification" and "ax_specification" blanchet@36928: constructs. blanchet@36928: - Added support for local definitions (for "function" and blanchet@36928: "termination" proofs). blanchet@36928: - Added support for term postprocessors. blanchet@36928: - Optimized "Multiset.multiset" and "FinFun.finfun". blanchet@36928: - Improved efficiency of "destroy_constrs" optimization. blanchet@36928: - Fixed soundness bugs related to "destroy_constrs" optimization and blanchet@36928: record getters. blanchet@37272: - Fixed soundness bug related to higher-order constructors. blanchet@37272: - Fixed soundness bug when "full_descrs" is enabled. blanchet@36928: - Improved precision of set constructs. blanchet@37260: - Added "atoms" option. blanchet@36928: - Added cache to speed up repeated Kodkod invocations on the same blanchet@36928: problems. blanchet@36928: - Renamed "MiniSatJNI", "zChaffJNI", "BerkMinAlloy", and blanchet@36928: "SAT4JLight" to "MiniSat_JNI", "zChaff_JNI", "BerkMin_Alloy", and blanchet@36928: "SAT4J_Light". INCOMPATIBILITY. blanchet@36928: - Removed "skolemize", "uncurry", "sym_break", "flatten_prop", blanchet@36928: "sharing_depth", and "show_skolems" options. INCOMPATIBILITY. blanchet@37264: - Removed "nitpick_intro" attribute. INCOMPATIBILITY. blanchet@36928: berghofe@37361: * Method "induct" now takes instantiations of the form t, where t is not berghofe@37361: a variable, as a shorthand for "x == t", where x is a fresh variable. berghofe@37361: If this is not intended, t has to be enclosed in parentheses. berghofe@37361: By default, the equalities generated by definitional instantiations berghofe@37361: are pre-simplified, which may cause parameters of inductive cases berghofe@37361: to disappear, or may even delete some of the inductive cases. berghofe@37361: Use "induct (no_simp)" instead of "induct" to restore the old berghofe@37361: behaviour. The (no_simp) option is also understood by the "cases" berghofe@37361: and "nominal_induct" methods, which now perform pre-simplification, too. berghofe@37361: INCOMPATIBILITY. berghofe@37361: haftmann@33993: huffman@36828: *** HOLCF *** huffman@36828: huffman@36828: * Variable names in lemmas generated by the domain package have huffman@36828: changed; the naming scheme is now consistent with the HOL datatype huffman@36828: package. Some proof scripts may be affected, INCOMPATIBILITY. huffman@36828: huffman@36828: * The domain package no longer defines the function "foo_copy" for huffman@36828: recursive domain "foo". The reach lemma is now stated directly in huffman@36828: terms of "foo_take". Lemmas and proofs that mention "foo_copy" must huffman@36828: be reformulated in terms of "foo_take", INCOMPATIBILITY. huffman@36828: huffman@36828: * Most definedness lemmas generated by the domain package (previously huffman@36828: of the form "x ~= UU ==> foo$x ~= UU") now have an if-and-only-if form huffman@36828: like "foo$x = UU <-> x = UU", which works better as a simp rule. wenzelm@37351: Proofs that used definedness lemmas as intro rules may break, huffman@36828: potential INCOMPATIBILITY. huffman@36828: huffman@36828: * Induction and casedist rules generated by the domain package now huffman@36828: declare proper case_names (one called "bottom", and one named for each huffman@36828: constructor). INCOMPATIBILITY. huffman@36828: huffman@36828: * For mutually-recursive domains, separate "reach" and "take_lemma" huffman@36828: rules are generated for each domain, INCOMPATIBILITY. huffman@36828: huffman@36828: foo_bar.reach ~> foo.reach bar.reach huffman@36828: foo_bar.take_lemmas ~> foo.take_lemma bar.take_lemma huffman@36828: huffman@36828: * Some lemmas generated by the domain package have been renamed for huffman@36828: consistency with the datatype package, INCOMPATIBILITY. huffman@36828: huffman@36828: foo.ind ~> foo.induct huffman@36828: foo.finite_ind ~> foo.finite_induct huffman@36828: foo.coind ~> foo.coinduct huffman@36828: foo.casedist ~> foo.exhaust huffman@36828: foo.exhaust ~> foo.nchotomy huffman@36828: huffman@36828: * For consistency with other definition packages, the fixrec package huffman@36828: now generates qualified theorem names, INCOMPATIBILITY. huffman@36828: huffman@36828: foo_simps ~> foo.simps huffman@36828: foo_unfold ~> foo.unfold huffman@36828: foo_induct ~> foo.induct huffman@36828: huffman@37087: * The "fixrec_simp" attribute has been removed. The "fixrec_simp" huffman@37087: method and internal fixrec proofs now use the default simpset instead. huffman@37087: INCOMPATIBILITY. huffman@37087: huffman@36828: * The "contlub" predicate has been removed. Proof scripts should use huffman@36828: lemma contI2 in place of monocontlub2cont, INCOMPATIBILITY. huffman@36828: huffman@36828: * The "admw" predicate has been removed, INCOMPATIBILITY. huffman@36828: huffman@36828: * The constants cpair, cfst, and csnd have been removed in favor of huffman@36828: Pair, fst, and snd from Isabelle/HOL, INCOMPATIBILITY. huffman@36828: huffman@36828: haftmann@33993: *** ML *** haftmann@33993: wenzelm@37351: * Antiquotations for basic formal entities: wenzelm@37351: wenzelm@37351: @{class NAME} -- type class wenzelm@37351: @{class_syntax NAME} -- syntax representation of the above wenzelm@37351: wenzelm@37351: @{type_name NAME} -- logical type wenzelm@37351: @{type_abbrev NAME} -- type abbreviation wenzelm@37351: @{nonterminal NAME} -- type of concrete syntactic category wenzelm@37351: @{type_syntax NAME} -- syntax representation of any of the above wenzelm@37351: wenzelm@37351: @{const_name NAME} -- logical constant (INCOMPATIBILITY) wenzelm@37351: @{const_abbrev NAME} -- abbreviated constant wenzelm@37351: @{const_syntax NAME} -- syntax representation of any of the above wenzelm@37351: wenzelm@37351: * Antiquotation @{syntax_const NAME} ensures that NAME refers to a raw wenzelm@37351: syntax constant (cf. 'syntax' command). wenzelm@37351: wenzelm@37351: * Antiquotation @{make_string} inlines a function to print arbitrary wenzelm@37351: values similar to the ML toplevel. The result is compiler dependent wenzelm@37351: and may fall back on "?" in certain situations. wenzelm@37351: wenzelm@37351: * Diagnostic commands 'ML_val' and 'ML_command' may refer to wenzelm@37351: antiquotations @{Isar.state} and @{Isar.goal}. This replaces impure wenzelm@37351: Isar.state() and Isar.goal(), which belong to the old TTY loop and do wenzelm@37351: not work with the asynchronous Isar document model. wenzelm@37351: wenzelm@37351: * Configuration options now admit dynamic default values, depending on wenzelm@37351: the context or even global references. wenzelm@37351: wenzelm@37351: * SHA1.digest digests strings according to SHA-1 (see RFC 3174). It wenzelm@37351: uses an efficient external library if available (for Poly/ML). wenzelm@37351: wenzelm@37144: * Renamed some important ML structures, while keeping the old names wenzelm@37144: for some time as aliases within the structure Legacy: wenzelm@37144: wenzelm@37144: OuterKeyword ~> Keyword wenzelm@37144: OuterLex ~> Token wenzelm@37144: OuterParse ~> Parse wenzelm@37144: OuterSyntax ~> Outer_Syntax wenzelm@37216: PrintMode ~> Print_Mode wenzelm@37144: SpecParse ~> Parse_Spec wenzelm@37216: ThyInfo ~> Thy_Info wenzelm@37216: ThyLoad ~> Thy_Load wenzelm@37216: ThyOutput ~> Thy_Output wenzelm@37145: TypeInfer ~> Type_Infer wenzelm@37144: wenzelm@37144: Note that "open Legacy" simplifies porting of sources, but forgetting wenzelm@37144: to remove it again will complicate porting again in the future. wenzelm@37144: wenzelm@37144: * Most operations that refer to a global context are named wenzelm@37144: accordingly, e.g. Simplifier.global_context or wenzelm@37144: ProofContext.init_global. There are some situations where a global wenzelm@37144: context actually works, but under normal circumstances one needs to wenzelm@37144: pass the proper local context through the code! wenzelm@37144: wenzelm@37144: * Discontinued old TheoryDataFun with its copy/init operation -- data wenzelm@37144: needs to be pure. Functor Theory_Data_PP retains the traditional wenzelm@37144: Pretty.pp argument to merge, which is absent in the standard wenzelm@37144: Theory_Data version. wenzelm@36429: wenzelm@37144: * Sorts.certify_sort and derived "cert" operations for types and terms wenzelm@37144: no longer minimize sorts. Thus certification at the boundary of the wenzelm@37144: inference kernel becomes invariant under addition of class relations, wenzelm@37144: which is an important monotonicity principle. Sorts are now minimized wenzelm@37144: in the syntax layer only, at the boundary between the end-user and the wenzelm@37144: system. Subtle INCOMPATIBILITY, may have to use Sign.minimize_sort wenzelm@37144: explicitly in rare situations. wenzelm@37144: wenzelm@35021: * Renamed old-style Drule.standard to Drule.export_without_context, to wenzelm@35021: emphasize that this is in no way a standard operation. wenzelm@35021: INCOMPATIBILITY. wenzelm@35021: wenzelm@34076: * Subgoal.FOCUS (and variants): resulting goal state is normalized as wenzelm@34076: usual for resolution. Rare INCOMPATIBILITY. wenzelm@34076: wenzelm@35845: * Renamed varify/unvarify operations to varify_global/unvarify_global wenzelm@35845: to emphasize that these only work in a global situation (which is wenzelm@35845: quite rare). wenzelm@35845: wenzelm@37144: * Curried take and drop in library.ML; negative length is interpreted wenzelm@37144: as infinity (as in chop). Subtle INCOMPATIBILITY. wenzelm@36961: wenzelm@37351: * Proof terms: type substitutions on proof constants now use canonical wenzelm@37351: order of type variables. INCOMPATIBILITY for tools working with proof wenzelm@37351: terms. wenzelm@37351: wenzelm@37351: * Raw axioms/defs may no longer carry sort constraints, and raw defs wenzelm@37351: may no longer carry premises. User-level specifications are wenzelm@37351: transformed accordingly by Thm.add_axiom/add_def. wenzelm@37351: haftmann@33993: wenzelm@34238: *** System *** wenzelm@34238: wenzelm@34238: * Discontinued special HOL_USEDIR_OPTIONS for the main HOL image; wenzelm@34238: ISABELLE_USEDIR_OPTIONS applies uniformly to all sessions. Note that wenzelm@34238: proof terms are enabled unconditionally in the new HOL-Proofs image. wenzelm@34238: wenzelm@34255: * Discontinued old ISABELLE and ISATOOL environment settings (legacy wenzelm@34255: feature since Isabelle2009). Use ISABELLE_PROCESS and ISABELLE_TOOL, wenzelm@34255: respectively. wenzelm@34255: wenzelm@36201: * Old lib/scripts/polyml-platform is superseded by the wenzelm@36201: ISABELLE_PLATFORM setting variable, which defaults to the 32 bit wenzelm@36201: variant, even on a 64 bit machine. The following example setting wenzelm@36201: prefers 64 bit if available: wenzelm@36201: wenzelm@36201: ML_PLATFORM="${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}" wenzelm@36201: wenzelm@37218: * The preliminary Isabelle/jEdit application demonstrates the emerging wenzelm@37218: Isabelle/Scala layer for advanced prover interaction and integration. wenzelm@37218: See src/Tools/jEdit or "isabelle jedit" provided by the properly built wenzelm@37218: component. wenzelm@37218: wenzelm@37375: * "IsabelleText" is a Unicode font derived from Bitstream Vera Mono wenzelm@37375: and Bluesky TeX fonts. It provides the usual Isabelle symbols, wenzelm@37375: similar to the default assignment of the document preparation system wenzelm@37375: (cf. isabellesym.sty). The Isabelle/Scala class Isabelle_System wenzelm@37375: provides some operations for direct access to the font without asking wenzelm@37375: the user for manual installation. wenzelm@37375: wenzelm@34238: haftmann@33993: wenzelm@33842: New in Isabelle2009-1 (December 2009) wenzelm@33842: ------------------------------------- wenzelm@30904: wenzelm@31547: *** General *** wenzelm@31547: wenzelm@31547: * Discontinued old form of "escaped symbols" such as \\. Only wenzelm@31547: one backslash should be used, even in ML sources. wenzelm@31547: wenzelm@31547: haftmann@30951: *** Pure *** haftmann@30951: ballarin@32846: * Locale interpretation propagates mixins along the locale hierarchy. ballarin@32846: The currently only available mixins are the equations used to map ballarin@32846: local definitions to terms of the target domain of an interpretation. ballarin@32846: wenzelm@33842: * Reactivated diagnostic command 'print_interps'. Use "print_interps wenzelm@33842: loc" to print all interpretations of locale "loc" in the theory. wenzelm@33842: Interpretations in proofs are not shown. ballarin@32846: ballarin@32983: * Thoroughly revised locales tutorial. New section on conditional ballarin@32983: interpretation. ballarin@32983: wenzelm@33843: * On instantiation of classes, remaining undefined class parameters wenzelm@33843: are formally declared. INCOMPATIBILITY. wenzelm@33843: haftmann@30951: wenzelm@33842: *** Document preparation *** wenzelm@33842: wenzelm@33842: * New generalized style concept for printing terms: @{foo (style) ...} wenzelm@33842: instead of @{foo_style style ...} (old form is still retained for wenzelm@33842: backward compatibility). Styles can be also applied for wenzelm@33842: antiquotations prop, term_type and typeof. haftmann@32891: haftmann@32891: haftmann@30930: *** HOL *** haftmann@30930: wenzelm@33842: * New proof method "smt" for a combination of first-order logic with wenzelm@33842: equality, linear and nonlinear (natural/integer/real) arithmetic, and wenzelm@33842: fixed-size bitvectors; there is also basic support for higher-order wenzelm@33842: features (esp. lambda abstractions). It is an incomplete decision wenzelm@33842: procedure based on external SMT solvers using the oracle mechanism; wenzelm@33842: for the SMT solver Z3, this method is proof-producing. Certificates wenzelm@33842: are provided to avoid calling the external solvers solely for wenzelm@33842: re-checking proofs. Due to a remote SMT service there is no need for wenzelm@33842: installing SMT solvers locally. See src/HOL/SMT. wenzelm@33842: wenzelm@33842: * New commands to load and prove verification conditions generated by wenzelm@33842: the Boogie program verifier or derived systems (e.g. the Verifying C wenzelm@33842: Compiler (VCC) or Spec#). See src/HOL/Boogie. wenzelm@33842: wenzelm@33842: * New counterexample generator tool 'nitpick' based on the Kodkod wenzelm@33842: relational model finder. See src/HOL/Tools/Nitpick and wenzelm@33842: src/HOL/Nitpick_Examples. wenzelm@33842: haftmann@33860: * New commands 'code_pred' and 'values' to invoke the predicate haftmann@33860: compiler and to enumerate values of inductive predicates. haftmann@33860: haftmann@33860: * A tabled implementation of the reflexive transitive closure. haftmann@33860: haftmann@33860: * New implementation of quickcheck uses generic code generator; haftmann@33860: default generators are provided for all suitable HOL types, records haftmann@33860: and datatypes. Old quickcheck can be re-activated importing theory haftmann@33860: Library/SML_Quickcheck. haftmann@33860: wenzelm@33843: * New testing tool Mirabelle for automated proof tools. Applies wenzelm@33843: several tools and tactics like sledgehammer, metis, or quickcheck, to wenzelm@33843: every proof step in a theory. To be used in batch mode via the wenzelm@33843: "mirabelle" utility. wenzelm@33843: wenzelm@33843: * New proof method "sos" (sum of squares) for nonlinear real wenzelm@33843: arithmetic (originally due to John Harison). It requires theory wenzelm@33843: Library/Sum_Of_Squares. It is not a complete decision procedure but wenzelm@33843: works well in practice on quantifier-free real arithmetic with +, -, wenzelm@33843: *, ^, =, <= and <, i.e. boolean combinations of equalities and wenzelm@33843: inequalities between polynomials. It makes use of external wenzelm@33843: semidefinite programming solvers. Method "sos" generates a wenzelm@33843: certificate that can be pasted into the proof thus avoiding the need wenzelm@33843: to call an external tool every time the proof is checked. See wenzelm@33843: src/HOL/Library/Sum_Of_Squares. wenzelm@33843: wenzelm@33843: * New method "linarith" invokes existing linear arithmetic decision wenzelm@33843: procedure only. wenzelm@33843: wenzelm@33843: * New command 'atp_minimal' reduces result produced by Sledgehammer. wenzelm@33843: wenzelm@33843: * New Sledgehammer option "Full Types" in Proof General settings menu. wenzelm@33843: Causes full type information to be output to the ATPs. This slows wenzelm@33843: ATPs down considerably but eliminates a source of unsound "proofs" wenzelm@33843: that fail later. wenzelm@33843: wenzelm@33843: * New method "metisFT": A version of metis that uses full type wenzelm@33843: information in order to avoid failures of proof reconstruction. wenzelm@33843: wenzelm@33843: * New evaluator "approximate" approximates an real valued term using wenzelm@33843: the same method as the approximation method. wenzelm@33843: wenzelm@33843: * Method "approximate" now supports arithmetic expressions as wenzelm@33843: boundaries of intervals and implements interval splitting and Taylor wenzelm@33843: series expansion. wenzelm@33843: wenzelm@33843: * ML antiquotation @{code_datatype} inserts definition of a datatype wenzelm@33843: generated by the code generator; e.g. see src/HOL/Predicate.thy. wenzelm@33843: haftmann@33860: * New theory SupInf of the supremum and infimum operators for sets of haftmann@33860: reals. haftmann@33860: haftmann@33860: * New theory Probability, which contains a development of measure haftmann@33860: theory, eventually leading to Lebesgue integration and probability. haftmann@33860: haftmann@33860: * Extended Multivariate Analysis to include derivation and Brouwer's haftmann@33860: fixpoint theorem. wenzelm@33843: wenzelm@33842: * Reorganization of number theory, INCOMPATIBILITY: wenzelm@33873: - new number theory development for nat and int, in theories Divides wenzelm@33873: and GCD as well as in new session Number_Theory wenzelm@33873: - some constants and facts now suffixed with _nat and _int wenzelm@33873: accordingly wenzelm@33873: - former session NumberTheory now named Old_Number_Theory, including wenzelm@33873: theories Legacy_GCD and Primes (prefer Number_Theory if possible) wenzelm@33842: - moved theory Pocklington from src/HOL/Library to wenzelm@33842: src/HOL/Old_Number_Theory haftmann@32479: wenzelm@33873: * Theory GCD includes functions Gcd/GCD and Lcm/LCM for the gcd and wenzelm@33873: lcm of finite and infinite sets. It is shown that they form a complete haftmann@32600: lattice. haftmann@32600: haftmann@32600: * Class semiring_div requires superclass no_zero_divisors and proof of haftmann@32600: div_mult_mult1; theorems div_mult_mult1, div_mult_mult2, haftmann@32600: div_mult_mult1_if, div_mult_mult1 and div_mult_mult2 have been haftmann@32600: generalized to class semiring_div, subsuming former theorems haftmann@32600: zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and haftmann@32600: zdiv_zmult_zmult2. div_mult_mult1 is now [simp] by default. haftmann@32600: INCOMPATIBILITY. haftmann@32600: haftmann@32588: * Refinements to lattice classes and sets: haftmann@32064: - less default intro/elim rules in locale variant, more default haftmann@32064: intro/elim rules in class variant: more uniformity wenzelm@33842: - lemma ge_sup_conv renamed to le_sup_iff, in accordance with wenzelm@33842: le_inf_iff wenzelm@33842: - dropped lemma alias inf_ACI for inf_aci (same for sup_ACI and wenzelm@33842: sup_aci) haftmann@32064: - renamed ACI to inf_sup_aci haftmann@32600: - new class "boolean_algebra" wenzelm@33842: - class "complete_lattice" moved to separate theory haftmann@33860: "Complete_Lattice"; corresponding constants (and abbreviations) wenzelm@33842: renamed and with authentic syntax: haftmann@33860: Set.Inf ~> Complete_Lattice.Inf haftmann@33860: Set.Sup ~> Complete_Lattice.Sup haftmann@33860: Set.INFI ~> Complete_Lattice.INFI haftmann@33860: Set.SUPR ~> Complete_Lattice.SUPR haftmann@33860: Set.Inter ~> Complete_Lattice.Inter haftmann@33860: Set.Union ~> Complete_Lattice.Union haftmann@33860: Set.INTER ~> Complete_Lattice.INTER haftmann@33860: Set.UNION ~> Complete_Lattice.UNION haftmann@32600: - authentic syntax for haftmann@32600: Set.Pow haftmann@32600: Set.image haftmann@32588: - mere abbreviations: haftmann@32588: Set.empty (for bot) haftmann@32588: Set.UNIV (for top) haftmann@33860: Set.inter (for inf, formerly Set.Int) haftmann@33860: Set.union (for sup, formerly Set.Un) haftmann@32588: Complete_Lattice.Inter (for Inf) haftmann@32588: Complete_Lattice.Union (for Sup) haftmann@32606: Complete_Lattice.INTER (for INFI) haftmann@32606: Complete_Lattice.UNION (for SUPR) haftmann@32600: - object-logic definitions as far as appropriate haftmann@32217: haftmann@32691: INCOMPATIBILITY. Care is required when theorems Int_subset_iff or wenzelm@33842: Un_subset_iff are explicitly deleted as default simp rules; then also wenzelm@33842: their lattice counterparts le_inf_iff and le_sup_iff have to be haftmann@32691: deleted to achieve the desired effect. haftmann@32064: wenzelm@33842: * Rules inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp wenzelm@33842: rules by default any longer; the same applies to min_max.inf_absorb1 wenzelm@33842: etc. INCOMPATIBILITY. wenzelm@33842: wenzelm@33842: * Rules sup_Int_eq and sup_Un_eq are no longer declared as wenzelm@33842: pred_set_conv by default. INCOMPATIBILITY. wenzelm@33842: wenzelm@33842: * Power operations on relations and functions are now one dedicated haftmann@32706: constant "compow" with infix syntax "^^". Power operation on wenzelm@31547: multiplicative monoids retains syntax "^" and is now defined generic wenzelm@31547: in class power. INCOMPATIBILITY. wenzelm@31547: wenzelm@33842: * Relation composition "R O S" now has a more standard argument order: wenzelm@33842: "R O S = {(x, z). EX y. (x, y) : R & (y, z) : S}". INCOMPATIBILITY, wenzelm@33842: rewrite propositions with "S O R" --> "R O S". Proofs may occasionally wenzelm@33842: break, since the O_assoc rule was not rewritten like this. Fix using wenzelm@33842: O_assoc[symmetric]. The same applies to the curried version "R OO S". wenzelm@32427: nipkow@33057: * Function "Inv" is renamed to "inv_into" and function "inv" is now an wenzelm@33842: abbreviation for "inv_into UNIV". Lemmas are renamed accordingly. nipkow@32988: INCOMPATIBILITY. nipkow@32988: haftmann@33860: * Most rules produced by inductive and datatype package have mandatory haftmann@33860: prefixes. INCOMPATIBILITY. nipkow@31790: wenzelm@33842: * Changed "DERIV_intros" to a dynamic fact, which can be augmented by wenzelm@33842: the attribute of the same name. Each of the theorems in the list wenzelm@33842: DERIV_intros assumes composition with an additional function and wenzelm@33842: matches a variable to the derivative, which has to be solved by the wenzelm@33842: Simplifier. Hence (auto intro!: DERIV_intros) computes the derivative wenzelm@33873: of most elementary terms. Former Maclauren.DERIV_tac and wenzelm@33873: Maclauren.deriv_tac should be replaced by (auto intro!: DERIV_intros). wenzelm@33873: INCOMPATIBILITY. haftmann@33860: haftmann@33860: * Code generator attributes follow the usual underscore convention: haftmann@33860: code_unfold replaces code unfold haftmann@33860: code_post replaces code post haftmann@33860: etc. haftmann@33860: INCOMPATIBILITY. wenzelm@31900: krauss@33471: * Renamed methods: krauss@33471: sizechange -> size_change krauss@33471: induct_scheme -> induction_schema haftmann@33860: INCOMPATIBILITY. nipkow@33673: wenzelm@33843: * Discontinued abbreviation "arbitrary" of constant "undefined". wenzelm@33843: INCOMPATIBILITY, use "undefined" directly. wenzelm@33843: haftmann@33860: * Renamed theorems: haftmann@33860: Suc_eq_add_numeral_1 -> Suc_eq_plus1 haftmann@33860: Suc_eq_add_numeral_1_left -> Suc_eq_plus1_left haftmann@33860: Suc_plus1 -> Suc_eq_plus1 haftmann@33860: *anti_sym -> *antisym* haftmann@33860: vector_less_eq_def -> vector_le_def haftmann@33860: INCOMPATIBILITY. haftmann@33860: haftmann@33860: * Added theorem List.map_map as [simp]. Removed List.map_compose. haftmann@33860: INCOMPATIBILITY. haftmann@33860: haftmann@33860: * Removed predicate "M hassize n" (<--> card M = n & finite M). haftmann@33860: INCOMPATIBILITY. haftmann@33860: hoelzl@31812: huffman@33825: *** HOLCF *** huffman@33825: wenzelm@33842: * Theory Representable defines a class "rep" of domains that are wenzelm@33842: representable (via an ep-pair) in the universal domain type "udom". huffman@33825: Instances are provided for all type constructors defined in HOLCF. huffman@33825: huffman@33825: * The 'new_domain' command is a purely definitional version of the huffman@33825: domain package, for representable domains. Syntax is identical to the huffman@33825: old domain package. The 'new_domain' package also supports indirect huffman@33825: recursion using previously-defined type constructors. See wenzelm@33842: src/HOLCF/ex/New_Domain.thy for examples. wenzelm@33842: wenzelm@33842: * Method "fixrec_simp" unfolds one step of a fixrec-defined constant huffman@33825: on the left-hand side of an equation, and then performs huffman@33825: simplification. Rewriting is done using rules declared with the wenzelm@33842: "fixrec_simp" attribute. The "fixrec_simp" method is intended as a wenzelm@33842: replacement for "fixpat"; see src/HOLCF/ex/Fixrec_ex.thy for examples. huffman@33825: huffman@33825: * The pattern-match compiler in 'fixrec' can now handle constructors huffman@33825: with HOL function types. Pattern-match combinators for the Pair huffman@33825: constructor are pre-configured. huffman@33825: huffman@33825: * The 'fixrec' package now produces better fixed-point induction rules huffman@33825: for mutually-recursive definitions: Induction rules have conclusions huffman@33825: of the form "P foo bar" instead of "P ". huffman@33825: huffman@33825: * The constant "sq_le" (with infix syntax "<<" or "\") has huffman@33825: been renamed to "below". The name "below" now replaces "less" in many wenzelm@33842: theorem names. (Legacy theorem names using "less" are still supported wenzelm@33842: as well.) huffman@33825: huffman@33825: * The 'fixrec' package now supports "bottom patterns". Bottom huffman@33825: patterns can be used to generate strictness rules, or to make huffman@33825: functions more strict (much like the bang-patterns supported by the wenzelm@33873: Glasgow Haskell Compiler). See src/HOLCF/ex/Fixrec_ex.thy for wenzelm@33873: examples. huffman@33825: huffman@33825: wenzelm@31304: *** ML *** wenzelm@31304: wenzelm@33843: * Support for Poly/ML 5.3.0, with improved reporting of compiler wenzelm@33843: errors and run-time exceptions, including detailed source positions. wenzelm@33843: wenzelm@33843: * Structure Name_Space (formerly NameSpace) now manages uniquely wenzelm@33843: identified entries, with some additional information such as source wenzelm@33843: position, logical grouping etc. wenzelm@33843: wenzelm@33524: * Theory and context data is now introduced by the simplified and wenzelm@33524: modernized functors Theory_Data, Proof_Data, Generic_Data. Data needs wenzelm@33524: to be pure, but the old TheoryDataFun for mutable data (with explicit wenzelm@33524: copy operation) is still available for some time. wenzelm@33524: wenzelm@32742: * Structure Synchronized (cf. src/Pure/Concurrent/synchronized.ML) wenzelm@32742: provides a high-level programming interface to synchronized state wenzelm@32742: variables with atomic update. This works via pure function wenzelm@32742: application within a critical section -- its runtime should be as wenzelm@32742: short as possible; beware of deadlocks if critical code is nested, wenzelm@32742: either directly or indirectly via other synchronized variables! wenzelm@32742: wenzelm@32742: * Structure Unsynchronized (cf. src/Pure/ML-Systems/unsynchronized.ML) wenzelm@32742: wraps raw ML references, explicitly indicating their non-thread-safe wenzelm@32742: behaviour. The Isar toplevel keeps this structure open, to wenzelm@32742: accommodate Proof General as well as quick and dirty interactive wenzelm@32742: experiments with references. wenzelm@32742: wenzelm@32365: * PARALLEL_CHOICE and PARALLEL_GOALS provide basic support for wenzelm@32365: parallel tactical reasoning. wenzelm@32365: wenzelm@32427: * Tacticals Subgoal.FOCUS, Subgoal.FOCUS_PREMS, Subgoal.FOCUS_PARAMS wenzelm@32427: are similar to SUBPROOF, but are slightly more flexible: only the wenzelm@32427: specified parts of the subgoal are imported into the context, and the wenzelm@32427: body tactic may introduce new subgoals and schematic variables. wenzelm@32427: wenzelm@32427: * Old tactical METAHYPS, which does not observe the proof context, has wenzelm@32427: been renamed to Old_Goals.METAHYPS and awaits deletion. Use SUBPROOF wenzelm@32427: or Subgoal.FOCUS etc. wenzelm@32216: wenzelm@31971: * Renamed functor TableFun to Table, and GraphFun to Graph. (Since wenzelm@31971: functors have their own ML name space there is no point to mark them wenzelm@31971: separately.) Minor INCOMPATIBILITY. wenzelm@31971: wenzelm@31901: * Renamed NamedThmsFun to Named_Thms. INCOMPATIBILITY. wenzelm@31901: wenzelm@33842: * Renamed several structures FooBar to Foo_Bar. Occasional, wenzelm@33842: INCOMPATIBILITY. wenzelm@33842: wenzelm@33843: * Operations of structure Skip_Proof no longer require quick_and_dirty wenzelm@33843: mode, which avoids critical setmp. wenzelm@33843: wenzelm@31306: * Eliminated old Attrib.add_attributes, Method.add_methods and related wenzelm@33842: combinators for "args". INCOMPATIBILITY, need to use simplified wenzelm@31306: Attrib/Method.setup introduced in Isabelle2009. wenzelm@31304: wenzelm@32151: * Proper context for simpset_of, claset_of, clasimpset_of. May fall wenzelm@32151: back on global_simpset_of, global_claset_of, global_clasimpset_of as wenzelm@32151: last resort. INCOMPATIBILITY. wenzelm@32151: wenzelm@32092: * Display.pretty_thm now requires a proper context (cf. former wenzelm@32092: ProofContext.pretty_thm). May fall back on Display.pretty_thm_global wenzelm@32092: or even Display.pretty_thm_without_context as last resort. wenzelm@32092: INCOMPATIBILITY. wenzelm@32092: wenzelm@32433: * Discontinued Display.pretty_ctyp/cterm etc. INCOMPATIBILITY, use wenzelm@32433: Syntax.pretty_typ/term directly, preferably with proper context wenzelm@32433: instead of global theory. wenzelm@32433: wenzelm@31304: wenzelm@31308: *** System *** wenzelm@31308: wenzelm@33842: * Further fine tuning of parallel proof checking, scales up to 8 cores wenzelm@33842: (max. speedup factor 5.0). See also Goal.parallel_proofs in ML and wenzelm@33842: usedir option -q. wenzelm@33842: wenzelm@32326: * Support for additional "Isabelle components" via etc/components, see wenzelm@32326: also the system manual. wenzelm@32326: wenzelm@32326: * The isabelle makeall tool now operates on all components with wenzelm@32326: IsaMakefile, not just hardwired "logics". wenzelm@32326: wenzelm@33842: * Removed "compress" option from isabelle-process and isabelle usedir; wenzelm@33842: this is always enabled. kleing@33818: wenzelm@31308: * Discontinued support for Poly/ML 4.x versions. wenzelm@31308: wenzelm@33842: * Isabelle tool "wwwfind" provides web interface for 'find_theorems' wenzelm@33842: on a given logic image. This requires the lighttpd webserver and is wenzelm@33842: currently supported on Linux only. wenzelm@32061: wenzelm@31308: wenzelm@31304: wenzelm@30845: New in Isabelle2009 (April 2009) wenzelm@30845: -------------------------------- haftmann@27104: wenzelm@27599: *** General *** wenzelm@27599: wenzelm@28504: * Simplified main Isabelle executables, with less surprises on wenzelm@28504: case-insensitive file-systems (such as Mac OS). wenzelm@28504: wenzelm@28504: - The main Isabelle tool wrapper is now called "isabelle" instead of wenzelm@28504: "isatool." wenzelm@28504: wenzelm@28504: - The former "isabelle" alias for "isabelle-process" has been wenzelm@28504: removed (should rarely occur to regular users). wenzelm@28504: wenzelm@28915: - The former "isabelle-interface" and its alias "Isabelle" have been wenzelm@28915: removed (interfaces are now regular Isabelle tools). wenzelm@28504: wenzelm@28504: Within scripts and make files, the Isabelle environment variables wenzelm@28504: ISABELLE_TOOL and ISABELLE_PROCESS replace old ISATOOL and ISABELLE, wenzelm@28504: respectively. (The latter are still available as legacy feature.) wenzelm@28504: wenzelm@28915: The old isabelle-interface wrapper could react in confusing ways if wenzelm@28915: the interface was uninstalled or changed otherwise. Individual wenzelm@28915: interface tool configuration is now more explicit, see also the wenzelm@28915: Isabelle system manual. In particular, Proof General is now available wenzelm@28915: via "isabelle emacs". wenzelm@28504: wenzelm@28504: INCOMPATIBILITY, need to adapt derivative scripts. Users may need to wenzelm@28504: purge installed copies of Isabelle executables and re-run "isabelle wenzelm@28504: install -p ...", or use symlinks. wenzelm@28504: wenzelm@28914: * The default for ISABELLE_HOME_USER is now ~/.isabelle instead of the wenzelm@30845: old ~/isabelle, which was slightly non-standard and apt to cause wenzelm@30845: surprises on case-insensitive file-systems (such as Mac OS). wenzelm@28914: wenzelm@28914: INCOMPATIBILITY, need to move existing ~/isabelle/etc, wenzelm@28914: ~/isabelle/heaps, ~/isabelle/browser_info to the new place. Special wenzelm@28914: care is required when using older releases of Isabelle. Note that wenzelm@28914: ISABELLE_HOME_USER can be changed in Isabelle/etc/settings of any wenzelm@30845: Isabelle distribution, in order to use the new ~/.isabelle uniformly. wenzelm@28914: wenzelm@29161: * Proofs of fully specified statements are run in parallel on wenzelm@30845: multi-core systems. A speedup factor of 2.5 to 3.2 can be expected on wenzelm@30845: a regular 4-core machine, if the initial heap space is made reasonably wenzelm@30845: large (cf. Poly/ML option -H). (Requires Poly/ML 5.2.1 or later.) wenzelm@30845: wenzelm@30845: * The main reference manuals ("isar-ref", "implementation", and wenzelm@30845: "system") have been updated and extended. Formally checked references wenzelm@30845: as hyperlinks are now available uniformly. wenzelm@30845: wenzelm@30163: wenzelm@27599: *** Pure *** wenzelm@27599: wenzelm@30845: * Complete re-implementation of locales. INCOMPATIBILITY in several wenzelm@30845: respects. The most important changes are listed below. See the wenzelm@30845: Tutorial on Locales ("locales" manual) for details. ballarin@29253: ballarin@29253: - In locale expressions, instantiation replaces renaming. Parameters ballarin@29253: must be declared in a for clause. To aid compatibility with previous ballarin@29253: parameter inheritance, in locale declarations, parameters that are not ballarin@29253: 'touched' (instantiation position "_" or omitted) are implicitly added ballarin@29253: with their syntax at the beginning of the for clause. ballarin@29253: ballarin@29253: - Syntax from abbreviations and definitions in locales is available in ballarin@29253: locale expressions and context elements. The latter is particularly ballarin@29253: useful in locale declarations. ballarin@29253: ballarin@29253: - More flexible mechanisms to qualify names generated by locale ballarin@29253: expressions. Qualifiers (prefixes) may be specified in locale wenzelm@30728: expressions, and can be marked as mandatory (syntax: "name!:") or wenzelm@30728: optional (syntax "name?:"). The default depends for plain "name:" wenzelm@30728: depends on the situation where a locale expression is used: in wenzelm@30728: commands 'locale' and 'sublocale' prefixes are optional, in wenzelm@30845: 'interpretation' and 'interpret' prefixes are mandatory. The old wenzelm@30728: implicit qualifiers derived from the parameter names of a locale are wenzelm@30728: no longer generated. wenzelm@30106: wenzelm@30845: - Command "sublocale l < e" replaces "interpretation l < e". The wenzelm@30106: instantiation clause in "interpretation" and "interpret" (square wenzelm@30106: brackets) is no longer available. Use locale expressions. ballarin@29253: wenzelm@30845: - When converting proof scripts, mandatory qualifiers in wenzelm@30728: 'interpretation' and 'interpret' should be retained by default, even wenzelm@30845: if this is an INCOMPATIBILITY compared to former behavior. In the wenzelm@30845: worst case, use the "name?:" form for non-mandatory ones. Qualifiers wenzelm@30845: in locale expressions range over a single locale instance only. wenzelm@30845: wenzelm@30845: - Dropped locale element "includes". This is a major INCOMPATIBILITY. wenzelm@30845: In existing theorem specifications replace the includes element by the wenzelm@30845: respective context elements of the included locale, omitting those wenzelm@30845: that are already present in the theorem specification. Multiple wenzelm@30845: assume elements of a locale should be replaced by a single one wenzelm@30845: involving the locale predicate. In the proof body, declarations (most wenzelm@30845: notably theorems) may be regained by interpreting the respective wenzelm@30845: locales in the proof context as required (command "interpret"). wenzelm@30845: wenzelm@30845: If using "includes" in replacement of a target solely because the wenzelm@30845: parameter types in the theorem are not as general as in the target, wenzelm@30845: consider declaring a new locale with additional type constraints on wenzelm@30845: the parameters (context element "constrains"). wenzelm@30845: wenzelm@30845: - Discontinued "locale (open)". INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: - Locale interpretation commands no longer attempt to simplify goal. wenzelm@30845: INCOMPATIBILITY: in rare situations the generated goal differs. Use wenzelm@30845: methods intro_locales and unfold_locales to clarify. wenzelm@30845: wenzelm@30845: - Locale interpretation commands no longer accept interpretation wenzelm@30845: attributes. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Class declaration: so-called "base sort" must not be given in import wenzelm@30845: list any longer, but is inferred from the specification. Particularly wenzelm@30845: in HOL, write wenzelm@30845: wenzelm@30845: class foo = ... wenzelm@30845: wenzelm@30845: instead of wenzelm@30845: wenzelm@30845: class foo = type + ... wenzelm@30845: wenzelm@30845: * Class target: global versions of theorems stemming do not carry a wenzelm@30845: parameter prefix any longer. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Class 'instance' command no longer accepts attached definitions. wenzelm@30845: INCOMPATIBILITY, use proper 'instantiation' target instead. wenzelm@30845: wenzelm@30845: * Recovered hiding of consts, which was accidentally broken in wenzelm@30845: Isabelle2007. Potential INCOMPATIBILITY, ``hide const c'' really wenzelm@30845: makes c inaccessible; consider using ``hide (open) const c'' instead. wenzelm@30845: wenzelm@30845: * Slightly more coherent Pure syntax, with updated documentation in wenzelm@30845: isar-ref manual. Removed locales meta_term_syntax and wenzelm@30845: meta_conjunction_syntax: TERM and &&& (formerly &&) are now permanent, wenzelm@30845: INCOMPATIBILITY in rare situations. Note that &&& should not be used wenzelm@30845: directly in regular applications. wenzelm@30845: wenzelm@30845: * There is a new syntactic category "float_const" for signed decimal wenzelm@30845: fractions (e.g. 123.45 or -123.45). wenzelm@30845: wenzelm@30845: * Removed exotic 'token_translation' command. INCOMPATIBILITY, use ML wenzelm@30845: interface with 'setup' command instead. wenzelm@30845: wenzelm@30845: * Command 'local_setup' is similar to 'setup', but operates on a local wenzelm@30845: theory context. haftmann@27104: wenzelm@28114: * The 'axiomatization' command now only works within a global theory wenzelm@28114: context. INCOMPATIBILITY. wenzelm@28114: wenzelm@30845: * Goal-directed proof now enforces strict proof irrelevance wrt. sort wenzelm@30845: hypotheses. Sorts required in the course of reasoning need to be wenzelm@30845: covered by the constraints in the initial statement, completed by the wenzelm@30845: type instance information of the background theory. Non-trivial sort wenzelm@30845: hypotheses, which rarely occur in practice, may be specified via wenzelm@30845: vacuous propositions of the form SORT_CONSTRAINT('a::c). For example: wenzelm@30845: wenzelm@30845: lemma assumes "SORT_CONSTRAINT('a::empty)" shows False ... wenzelm@30845: wenzelm@30845: The result contains an implicit sort hypotheses as before -- wenzelm@30845: SORT_CONSTRAINT premises are eliminated as part of the canonical rule wenzelm@30845: normalization. wenzelm@30845: wenzelm@30845: * Generalized Isar history, with support for linear undo, direct state wenzelm@30845: addressing etc. wenzelm@30845: wenzelm@30845: * Changed defaults for unify configuration options: wenzelm@30845: wenzelm@30845: unify_trace_bound = 50 (formerly 25) wenzelm@30845: unify_search_bound = 60 (formerly 30) wenzelm@30845: wenzelm@30845: * Different bookkeeping for code equations (INCOMPATIBILITY): wenzelm@30845: wenzelm@30845: a) On theory merge, the last set of code equations for a particular wenzelm@30845: constant is taken (in accordance with the policy applied by other wenzelm@30845: parts of the code generator framework). wenzelm@30845: wenzelm@30845: b) Code equations stemming from explicit declarations (e.g. code wenzelm@30845: attribute) gain priority over default code equations stemming wenzelm@30845: from definition, primrec, fun etc. wenzelm@30845: wenzelm@30845: * Keyword 'code_exception' now named 'code_abort'. INCOMPATIBILITY. wenzelm@30845: haftmann@30965: * Unified theorem tables for both code generators. Thus [code wenzelm@30845: func] has disappeared and only [code] remains. INCOMPATIBILITY. wenzelm@30577: wenzelm@30577: * Command 'find_consts' searches for constants based on type and name wenzelm@30577: patterns, e.g. kleing@29883: kleing@29883: find_consts "_ => bool" kleing@29883: wenzelm@30106: By default, matching is against subtypes, but it may be restricted to wenzelm@30728: the whole type. Searching by name is possible. Multiple queries are wenzelm@30106: conjunctive and queries may be negated by prefixing them with a wenzelm@30106: hyphen: kleing@29883: kleing@29883: find_consts strict: "_ => bool" name: "Int" -"int => int" kleing@29861: wenzelm@30845: * New 'find_theorems' criterion "solves" matches theorems that wenzelm@30845: directly solve the current goal (modulo higher-order unification). wenzelm@30845: wenzelm@30845: * Auto solve feature for main theorem statements: whenever a new goal wenzelm@30845: is stated, "find_theorems solves" is called; any theorems that could wenzelm@30845: solve the lemma directly are listed as part of the goal state. wenzelm@30845: Cf. associated options in Proof General Isabelle settings menu, wenzelm@30845: enabled by default, with reasonable timeout for pathological cases of wenzelm@30845: higher-order unification. webertj@30415: haftmann@27104: wenzelm@27381: *** Document preparation *** wenzelm@27381: wenzelm@27381: * Antiquotation @{lemma} now imitates a regular terminal proof, wenzelm@27392: demanding keyword 'by' and supporting the full method expression wenzelm@27519: syntax just like the Isar command 'by'. wenzelm@27381: wenzelm@27381: haftmann@27104: *** HOL *** haftmann@27104: wenzelm@30845: * Integrated main parts of former image HOL-Complex with HOL. Entry wenzelm@30845: points Main and Complex_Main remain as before. wenzelm@30845: wenzelm@30845: * Logic image HOL-Plain provides a minimal HOL with the most important wenzelm@30845: tools available (inductive, datatype, primrec, ...). This facilitates wenzelm@30845: experimentation and tool development. Note that user applications wenzelm@30845: (and library theories) should never refer to anything below theory wenzelm@30845: Main, as before. wenzelm@30845: wenzelm@30845: * Logic image HOL-Main stops at theory Main, and thus facilitates wenzelm@30845: experimentation due to shorter build times. wenzelm@30845: wenzelm@30845: * Logic image HOL-NSA contains theories of nonstandard analysis which wenzelm@30845: were previously part of former HOL-Complex. Entry point Hyperreal wenzelm@30845: remains valid, but theories formerly using Complex_Main should now use wenzelm@30845: new entry point Hypercomplex. wenzelm@30845: wenzelm@30845: * Generic ATP manager for Sledgehammer, based on ML threads instead of wenzelm@30845: Posix processes. Avoids potentially expensive forking of the ML wenzelm@30845: process. New thread-based implementation also works on non-Unix wenzelm@30845: platforms (Cygwin). Provers are no longer hardwired, but defined wenzelm@30845: within the theory via plain ML wrapper functions. Basic Sledgehammer wenzelm@30845: commands are covered in the isar-ref manual. wenzelm@30845: wenzelm@30845: * Wrapper scripts for remote SystemOnTPTP service allows to use wenzelm@30845: sledgehammer without local ATP installation (Vampire etc.). Other wenzelm@30845: provers may be included via suitable ML wrappers, see also wenzelm@30845: src/HOL/ATP_Linkup.thy. wenzelm@30845: wenzelm@30845: * ATP selection (E/Vampire/Spass) is now via Proof General's settings wenzelm@30845: menu. wenzelm@30845: wenzelm@30845: * The metis method no longer fails because the theorem is too trivial wenzelm@30845: (contains the empty clause). wenzelm@30845: wenzelm@30845: * The metis method now fails in the usual manner, rather than raising wenzelm@30845: an exception, if it determines that it cannot prove the theorem. wenzelm@30845: wenzelm@30845: * Method "coherent" implements a prover for coherent logic (see also wenzelm@30845: src/Tools/coherent.ML). wenzelm@30845: wenzelm@30845: * Constants "undefined" and "default" replace "arbitrary". Usually wenzelm@30845: "undefined" is the right choice to replace "arbitrary", though wenzelm@30845: logically there is no difference. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Command "value" now integrates different evaluation mechanisms. The wenzelm@30845: result of the first successful evaluation mechanism is printed. In wenzelm@30845: square brackets a particular named evaluation mechanisms may be wenzelm@30845: specified (currently, [SML], [code] or [nbe]). See further wenzelm@30845: src/HOL/ex/Eval_Examples.thy. wenzelm@30845: wenzelm@30845: * Normalization by evaluation now allows non-leftlinear equations. wenzelm@30845: Declare with attribute [code nbe]. wenzelm@30845: wenzelm@30845: * Methods "case_tac" and "induct_tac" now refer to the very same rules wenzelm@30845: as the structured Isar versions "cases" and "induct", cf. the wenzelm@30845: corresponding "cases" and "induct" attributes. Mutual induction rules wenzelm@30845: are now presented as a list of individual projections wenzelm@30845: (e.g. foo_bar.inducts for types foo and bar); the old format with wenzelm@30845: explicit HOL conjunction is no longer supported. INCOMPATIBILITY, in wenzelm@30845: rare situations a different rule is selected --- notably nested tuple wenzelm@30845: elimination instead of former prod.exhaust: use explicit (case_tac t wenzelm@30845: rule: prod.exhaust) here. wenzelm@30845: wenzelm@30845: * Attributes "cases", "induct", "coinduct" support "del" option. wenzelm@30845: wenzelm@30845: * Removed fact "case_split_thm", which duplicates "case_split". wenzelm@30845: wenzelm@30845: * The option datatype has been moved to a new theory Option. Renamed wenzelm@30845: option_map to Option.map, and o2s to Option.set, INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * New predicate "strict_mono" classifies strict functions on partial wenzelm@30845: orders. With strict functions on linear orders, reasoning about wenzelm@30845: (in)equalities is facilitated by theorems "strict_mono_eq", wenzelm@30845: "strict_mono_less_eq" and "strict_mono_less". wenzelm@30845: wenzelm@30845: * Some set operations are now proper qualified constants with wenzelm@30845: authentic syntax. INCOMPATIBILITY: haftmann@30304: haftmann@30304: op Int ~> Set.Int haftmann@30304: op Un ~> Set.Un haftmann@30304: INTER ~> Set.INTER haftmann@30304: UNION ~> Set.UNION haftmann@30304: Inter ~> Set.Inter haftmann@30304: Union ~> Set.Union haftmann@30304: {} ~> Set.empty haftmann@30304: UNIV ~> Set.UNIV haftmann@30304: wenzelm@30845: * Class complete_lattice with operations Inf, Sup, INFI, SUPR now in wenzelm@30845: theory Set. wenzelm@30845: wenzelm@30845: * Auxiliary class "itself" has disappeared -- classes without any wenzelm@30845: parameter are treated as expected by the 'class' command. haftmann@29797: haftmann@29823: * Leibnitz's Series for Pi and the arcus tangens and logarithm series. haftmann@29823: wenzelm@30845: * Common decision procedures (Cooper, MIR, Ferrack, Approximation, wenzelm@30845: Dense_Linear_Order) are now in directory HOL/Decision_Procs. wenzelm@30845: wenzelm@30845: * Theory src/HOL/Decision_Procs/Approximation provides the new proof wenzelm@30845: method "approximation". It proves formulas on real values by using wenzelm@30845: interval arithmetic. In the formulas are also the transcendental wenzelm@30845: functions sin, cos, tan, atan, ln, exp and the constant pi are wenzelm@30845: allowed. For examples see wenzelm@30845: src/HOL/Descision_Procs/ex/Approximation_Ex.thy. haftmann@29823: haftmann@29823: * Theory "Reflection" now resides in HOL/Library. haftmann@29650: wenzelm@30845: * Entry point to Word library now simply named "Word". wenzelm@30845: INCOMPATIBILITY. haftmann@29628: haftmann@29197: * Made source layout more coherent with logical distribution haftmann@29197: structure: haftmann@28952: haftmann@28952: src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy haftmann@28952: src/HOL/Library/Code_Message.thy ~> src/HOL/ haftmann@28952: src/HOL/Library/GCD.thy ~> src/HOL/ haftmann@28952: src/HOL/Library/Order_Relation.thy ~> src/HOL/ haftmann@28952: src/HOL/Library/Parity.thy ~> src/HOL/ haftmann@28952: src/HOL/Library/Univ_Poly.thy ~> src/HOL/ huffman@30176: src/HOL/Real/ContNotDenum.thy ~> src/HOL/Library/ haftmann@28952: src/HOL/Real/Lubs.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/PReal.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/Rational.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/RComplete.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/RealDef.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/RealPow.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/Real.thy ~> src/HOL/ haftmann@28952: src/HOL/Complex/Complex_Main.thy ~> src/HOL/ haftmann@28952: src/HOL/Complex/Complex.thy ~> src/HOL/ huffman@30176: src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/Library/ huffman@30176: src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/Library/ haftmann@28952: src/HOL/Hyperreal/Deriv.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Fact.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Integration.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Lim.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Ln.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Log.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Series.thy ~> src/HOL/ haftmann@29197: src/HOL/Hyperreal/SEQ.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Taylor.thy ~> src/HOL/ haftmann@28952: src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/ haftmann@28952: src/HOL/Real/Float ~> src/HOL/Library/ haftmann@29197: src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach haftmann@29197: src/HOL/Real/RealVector.thy ~> src/HOL/ haftmann@28952: haftmann@28952: src/HOL/arith_data.ML ~> src/HOL/Tools haftmann@28952: src/HOL/hologic.ML ~> src/HOL/Tools haftmann@28952: src/HOL/simpdata.ML ~> src/HOL/Tools haftmann@28952: src/HOL/int_arith1.ML ~> src/HOL/Tools/int_arith.ML haftmann@28952: src/HOL/int_factor_simprocs.ML ~> src/HOL/Tools haftmann@28952: src/HOL/nat_simprocs.ML ~> src/HOL/Tools haftmann@28952: src/HOL/Real/float_arith.ML ~> src/HOL/Tools haftmann@28952: src/HOL/Real/float_syntax.ML ~> src/HOL/Tools haftmann@28952: src/HOL/Real/rat_arith.ML ~> src/HOL/Tools haftmann@28952: src/HOL/Real/real_arith.ML ~> src/HOL/Tools haftmann@28952: haftmann@29398: src/HOL/Library/Array.thy ~> src/HOL/Imperative_HOL haftmann@29398: src/HOL/Library/Heap_Monad.thy ~> src/HOL/Imperative_HOL haftmann@29398: src/HOL/Library/Heap.thy ~> src/HOL/Imperative_HOL haftmann@29398: src/HOL/Library/Imperative_HOL.thy ~> src/HOL/Imperative_HOL haftmann@29398: src/HOL/Library/Ref.thy ~> src/HOL/Imperative_HOL haftmann@29398: src/HOL/Library/Relational.thy ~> src/HOL/Imperative_HOL haftmann@29398: wenzelm@30845: * If methods "eval" and "evaluation" encounter a structured proof wenzelm@30845: state with !!/==>, only the conclusion is evaluated to True (if wenzelm@30845: possible), avoiding strange error messages. wenzelm@30845: wenzelm@30845: * Method "sizechange" automates termination proofs using (a wenzelm@30845: modification of) the size-change principle. Requires SAT solver. See wenzelm@30845: src/HOL/ex/Termination.thy for examples. wenzelm@30845: wenzelm@30845: * Simplifier: simproc for let expressions now unfolds if bound wenzelm@30845: variable occurs at most once in let expression body. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Method "arith": Linear arithmetic now ignores all inequalities when wenzelm@30845: fast_arith_neq_limit is exceeded, instead of giving up entirely. wenzelm@30845: wenzelm@30845: * New attribute "arith" for facts that should always be used wenzelm@30845: automatically by arithmetic. It is intended to be used locally in wenzelm@30845: proofs, e.g. wenzelm@30845: wenzelm@30845: assumes [arith]: "x > 0" wenzelm@30845: nipkow@30706: Global usage is discouraged because of possible performance impact. nipkow@30706: wenzelm@30845: * New classes "top" and "bot" with corresponding operations "top" and wenzelm@30845: "bot" in theory Orderings; instantiation of class "complete_lattice" wenzelm@30845: requires instantiation of classes "top" and "bot". INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Changed definition lemma "less_fun_def" in order to provide an wenzelm@30845: instance for preorders on functions; use lemma "less_le" instead. wenzelm@30845: INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Theory Orderings: class "wellorder" moved here, with explicit wenzelm@30845: induction rule "less_induct" as assumption. For instantiation of wenzelm@30845: "wellorder" by means of predicate "wf", use rule wf_wellorderI. wenzelm@30845: INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Theory Orderings: added class "preorder" as superclass of "order". wenzelm@27793: INCOMPATIBILITY: Instantiation proofs for order, linorder wenzelm@27793: etc. slightly changed. Some theorems named order_class.* now named wenzelm@27793: preorder_class.*. wenzelm@27793: wenzelm@30845: * Theory Relation: renamed "refl" to "refl_on", "reflexive" to "refl, wenzelm@30845: "diag" to "Id_on". wenzelm@30845: wenzelm@30845: * Theory Finite_Set: added a new fold combinator of type wenzelm@30845: nipkow@28855: ('a => 'b => 'b) => 'b => 'a set => 'b wenzelm@30845: wenzelm@30845: Occasionally this is more convenient than the old fold combinator wenzelm@30845: which is now defined in terms of the new one and renamed to wenzelm@30845: fold_image. wenzelm@30845: wenzelm@30845: * Theories Ring_and_Field and OrderedGroup: The lemmas "group_simps" wenzelm@30845: and "ring_simps" have been replaced by "algebra_simps" (which can be wenzelm@30845: extended with further lemmas!). At the moment both still exist but wenzelm@30845: the former will disappear at some point. wenzelm@30845: wenzelm@30845: * Theory Power: Lemma power_Suc is now declared as a simp rule in wenzelm@30845: class recpower. Type-specific simp rules for various recpower types wenzelm@30845: have been removed. INCOMPATIBILITY, rename old lemmas as follows: huffman@30273: huffman@30273: rat_power_0 -> power_0 huffman@30273: rat_power_Suc -> power_Suc huffman@30273: realpow_0 -> power_0 huffman@30273: realpow_Suc -> power_Suc huffman@30273: complexpow_0 -> power_0 huffman@30273: complexpow_Suc -> power_Suc huffman@30273: power_poly_0 -> power_0 huffman@30273: power_poly_Suc -> power_Suc huffman@30273: wenzelm@30845: * Theories Ring_and_Field and Divides: Definition of "op dvd" has been wenzelm@27793: moved to separate class dvd in Ring_and_Field; a couple of lemmas on wenzelm@27793: dvd has been generalized to class comm_semiring_1. Likewise a bunch wenzelm@27793: of lemmas from Divides has been generalized from nat to class wenzelm@27793: semiring_div. INCOMPATIBILITY. This involves the following theorem wenzelm@27793: renames resulting from duplicate elimination: haftmann@27651: haftmann@27651: dvd_def_mod ~> dvd_eq_mod_eq_0 haftmann@27651: zero_dvd_iff ~> dvd_0_left_iff haftmann@28559: dvd_0 ~> dvd_0_right haftmann@27651: DIVISION_BY_ZERO_DIV ~> div_by_0 haftmann@27651: DIVISION_BY_ZERO_MOD ~> mod_by_0 haftmann@27651: mult_div ~> div_mult_self2_is_id haftmann@27651: mult_mod ~> mod_mult_self2_is_0 haftmann@27651: wenzelm@30845: * Theory IntDiv: removed many lemmas that are instances of class-based wenzelm@30845: generalizations (from Divides and Ring_and_Field). INCOMPATIBILITY, wenzelm@30845: rename old lemmas as follows: nipkow@30044: nipkow@30044: dvd_diff -> nat_dvd_diff nipkow@30044: dvd_zminus_iff -> dvd_minus_iff nipkow@30224: mod_add1_eq -> mod_add_eq nipkow@30224: mod_mult1_eq -> mod_mult_right_eq nipkow@30224: mod_mult1_eq' -> mod_mult_left_eq nipkow@30224: mod_mult_distrib_mod -> mod_mult_eq nipkow@30044: nat_mod_add_left_eq -> mod_add_left_eq nipkow@30044: nat_mod_add_right_eq -> mod_add_right_eq nipkow@30044: nat_mod_div_trivial -> mod_div_trivial nipkow@30044: nat_mod_mod_trivial -> mod_mod_trivial nipkow@30044: zdiv_zadd_self1 -> div_add_self1 nipkow@30044: zdiv_zadd_self2 -> div_add_self2 nipkow@30181: zdiv_zmult_self1 -> div_mult_self2_is_id nipkow@30044: zdiv_zmult_self2 -> div_mult_self1_is_id nipkow@30044: zdvd_triv_left -> dvd_triv_left nipkow@30044: zdvd_triv_right -> dvd_triv_right nipkow@30044: zdvd_zmult_cancel_disj -> dvd_mult_cancel_left nipkow@30085: zmod_eq0_zdvd_iff -> dvd_eq_mod_eq_0[symmetric] nipkow@30044: zmod_zadd_left_eq -> mod_add_left_eq nipkow@30044: zmod_zadd_right_eq -> mod_add_right_eq nipkow@30044: zmod_zadd_self1 -> mod_add_self1 nipkow@30044: zmod_zadd_self2 -> mod_add_self2 nipkow@30224: zmod_zadd1_eq -> mod_add_eq nipkow@30044: zmod_zdiff1_eq -> mod_diff_eq nipkow@30044: zmod_zdvd_zmod -> mod_mod_cancel nipkow@30044: zmod_zmod_cancel -> mod_mod_cancel nipkow@30044: zmod_zmult_self1 -> mod_mult_self2_is_0 nipkow@30044: zmod_zmult_self2 -> mod_mult_self1_is_0 nipkow@30044: zmod_1 -> mod_by_1 nipkow@30044: zdiv_1 -> div_by_1 nipkow@30044: zdvd_abs1 -> abs_dvd_iff nipkow@30044: zdvd_abs2 -> dvd_abs_iff nipkow@30044: zdvd_refl -> dvd_refl nipkow@30044: zdvd_trans -> dvd_trans nipkow@30044: zdvd_zadd -> dvd_add nipkow@30044: zdvd_zdiff -> dvd_diff nipkow@30044: zdvd_zminus_iff -> dvd_minus_iff nipkow@30044: zdvd_zminus2_iff -> minus_dvd_iff nipkow@30044: zdvd_zmultD -> dvd_mult_right nipkow@30044: zdvd_zmultD2 -> dvd_mult_left nipkow@30044: zdvd_zmult_mono -> mult_dvd_mono nipkow@30044: zdvd_0_right -> dvd_0_right nipkow@30044: zdvd_0_left -> dvd_0_left_iff nipkow@30044: zdvd_1_left -> one_dvd nipkow@30044: zminus_dvd_iff -> minus_dvd_iff nipkow@30044: wenzelm@30845: * Theory Rational: 'Fract k 0' now equals '0'. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * The real numbers offer decimal input syntax: 12.34 is translated wenzelm@30845: into 1234/10^2. This translation is not reversed upon output. wenzelm@30845: wenzelm@30845: * Theory Library/Polynomial defines an abstract type 'a poly of wenzelm@30845: univariate polynomials with coefficients of type 'a. In addition to wenzelm@30845: the standard ring operations, it also supports div and mod. Code wenzelm@30845: generation is also supported, using list-style constructors. wenzelm@30845: wenzelm@30845: * Theory Library/Inner_Product defines a class of real_inner for real wenzelm@30845: inner product spaces, with an overloaded operation inner :: 'a => 'a wenzelm@30845: => real. Class real_inner is a subclass of real_normed_vector from wenzelm@30845: theory RealVector. wenzelm@30845: wenzelm@30845: * Theory Library/Product_Vector provides instances for the product wenzelm@30845: type 'a * 'b of several classes from RealVector and Inner_Product. wenzelm@30845: Definitions of addition, subtraction, scalar multiplication, norms, wenzelm@30845: and inner products are included. wenzelm@30845: wenzelm@30845: * Theory Library/Bit defines the field "bit" of integers modulo 2. In wenzelm@30845: addition to the field operations, numerals and case syntax are also wenzelm@30845: supported. wenzelm@30845: wenzelm@30845: * Theory Library/Diagonalize provides constructive version of Cantor's wenzelm@30845: first diagonalization argument. wenzelm@30845: wenzelm@30845: * Theory Library/GCD: Curried operations gcd, lcm (for nat) and zgcd, wenzelm@27599: zlcm (for int); carried together from various gcd/lcm developements in wenzelm@30845: the HOL Distribution. Constants zgcd and zlcm replace former igcd and wenzelm@30845: ilcm; corresponding theorems renamed accordingly. INCOMPATIBILITY, wenzelm@30845: may recover tupled syntax as follows: haftmann@27556: haftmann@27556: hide (open) const gcd haftmann@27556: abbreviation gcd where haftmann@27556: "gcd == (%(a, b). GCD.gcd a b)" haftmann@27556: notation (output) haftmann@27556: GCD.gcd ("gcd '(_, _')") haftmann@27556: wenzelm@30845: The same works for lcm, zgcd, zlcm. wenzelm@30845: wenzelm@30845: * Theory Library/Nat_Infinity: added addition, numeral syntax and more wenzelm@30845: instantiations for algebraic structures. Removed some duplicate wenzelm@30845: theorems. Changes in simp rules. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * ML antiquotation @{code} takes a constant as argument and generates haftmann@27651: corresponding code in background and inserts name of the corresponding haftmann@27651: resulting ML value/function/datatype constructor binding in place. haftmann@27651: All occurrences of @{code} with a single ML block are generated haftmann@27651: simultaneously. Provides a generic and safe interface for wenzelm@30845: instrumentalizing code generation. See wenzelm@30845: src/HOL/Decision_Procs/Ferrack.thy for a more ambitious application. wenzelm@30845: In future you ought to refrain from ad-hoc compiling generated SML wenzelm@30845: code on the ML toplevel. Note that (for technical reasons) @{code} wenzelm@30845: cannot refer to constants for which user-defined serializations are wenzelm@30845: set. Refer to the corresponding ML counterpart directly in that wenzelm@30845: cases. wenzelm@27122: wenzelm@27122: * Command 'rep_datatype': instead of theorem names the command now wenzelm@27122: takes a list of terms denoting the constructors of the type to be wenzelm@27122: represented as datatype. The characteristic theorems have to be wenzelm@27122: proven. INCOMPATIBILITY. Also observe that the following theorems wenzelm@27122: have disappeared in favour of existing ones: wenzelm@27122: haftmann@27104: unit_induct ~> unit.induct haftmann@27104: prod_induct ~> prod.induct haftmann@27104: sum_induct ~> sum.induct haftmann@27104: Suc_Suc_eq ~> nat.inject haftmann@27104: Suc_not_Zero Zero_not_Suc ~> nat.distinct haftmann@27104: haftmann@27104: ballarin@27696: *** HOL-Algebra *** ballarin@27696: ballarin@27713: * New locales for orders and lattices where the equivalence relation wenzelm@30106: is not restricted to equality. INCOMPATIBILITY: all order and lattice wenzelm@30106: locales use a record structure with field eq for the equivalence. ballarin@27713: ballarin@27713: * New theory of factorial domains. ballarin@27713: wenzelm@30845: * Units_l_inv and Units_r_inv are now simp rules by default. ballarin@27696: INCOMPATIBILITY. Simplifier proof that require deletion of l_inv ballarin@27696: and/or r_inv will now also require deletion of these lemmas. ballarin@27696: wenzelm@30845: * Renamed the following theorems, INCOMPATIBILITY: wenzelm@30845: ballarin@27696: UpperD ~> Upper_memD ballarin@27696: LowerD ~> Lower_memD ballarin@27696: least_carrier ~> least_closed ballarin@27696: greatest_carrier ~> greatest_closed ballarin@27696: greatest_Lower_above ~> greatest_Lower_below ballarin@27717: one_zero ~> carrier_one_zero ballarin@27717: one_not_zero ~> carrier_one_not_zero (collision with assumption) ballarin@27696: wenzelm@27793: wenzelm@30849: *** HOL-Nominal *** wenzelm@30849: wenzelm@30855: * Nominal datatypes can now contain type-variables. wenzelm@30855: wenzelm@30855: * Commands 'nominal_inductive' and 'equivariance' work with local wenzelm@30855: theory targets. wenzelm@30855: wenzelm@30855: * Nominal primrec can now works with local theory targets and its wenzelm@30855: specification syntax now conforms to the general format as seen in wenzelm@30855: 'inductive' etc. wenzelm@30855: wenzelm@30855: * Method "perm_simp" honours the standard simplifier attributes wenzelm@30855: (no_asm), (no_asm_use) etc. wenzelm@30855: wenzelm@30855: * The new predicate #* is defined like freshness, except that on the wenzelm@30855: left hand side can be a set or list of atoms. wenzelm@30855: wenzelm@30855: * Experimental command 'nominal_inductive2' derives strong induction wenzelm@30855: principles for inductive definitions. In contrast to wenzelm@30855: 'nominal_inductive', which can only deal with a fixed number of wenzelm@30855: binders, it can deal with arbitrary expressions standing for sets of wenzelm@30855: atoms to be avoided. The only inductive definition we have at the wenzelm@30855: moment that needs this generalisation is the typing rule for Lets in wenzelm@30855: the algorithm W: wenzelm@30855: wenzelm@30855: Gamma |- t1 : T1 (x,close Gamma T1)::Gamma |- t2 : T2 x#Gamma wenzelm@30855: ----------------------------------------------------------------- wenzelm@30855: Gamma |- Let x be t1 in t2 : T2 wenzelm@30855: wenzelm@30855: In this rule one wants to avoid all the binders that are introduced by wenzelm@30855: "close Gamma T1". We are looking for other examples where this wenzelm@30855: feature might be useful. Please let us know. wenzelm@30849: wenzelm@30849: huffman@30176: *** HOLCF *** huffman@30176: huffman@30176: * Reimplemented the simplification procedure for proving continuity huffman@30176: subgoals. The new simproc is extensible; users can declare additional huffman@30176: continuity introduction rules with the attribute [cont2cont]. huffman@30176: huffman@30176: * The continuity simproc now uses a different introduction rule for huffman@30176: solving continuity subgoals on terms with lambda abstractions. In huffman@30176: some rare cases the new simproc may fail to solve subgoals that the huffman@30176: old one could solve, and "simp add: cont2cont_LAM" may be necessary. huffman@30176: Potential INCOMPATIBILITY. huffman@30176: wenzelm@30847: * Command 'fixrec': specification syntax now conforms to the general wenzelm@30855: format as seen in 'inductive' etc. See src/HOLCF/ex/Fixrec_ex.thy for wenzelm@30855: examples. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: wenzelm@30845: *** ZF *** wenzelm@30845: wenzelm@30845: * Proof of Zorn's Lemma for partial orders. huffman@30176: huffman@30176: wenzelm@27246: *** ML *** wenzelm@28088: wenzelm@30845: * Multithreading for Poly/ML 5.1/5.2 is no longer supported, only for wenzelm@30845: Poly/ML 5.2.1 or later. Important note: the TimeLimit facility wenzelm@30845: depends on multithreading, so timouts will not work before Poly/ML wenzelm@30845: 5.2.1! wenzelm@30845: wenzelm@29161: * High-level support for concurrent ML programming, see wenzelm@29161: src/Pure/Cuncurrent. The data-oriented model of "future values" is wenzelm@29161: particularly convenient to organize independent functional wenzelm@29161: computations. The concept of "synchronized variables" provides a wenzelm@29161: higher-order interface for components with shared state, avoiding the wenzelm@30845: delicate details of mutexes and condition variables. (Requires wenzelm@30845: Poly/ML 5.2.1 or later.) wenzelm@30845: wenzelm@30845: * ML bindings produced via Isar commands are stored within the Isar wenzelm@30845: context (theory or proof). Consequently, commands like 'use' and 'ML' wenzelm@30845: become thread-safe and work with undo as expected (concerning wenzelm@30845: top-level bindings, not side-effects on global references). wenzelm@30845: INCOMPATIBILITY, need to provide proper Isar context when invoking the wenzelm@30845: compiler at runtime; really global bindings need to be given outside a wenzelm@30845: theory. (Requires Poly/ML 5.2 or later.) wenzelm@30845: wenzelm@30845: * Command 'ML_prf' is analogous to 'ML' but works within a proof wenzelm@30845: context. Top-level ML bindings are stored within the proof context in wenzelm@30845: a purely sequential fashion, disregarding the nested proof structure. wenzelm@30845: ML bindings introduced by 'ML_prf' are discarded at the end of the wenzelm@30845: proof. (Requires Poly/ML 5.2 or later.) wenzelm@29161: wenzelm@30530: * Simplified ML attribute and method setup, cf. functions Attrib.setup wenzelm@30845: and Method.setup, as well as Isar commands 'attribute_setup' and wenzelm@30547: 'method_setup'. INCOMPATIBILITY for 'method_setup', need to simplify wenzelm@30547: existing code accordingly, or use plain 'setup' together with old wenzelm@30547: Method.add_method. wenzelm@30530: wenzelm@28294: * Simplified ML oracle interface Thm.add_oracle promotes 'a -> cterm wenzelm@28294: to 'a -> thm, while results are always tagged with an authentic oracle wenzelm@28294: name. The Isar command 'oracle' is now polymorphic, no argument type wenzelm@28294: is specified. INCOMPATIBILITY, need to simplify existing oracle code wenzelm@28294: accordingly. Note that extra performance may be gained by producing wenzelm@28294: the cterm carefully, avoiding slow Thm.cterm_of. wenzelm@28294: wenzelm@30845: * Simplified interface for defining document antiquotations via wenzelm@30845: ThyOutput.antiquotation, ThyOutput.output, and optionally wenzelm@30845: ThyOutput.maybe_pretty_source. INCOMPATIBILITY, need to simplify user wenzelm@30845: antiquotations accordingly, see src/Pure/Thy/thy_output.ML for common wenzelm@30845: examples. wenzelm@28099: wenzelm@30395: * More systematic treatment of long names, abstract name bindings, and wenzelm@30395: name space operations. Basic operations on qualified names have been wenzelm@30399: move from structure NameSpace to Long_Name, e.g. Long_Name.base_name, wenzelm@30395: Long_Name.append. Old type bstring has been mostly replaced by wenzelm@30395: abstract type binding (see structure Binding), which supports precise wenzelm@30845: qualification by packages and local theory targets, as well as proper wenzelm@30845: tracking of source positions. INCOMPATIBILITY, need to wrap old wenzelm@30845: bstring values into Binding.name, or better pass through abstract wenzelm@30399: bindings everywhere. See further src/Pure/General/long_name.ML, wenzelm@30395: src/Pure/General/binding.ML and src/Pure/General/name_space.ML wenzelm@30395: wenzelm@28089: * Result facts (from PureThy.note_thms, ProofContext.note_thms, wenzelm@28089: LocalTheory.note etc.) now refer to the *full* internal name, not the wenzelm@28089: bstring as before. INCOMPATIBILITY, not detected by ML type-checking! wenzelm@28089: wenzelm@27287: * Disposed old type and term read functions (Sign.read_def_typ, wenzelm@27287: Sign.read_typ, Sign.read_def_terms, Sign.read_term, wenzelm@27287: Thm.read_def_cterms, Thm.read_cterm etc.). INCOMPATIBILITY, should wenzelm@27287: use regular Syntax.read_typ, Syntax.read_term, Syntax.read_typ_global, wenzelm@27269: Syntax.read_term_global etc.; see also OldGoals.read_term as last wenzelm@27269: resort for legacy applications. wenzelm@27269: wenzelm@30609: * Disposed old declarations, tactics, tactic combinators that refer to wenzelm@30609: the simpset or claset of an implicit theory (such as Addsimps, wenzelm@30609: Simp_tac, SIMPSET). INCOMPATIBILITY, should use @{simpset} etc. in wenzelm@30609: embedded ML text, or local_simpset_of with a proper context passed as wenzelm@30609: explicit runtime argument. wenzelm@30609: wenzelm@30845: * Rules and tactics that read instantiations (read_instantiate, wenzelm@30845: res_inst_tac, thin_tac, subgoal_tac etc.) now demand a proper proof wenzelm@30845: context, which is required for parsing and type-checking. Moreover, wenzelm@30845: the variables are specified as plain indexnames, not string encodings wenzelm@30845: thereof. INCOMPATIBILITY. wenzelm@30845: wenzelm@30845: * Generic Toplevel.add_hook interface allows to analyze the result of wenzelm@30845: transactions. E.g. see src/Pure/ProofGeneral/proof_general_pgip.ML wenzelm@30845: for theorem dependency output of transactions resulting in a new wenzelm@30845: theory state. wenzelm@30845: wenzelm@30845: * ML antiquotations: block-structured compilation context indicated by wenzelm@27391: \ ... \; additional antiquotation forms: wenzelm@27391: wenzelm@30845: @{binding name} - basic name binding wenzelm@27519: @{let ?pat = term} - term abbreviation (HO matching) wenzelm@27519: @{note name = fact} - fact abbreviation wenzelm@27519: @{thm fact} - singleton fact (with attributes) wenzelm@27519: @{thms fact} - general fact (with attributes) wenzelm@27519: @{lemma prop by method} - singleton goal wenzelm@27519: @{lemma prop by meth1 meth2} - singleton goal wenzelm@27519: @{lemma prop1 ... propN by method} - general goal wenzelm@27519: @{lemma prop1 ... propN by meth1 meth2} - general goal wenzelm@27519: @{lemma (open) ...} - open derivation wenzelm@27380: wenzelm@27246: wenzelm@27979: *** System *** wenzelm@27979: wenzelm@28248: * The Isabelle "emacs" tool provides a specific interface to invoke wenzelm@28248: Proof General / Emacs, with more explicit failure if that is not wenzelm@28248: installed (the old isabelle-interface script silently falls back on wenzelm@28248: isabelle-process). The PROOFGENERAL_HOME setting determines the wenzelm@28248: installation location of the Proof General distribution. wenzelm@28248: wenzelm@27979: * Isabelle/lib/classes/Pure.jar provides basic support to integrate wenzelm@27979: the Isabelle process into a JVM/Scala application. See wenzelm@27979: Isabelle/lib/jedit/plugin for a minimal example. (The obsolete Java wenzelm@27979: process wrapper has been discontinued.) wenzelm@27979: wenzelm@30845: * Added homegrown Isabelle font with unicode layout, see lib/fonts. wenzelm@30845: wenzelm@30845: * Various status messages (with exact source position information) are wenzelm@27979: emitted, if proper markup print mode is enabled. This allows wenzelm@27979: user-interface components to provide detailed feedback on internal wenzelm@27979: prover operations. wenzelm@27979: wenzelm@27979: wenzelm@27143: wenzelm@27008: New in Isabelle2008 (June 2008) wenzelm@27008: ------------------------------- wenzelm@25464: wenzelm@25522: *** General *** wenzelm@25522: wenzelm@27061: * The Isabelle/Isar Reference Manual (isar-ref) has been reorganized wenzelm@27061: and updated, with formally checked references as hyperlinks. wenzelm@27061: wenzelm@25994: * Theory loader: use_thy (and similar operations) no longer set the wenzelm@25994: implicit ML context, which was occasionally hard to predict and in wenzelm@25994: conflict with concurrency. INCOMPATIBILITY, use ML within Isar which wenzelm@25994: provides a proper context already. wenzelm@25994: wenzelm@26323: * Theory loader: old-style ML proof scripts being *attached* to a thy wenzelm@26323: file are no longer supported. INCOMPATIBILITY, regular 'uses' and wenzelm@26323: 'use' within a theory file will do the job. wenzelm@26323: wenzelm@26650: * Name space merge now observes canonical order, i.e. the second space wenzelm@26650: is inserted into the first one, while existing entries in the first wenzelm@26659: space take precedence. INCOMPATIBILITY in rare situations, may try to wenzelm@26650: swap theory imports. wenzelm@26650: wenzelm@27067: * Syntax: symbol \ is now considered a letter. Potential wenzelm@27067: INCOMPATIBILITY in identifier syntax etc. wenzelm@27067: wenzelm@27067: * Outer syntax: string tokens no longer admit escaped white space, wenzelm@27067: which was an accidental (undocumented) feature. INCOMPATIBILITY, use wenzelm@27067: white space without escapes. wenzelm@27067: wenzelm@27067: * Outer syntax: string tokens may contain arbitrary character codes wenzelm@27067: specified via 3 decimal digits (as in SML). E.g. "foo\095bar" for wenzelm@27067: "foo_bar". wenzelm@27067: wenzelm@25522: haftmann@25502: *** Pure *** haftmann@25502: wenzelm@26718: * Context-dependent token translations. Default setup reverts locally wenzelm@26718: fixed variables, and adds hilite markup for undeclared frees. wenzelm@26718: berghofe@26681: * Unused theorems can be found using the new command 'unused_thms'. berghofe@26681: There are three ways of invoking it: berghofe@26681: berghofe@26681: (1) unused_thms berghofe@26681: Only finds unused theorems in the current theory. berghofe@26681: berghofe@26681: (2) unused_thms thy_1 ... thy_n - berghofe@26681: Finds unused theorems in the current theory and all of its ancestors, berghofe@26681: excluding the theories thy_1 ... thy_n and all of their ancestors. berghofe@26681: berghofe@26681: (3) unused_thms thy_1 ... thy_n - thy'_1 ... thy'_m berghofe@26681: Finds unused theorems in the theories thy'_1 ... thy'_m and all of berghofe@26681: their ancestors, excluding the theories thy_1 ... thy_n and all of berghofe@26681: their ancestors. berghofe@26681: wenzelm@26718: In order to increase the readability of the list produced by wenzelm@26718: unused_thms, theorems that have been created by a particular instance wenzelm@26874: of a theory command such as 'inductive' or 'function' are considered wenzelm@26874: to belong to the same "group", meaning that if at least one theorem in wenzelm@26718: this group is used, the other theorems in the same group are no longer wenzelm@26718: reported as unused. Moreover, if all theorems in the group are wenzelm@26718: unused, only one theorem in the group is displayed. wenzelm@26718: wenzelm@26718: Note that proof objects have to be switched on in order for wenzelm@26718: unused_thms to work properly (i.e. !proofs must be >= 1, which is wenzelm@26874: usually the case when using Proof General with the default settings). berghofe@26681: wenzelm@26650: * Authentic naming of facts disallows ad-hoc overwriting of previous wenzelm@26650: theorems within the same name space. INCOMPATIBILITY, need to remove wenzelm@26650: duplicate fact bindings, or even accidental fact duplications. Note wenzelm@26650: that tools may maintain dynamically scoped facts systematically, using wenzelm@26650: PureThy.add_thms_dynamic. wenzelm@26650: wenzelm@26660: * Command 'hide' now allows to hide from "fact" name space as well. wenzelm@26660: wenzelm@26496: * Eliminated destructive theorem database, simpset, claset, and wenzelm@26496: clasimpset. Potential INCOMPATIBILITY, really need to observe linear wenzelm@26496: update of theories within ML code. wenzelm@26479: wenzelm@26955: * Eliminated theory ProtoPure and CPure, leaving just one Pure theory. wenzelm@26955: INCOMPATIBILITY, object-logics depending on former Pure require wenzelm@26955: additional setup PureThy.old_appl_syntax_setup; object-logics wenzelm@26955: depending on former CPure need to refer to Pure. wenzelm@26650: wenzelm@26495: * Commands 'use' and 'ML' are now purely functional, operating on wenzelm@26479: theory/local_theory. Removed former 'ML_setup' (on theory), use 'ML' wenzelm@26479: instead. Added 'ML_val' as mere diagnostic replacement for 'ML'. wenzelm@26479: INCOMPATIBILITY. wenzelm@26479: wenzelm@26874: * Command 'setup': discontinued implicit version with ML reference. wenzelm@26434: wenzelm@25970: * Instantiation target allows for simultaneous specification of class wenzelm@25970: instance operations together with an instantiation proof. wenzelm@25970: Type-checking phase allows to refer to class operations uniformly. wenzelm@27067: See src/HOL/Complex/Complex.thy for an Isar example and wenzelm@27067: src/HOL/Library/Eval.thy for an ML example. haftmann@25502: wenzelm@26201: * Indexing of literal facts: be more serious about including only wenzelm@26201: facts from the visible specification/proof context, but not the wenzelm@26201: background context (locale etc.). Affects `prop` notation and method wenzelm@26201: "fact". INCOMPATIBILITY: need to name facts explicitly in rare wenzelm@26201: situations. wenzelm@26201: wenzelm@26925: * Method "cases", "induct", "coinduct": removed obsolete/undocumented wenzelm@26925: "(open)" option, which used to expose internal bound variables to the wenzelm@26925: proof text. wenzelm@26925: wenzelm@26925: * Isar statements: removed obsolete case "rule_context". wenzelm@26925: INCOMPATIBILITY, better use explicit fixes/assumes. wenzelm@26925: wenzelm@26874: * Locale proofs: default proof step now includes 'unfold_locales'; wenzelm@26874: hence 'proof' without argument may be used to unfold locale wenzelm@26874: predicates. ballarin@26765: ballarin@26765: haftmann@26762: *** Document preparation *** haftmann@26762: wenzelm@26914: * Simplified pdfsetup.sty: color/hyperref is used unconditionally for wenzelm@26914: both pdf and dvi (hyperlinks usually work in xdvi as well); removed wenzelm@26914: obsolete thumbpdf setup (contemporary PDF viewers do this on the wenzelm@26914: spot); renamed link color from "darkblue" to "linkcolor" (default wenzelm@26920: value unchanged, can be redefined via \definecolor); no longer sets wenzelm@26920: "a4paper" option (unnecessary or even intrusive). wenzelm@26914: wenzelm@27008: * Antiquotation @{lemma A method} proves proposition A by the given wenzelm@27008: method (either a method name or a method name plus (optional) method wenzelm@27008: arguments in parentheses) and prints A just like @{prop A}. haftmann@26762: haftmann@26762: wenzelm@25464: *** HOL *** wenzelm@25464: wenzelm@27067: * New primrec package. Specification syntax conforms in style to wenzelm@27067: definition/function/.... No separate induction rule is provided. The wenzelm@27067: "primrec" command distinguishes old-style and new-style specifications wenzelm@27067: by syntax. The former primrec package is now named OldPrimrecPackage. wenzelm@27067: When adjusting theories, beware: constants stemming from new-style wenzelm@27067: primrec specifications have authentic syntax. wenzelm@27067: wenzelm@27067: * Metis prover is now an order of magnitude faster, and also works wenzelm@27067: with multithreading. wenzelm@27067: wenzelm@27067: * Metis: the maximum number of clauses that can be produced from a wenzelm@27067: theorem is now given by the attribute max_clauses. Theorems that wenzelm@27067: exceed this number are ignored, with a warning printed. wenzelm@27067: wenzelm@27067: * Sledgehammer no longer produces structured proofs by default. To wenzelm@27067: enable, declare [[sledgehammer_full = true]]. Attributes wenzelm@27067: reconstruction_modulus, reconstruction_sorts renamed wenzelm@27067: sledgehammer_modulus, sledgehammer_sorts. INCOMPATIBILITY. wenzelm@27067: haftmann@27104: * Method "induct_scheme" derives user-specified induction rules wenzelm@27067: from well-founded induction and completeness of patterns. This factors wenzelm@27067: out some operations that are done internally by the function package wenzelm@27067: and makes them available separately. See wenzelm@27067: src/HOL/ex/Induction_Scheme.thy for examples. wenzelm@27067: wenzelm@27067: * More flexible generation of measure functions for termination wenzelm@27067: proofs: Measure functions can be declared by proving a rule of the wenzelm@27067: form "is_measure f" and giving it the [measure_function] attribute. wenzelm@27067: The "is_measure" predicate is logically meaningless (always true), and wenzelm@27067: just guides the heuristic. To find suitable measure functions, the wenzelm@27067: termination prover sets up the goal "is_measure ?f" of the appropriate blanchet@56846: type and generates all solutions by Prolog-style backward proof using wenzelm@27067: the declared rules. wenzelm@27067: lp15@57253: This setup also deals with rules like wenzelm@27067: wenzelm@27067: "is_measure f ==> is_measure (list_size f)" wenzelm@27067: wenzelm@27067: which accommodates nested datatypes that recurse through lists. wenzelm@27067: Similar rules are predeclared for products and option types. wenzelm@27067: berghofe@26964: * Turned the type of sets "'a set" into an abbreviation for "'a => bool" berghofe@26964: berghofe@26964: INCOMPATIBILITIES: berghofe@26964: wenzelm@27008: - Definitions of overloaded constants on sets have to be replaced by wenzelm@27008: definitions on => and bool. berghofe@26964: berghofe@26964: - Some definitions of overloaded operators on sets can now be proved wenzelm@27008: using the definitions of the operators on => and bool. Therefore, wenzelm@27008: the following theorems have been renamed: berghofe@26964: berghofe@26964: subset_def -> subset_eq berghofe@26964: psubset_def -> psubset_eq berghofe@26964: set_diff_def -> set_diff_eq berghofe@26964: Compl_def -> Compl_eq berghofe@26964: Sup_set_def -> Sup_set_eq berghofe@26964: Inf_set_def -> Inf_set_eq berghofe@26964: sup_set_def -> sup_set_eq berghofe@26964: inf_set_def -> inf_set_eq berghofe@26964: berghofe@26964: - Due to the incompleteness of the HO unification algorithm, some berghofe@26964: rules such as subst may require manual instantiation, if some of berghofe@26964: the unknowns in the rule is a set. berghofe@26964: berghofe@26964: - Higher order unification and forward proofs: berghofe@26964: The proof pattern berghofe@26964: berghofe@26964: have "P (S::'a set)" <...> berghofe@26964: then have "EX S. P S" .. berghofe@26964: wenzelm@27008: no longer works (due to the incompleteness of the HO unification wenzelm@27008: algorithm) and must be replaced by the pattern berghofe@26964: berghofe@26964: have "EX S. P S" berghofe@26964: proof berghofe@26964: show "P S" <...> berghofe@26964: qed berghofe@26964: berghofe@26964: - Calculational reasoning with subst (or similar rules): berghofe@26964: The proof pattern berghofe@26964: berghofe@26964: have "P (S::'a set)" <...> berghofe@26964: also have "S = T" <...> berghofe@26964: finally have "P T" . berghofe@26964: wenzelm@27008: no longer works (for similar reasons as the previous example) and wenzelm@27008: must be replaced by something like berghofe@26964: berghofe@26964: have "P (S::'a set)" <...> berghofe@26964: moreover have "S = T" <...> berghofe@26964: ultimately have "P T" by simp berghofe@26964: berghofe@26964: - Tactics or packages written in ML code: berghofe@26964: Code performing pattern matching on types via berghofe@26964: berghofe@26964: Type ("set", [T]) => ... berghofe@26964: wenzelm@27008: must be rewritten. Moreover, functions like strip_type or wenzelm@27008: binder_types no longer return the right value when applied to a wenzelm@27008: type of the form berghofe@26964: berghofe@26964: T1 => ... => Tn => U => bool berghofe@26964: berghofe@26964: rather than berghofe@26964: berghofe@26964: T1 => ... => Tn => U set berghofe@26964: wenzelm@26874: * Merged theories Wellfounded_Recursion, Accessible_Part and wenzelm@27067: Wellfounded_Relations to theory Wellfounded. krauss@26748: haftmann@26513: * Explicit class "eq" for executable equality. INCOMPATIBILITY. haftmann@26513: wenzelm@26874: * Class finite no longer treats UNIV as class parameter. Use class wenzelm@26874: enum from theory Library/Enum instead to achieve a similar effect. haftmann@26445: INCOMPATIBILITY. haftmann@26445: wenzelm@26874: * Theory List: rule list_induct2 now has explicitly named cases "Nil" wenzelm@26874: and "Cons". INCOMPATIBILITY. wenzelm@26874: wenzelm@26422: * HOL (and FOL): renamed variables in rules imp_elim and swap. wenzelm@26422: Potential INCOMPATIBILITY. wenzelm@26422: wenzelm@26874: * Theory Product_Type: duplicated lemmas split_Pair_apply and wenzelm@26874: injective_fst_snd removed, use split_eta and prod_eqI instead. wenzelm@26874: Renamed upd_fst to apfst and upd_snd to apsnd. INCOMPATIBILITY. haftmann@26355: wenzelm@26335: * Theory Nat: removed redundant lemmas that merely duplicate lemmas of wenzelm@26335: the same name in theory Orderings: wenzelm@26335: wenzelm@26335: less_trans wenzelm@26335: less_linear wenzelm@26335: le_imp_less_or_eq wenzelm@26335: le_less_trans wenzelm@26335: less_le_trans wenzelm@26335: less_not_sym wenzelm@26335: less_asym wenzelm@26335: wenzelm@26335: Renamed less_imp_le to less_imp_le_nat, and less_irrefl to wenzelm@26335: less_irrefl_nat. Potential INCOMPATIBILITY due to more general types wenzelm@26335: and different variable names. wenzelm@26315: haftmann@26231: * Library/Option_ord.thy: Canonical order on option type. haftmann@26231: wenzelm@27008: * Library/RBT.thy: Red-black trees, an efficient implementation of wenzelm@27008: finite maps. krauss@26197: haftmann@26231: * Library/Countable.thy: Type class for countable types. haftmann@26231: wenzelm@26180: * Theory Int: The representation of numerals has changed. The infix wenzelm@26180: operator BIT and the bit datatype with constructors B0 and B1 have wenzelm@26180: disappeared. INCOMPATIBILITY, use "Int.Bit0 x" and "Int.Bit1 y" in wenzelm@26180: place of "x BIT bit.B0" and "y BIT bit.B1", respectively. Theorems wenzelm@26180: involving BIT, B0, or B1 have been renamed with "Bit0" or "Bit1" wenzelm@26180: accordingly. wenzelm@26180: wenzelm@26180: * Theory Nat: definition of <= and < on natural numbers no longer wenzelm@26180: depend on well-founded relations. INCOMPATIBILITY. Definitions wenzelm@26180: le_def and less_def have disappeared. Consider lemmas not_less wenzelm@26180: [symmetric, where ?'a = nat] and less_eq [symmetric] instead. wenzelm@26180: wenzelm@26180: * Theory Finite_Set: locales ACf, ACe, ACIf, ACIfSL and ACIfSLlin wenzelm@26180: (whose purpose mainly is for various fold_set functionals) have been wenzelm@26874: abandoned in favor of the existing algebraic classes wenzelm@26180: ab_semigroup_mult, comm_monoid_mult, ab_semigroup_idem_mult, wenzelm@26180: lower_semilattice (resp. upper_semilattice) and linorder. haftmann@26139: INCOMPATIBILITY. haftmann@26041: wenzelm@26180: * Theory Transitive_Closure: induct and cases rules now declare proper wenzelm@26180: case_names ("base" and "step"). INCOMPATIBILITY. wenzelm@26180: wenzelm@26180: * Theorem Inductive.lfp_ordinal_induct generalized to complete wenzelm@26180: lattices. The form set-specific version is available as wenzelm@26180: Inductive.lfp_ordinal_induct_set. haftmann@26013: wenzelm@26874: * Renamed theorems "power.simps" to "power_int.simps". wenzelm@27067: INCOMPATIBILITY. haftmann@25961: wenzelm@26180: * Class semiring_div provides basic abstract properties of semirings haftmann@25942: with division and modulo operations. Subsumes former class dvd_mod. haftmann@25942: wenzelm@26180: * Merged theories IntDef, Numeral and IntArith into unified theory wenzelm@26180: Int. INCOMPATIBILITY. wenzelm@26180: wenzelm@26180: * Theory Library/Code_Index: type "index" now represents natural wenzelm@26180: numbers rather than integers. INCOMPATIBILITY. wenzelm@26180: wenzelm@26180: * New class "uminus" with operation "uminus" (split of from class wenzelm@26180: "minus" which now only has operation "minus", binary). haftmann@25919: INCOMPATIBILITY. haftmann@25919: wenzelm@25522: * Constants "card", "internal_split", "option_map" now with authentic haftmann@25919: syntax. INCOMPATIBILITY. wenzelm@25522: wenzelm@25522: * Definitions subset_def, psubset_def, set_diff_def, Compl_def, wenzelm@25522: le_bool_def, less_bool_def, le_fun_def, less_fun_def, inf_bool_def, wenzelm@25522: sup_bool_def, Inf_bool_def, Sup_bool_def, inf_fun_def, sup_fun_def, wenzelm@25522: Inf_fun_def, Sup_fun_def, inf_set_def, sup_set_def, Inf_set_def, wenzelm@25522: Sup_set_def, le_def, less_def, option_map_def now with object haftmann@25919: equality. INCOMPATIBILITY. wenzelm@25464: schirmer@25705: * Records. Removed K_record, and replaced it by pure lambda term wenzelm@25726: %x. c. The simplifier setup is now more robust against eta expansion. schirmer@25705: INCOMPATIBILITY: in cases explicitly referring to K_record. wenzelm@25464: wenzelm@27067: * Library/Multiset: {#a, b, c#} abbreviates {#a#} + {#b#} + {#c#}. wenzelm@27067: wenzelm@27067: * Library/ListVector: new theory of arithmetic vector operations. wenzelm@27067: wenzelm@27067: * Library/Order_Relation: new theory of various orderings as sets of wenzelm@27067: pairs. Defines preorders, partial orders, linear orders and wenzelm@27067: well-orders on sets and on types. krauss@26877: wenzelm@25726: krauss@26197: *** ZF *** krauss@26197: wenzelm@26874: * Renamed some theories to allow to loading both ZF and HOL in the wenzelm@26874: same session: wenzelm@26874: wenzelm@26874: Datatype -> Datatype_ZF wenzelm@26874: Inductive -> Inductive_ZF wenzelm@26874: Int -> Int_ZF wenzelm@26874: IntDiv -> IntDiv_ZF wenzelm@26874: Nat -> Nat_ZF wenzelm@26874: List -> List_ZF wenzelm@26874: Main -> Main_ZF wenzelm@26874: wenzelm@26874: INCOMPATIBILITY: ZF theories that import individual theories below wenzelm@26874: Main might need to be adapted. Regular theory Main is still wenzelm@26874: available, as trivial extension of Main_ZF. krauss@26197: krauss@26197: wenzelm@25737: *** ML *** wenzelm@25737: wenzelm@27067: * ML within Isar: antiquotation @{const name} or @{const wenzelm@27067: name(typargs)} produces statically-checked Const term. wenzelm@27067: wenzelm@26401: * Functor NamedThmsFun: data is available to the user as dynamic fact wenzelm@26724: (of the same name). Removed obsolete print command. wenzelm@26401: wenzelm@27067: * Removed obsolete "use_legacy_bindings" function. wenzelm@26188: wenzelm@25737: * The ``print mode'' is now a thread-local value derived from a global wenzelm@25737: template (the former print_mode reference), thus access becomes wenzelm@25737: non-critical. The global print_mode reference is for session wenzelm@25737: management only; user-code should use print_mode_value, wenzelm@25737: print_mode_active, PrintMode.setmp etc. INCOMPATIBILITY. wenzelm@25737: wenzelm@26874: * Functions system/system_out provide a robust way to invoke external wenzelm@29161: shell commands, with propagation of interrupts (requires Poly/ML wenzelm@29161: 5.2.1). Do not use OS.Process.system etc. from the basis library! wenzelm@26222: wenzelm@25737: wenzelm@25626: *** System *** wenzelm@25626: wenzelm@25971: * Default settings: PROOFGENERAL_OPTIONS no longer impose xemacs --- wenzelm@25971: in accordance with Proof General 3.7, which prefers GNU emacs. wenzelm@25970: wenzelm@25626: * isatool tty runs Isabelle process with plain tty interaction; wenzelm@25626: optional line editor may be specified via ISABELLE_LINE_EDITOR wenzelm@25626: setting, the default settings attempt to locate "ledit" and "rlwrap". wenzelm@25626: wenzelm@25651: * isatool browser now works with Cygwin as well, using general wenzelm@25651: "javapath" function defined in Isabelle process environment. wenzelm@25651: wenzelm@27067: * YXML notation provides a simple and efficient alternative to wenzelm@27067: standard XML transfer syntax. See src/Pure/General/yxml.ML and wenzelm@27067: isatool yxml as described in the Isabelle system manual. wenzelm@25651: wenzelm@25652: * JVM class isabelle.IsabelleProcess (located in Isabelle/lib/classes) wenzelm@25651: provides general wrapper for managing an Isabelle process in a robust wenzelm@25651: fashion, with ``cooked'' output from stdin/stderr. wenzelm@25651: wenzelm@25855: * Rudimentary Isabelle plugin for jEdit (see Isabelle/lib/jedit), wenzelm@25855: based on Isabelle/JVM process wrapper (see Isabelle/lib/classes). wenzelm@25855: wenzelm@27067: * Removed obsolete THIS_IS_ISABELLE_BUILD feature. NB: the documented wenzelm@27067: way of changing the user's settings is via wenzelm@27067: ISABELLE_HOME_USER/etc/settings, which is a fully featured bash wenzelm@27067: script. wenzelm@27067: wenzelm@27067: * Multithreading.max_threads := 0 refers to the number of actual CPU wenzelm@27067: cores of the underlying machine, which is a good starting point for wenzelm@27067: optimal performance tuning. The corresponding usedir option -M allows wenzelm@27067: "max" as an alias for "0". WARNING: does not work on certain versions wenzelm@27067: of Mac OS (with Poly/ML 5.1). wenzelm@27067: wenzelm@27067: * isabelle-process: non-ML sessions are run with "nice", to reduce the wenzelm@27067: adverse effect of Isabelle flooding interactive front-ends (notably wenzelm@27067: ProofGeneral / XEmacs). wenzelm@27067: wenzelm@25626: wenzelm@25464: wenzelm@25429: New in Isabelle2007 (November 2007) wenzelm@25429: ----------------------------------- wenzelm@17754: wenzelm@17754: *** General *** wenzelm@17754: wenzelm@22826: * More uniform information about legacy features, notably a wenzelm@22826: warning/error of "Legacy feature: ...", depending on the state of the wenzelm@23367: tolerate_legacy_features flag (default true). FUTURE INCOMPATIBILITY: wenzelm@23367: legacy features will disappear eventually. wenzelm@22826: wenzelm@17918: * Theory syntax: the header format ``theory A = B + C:'' has been wenzelm@17918: discontinued in favour of ``theory A imports B C begin''. Use isatool wenzelm@17918: fixheaders to convert existing theory files. INCOMPATIBILITY. wenzelm@17918: wenzelm@17918: * Theory syntax: the old non-Isar theory file format has been wenzelm@17918: discontinued altogether. Note that ML proof scripts may still be used wenzelm@17918: with Isar theories; migration is usually quite simple with the ML wenzelm@17918: function use_legacy_bindings. INCOMPATIBILITY. wenzelm@17918: wenzelm@22871: * Theory syntax: some popular names (e.g. 'class', 'declaration', wenzelm@22871: 'fun', 'help', 'if') are now keywords. INCOMPATIBILITY, use double wenzelm@22871: quotes. wenzelm@19814: wenzelm@23888: * Theory loader: be more serious about observing the static theory wenzelm@23888: header specifications (including optional directories), but not the wenzelm@24172: accidental file locations of previously successful loads. The strict wenzelm@24172: update policy of former update_thy is now already performed by wenzelm@24172: use_thy, so the former has been removed; use_thys updates several wenzelm@24172: theories simultaneously, just as 'imports' within a theory header wenzelm@24172: specification, but without merging the results. Potential wenzelm@24172: INCOMPATIBILITY: may need to refine theory headers and commands wenzelm@24172: ROOT.ML which depend on load order. wenzelm@23888: wenzelm@23888: * Theory loader: optional support for content-based file wenzelm@23888: identification, instead of the traditional scheme of full physical wenzelm@23889: path plus date stamp; configured by the ISABELLE_FILE_IDENT setting wenzelm@23888: (cf. the system manual). The new scheme allows to work with wenzelm@23888: non-finished theories in persistent session images, such that source wenzelm@23888: files may be moved later on without requiring reloads. wenzelm@23888: wenzelm@24187: * Theory loader: old-style ML proof scripts being *attached* to a thy wenzelm@24187: file (with the same base name as the theory) are considered a legacy wenzelm@24800: feature, which will disappear eventually. Even now, the theory loader wenzelm@24800: no longer maintains dependencies on such files. wenzelm@24800: wenzelm@24800: * Syntax: the scope for resolving ambiguities via type-inference is wenzelm@24800: now limited to individual terms, instead of whole simultaneous wenzelm@24234: specifications as before. This greatly reduces the complexity of the wenzelm@24234: syntax module and improves flexibility by separating parsing and wenzelm@24234: type-checking. INCOMPATIBILITY: additional type-constraints (explicit wenzelm@24234: 'fixes' etc.) are required in rare situations. wenzelm@24234: wenzelm@25034: * Syntax: constants introduced by new-style packages ('definition', wenzelm@25034: 'abbreviation' etc.) are passed through the syntax module in wenzelm@25034: ``authentic mode''. This means that associated mixfix annotations wenzelm@25034: really stick to such constants, independently of potential name space wenzelm@25034: ambiguities introduced later on. INCOMPATIBILITY: constants in parse wenzelm@25034: trees are represented slightly differently, may need to adapt syntax wenzelm@25034: translations accordingly. Use CONST marker in 'translations' and wenzelm@25034: @{const_syntax} antiquotation in 'parse_translation' etc. wenzelm@25034: wenzelm@17981: * Legacy goal package: reduced interface to the bare minimum required wenzelm@17981: to keep existing proof scripts running. Most other user-level wenzelm@17981: functions are now part of the OldGoals structure, which is *not* open wenzelm@17981: by default (consider isatool expandshort before open OldGoals). wenzelm@17981: Removed top_sg, prin, printyp, pprint_term/typ altogether, because wenzelm@17981: these tend to cause confusion about the actual goal (!) context being wenzelm@17981: used here, which is not necessarily the same as the_context(). wenzelm@17918: wenzelm@23379: * Command 'find_theorems': supports "*" wild-card in "name:" wenzelm@23379: criterion; "with_dups" option. Certain ProofGeneral versions might wenzelm@23379: support a specific search form (see ProofGeneral/CHANGES). webertj@22965: wenzelm@20370: * The ``prems limit'' option (cf. ProofContext.prems_limit) is now -1 wenzelm@20370: by default, which means that "prems" (and also "fixed variables") are wenzelm@20370: suppressed from proof state output. Note that the ProofGeneral wenzelm@20370: settings mechanism allows to change and save options persistently, but wenzelm@20370: older versions of Isabelle will fail to start up if a negative prems wenzelm@20370: limit is imposed. wenzelm@20370: wenzelm@21308: * Local theory targets may be specified by non-nested blocks of wenzelm@21308: ``context/locale/class ... begin'' followed by ``end''. The body may wenzelm@21308: contain definitions, theorems etc., including any derived mechanism wenzelm@21308: that has been implemented on top of these primitives. This concept wenzelm@21308: generalizes the existing ``theorem (in ...)'' towards more versatility wenzelm@21308: and scalability. wenzelm@21308: wenzelm@21960: * Proof General interface: proper undo of final 'end' command; wenzelm@21960: discontinued Isabelle/classic mode (ML proof scripts). wenzelm@21960: wenzelm@17754: wenzelm@17865: *** Document preparation *** wenzelm@17865: wenzelm@21717: * Added antiquotation @{theory name} which prints the given name, wenzelm@21717: after checking that it refers to a valid ancestor theory in the wenzelm@21717: current context. haftmann@21339: wenzelm@17869: * Added antiquotations @{ML_type text} and @{ML_struct text} which wenzelm@17869: check the given source text as ML type/structure, printing verbatim. wenzelm@17865: wenzelm@21717: * Added antiquotation @{abbrev "c args"} which prints the abbreviation wenzelm@21717: "c args == rhs" given in the current context. (Any number of wenzelm@21735: arguments may be given on the LHS.) wenzelm@21717: wenzelm@21717: wenzelm@17779: *** Pure *** wenzelm@17779: wenzelm@24800: * The 'class' package offers a combination of axclass and locale to wenzelm@25129: achieve Haskell-like type classes in Isabelle. Definitions and wenzelm@25129: theorems within a class context produce both relative results (with wenzelm@25129: implicit parameters according to the locale context), and polymorphic wenzelm@25129: constants with qualified polymorphism (according to the class wenzelm@25129: context). Within the body context of a 'class' target, a separate wenzelm@25129: syntax layer ("user space type system") takes care of converting wenzelm@25129: between global polymorphic consts and internal locale representation. wenzelm@25177: See src/HOL/ex/Classpackage.thy for examples (as well as main HOL). haftmann@25184: "isatool doc classes" provides a tutorial. wenzelm@20807: haftmann@25199: * Generic code generator framework allows to generate executable wenzelm@24800: code for ML and Haskell (including Isabelle classes). A short usage wenzelm@24800: sketch: haftmann@20188: haftmann@20188: internal compilation: haftmann@25199: export_code in SML haftmann@20453: writing SML code to a file: haftmann@25199: export_code in SML haftmann@22735: writing OCaml code to a file: haftmann@25199: export_code in OCaml haftmann@20188: writing Haskell code to a bunch of files: haftmann@25199: export_code in Haskell haftmann@25199: haftmann@25199: evaluating closed propositions to True/False using code generation: haftmann@25184: method ``eval'' haftmann@25184: haftmann@25184: Reasonable default setup of framework in HOL. haftmann@20453: haftmann@20453: Theorem attributs for selecting and transforming function equations theorems: haftmann@20453: haftmann@22845: [code fun]: select a theorem as function equation for a specific constant haftmann@22845: [code fun del]: deselect a theorem as function equation for a specific constant haftmann@22845: [code inline]: select an equation theorem for unfolding (inlining) in place haftmann@22845: [code inline del]: deselect an equation theorem for unfolding (inlining) in place haftmann@20453: haftmann@22735: User-defined serializations (target in {SML, OCaml, Haskell}): haftmann@20453: haftmann@20453: code_const haftmann@20453: {(target) }+ haftmann@20453: haftmann@20453: code_type haftmann@20453: {(target) }+ haftmann@20453: haftmann@20453: code_instance haftmann@20453: {(target)}+ haftmann@20453: where instance ::= :: haftmann@20453: haftmann@20453: code_class haftmann@20453: {(target) }+ haftmann@20453: where class target syntax ::= {where { == }+}? haftmann@20453: haftmann@25199: code_instance and code_class only are effective to target Haskell. haftmann@22735: wenzelm@25177: For example usage see src/HOL/ex/Codegenerator.thy and wenzelm@25177: src/HOL/ex/Codegenerator_Pretty.thy. A separate tutorial on code wenzelm@24800: generation from Isabelle/HOL theories is available via "isatool doc wenzelm@24800: codegen". haftmann@20188: wenzelm@25129: * Code generator: consts in 'consts_code' Isar commands are now wenzelm@25129: referred to by usual term syntax (including optional type wenzelm@25129: annotations). wenzelm@25129: wenzelm@19254: * Command 'no_translations' removes translation rules from theory wenzelm@19254: syntax. wenzelm@19254: wenzelm@19625: * Overloaded definitions are now actually checked for acyclic wenzelm@19714: dependencies. The overloading scheme is slightly more general than wenzelm@19714: that of Haskell98, although Isabelle does not demand an exact wenzelm@19714: correspondence to type class and instance declarations. wenzelm@19714: INCOMPATIBILITY, use ``defs (unchecked overloaded)'' to admit more wenzelm@19714: exotic versions of overloading -- at the discretion of the user! wenzelm@19711: wenzelm@19711: Polymorphic constants are represented via type arguments, i.e. the wenzelm@19711: instantiation that matches an instance against the most general wenzelm@19711: declaration given in the signature. For example, with the declaration wenzelm@19711: c :: 'a => 'a => 'a, an instance c :: nat => nat => nat is represented wenzelm@19711: as c(nat). Overloading is essentially simultaneous structural wenzelm@19711: recursion over such type arguments. Incomplete specification patterns wenzelm@19714: impose global constraints on all occurrences, e.g. c('a * 'a) on the wenzelm@19715: LHS means that more general c('a * 'b) will be disallowed on any RHS. wenzelm@19714: Command 'print_theory' outputs the normalized system of recursive wenzelm@19714: equations, see section "definitions". wenzelm@19625: wenzelm@24086: * Configuration options are maintained within the theory or proof wenzelm@24086: context (with name and type bool/int/string), providing a very simple wenzelm@24086: interface to a poor-man's version of general context data. Tools may wenzelm@24110: declare options in ML (e.g. using Attrib.config_int) and then refer to wenzelm@24110: these values using Config.get etc. Users may change options via an wenzelm@24110: associated attribute of the same name. This form of context wenzelm@24110: declaration works particularly well with commands 'declare' or wenzelm@24110: 'using', for example ``declare [[foo = 42]]''. Thus it has become wenzelm@24110: very easy to avoid global references, which would not observe Isar wenzelm@24110: toplevel undo/redo and fail to work with multithreading. wenzelm@24086: wenzelm@24172: Various global ML references of Pure and HOL have been turned into wenzelm@24172: configuration options: wenzelm@24172: wenzelm@24172: Unify.search_bound unify_search_bound wenzelm@24172: Unify.trace_bound unify_trace_bound wenzelm@24172: Unify.trace_simp unify_trace_simp wenzelm@24172: Unify.trace_types unify_trace_types wenzelm@24172: Simplifier.simp_depth_limit simp_depth_limit wenzelm@24172: Blast.depth_limit blast_depth_limit wenzelm@24172: DatatypeProp.dtK datatype_distinctness_limit wenzelm@24172: fast_arith_neq_limit fast_arith_neq_limit wenzelm@24172: fast_arith_split_limit fast_arith_split_limit wenzelm@24172: wenzelm@24086: * Named collections of theorems may be easily installed as context wenzelm@24800: data using the functor NamedThmsFun (see also wenzelm@24086: src/Pure/Tools/named_thms.ML). The user may add or delete facts via wenzelm@24110: attributes; there is also a toplevel print command. This facility is wenzelm@24110: just a common case of general context data, which is the preferred way wenzelm@24110: for anything more complex than just a list of facts in canonical wenzelm@24110: order. wenzelm@24086: wenzelm@24032: * Isar: command 'declaration' augments a local theory by generic wenzelm@24032: declaration functions written in ML. This enables arbitrary content wenzelm@24032: being added to the context, depending on a morphism that tells the wenzelm@24032: difference of the original declaration context wrt. the application wenzelm@24032: context encountered later on. wenzelm@24032: wenzelm@24032: * Isar: proper interfaces for simplification procedures. Command wenzelm@24032: 'simproc_setup' declares named simprocs (with match patterns, and body wenzelm@24032: text in ML). Attribute "simproc" adds/deletes simprocs in the current wenzelm@24032: context. ML antiquotation @{simproc name} retrieves named simprocs. wenzelm@24032: wenzelm@24032: * Isar: an extra pair of brackets around attribute declarations wenzelm@24032: abbreviates a theorem reference involving an internal dummy fact, wenzelm@24032: which will be ignored later --- only the effect of the attribute on wenzelm@24032: the background context will persist. This form of in-place wenzelm@24032: declarations is particularly useful with commands like 'declare' and wenzelm@24032: 'using', for example ``have A using [[simproc a]] by simp''. wenzelm@24032: wenzelm@23369: * Isar: method "assumption" (and implicit closing of subproofs) now wenzelm@23369: takes simple non-atomic goal assumptions into account: after applying wenzelm@23369: an assumption as a rule the resulting subgoals are solved by atomic wenzelm@23369: assumption steps. This is particularly useful to finish 'obtain' wenzelm@23369: goals, such as "!!x. (!!x. P x ==> thesis) ==> P x ==> thesis", wenzelm@23369: without referring to the original premise "!!x. P x ==> thesis" in the wenzelm@23369: Isar proof context. POTENTIAL INCOMPATIBILITY: method "assumption" is wenzelm@23369: more permissive. wenzelm@23369: wenzelm@23369: * Isar: implicit use of prems from the Isar proof context is wenzelm@23369: considered a legacy feature. Common applications like ``have A .'' wenzelm@23369: may be replaced by ``have A by fact'' or ``note `A`''. In general, wenzelm@23369: referencing facts explicitly here improves readability and wenzelm@23369: maintainability of proof texts. wenzelm@23369: wenzelm@17865: * Isar: improper proof element 'guess' is like 'obtain', but derives wenzelm@17865: the obtained context from the course of reasoning! For example: wenzelm@17865: wenzelm@17865: assume "EX x y. A x & B y" -- "any previous fact" wenzelm@17865: then guess x and y by clarify wenzelm@17865: wenzelm@17865: This technique is potentially adventurous, depending on the facts and wenzelm@17865: proof tools being involved here. wenzelm@17865: wenzelm@18020: * Isar: known facts from the proof context may be specified as literal wenzelm@18020: propositions, using ASCII back-quote syntax. This works wherever wenzelm@18020: named facts used to be allowed so far, in proof commands, proof wenzelm@18020: methods, attributes etc. Literal facts are retrieved from the context wenzelm@18020: according to unification of type and term parameters. For example, wenzelm@18020: provided that "A" and "A ==> B" and "!!x. P x ==> Q x" are known wenzelm@18020: theorems in the current context, then these are valid literal facts: wenzelm@18020: `A` and `A ==> B` and `!!x. P x ==> Q x" as well as `P a ==> Q a` etc. wenzelm@18020: wenzelm@18020: There is also a proof method "fact" which does the same composition wenzelm@18044: for explicit goal states, e.g. the following proof texts coincide with wenzelm@18044: certain special cases of literal facts: wenzelm@18020: wenzelm@18020: have "A" by fact == note `A` wenzelm@18020: have "A ==> B" by fact == note `A ==> B` wenzelm@18020: have "!!x. P x ==> Q x" by fact == note `!!x. P x ==> Q x` wenzelm@18020: have "P a ==> Q a" by fact == note `P a ==> Q a` wenzelm@18020: wenzelm@20118: * Isar: ":" (colon) is no longer a symbolic identifier character in wenzelm@20118: outer syntax. Thus symbolic identifiers may be used without wenzelm@20118: additional white space in declarations like this: ``assume *: A''. wenzelm@20118: wenzelm@20013: * Isar: 'print_facts' prints all local facts of the current context, wenzelm@20013: both named and unnamed ones. wenzelm@20013: wenzelm@18308: * Isar: 'def' now admits simultaneous definitions, e.g.: wenzelm@18308: wenzelm@18308: def x == "t" and y == "u" wenzelm@18308: wenzelm@18540: * Isar: added command 'unfolding', which is structurally similar to wenzelm@18540: 'using', but affects both the goal state and facts by unfolding given wenzelm@18815: rewrite rules. Thus many occurrences of the 'unfold' method or wenzelm@18540: 'unfolded' attribute may be replaced by first-class proof text. wenzelm@18540: wenzelm@18815: * Isar: methods 'unfold' / 'fold', attributes 'unfolded' / 'folded', wenzelm@18815: and command 'unfolding' now all support object-level equalities wenzelm@18815: (potentially conditional). The underlying notion of rewrite rule is wenzelm@18815: analogous to the 'rule_format' attribute, but *not* that of the wenzelm@18815: Simplifier (which is usually more generous). wenzelm@18815: kleing@24238: * Isar: the new attribute [rotated n] (default n = 1) rotates the kleing@24238: premises of a theorem by n. Useful in conjunction with drule. kleing@24238: wenzelm@19220: * Isar: the goal restriction operator [N] (default N = 1) evaluates a wenzelm@19220: method expression within a sandbox consisting of the first N wenzelm@19240: sub-goals, which need to exist. For example, ``simp_all [3]'' wenzelm@19240: simplifies the first three sub-goals, while (rule foo, simp_all)[] wenzelm@19240: simplifies all new goals that emerge from applying rule foo to the wenzelm@19240: originally first one. wenzelm@19220: wenzelm@19814: * Isar: schematic goals are no longer restricted to higher-order wenzelm@19814: patterns; e.g. ``lemma "?P(?x)" by (rule TrueI)'' now works as wenzelm@19814: expected. wenzelm@19814: wenzelm@18901: * Isar: the conclusion of a long theorem statement is now either wenzelm@18901: 'shows' (a simultaneous conjunction, as before), or 'obtains' wenzelm@18901: (essentially a disjunction of cases with local parameters and wenzelm@18901: assumptions). The latter allows to express general elimination rules wenzelm@18910: adequately; in this notation common elimination rules look like this: wenzelm@18901: wenzelm@18901: lemma exE: -- "EX x. P x ==> (!!x. P x ==> thesis) ==> thesis" wenzelm@18901: assumes "EX x. P x" wenzelm@18901: obtains x where "P x" wenzelm@18901: wenzelm@18901: lemma conjE: -- "A & B ==> (A ==> B ==> thesis) ==> thesis" wenzelm@18901: assumes "A & B" wenzelm@18901: obtains A and B wenzelm@18901: wenzelm@18901: lemma disjE: -- "A | B ==> (A ==> thesis) ==> (B ==> thesis) ==> thesis" wenzelm@18901: assumes "A | B" wenzelm@18901: obtains wenzelm@18901: A wenzelm@18901: | B wenzelm@18901: wenzelm@18910: The subsequent classical rules even refer to the formal "thesis" wenzelm@18901: explicitly: wenzelm@18901: wenzelm@18901: lemma classical: -- "(~ thesis ==> thesis) ==> thesis" wenzelm@18901: obtains "~ thesis" wenzelm@18901: wenzelm@18910: lemma Peirce's_Law: -- "((thesis ==> something) ==> thesis) ==> thesis" wenzelm@18910: obtains "thesis ==> something" wenzelm@18901: wenzelm@18901: The actual proof of an 'obtains' statement is analogous to that of the wenzelm@18910: Isar proof element 'obtain', only that there may be several cases. wenzelm@18910: Optional case names may be specified in parentheses; these will be wenzelm@18910: available both in the present proof and as annotations in the wenzelm@18910: resulting rule, for later use with the 'cases' method (cf. attribute wenzelm@18910: case_names). wenzelm@18901: wenzelm@21447: * Isar: the assumptions of a long theorem statement are available as wenzelm@21447: "assms" fact in the proof context. This is more appropriate than the wenzelm@21447: (historical) "prems", which refers to all assumptions of the current wenzelm@21447: context, including those from the target locale, proof body etc. wenzelm@21447: wenzelm@19263: * Isar: 'print_statement' prints theorems from the current theory or wenzelm@19263: proof context in long statement form, according to the syntax of a wenzelm@19263: top-level lemma. wenzelm@19263: wenzelm@18901: * Isar: 'obtain' takes an optional case name for the local context wenzelm@18901: introduction rule (default "that"). wenzelm@18901: wenzelm@19587: * Isar: removed obsolete 'concl is' patterns. INCOMPATIBILITY, use wenzelm@19587: explicit (is "_ ==> ?foo") in the rare cases where this still happens wenzelm@19587: to occur. wenzelm@19587: wenzelm@19682: * Pure: syntax "CONST name" produces a fully internalized constant wenzelm@19682: according to the current context. This is particularly useful for wenzelm@19682: syntax translations that should refer to internal constant wenzelm@19682: representations independently of name spaces. wenzelm@19682: wenzelm@21537: * Pure: syntax constant for foo (binder "FOO ") is called "foo_binder" wenzelm@21537: instead of "FOO ". This allows multiple binder declarations to coexist wenzelm@21537: in the same context. INCOMPATIBILITY. wenzelm@21537: wenzelm@21209: * Isar/locales: 'notation' provides a robust interface to the 'syntax' wenzelm@21209: primitive that also works in a locale context (both for constants and wenzelm@24950: fixed variables). Type declaration and internal syntactic representation wenzelm@24950: of given constants retrieved from the context. Likewise, the wenzelm@24950: 'no_notation' command allows to remove given syntax annotations from the wenzelm@24950: current context. wenzelm@19682: wenzelm@19665: * Isar/locales: new derived specification elements 'axiomatization', wenzelm@19665: 'definition', 'abbreviation', which support type-inference, admit wenzelm@19083: object-level specifications (equality, equivalence). See also the wenzelm@19083: isar-ref manual. Examples: wenzelm@19081: wenzelm@19665: axiomatization wenzelm@21595: eq (infix "===" 50) where wenzelm@21595: eq_refl: "x === x" and eq_subst: "x === y ==> P x ==> P y" wenzelm@21595: wenzelm@21595: definition "f x y = x + y + 1" wenzelm@21595: definition g where "g x = f x x" wenzelm@19081: wenzelm@19363: abbreviation wenzelm@21595: neq (infix "=!=" 50) where wenzelm@19363: "x =!= y == ~ (x === y)" wenzelm@19081: wenzelm@19083: These specifications may be also used in a locale context. Then the wenzelm@19083: constants being introduced depend on certain fixed parameters, and the wenzelm@19083: constant name is qualified by the locale base name. An internal wenzelm@19083: abbreviation takes care for convenient input and output, making the wenzelm@19088: parameters implicit and using the original short name. See also wenzelm@25177: src/HOL/ex/Abstract_NAT.thy for an example of deriving polymorphic wenzelm@19083: entities from a monomorphic theory. wenzelm@19083: wenzelm@19083: Presently, abbreviations are only available 'in' a target locale, but wenzelm@19363: not inherited by general import expressions. Also note that wenzelm@19363: 'abbreviation' may be used as a type-safe replacement for 'syntax' + wenzelm@24735: 'translations' in common applications. The "no_abbrevs" print mode wenzelm@24735: prevents folding of abbreviations in term output. wenzelm@19084: wenzelm@19682: Concrete syntax is attached to specified constants in internal form, wenzelm@19682: independently of name spaces. The parse tree representation is wenzelm@21209: slightly different -- use 'notation' instead of raw 'syntax', and wenzelm@19682: 'translations' with explicit "CONST" markup to accommodate this. wenzelm@19665: wenzelm@24800: * Pure/Isar: unified syntax for new-style specification mechanisms wenzelm@24800: (e.g. 'definition', 'abbreviation', or 'inductive' in HOL) admits wenzelm@24800: full type inference and dummy patterns ("_"). For example: wenzelm@24735: wenzelm@24735: definition "K x _ = x" wenzelm@24735: wenzelm@24738: inductive conj for A B wenzelm@24738: where "A ==> B ==> conj A B" wenzelm@24738: wenzelm@21735: * Pure: command 'print_abbrevs' prints all constant abbreviations of wenzelm@21735: the current context. Print mode "no_abbrevs" prevents inversion of wenzelm@21735: abbreviations on output. wenzelm@21735: wenzelm@24800: * Isar/locales: improved parameter handling: use of locales "var" and wenzelm@24800: "struct" no longer necessary; - parameter renamings are no longer wenzelm@24800: required to be injective. For example, this allows to define wenzelm@24800: endomorphisms as locale endom = homom mult mult h. ballarin@19783: ballarin@19931: * Isar/locales: changed the way locales with predicates are defined. ballarin@19931: Instead of accumulating the specification, the imported expression is wenzelm@22126: now an interpretation. INCOMPATIBILITY: different normal form of wenzelm@22126: locale expressions. In particular, in interpretations of locales with wenzelm@22126: predicates, goals repesenting already interpreted fragments are not wenzelm@22126: removed automatically. Use methods `intro_locales' and wenzelm@22126: `unfold_locales'; see below. wenzelm@22126: wenzelm@22126: * Isar/locales: new methods `intro_locales' and `unfold_locales' wenzelm@22126: provide backward reasoning on locales predicates. The methods are wenzelm@22126: aware of interpretations and discharge corresponding goals. wenzelm@22126: `intro_locales' is less aggressive then `unfold_locales' and does not wenzelm@22126: unfold predicates to assumptions. ballarin@19931: ballarin@19931: * Isar/locales: the order in which locale fragments are accumulated wenzelm@22126: has changed. This enables to override declarations from fragments due wenzelm@22126: to interpretations -- for example, unwanted simp rules. ballarin@19931: ballarin@23920: * Isar/locales: interpretation in theories and proof contexts has been ballarin@23920: extended. One may now specify (and prove) equations, which are ballarin@23920: unfolded in interpreted theorems. This is useful for replacing ballarin@23920: defined concepts (constants depending on locale parameters) by ballarin@23920: concepts already existing in the target context. Example: ballarin@23920: ballarin@23920: interpretation partial_order ["op <= :: [int, int] => bool"] ballarin@23920: where "partial_order.less (op <=) (x::int) y = (x < y)" ballarin@23920: wenzelm@24800: Typically, the constant `partial_order.less' is created by a wenzelm@24800: definition specification element in the context of locale wenzelm@24800: partial_order. wenzelm@24800: wenzelm@24859: * Method "induct": improved internal context management to support wenzelm@24800: local fixes and defines on-the-fly. Thus explicit meta-level wenzelm@24800: connectives !! and ==> are rarely required anymore in inductive goals wenzelm@24800: (using object-logic connectives for this purpose has been long wenzelm@24800: obsolete anyway). Common proof patterns are explained in wenzelm@25177: src/HOL/Induct/Common_Patterns.thy, see also wenzelm@25177: src/HOL/Isar_examples/Puzzle.thy and src/HOL/Lambda for realistic wenzelm@25177: examples. wenzelm@24606: wenzelm@24859: * Method "induct": improved handling of simultaneous goals. Instead of wenzelm@24606: introducing object-level conjunction, the statement is now split into wenzelm@24606: several conclusions, while the corresponding symbolic cases are nested wenzelm@24606: accordingly. INCOMPATIBILITY, proofs need to be structured explicitly, wenzelm@25177: see src/HOL/Induct/Common_Patterns.thy, for example. wenzelm@24606: wenzelm@24859: * Method "induct": mutual induction rules are now specified as a list wenzelm@24800: of rule sharing the same induction cases. HOL packages usually provide wenzelm@24606: foo_bar.inducts for mutually defined items foo and bar (e.g. inductive wenzelm@24859: predicates/sets or datatypes). INCOMPATIBILITY, users need to specify wenzelm@24859: mutual induction rules differently, i.e. like this: wenzelm@18506: wenzelm@18506: (induct rule: foo_bar.inducts) wenzelm@18506: (induct set: foo bar) wenzelm@24859: (induct pred: foo bar) wenzelm@18506: (induct type: foo bar) wenzelm@18506: wenzelm@18506: The ML function ProjectRule.projections turns old-style rules into the wenzelm@18506: new format. wenzelm@18506: wenzelm@24859: * Method "coinduct": dual of induction, see wenzelm@18399: src/HOL/Library/Coinductive_List.thy for various examples. wenzelm@18399: wenzelm@24859: * Method "cases", "induct", "coinduct": the ``(open)'' option is wenzelm@24859: considered a legacy feature. wenzelm@24859: wenzelm@20919: * Attribute "symmetric" produces result with standardized schematic wenzelm@20919: variables (index 0). Potential INCOMPATIBILITY. wenzelm@20919: wenzelm@22126: * Simplifier: by default the simplifier trace only shows top level wenzelm@22126: rewrites now. That is, trace_simp_depth_limit is set to 1 by wenzelm@22126: default. Thus there is less danger of being flooded by the trace. The wenzelm@22126: trace indicates where parts have been suppressed. lp15@57253: wenzelm@18536: * Provers/classical: removed obsolete classical version of elim_format wenzelm@18536: attribute; classical elim/dest rules are now treated uniformly when wenzelm@18536: manipulating the claset. wenzelm@18536: wenzelm@18694: * Provers/classical: stricter checks to ensure that supplied intro, wenzelm@18694: dest and elim rules are well-formed; dest and elim rules must have at wenzelm@18694: least one premise. wenzelm@18694: wenzelm@18694: * Provers/classical: attributes dest/elim/intro take an optional wenzelm@18695: weight argument for the rule (just as the Pure versions). Weights are wenzelm@18696: ignored by automated tools, but determine the search order of single wenzelm@18694: rule steps. paulson@18557: wenzelm@18536: * Syntax: input syntax now supports dummy variable binding "%_. b", wenzelm@18536: where the body does not mention the bound variable. Note that dummy wenzelm@18536: patterns implicitly depend on their context of bounds, which makes wenzelm@18536: "{_. _}" match any set comprehension as expected. Potential wenzelm@18536: INCOMPATIBILITY -- parse translations need to cope with syntactic wenzelm@18536: constant "_idtdummy" in the binding position. wenzelm@18536: wenzelm@18536: * Syntax: removed obsolete syntactic constant "_K" and its associated wenzelm@18536: parse translation. INCOMPATIBILITY -- use dummy abstraction instead, wenzelm@18536: for example "A -> B" => "Pi A (%_. B)". wenzelm@17779: wenzelm@20582: * Pure: 'class_deps' command visualizes the subclass relation, using wenzelm@20582: the graph browser tool. wenzelm@20582: wenzelm@24800: * Pure: 'print_theory' now suppresses certain internal declarations by wenzelm@24800: default; use '!' option for full details. wenzelm@20620: wenzelm@17865: nipkow@17806: *** HOL *** nipkow@17806: wenzelm@25129: * Method "metis" proves goals by applying the Metis general-purpose wenzelm@25129: resolution prover (see also http://gilith.com/software/metis/). wenzelm@25129: Examples are in the directory MetisExamples. WARNING: the wenzelm@25129: Isabelle/HOL-Metis integration does not yet work properly with wenzelm@25129: multi-threading. lp15@57253: wenzelm@25129: * Command 'sledgehammer' invokes external automatic theorem provers as wenzelm@25129: background processes. It generates calls to the "metis" method if wenzelm@25129: successful. These can be pasted into the proof. Users do not have to wenzelm@25129: wait for the automatic provers to return. WARNING: does not really wenzelm@25129: work with multi-threading. wenzelm@25129: wenzelm@24804: * New "auto_quickcheck" feature tests outermost goal statements for wenzelm@24804: potential counter-examples. Controlled by ML references wenzelm@24804: auto_quickcheck (default true) and auto_quickcheck_time_limit (default wenzelm@25129: 5000 milliseconds). Fails silently if statements is outside of wenzelm@25129: executable fragment, or any other codgenerator problem occurs. wenzelm@24804: haftmann@25184: * New constant "undefined" with axiom "undefined x = undefined". haftmann@25184: haftmann@25184: * Added class "HOL.eq", allowing for code generation with polymorphic haftmann@25184: equality. haftmann@25184: haftmann@25184: * Some renaming of class constants due to canonical name prefixing in haftmann@25184: the new 'class' package: haftmann@25184: haftmann@25184: HOL.abs ~> HOL.abs_class.abs haftmann@25184: HOL.divide ~> HOL.divide_class.divide haftmann@25184: 0 ~> HOL.zero_class.zero haftmann@25184: 1 ~> HOL.one_class.one haftmann@25184: op + ~> HOL.plus_class.plus haftmann@25184: op - ~> HOL.minus_class.minus haftmann@25184: uminus ~> HOL.minus_class.uminus haftmann@25184: op * ~> HOL.times_class.times haftmann@25184: op < ~> HOL.ord_class.less haftmann@25184: op <= > HOL.ord_class.less_eq haftmann@25184: Nat.power ~> Power.power_class.power haftmann@25184: Nat.size ~> Nat.size_class.size haftmann@25184: Numeral.number_of ~> Numeral.number_class.number_of haftmann@25184: FixedPoint.Inf ~> Lattices.complete_lattice_class.Inf haftmann@25184: FixedPoint.Sup ~> Lattices.complete_lattice_class.Sup haftmann@25184: Orderings.min ~> Orderings.ord_class.min haftmann@25184: Orderings.max ~> Orderings.ord_class.max haftmann@25184: Divides.op div ~> Divides.div_class.div haftmann@25184: Divides.op mod ~> Divides.div_class.mod haftmann@25184: Divides.op dvd ~> Divides.div_class.dvd haftmann@25184: haftmann@25184: INCOMPATIBILITY. Adaptions may be required in the following cases: haftmann@25184: haftmann@25184: a) User-defined constants using any of the names "plus", "minus", haftmann@25184: "times", "less" or "less_eq". The standard syntax translations for haftmann@25184: "+", "-" and "*" may go wrong. INCOMPATIBILITY: use more specific haftmann@25184: names. haftmann@25184: haftmann@25184: b) Variables named "plus", "minus", "times", "less", "less_eq" haftmann@25184: INCOMPATIBILITY: use more specific names. haftmann@25184: haftmann@25184: c) Permutative equations (e.g. "a + b = b + a") haftmann@25184: Since the change of names also changes the order of terms, permutative haftmann@25184: rewrite rules may get applied in a different order. Experience shows haftmann@25184: that this is rarely the case (only two adaptions in the whole Isabelle haftmann@25184: distribution). INCOMPATIBILITY: rewrite proofs haftmann@25184: haftmann@25184: d) ML code directly refering to constant names haftmann@25184: This in general only affects hand-written proof tactics, simprocs and haftmann@25184: so on. INCOMPATIBILITY: grep your sourcecode and replace names. haftmann@25184: Consider using @{const_name} antiquotation. haftmann@25184: haftmann@25184: * New class "default" with associated constant "default". haftmann@25184: haftmann@25184: * Function "sgn" is now overloaded and available on int, real, complex haftmann@25184: (and other numeric types), using class "sgn". Two possible defs of haftmann@25184: sgn are given as equational assumptions in the classes sgn_if and haftmann@25184: sgn_div_norm; ordered_idom now also inherits from sgn_if. haftmann@25184: INCOMPATIBILITY. haftmann@25184: haftmann@25184: * Locale "partial_order" now unified with class "order" (cf. theory haftmann@25184: Orderings), added parameter "less". INCOMPATIBILITY. haftmann@25184: haftmann@25184: * Renamings in classes "order" and "linorder": facts "refl", "trans" and haftmann@25184: "cases" to "order_refl", "order_trans" and "linorder_cases", to avoid haftmann@25184: clashes with HOL "refl" and "trans". INCOMPATIBILITY. haftmann@25184: haftmann@25184: * Classes "order" and "linorder": potential INCOMPATIBILITY due to haftmann@25184: changed order of proof goals in instance proofs. haftmann@25184: haftmann@25184: * The transitivity reasoner for partial and linear orders is set up haftmann@25184: for classes "order" and "linorder". Instances of the reasoner are available haftmann@25184: in all contexts importing or interpreting the corresponding locales. haftmann@25184: Method "order" invokes the reasoner separately; the reasoner haftmann@25184: is also integrated with the Simplifier as a solver. Diagnostic haftmann@25184: command 'print_orders' shows the available instances of the reasoner haftmann@25184: in the current context. haftmann@25184: haftmann@25184: * Localized monotonicity predicate in theory "Orderings"; integrated haftmann@25184: lemmas max_of_mono and min_of_mono with this predicate. haftmann@25184: INCOMPATIBILITY. haftmann@25184: haftmann@25184: * Formulation of theorem "dense" changed slightly due to integration haftmann@25184: with new class dense_linear_order. haftmann@25184: haftmann@25184: * Uniform lattice theory development in HOL. haftmann@25184: haftmann@25184: constants "meet" and "join" now named "inf" and "sup" haftmann@25184: constant "Meet" now named "Inf" haftmann@25184: haftmann@25184: classes "meet_semilorder" and "join_semilorder" now named haftmann@25184: "lower_semilattice" and "upper_semilattice" haftmann@25184: class "lorder" now named "lattice" haftmann@25184: class "comp_lat" now named "complete_lattice" haftmann@25184: haftmann@25184: Instantiation of lattice classes allows explicit definitions haftmann@25184: for "inf" and "sup" operations (or "Inf" and "Sup" for complete lattices). haftmann@25184: haftmann@25184: INCOMPATIBILITY. Theorem renames: haftmann@25184: haftmann@25184: meet_left_le ~> inf_le1 haftmann@25184: meet_right_le ~> inf_le2 haftmann@25184: join_left_le ~> sup_ge1 haftmann@25184: join_right_le ~> sup_ge2 haftmann@25184: meet_join_le ~> inf_sup_ord haftmann@25184: le_meetI ~> le_infI haftmann@25184: join_leI ~> le_supI haftmann@25184: le_meet ~> le_inf_iff haftmann@25184: le_join ~> ge_sup_conv haftmann@25184: meet_idempotent ~> inf_idem haftmann@25184: join_idempotent ~> sup_idem haftmann@25184: meet_comm ~> inf_commute haftmann@25184: join_comm ~> sup_commute haftmann@25184: meet_leI1 ~> le_infI1 haftmann@25184: meet_leI2 ~> le_infI2 haftmann@25184: le_joinI1 ~> le_supI1 haftmann@25184: le_joinI2 ~> le_supI2 haftmann@25184: meet_assoc ~> inf_assoc haftmann@25184: join_assoc ~> sup_assoc haftmann@25184: meet_left_comm ~> inf_left_commute haftmann@25184: meet_left_idempotent ~> inf_left_idem haftmann@25184: join_left_comm ~> sup_left_commute haftmann@25184: join_left_idempotent ~> sup_left_idem haftmann@25184: meet_aci ~> inf_aci haftmann@25184: join_aci ~> sup_aci haftmann@25184: le_def_meet ~> le_iff_inf haftmann@25184: le_def_join ~> le_iff_sup haftmann@25184: join_absorp2 ~> sup_absorb2 haftmann@25184: join_absorp1 ~> sup_absorb1 haftmann@25184: meet_absorp1 ~> inf_absorb1 haftmann@25184: meet_absorp2 ~> inf_absorb2 haftmann@25184: meet_join_absorp ~> inf_sup_absorb haftmann@25184: join_meet_absorp ~> sup_inf_absorb haftmann@25184: distrib_join_le ~> distrib_sup_le haftmann@25184: distrib_meet_le ~> distrib_inf_le haftmann@25184: haftmann@25184: add_meet_distrib_left ~> add_inf_distrib_left haftmann@25184: add_join_distrib_left ~> add_sup_distrib_left haftmann@25184: is_join_neg_meet ~> is_join_neg_inf haftmann@25184: is_meet_neg_join ~> is_meet_neg_sup haftmann@25184: add_meet_distrib_right ~> add_inf_distrib_right haftmann@25184: add_join_distrib_right ~> add_sup_distrib_right haftmann@25184: add_meet_join_distribs ~> add_sup_inf_distribs haftmann@25184: join_eq_neg_meet ~> sup_eq_neg_inf haftmann@25184: meet_eq_neg_join ~> inf_eq_neg_sup haftmann@25184: add_eq_meet_join ~> add_eq_inf_sup haftmann@25184: meet_0_imp_0 ~> inf_0_imp_0 haftmann@25184: join_0_imp_0 ~> sup_0_imp_0 haftmann@25184: meet_0_eq_0 ~> inf_0_eq_0 haftmann@25184: join_0_eq_0 ~> sup_0_eq_0 haftmann@25184: neg_meet_eq_join ~> neg_inf_eq_sup haftmann@25184: neg_join_eq_meet ~> neg_sup_eq_inf haftmann@25184: join_eq_if ~> sup_eq_if haftmann@25184: haftmann@25184: mono_meet ~> mono_inf haftmann@25184: mono_join ~> mono_sup haftmann@25184: meet_bool_eq ~> inf_bool_eq haftmann@25184: join_bool_eq ~> sup_bool_eq haftmann@25184: meet_fun_eq ~> inf_fun_eq haftmann@25184: join_fun_eq ~> sup_fun_eq haftmann@25184: meet_set_eq ~> inf_set_eq haftmann@25184: join_set_eq ~> sup_set_eq haftmann@25184: meet1_iff ~> inf1_iff haftmann@25184: meet2_iff ~> inf2_iff haftmann@25184: meet1I ~> inf1I haftmann@25184: meet2I ~> inf2I haftmann@25184: meet1D1 ~> inf1D1 haftmann@25184: meet2D1 ~> inf2D1 haftmann@25184: meet1D2 ~> inf1D2 haftmann@25184: meet2D2 ~> inf2D2 haftmann@25184: meet1E ~> inf1E haftmann@25184: meet2E ~> inf2E haftmann@25184: join1_iff ~> sup1_iff haftmann@25184: join2_iff ~> sup2_iff haftmann@25184: join1I1 ~> sup1I1 haftmann@25184: join2I1 ~> sup2I1 haftmann@25184: join1I1 ~> sup1I1 haftmann@25184: join2I2 ~> sup1I2 haftmann@25184: join1CI ~> sup1CI haftmann@25184: join2CI ~> sup2CI haftmann@25184: join1E ~> sup1E haftmann@25184: join2E ~> sup2E haftmann@25184: haftmann@25184: is_meet_Meet ~> is_meet_Inf haftmann@25184: Meet_bool_def ~> Inf_bool_def haftmann@25184: Meet_fun_def ~> Inf_fun_def haftmann@25184: Meet_greatest ~> Inf_greatest haftmann@25184: Meet_lower ~> Inf_lower haftmann@25184: Meet_set_def ~> Inf_set_def haftmann@25184: haftmann@25184: Sup_def ~> Sup_Inf haftmann@25184: Sup_bool_eq ~> Sup_bool_def haftmann@25184: Sup_fun_eq ~> Sup_fun_def haftmann@25184: Sup_set_eq ~> Sup_set_def haftmann@25184: haftmann@25184: listsp_meetI ~> listsp_infI haftmann@25184: listsp_meet_eq ~> listsp_inf_eq haftmann@25184: haftmann@25184: meet_min ~> inf_min haftmann@25184: join_max ~> sup_max haftmann@25184: haftmann@25184: * Added syntactic class "size"; overloaded constant "size" now has haftmann@25184: type "'a::size ==> bool" haftmann@25184: wenzelm@24800: * Internal reorganisation of `size' of datatypes: size theorems wenzelm@24800: "foo.size" are no longer subsumed by "foo.simps" (but are still wenzelm@24800: simplification rules by default!); theorems "prod.size" now named haftmann@25184: "*.size". haftmann@25184: haftmann@25184: * Class "div" now inherits from class "times" rather than "type". haftmann@25184: INCOMPATIBILITY. wenzelm@24800: wenzelm@24800: * HOL/Finite_Set: "name-space" locales Lattice, Distrib_lattice, wenzelm@24800: Linorder etc. have disappeared; operations defined in terms of wenzelm@24800: fold_set now are named Inf_fin, Sup_fin. INCOMPATIBILITY. wenzelm@24800: wenzelm@25129: * HOL/Nat: neq0_conv no longer declared as iff. INCOMPATIBILITY. wenzelm@25129: wenzelm@24800: * HOL-Word: New extensive library and type for generic, fixed size blanchet@47672: machine words, with arithmetic, bit-wise, shifting and rotating wenzelm@24800: operations, reflection into int, nat, and bool lists, automation for wenzelm@24800: linear arithmetic (by automatic reflection into nat or int), including wenzelm@24800: lemmas on overflow and monotonicity. Instantiated to all appropriate wenzelm@24800: arithmetic type classes, supporting automatic simplification of wenzelm@24800: numerals on all operations. kleing@24333: kleing@24333: * Library/Boolean_Algebra: locales for abstract boolean algebras. kleing@24333: kleing@24333: * Library/Numeral_Type: numbers as types, e.g. TYPE(32). kleing@24333: haftmann@23850: * Code generator library theories: haftmann@24993: - Code_Integer represents HOL integers by big integer literals in target haftmann@23850: languages. haftmann@24993: - Code_Char represents HOL characters by character literals in target haftmann@23850: languages. haftmann@24993: - Code_Char_chr like Code_Char, but also offers treatment of character haftmann@24993: codes; includes Code_Integer. wenzelm@24800: - Executable_Set allows to generate code for finite sets using lists. wenzelm@24800: - Executable_Rat implements rational numbers as triples (sign, enumerator, haftmann@23850: denominator). wenzelm@24800: - Executable_Real implements a subset of real numbers, namly those haftmann@23850: representable by rational numbers. wenzelm@24800: - Efficient_Nat implements natural numbers by integers, which in general will haftmann@23850: result in higher efficency; pattern matching with 0/Suc is eliminated; haftmann@24993: includes Code_Integer. haftmann@24993: - Code_Index provides an additional datatype index which is mapped to haftmann@24993: target-language built-in integers. haftmann@26355: - Code_Message provides an additional datatype message_string which is isomorphic to haftmann@24993: strings; messages are mapped to target-language strings. haftmann@23850: berghofe@23783: * New package for inductive predicates berghofe@23783: berghofe@23783: An n-ary predicate p with m parameters z_1, ..., z_m can now be defined via berghofe@23783: berghofe@23783: inductive berghofe@23783: p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" berghofe@23783: for z_1 :: U_1 and ... and z_n :: U_m berghofe@23783: where berghofe@23783: rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n" berghofe@23783: | ... berghofe@23783: wenzelm@24800: with full support for type-inference, rather than berghofe@23783: berghofe@23783: consts s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" berghofe@23783: berghofe@23783: abbreviation p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" berghofe@23783: where "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m" berghofe@23783: berghofe@23783: inductive "s z_1 ... z_m" berghofe@23783: intros berghofe@23783: rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m" berghofe@23783: ... berghofe@23783: berghofe@23783: For backward compatibility, there is a wrapper allowing inductive berghofe@23783: sets to be defined with the new package via berghofe@23783: berghofe@23783: inductive_set berghofe@23783: s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" berghofe@23783: for z_1 :: U_1 and ... and z_n :: U_m berghofe@23783: where berghofe@23783: rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m" berghofe@23783: | ... berghofe@23783: berghofe@23783: or berghofe@23783: berghofe@23783: inductive_set berghofe@23783: s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set" berghofe@23783: and p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool" berghofe@23783: for z_1 :: U_1 and ... and z_n :: U_m berghofe@23783: where berghofe@23783: "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m" berghofe@23783: | rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n" berghofe@23783: | ... berghofe@23783: berghofe@23783: if the additional syntax "p ..." is required. berghofe@23783: wenzelm@25177: Numerous examples can be found in the subdirectories src/HOL/Auth, wenzelm@25177: src/HOL/Bali, src/HOL/Induct, and src/HOL/MicroJava. berghofe@23783: berghofe@23783: INCOMPATIBILITIES: berghofe@23783: berghofe@23783: - Since declaration and definition of inductive sets or predicates wenzelm@24800: is no longer separated, abbreviations involving the newly wenzelm@24800: introduced sets or predicates must be specified together with the wenzelm@24800: introduction rules after the 'where' keyword (see above), rather wenzelm@24800: than before the actual inductive definition. wenzelm@24800: wenzelm@24800: - The variables in induction and elimination rules are now wenzelm@24800: quantified in the order of their occurrence in the introduction wenzelm@24800: rules, rather than in alphabetical order. Since this may break wenzelm@24800: some proofs, these proofs either have to be repaired, e.g. by wenzelm@24800: reordering the variables a_i_1 ... a_i_{k_i} in Isar 'case' wenzelm@24800: statements of the form berghofe@23783: berghofe@23783: case (rule_i a_i_1 ... a_i_{k_i}) berghofe@23783: berghofe@23783: or the old order of quantification has to be restored by explicitly adding berghofe@23783: meta-level quantifiers in the introduction rules, i.e. berghofe@23783: berghofe@23783: | rule_i: "!!a_i_1 ... a_i_{k_i}. ... ==> p z_1 ... z_m t_i_1 ... t_i_n" berghofe@23783: berghofe@23783: - The format of the elimination rules is now berghofe@23783: berghofe@23783: p z_1 ... z_m x_1 ... x_n ==> berghofe@23783: (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P) berghofe@23783: ==> ... ==> P berghofe@23783: berghofe@23783: for predicates and berghofe@23783: berghofe@23783: (x_1, ..., x_n) : s z_1 ... z_m ==> berghofe@23783: (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P) berghofe@23783: ==> ... ==> P berghofe@23783: berghofe@23783: for sets rather than berghofe@23783: berghofe@23783: x : s z_1 ... z_m ==> berghofe@23783: (!!a_1_1 ... a_1_{k_1}. x = (t_1_1, ..., t_1_n) ==> ... ==> P) berghofe@23783: ==> ... ==> P berghofe@23783: wenzelm@24800: This may require terms in goals to be expanded to n-tuples wenzelm@24800: (e.g. using case_tac or simplification with the split_paired_all wenzelm@24800: rule) before the above elimination rule is applicable. wenzelm@24800: wenzelm@24800: - The elimination or case analysis rules for (mutually) inductive wenzelm@24800: sets or predicates are now called "p_1.cases" ... "p_k.cases". The wenzelm@24800: list of rules "p_1_..._p_k.elims" is no longer available. berghofe@23783: krauss@25198: * New package "function"/"fun" for general recursive functions, krauss@25198: supporting mutual and nested recursion, definitions in local contexts, krauss@25198: more general pattern matching and partiality. See HOL/ex/Fundefs.thy krauss@25198: for small examples, and the separate tutorial on the function krauss@25198: package. The old recdef "package" is still available as before, but krauss@25198: users are encouraged to use the new package. krauss@25198: krauss@25198: * Method "lexicographic_order" automatically synthesizes termination lp15@57253: relations as lexicographic combinations of size measures. krauss@25198: wenzelm@24800: * Case-expressions allow arbitrary constructor-patterns (including wenzelm@24800: "_") and take their order into account, like in functional wenzelm@24800: programming. Internally, this is translated into nested wenzelm@24800: case-expressions; missing cases are added and mapped to the predefined wenzelm@24800: constant "undefined". In complicated cases printing may no longer show wenzelm@24800: the original input but the internal form. Lambda-abstractions allow wenzelm@24800: the same form of pattern matching: "% pat1 => e1 | ..." is an wenzelm@24800: abbreviation for "%x. case x of pat1 => e1 | ..." where x is a new wenzelm@24800: variable. nipkow@23564: huffman@23468: * IntDef: The constant "int :: nat => int" has been removed; now "int" wenzelm@24800: is an abbreviation for "of_nat :: nat => int". The simplification wenzelm@24800: rules for "of_nat" have been changed to work like "int" did wenzelm@24800: previously. Potential INCOMPATIBILITY: huffman@23468: - "of_nat (Suc m)" simplifies to "1 + of_nat m" instead of "of_nat m + 1" huffman@23468: - of_nat_diff and of_nat_mult are no longer default simp rules huffman@23377: chaieb@23295: * Method "algebra" solves polynomial equations over (semi)rings using wenzelm@24800: Groebner bases. The (semi)ring structure is defined by locales and the wenzelm@24800: tool setup depends on that generic context. Installing the method for wenzelm@24800: a specific type involves instantiating the locale and possibly adding wenzelm@24800: declarations for computation on the coefficients. The method is wenzelm@24800: already instantiated for natural numbers and for the axiomatic class wenzelm@24800: of idoms with numerals. See also the paper by Chaieb and Wenzel at wenzelm@24800: CALCULEMUS 2007 for the general principles underlying this wenzelm@24800: architecture of context-aware proof-tools. wenzelm@24800: wenzelm@25033: * Method "ferrack" implements quantifier elimination over wenzelm@25033: special-purpose dense linear orders using locales (analogous to wenzelm@25033: "algebra"). The method is already installed for class wenzelm@25033: {ordered_field,recpower,number_ring} which subsumes real, hyperreal, wenzelm@25033: rat, etc. wenzelm@25033: wenzelm@24800: * Former constant "List.op @" now named "List.append". Use ML wenzelm@24800: antiquotations @{const_name List.append} or @{term " ... @ ... "} to wenzelm@24800: circumvent possible incompatibilities when working on ML level. wenzelm@24800: haftmann@24996: * primrec: missing cases mapped to "undefined" instead of "arbitrary". haftmann@22845: wenzelm@24800: * New function listsum :: 'a list => 'a for arbitrary monoids. wenzelm@24800: Special syntax: "SUM x <- xs. f x" (and latex variants) wenzelm@24800: wenzelm@24800: * New syntax for Haskell-like list comprehension (input only), eg. wenzelm@25177: [(x,y). x <- xs, y <- ys, x ~= y], see also src/HOL/List.thy. wenzelm@24800: wenzelm@24800: * The special syntax for function "filter" has changed from [x : wenzelm@24800: xs. P] to [x <- xs. P] to avoid an ambiguity caused by list wenzelm@24800: comprehension syntax, and for uniformity. INCOMPATIBILITY. wenzelm@24800: wenzelm@24800: * [a..b] is now defined for arbitrary linear orders. It used to be wenzelm@24800: defined on nat only, as an abbreviation for [a.. B" for equality on bool (with priority wenzelm@17996: 25 like -->); output depends on the "iff" print_mode, the default is wenzelm@17996: "A = B" (with priority 50). wenzelm@17996: wenzelm@21265: * Relations less (<) and less_eq (<=) are also available on type bool. wenzelm@21265: Modified syntax to disallow nesting without explicit parentheses, wenzelm@24800: e.g. "(x < y) < z" or "x < (y < z)", but NOT "x < y < z". Potential wenzelm@24800: INCOMPATIBILITY. wenzelm@21265: nipkow@18674: * "LEAST x:A. P" expands to "LEAST x. x:A & P" (input only). nipkow@18674: krauss@20716: * Relation composition operator "op O" now has precedence 75 and binds krauss@20716: stronger than union and intersection. INCOMPATIBILITY. krauss@20716: wenzelm@22126: * The old set interval syntax "{m..n(}" (and relatives) has been wenzelm@22126: removed. Use "{m.. ==> False", equivalences wenzelm@24800: (i.e. "=" on type bool) are handled, variable names of the form wenzelm@24800: "lit_" are no longer reserved, significant speedup. wenzelm@24800: wenzelm@24800: * Methods "sat" and "satx" can now replay MiniSat proof traces. wenzelm@22126: zChaff is still supported as well. wenzelm@22126: wenzelm@22126: * 'inductive' and 'datatype': provide projections of mutual rules, wenzelm@22126: bundled as foo_bar.inducts; wenzelm@22126: wenzelm@22126: * Library: moved theories Parity, GCD, Binomial, Infinite_Set to wenzelm@22126: Library. wenzelm@21256: wenzelm@21256: * Library: moved theory Accessible_Part to main HOL. wenzelm@19572: wenzelm@18446: * Library: added theory Coinductive_List of potentially infinite lists wenzelm@18446: as greatest fixed-point. wenzelm@18399: wenzelm@19254: * Library: added theory AssocList which implements (finite) maps as schirmer@19252: association lists. webertj@17809: wenzelm@24800: * Method "evaluation" solves goals (i.e. a boolean expression) wenzelm@24800: efficiently by compiling it to ML. The goal is "proved" (via an wenzelm@24800: oracle) if it evaluates to True. wenzelm@20807: wenzelm@20807: * Linear arithmetic now splits certain operators (e.g. min, max, abs) wenzelm@24800: also when invoked by the simplifier. This results in the Simplifier wenzelm@24800: being more powerful on arithmetic goals. INCOMPATIBILITY. wenzelm@24800: Configuration option fast_arith_split_limit=0 recovers the old wenzelm@24800: behavior. webertj@20217: wenzelm@22126: * Support for hex (0x20) and binary (0b1001) numerals. wenzelm@19254: wenzelm@20807: * New method: reify eqs (t), where eqs are equations for an wenzelm@20807: interpretation I :: 'a list => 'b => 'c and t::'c is an optional wenzelm@20807: parameter, computes a term s::'b and a list xs::'a list and proves the wenzelm@20807: theorem I xs s = t. This is also known as reification or quoting. The wenzelm@20807: resulting theorem is applied to the subgoal to substitute t with I xs wenzelm@20807: s. If t is omitted, the subgoal itself is reified. wenzelm@20807: wenzelm@20807: * New method: reflection corr_thm eqs (t). The parameters eqs and (t) wenzelm@20807: are as explained above. corr_thm is a theorem for I vs (f t) = I vs t, wenzelm@20807: where f is supposed to be a computable function (in the sense of code wenzelm@20807: generattion). The method uses reify to compute s and xs as above then wenzelm@20807: applies corr_thm and uses normalization by evaluation to "prove" f s = wenzelm@20807: r and finally gets the theorem t = r, which is again applied to the wenzelm@25177: subgoal. An Example is available in src/HOL/ex/ReflectionEx.thy. wenzelm@25177: wenzelm@25177: * Reflection: Automatic reification now handels binding, an example is wenzelm@25177: available in src/HOL/ex/ReflectionEx.thy wenzelm@20807: wenzelm@25397: * HOL-Statespace: ``State Spaces: The Locale Way'' introduces a schirmer@25409: command 'statespace' that is similar to 'record', but introduces an wenzelm@25397: abstract specification based on the locale infrastructure instead of wenzelm@25397: HOL types. This leads to extra flexibility in composing state spaces, wenzelm@25397: in particular multiple inheritance and renaming of components. wenzelm@25397: wenzelm@25397: wenzelm@19653: *** HOL-Complex *** wenzelm@19653: huffman@22971: * Hyperreal: Functions root and sqrt are now defined on negative real huffman@22971: inputs so that root n (- x) = - root n x and sqrt (- x) = - sqrt x. huffman@22971: Nonnegativity side conditions have been removed from many lemmas, so huffman@22971: that more subgoals may now be solved by simplification; potential huffman@22971: INCOMPATIBILITY. huffman@22971: wenzelm@24800: * Real: new type classes formalize real normed vector spaces and huffman@21791: algebras, using new overloaded constants scaleR :: real => 'a => 'a huffman@21791: and norm :: 'a => real. huffman@21791: wenzelm@24800: * Real: constant of_real :: real => 'a::real_algebra_1 injects from wenzelm@24800: reals into other types. The overloaded constant Reals :: 'a set is now wenzelm@24800: defined as range of_real; potential INCOMPATIBILITY. wenzelm@24800: wenzelm@24800: * Real: proper support for ML code generation, including 'quickcheck'. nipkow@23013: Reals are implemented as arbitrary precision rationals. nipkow@23013: wenzelm@22126: * Hyperreal: Several constants that previously worked only for the wenzelm@22126: reals have been generalized, so they now work over arbitrary vector wenzelm@22126: spaces. Type annotations may need to be added in some cases; potential wenzelm@22126: INCOMPATIBILITY. huffman@21791: huffman@22972: Infinitesimal :: ('a::real_normed_vector) star set huffman@22972: HFinite :: ('a::real_normed_vector) star set huffman@22972: HInfinite :: ('a::real_normed_vector) star set huffman@21791: approx :: ('a::real_normed_vector) star => 'a star => bool huffman@21791: monad :: ('a::real_normed_vector) star => 'a star set huffman@21791: galaxy :: ('a::real_normed_vector) star => 'a star set huffman@22972: (NS)LIMSEQ :: [nat => 'a::real_normed_vector, 'a] => bool huffman@21791: (NS)convergent :: (nat => 'a::real_normed_vector) => bool huffman@21791: (NS)Bseq :: (nat => 'a::real_normed_vector) => bool huffman@21791: (NS)Cauchy :: (nat => 'a::real_normed_vector) => bool huffman@21791: (NS)LIM :: ['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool huffman@21791: is(NS)Cont :: ['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool huffman@21791: deriv :: ['a::real_normed_field => 'a, 'a, 'a] => bool huffman@22972: sgn :: 'a::real_normed_vector => 'a huffman@23116: exp :: 'a::{recpower,real_normed_field,banach} => 'a huffman@21791: huffman@21791: * Complex: Some complex-specific constants are now abbreviations for wenzelm@22126: overloaded ones: complex_of_real = of_real, cmod = norm, hcmod = wenzelm@22126: hnorm. Other constants have been entirely removed in favor of the wenzelm@22126: polymorphic versions (INCOMPATIBILITY): huffman@21791: huffman@21791: approx <-- capprox huffman@21791: HFinite <-- CFinite huffman@21791: HInfinite <-- CInfinite huffman@21791: Infinitesimal <-- CInfinitesimal huffman@21791: monad <-- cmonad huffman@21791: galaxy <-- cgalaxy huffman@21791: (NS)LIM <-- (NS)CLIM, (NS)CRLIM huffman@21791: is(NS)Cont <-- is(NS)Contc, is(NS)contCR huffman@21791: (ns)deriv <-- (ns)cderiv huffman@21791: wenzelm@19653: wenzelm@24801: *** HOL-Algebra *** wenzelm@24801: wenzelm@24801: * Formalisation of ideals and the quotient construction over rings. wenzelm@24801: wenzelm@24801: * Order and lattice theory no longer based on records. wenzelm@24801: INCOMPATIBILITY. wenzelm@24801: wenzelm@24801: * Renamed lemmas least_carrier -> least_closed and greatest_carrier -> wenzelm@24801: greatest_closed. INCOMPATIBILITY. wenzelm@24801: wenzelm@24801: * Method algebra is now set up via an attribute. For examples see wenzelm@24801: Ring.thy. INCOMPATIBILITY: the method is now weaker on combinations wenzelm@24801: of algebraic structures. wenzelm@24801: wenzelm@24801: * Renamed theory CRing to Ring. wenzelm@24801: wenzelm@24801: wenzelm@24801: *** HOL-Nominal *** wenzelm@24801: wenzelm@25148: * Substantial, yet incomplete support for nominal datatypes (binding wenzelm@25177: structures) based on HOL-Nominal logic. See src/HOL/Nominal and wenzelm@25177: src/HOL/Nominal/Examples. Prospective users should consult wenzelm@25148: http://isabelle.in.tum.de/nominal/ wenzelm@25148: wenzelm@24801: wenzelm@17878: *** ML *** wenzelm@17878: wenzelm@24643: * ML basics: just one true type int, which coincides with IntInf.int wenzelm@24643: (even on SML/NJ). wenzelm@24643: wenzelm@22138: * ML within Isar: antiquotations allow to embed statically-checked wenzelm@22138: formal entities in the source, referring to the context available at wenzelm@22138: compile-time. For example: wenzelm@22138: wenzelm@25142: ML {* @{sort "{zero,one}"} *} wenzelm@22138: ML {* @{typ "'a => 'b"} *} wenzelm@22138: ML {* @{term "%x. x"} *} wenzelm@22138: ML {* @{prop "x == y"} *} wenzelm@22138: ML {* @{ctyp "'a => 'b"} *} wenzelm@22138: ML {* @{cterm "%x. x"} *} wenzelm@22138: ML {* @{cprop "x == y"} *} wenzelm@22138: ML {* @{thm asm_rl} *} wenzelm@22138: ML {* @{thms asm_rl} *} wenzelm@24692: ML {* @{type_name c} *} wenzelm@25142: ML {* @{type_syntax c} *} wenzelm@22376: ML {* @{const_name c} *} wenzelm@22376: ML {* @{const_syntax c} *} wenzelm@22138: ML {* @{context} *} wenzelm@22138: ML {* @{theory} *} wenzelm@22138: ML {* @{theory Pure} *} wenzelm@24692: ML {* @{theory_ref} *} wenzelm@24692: ML {* @{theory_ref Pure} *} wenzelm@22138: ML {* @{simpset} *} wenzelm@22138: ML {* @{claset} *} wenzelm@22138: ML {* @{clasimpset} *} wenzelm@22138: wenzelm@22151: The same works for sources being ``used'' within an Isar context. wenzelm@22151: wenzelm@22152: * ML in Isar: improved error reporting; extra verbosity with wenzelm@24706: ML_Context.trace enabled. wenzelm@22152: wenzelm@19032: * Pure/General/table.ML: the join operations now works via exceptions wenzelm@24706: DUP/SAME instead of type option. This is simpler in simple cases, and wenzelm@19081: admits slightly more efficient complex applications. wenzelm@18446: wenzelm@24800: * Pure: 'advanced' translation functions (parse_translation etc.) now wenzelm@24800: use Context.generic instead of just theory. wenzelm@24800: wenzelm@18642: * Pure: datatype Context.generic joins theory/Proof.context and wenzelm@18644: provides some facilities for code that works in either kind of wenzelm@18642: context, notably GenericDataFun for uniform theory and proof data. wenzelm@18642: wenzelm@18737: * Pure: simplified internal attribute type, which is now always wenzelm@24706: Context.generic * thm -> Context.generic * thm. Global (theory) vs. wenzelm@24706: local (Proof.context) attributes have been discontinued, while wenzelm@24706: minimizing code duplication. Thm.rule_attribute and wenzelm@24706: Thm.declaration_attribute build canonical attributes; see also structure wenzelm@24706: Context for further operations on Context.generic, notably wenzelm@24706: GenericDataFun. INCOMPATIBILITY, need to adapt attribute type wenzelm@19006: declarations and definitions. wenzelm@19006: wenzelm@24800: * Context data interfaces (Theory/Proof/GenericDataFun): removed wenzelm@24800: name/print, uninitialized data defaults to ad-hoc copy of empty value, wenzelm@24800: init only required for impure data. INCOMPATIBILITY: empty really need wenzelm@24800: to be empty (no dependencies on theory content!) wenzelm@24800: wenzelm@19508: * Pure/kernel: consts certification ignores sort constraints given in wenzelm@24800: signature declarations. (This information is not relevant to the wenzelm@24800: logic, but only for type inference.) SIGNIFICANT INTERNAL CHANGE, wenzelm@24800: potential INCOMPATIBILITY. wenzelm@19508: wenzelm@19508: * Pure: axiomatic type classes are now purely definitional, with wenzelm@19508: explicit proofs of class axioms and super class relations performed wenzelm@24706: internally. See Pure/axclass.ML for the main internal interfaces -- wenzelm@36856: notably AxClass.define_class supercedes AxClass.add_axclass, and wenzelm@24706: AxClass.axiomatize_class/classrel/arity supersede wenzelm@19508: Sign.add_classes/classrel/arities. wenzelm@19508: wenzelm@19006: * Pure/Isar: Args/Attrib parsers operate on Context.generic -- wenzelm@19006: global/local versions on theory vs. Proof.context have been wenzelm@19006: discontinued; Attrib.syntax and Method.syntax have been adapted wenzelm@19006: accordingly. INCOMPATIBILITY, need to adapt parser expressions for wenzelm@19006: attributes, methods, etc. wenzelm@18642: wenzelm@18446: * Pure: several functions of signature "... -> theory -> theory * ..." wenzelm@18446: have been reoriented to "... -> theory -> ... * theory" in order to wenzelm@18446: allow natural usage in combination with the ||>, ||>>, |-> and wenzelm@18446: fold_map combinators. haftmann@18051: wenzelm@21647: * Pure: official theorem names (closed derivations) and additional wenzelm@21647: comments (tags) are now strictly separate. Name hints -- which are wenzelm@21647: maintained as tags -- may be attached any time without affecting the wenzelm@21647: derivation. wenzelm@21647: wenzelm@18020: * Pure: primitive rule lift_rule now takes goal cterm instead of an wenzelm@18145: actual goal state (thm). Use Thm.lift_rule (Thm.cprem_of st i) to wenzelm@18020: achieve the old behaviour. wenzelm@18020: wenzelm@18020: * Pure: the "Goal" constant is now called "prop", supporting a wenzelm@18020: slightly more general idea of ``protecting'' meta-level rule wenzelm@18020: statements. wenzelm@18020: wenzelm@20040: * Pure: Logic.(un)varify only works in a global context, which is now wenzelm@20040: enforced instead of silently assumed. INCOMPATIBILITY, may use wenzelm@20040: Logic.legacy_(un)varify as temporary workaround. wenzelm@20040: wenzelm@20090: * Pure: structure Name provides scalable operations for generating wenzelm@20090: internal variable names, notably Name.variants etc. This replaces wenzelm@20090: some popular functions from term.ML: wenzelm@20090: wenzelm@20090: Term.variant -> Name.variant wenzelm@24800: Term.variantlist -> Name.variant_list wenzelm@20090: Term.invent_names -> Name.invent_list wenzelm@20090: wenzelm@20090: Note that low-level renaming rarely occurs in new code -- operations wenzelm@20090: from structure Variable are used instead (see below). wenzelm@20090: wenzelm@20040: * Pure: structure Variable provides fundamental operations for proper wenzelm@20040: treatment of fixed/schematic variables in a context. For example, wenzelm@20040: Variable.import introduces fixes for schematics of given facts and wenzelm@20040: Variable.export reverses the effect (up to renaming) -- this replaces wenzelm@20040: various freeze_thaw operations. wenzelm@20040: wenzelm@18567: * Pure: structure Goal provides simple interfaces for wenzelm@17981: init/conclude/finish and tactical prove operations (replacing former wenzelm@20040: Tactic.prove). Goal.prove is the canonical way to prove results wenzelm@20040: within a given context; Goal.prove_global is a degraded version for wenzelm@20040: theory level goals, including a global Drule.standard. Note that wenzelm@20040: OldGoals.prove_goalw_cterm has long been obsolete, since it is wenzelm@20040: ill-behaved in a local proof context (e.g. with local fixes/assumes or wenzelm@20040: in a locale context). wenzelm@17981: wenzelm@24706: * Pure/Syntax: generic interfaces for parsing (Syntax.parse_term etc.) wenzelm@24706: and type checking (Syntax.check_term etc.), with common combinations wenzelm@24706: (Syntax.read_term etc.). These supersede former Sign.read_term etc. wenzelm@24706: which are considered legacy and await removal. wenzelm@24706: wenzelm@24920: * Pure/Syntax: generic interfaces for type unchecking wenzelm@24920: (Syntax.uncheck_terms etc.) and unparsing (Syntax.unparse_term etc.), wenzelm@24920: with common combinations (Syntax.pretty_term, Syntax.string_of_term wenzelm@24920: etc.). Former Sign.pretty_term, Sign.string_of_term etc. are still wenzelm@24924: available for convenience, but refer to the very same operations using wenzelm@24924: a mere theory instead of a full context. wenzelm@24920: wenzelm@18815: * Isar: simplified treatment of user-level errors, using exception wenzelm@18687: ERROR of string uniformly. Function error now merely raises ERROR, wenzelm@18686: without any side effect on output channels. The Isar toplevel takes wenzelm@18686: care of proper display of ERROR exceptions. ML code may use plain wenzelm@18686: handle/can/try; cat_error may be used to concatenate errors like this: wenzelm@18686: wenzelm@18686: ... handle ERROR msg => cat_error msg "..." wenzelm@18686: wenzelm@18686: Toplevel ML code (run directly or through the Isar toplevel) may be wenzelm@18687: embedded into the Isar toplevel with exception display/debug like wenzelm@18687: this: wenzelm@18686: wenzelm@18686: Isar.toplevel (fn () => ...) wenzelm@18686: wenzelm@18686: INCOMPATIBILITY, removed special transform_error facilities, removed wenzelm@18686: obsolete variants of user-level exceptions (ERROR_MESSAGE, wenzelm@18686: Context.PROOF, ProofContext.CONTEXT, Proof.STATE, ProofHistory.FAIL) wenzelm@18686: -- use plain ERROR instead. wenzelm@18686: wenzelm@18815: * Isar: theory setup now has type (theory -> theory), instead of a wenzelm@18722: list. INCOMPATIBILITY, may use #> to compose setup functions. wenzelm@18722: wenzelm@24706: * Isar: ML toplevel pretty printer for type Proof.context, subject to wenzelm@24706: ProofContext.debug/verbose flags. wenzelm@18815: wenzelm@18815: * Isar: Toplevel.theory_to_proof admits transactions that modify the wenzelm@18815: theory before entering a proof state. Transactions now always see a wenzelm@18815: quasi-functional intermediate checkpoint, both in interactive and wenzelm@18590: batch mode. wenzelm@18567: wenzelm@24867: * Isar: simplified interfaces for outer syntax. Renamed wenzelm@24867: OuterSyntax.add_keywords to OuterSyntax.keywords. Removed wenzelm@24867: OuterSyntax.add_parsers -- this functionality is now included in wenzelm@24867: OuterSyntax.command etc. INCOMPATIBILITY. wenzelm@24867: wenzelm@17878: * Simplifier: the simpset of a running simplification process now wenzelm@17878: contains a proof context (cf. Simplifier.the_context), which is the wenzelm@17878: very context that the initial simpset has been retrieved from (by wenzelm@17890: simpset_of/local_simpset_of). Consequently, all plug-in components wenzelm@17878: (solver, looper etc.) may depend on arbitrary proof data. wenzelm@17878: wenzelm@17878: * Simplifier.inherit_context inherits the proof context (plus the wenzelm@17878: local bounds) of the current simplification process; any simproc wenzelm@17878: etc. that calls the Simplifier recursively should do this! Removed wenzelm@17878: former Simplifier.inherit_bounds, which is already included here -- wenzelm@17890: INCOMPATIBILITY. Tools based on low-level rewriting may even have to wenzelm@17890: specify an explicit context using Simplifier.context/theory_context. wenzelm@17878: wenzelm@17878: * Simplifier/Classical Reasoner: more abstract interfaces wenzelm@17878: change_simpset/claset for modifying the simpset/claset reference of a wenzelm@17878: theory; raw versions simpset/claset_ref etc. have been discontinued -- wenzelm@17878: INCOMPATIBILITY. wenzelm@17878: wenzelm@18540: * Provers: more generic wrt. syntax of object-logics, avoid hardwired wenzelm@18540: "Trueprop" etc. wenzelm@18540: wenzelm@17878: wenzelm@20988: *** System *** wenzelm@20988: wenzelm@25433: * settings: the default heap location within ISABELLE_HOME_USER now wenzelm@25433: includes ISABELLE_IDENTIFIER. This simplifies use of multiple wenzelm@25433: Isabelle installations. wenzelm@21471: wenzelm@20988: * isabelle-process: option -S (secure mode) disables some critical wenzelm@20988: operations, notably runtime compilation and evaluation of ML source wenzelm@20988: code. wenzelm@20988: wenzelm@24891: * Basic Isabelle mode for jEdit, see Isabelle/lib/jedit/. wenzelm@24891: wenzelm@24801: * Support for parallel execution, using native multicore support of wenzelm@24800: Poly/ML 5.1. The theory loader exploits parallelism when processing wenzelm@24800: independent theories, according to the given theory header wenzelm@24800: specifications. The maximum number of worker threads is specified via wenzelm@24800: usedir option -M or the "max-threads" setting in Proof General. A wenzelm@24800: speedup factor of 1.5--3.5 can be expected on a 4-core machine, and up wenzelm@24800: to 6 on a 8-core machine. User-code needs to observe certain wenzelm@24800: guidelines for thread-safe programming, see appendix A in the Isar wenzelm@24800: Implementation manual. wenzelm@24210: wenzelm@17754: wenzelm@25448: wenzelm@17720: New in Isabelle2005 (October 2005) wenzelm@17720: ---------------------------------- wenzelm@14655: wenzelm@14655: *** General *** wenzelm@14655: nipkow@15130: * Theory headers: the new header syntax for Isar theories is nipkow@15130: nipkow@15130: theory wenzelm@16234: imports ... wenzelm@16234: uses ... nipkow@15130: begin nipkow@15130: wenzelm@16234: where the 'uses' part is optional. The previous syntax wenzelm@16234: wenzelm@16234: theory = + ... + : wenzelm@16234: wenzelm@16717: will disappear in the next release. Use isatool fixheaders to convert wenzelm@16717: existing theory files. Note that there is no change in ancient wenzelm@17371: non-Isar theories now, but these will disappear soon. nipkow@15130: berghofe@15475: * Theory loader: parent theories can now also be referred to via wenzelm@16234: relative and absolute paths. wenzelm@16234: wenzelm@17408: * Command 'find_theorems' searches for a list of criteria instead of a wenzelm@17408: list of constants. Known criteria are: intro, elim, dest, name:string, wenzelm@17408: simp:term, and any term. Criteria can be preceded by '-' to select wenzelm@17408: theorems that do not match. Intro, elim, dest select theorems that wenzelm@17408: match the current goal, name:s selects theorems whose fully qualified wenzelm@17408: name contain s, and simp:term selects all simplification rules whose wenzelm@17408: lhs match term. Any other term is interpreted as pattern and selects wenzelm@17408: all theorems matching the pattern. Available in ProofGeneral under wenzelm@17408: 'ProofGeneral -> Find Theorems' or C-c C-f. Example: wenzelm@16234: wenzelm@17275: C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL." wenzelm@16234: wenzelm@16234: prints the last 100 theorems matching the pattern "(_::nat) + _ + _", wenzelm@16234: matching the current goal as introduction rule and not having "HOL." wenzelm@16234: in their name (i.e. not being defined in theory HOL). wenzelm@16013: wenzelm@17408: * Command 'thms_containing' has been discontinued in favour of wenzelm@17408: 'find_theorems'; INCOMPATIBILITY. wenzelm@17408: wenzelm@17385: * Communication with Proof General is now 8bit clean, which means that wenzelm@17385: Unicode text in UTF-8 encoding may be used within theory texts (both wenzelm@17408: formal and informal parts). Cf. option -U of the Isabelle Proof wenzelm@17538: General interface. Here are some simple examples (cf. src/HOL/ex): wenzelm@17538: wenzelm@17538: http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html wenzelm@17538: http://isabelle.in.tum.de/library/HOL/ex/Chinese.html wenzelm@17385: wenzelm@17425: * Improved efficiency of the Simplifier and, to a lesser degree, the wenzelm@17425: Classical Reasoner. Typical big applications run around 2 times wenzelm@17425: faster. wenzelm@17425: wenzelm@15703: wenzelm@15703: *** Document preparation *** wenzelm@15703: wenzelm@16234: * Commands 'display_drafts' and 'print_drafts' perform simple output wenzelm@16234: of raw sources. Only those symbols that do not require additional wenzelm@16234: LaTeX packages (depending on comments in isabellesym.sty) are wenzelm@16234: displayed properly, everything else is left verbatim. isatool display wenzelm@16234: and isatool print are used as front ends (these are subject to the wenzelm@16234: DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively). wenzelm@16234: wenzelm@17047: * Command tags control specific markup of certain regions of text, wenzelm@17047: notably folding and hiding. Predefined tags include "theory" (for wenzelm@17047: theory begin and end), "proof" for proof commands, and "ML" for wenzelm@17047: commands involving ML code; the additional tags "visible" and wenzelm@17047: "invisible" are unused by default. Users may give explicit tag wenzelm@17047: specifications in the text, e.g. ''by %invisible (auto)''. The wenzelm@17047: interpretation of tags is determined by the LaTeX job during document wenzelm@17047: preparation: see option -V of isatool usedir, or options -n and -t of wenzelm@17047: isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag, wenzelm@17047: \isadroptag. wenzelm@17047: wenzelm@17047: Several document versions may be produced at the same time via isatool wenzelm@17047: usedir (the generated index.html will link all of them). Typical wenzelm@17047: specifications include ''-V document=theory,proof,ML'' to present wenzelm@17047: theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold wenzelm@17047: proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit wenzelm@17047: these parts without any formal replacement text. The Isabelle site wenzelm@17047: default settings produce ''document'' and ''outline'' versions as wenzelm@17047: specified above. wenzelm@16234: haftmann@17402: * Several new antiquotations: wenzelm@15979: wenzelm@15979: @{term_type term} prints a term with its type annotated; wenzelm@15979: wenzelm@15979: @{typeof term} prints the type of a term; wenzelm@15979: wenzelm@16234: @{const const} is the same as @{term const}, but checks that the wenzelm@16234: argument is a known logical constant; wenzelm@15979: wenzelm@15979: @{term_style style term} and @{thm_style style thm} print a term or wenzelm@16234: theorem applying a "style" to it wenzelm@16234: wenzelm@17117: @{ML text} wenzelm@17117: wenzelm@16234: Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of wenzelm@16234: definitions, equations, inequations etc., 'concl' printing only the schirmer@17393: conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19' wenzelm@16234: to print the specified premise. TermStyle.add_style provides an ML wenzelm@16234: interface for introducing further styles. See also the "LaTeX Sugar" wenzelm@17117: document practical applications. The ML antiquotation prints wenzelm@17117: type-checked ML expressions verbatim. wenzelm@16234: wenzelm@17259: * Markup commands 'chapter', 'section', 'subsection', 'subsubsection', wenzelm@17259: and 'text' support optional locale specification '(in loc)', which wenzelm@17269: specifies the default context for interpreting antiquotations. For wenzelm@17269: example: 'text (in lattice) {* @{thm inf_assoc}*}'. wenzelm@17259: wenzelm@17259: * Option 'locale=NAME' of antiquotations specifies an alternative wenzelm@17259: context interpreting the subsequent argument. For example: @{thm wenzelm@17269: [locale=lattice] inf_assoc}. wenzelm@17259: wenzelm@17097: * Proper output of proof terms (@{prf ...} and @{full_prf ...}) within wenzelm@17097: a proof context. wenzelm@17097: wenzelm@17097: * Proper output of antiquotations for theory commands involving a wenzelm@17097: proof context (such as 'locale' or 'theorem (in loc) ...'). wenzelm@17097: wenzelm@17193: * Delimiters of outer tokens (string etc.) now produce separate LaTeX wenzelm@17193: macros (\isachardoublequoteopen, isachardoublequoteclose etc.). wenzelm@17193: wenzelm@17193: * isatool usedir: new option -C (default true) controls whether option wenzelm@17193: -D should include a copy of the original document directory; -C false wenzelm@17193: prevents unwanted effects such as copying of administrative CVS data. wenzelm@17193: wenzelm@16234: wenzelm@16234: *** Pure *** wenzelm@16234: wenzelm@16234: * Considerably improved version of 'constdefs' command. Now performs wenzelm@16234: automatic type-inference of declared constants; additional support for wenzelm@16234: local structure declarations (cf. locales and HOL records), see also wenzelm@16234: isar-ref manual. Potential INCOMPATIBILITY: need to observe strictly wenzelm@16234: sequential dependencies of definitions within a single 'constdefs' wenzelm@16234: section; moreover, the declared name needs to be an identifier. If wenzelm@16234: all fails, consider to fall back on 'consts' and 'defs' separately. wenzelm@16234: wenzelm@16234: * Improved indexed syntax and implicit structures. First of all, wenzelm@16234: indexed syntax provides a notational device for subscripted wenzelm@16234: application, using the new syntax \<^bsub>term\<^esub> for arbitrary wenzelm@16234: expressions. Secondly, in a local context with structure wenzelm@16234: declarations, number indexes \<^sub>n or the empty index (default wenzelm@16234: number 1) refer to a certain fixed variable implicitly; option wenzelm@16234: show_structs controls printing of implicit structures. Typical wenzelm@16234: applications of these concepts involve record types and locales. wenzelm@16234: wenzelm@16234: * New command 'no_syntax' removes grammar declarations (and wenzelm@16234: translations) resulting from the given syntax specification, which is wenzelm@16234: interpreted in the same manner as for the 'syntax' command. wenzelm@16234: wenzelm@16234: * 'Advanced' translation functions (parse_translation etc.) may depend wenzelm@16234: on the signature of the theory context being presently used for wenzelm@16234: parsing/printing, see also isar-ref manual. wenzelm@16234: wenzelm@16856: * Improved 'oracle' command provides a type-safe interface to turn an wenzelm@16856: ML expression of type theory -> T -> term into a primitive rule of wenzelm@16856: type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle wenzelm@16856: is already included here); see also FOL/ex/IffExample.thy; wenzelm@16856: INCOMPATIBILITY. wenzelm@16856: wenzelm@17275: * axclass: name space prefix for class "c" is now "c_class" (was "c" wenzelm@17275: before); "cI" is no longer bound, use "c.intro" instead. wenzelm@17275: INCOMPATIBILITY. This change avoids clashes of fact bindings for wenzelm@17275: axclasses vs. locales. wenzelm@17275: wenzelm@16234: * Improved internal renaming of symbolic identifiers -- attach primes wenzelm@16234: instead of base 26 numbers. wenzelm@16234: wenzelm@16234: * New flag show_question_marks controls printing of leading question wenzelm@16234: marks in schematic variable names. wenzelm@16234: wenzelm@16234: * In schematic variable names, *any* symbol following \<^isub> or wenzelm@16234: \<^isup> is now treated as part of the base name. For example, the wenzelm@16234: following works without printing of awkward ".0" indexes: wenzelm@16234: wenzelm@16234: lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1" wenzelm@16234: by simp wenzelm@16234: wenzelm@16234: * Inner syntax includes (*(*nested*) comments*). wenzelm@16234: wenzelm@17548: * Pretty printer now supports unbreakable blocks, specified in mixfix wenzelm@16234: annotations as "(00...)". wenzelm@16234: wenzelm@16234: * Clear separation of logical types and nonterminals, where the latter wenzelm@16234: may only occur in 'syntax' specifications or type abbreviations. wenzelm@16234: Before that distinction was only partially implemented via type class wenzelm@16234: "logic" vs. "{}". Potential INCOMPATIBILITY in rare cases of improper wenzelm@16234: use of 'types'/'consts' instead of 'nonterminals'/'syntax'. Some very wenzelm@16234: exotic syntax specifications may require further adaption wenzelm@17691: (e.g. Cube/Cube.thy). wenzelm@16234: wenzelm@16234: * Removed obsolete type class "logic", use the top sort {} instead. wenzelm@16234: Note that non-logical types should be declared as 'nonterminals' wenzelm@16234: rather than 'types'. INCOMPATIBILITY for new object-logic wenzelm@16234: specifications. wenzelm@16234: ballarin@17095: * Attributes 'induct' and 'cases': type or set names may now be ballarin@17095: locally fixed variables as well. ballarin@17095: wenzelm@16234: * Simplifier: can now control the depth to which conditional rewriting wenzelm@16234: is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth wenzelm@16234: Limit. wenzelm@16234: wenzelm@16234: * Simplifier: simplification procedures may now take the current wenzelm@16234: simpset into account (cf. Simplifier.simproc(_i) / mk_simproc wenzelm@16234: interface), which is very useful for calling the Simplifier wenzelm@16234: recursively. Minor INCOMPATIBILITY: the 'prems' argument of simprocs wenzelm@16234: is gone -- use prems_of_ss on the simpset instead. Moreover, the wenzelm@16234: low-level mk_simproc no longer applies Logic.varify internally, to wenzelm@16234: allow for use in a context of fixed variables. wenzelm@16234: wenzelm@16234: * thin_tac now works even if the assumption being deleted contains !! wenzelm@16234: or ==>. More generally, erule now works even if the major premise of wenzelm@16234: the elimination rule contains !! or ==>. wenzelm@16234: wenzelm@17597: * Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY. nipkow@17590: wenzelm@16234: * Reorganized bootstrapping of the Pure theories; CPure is now derived wenzelm@16234: from Pure, which contains all common declarations already. Both wenzelm@16234: theories are defined via plain Isabelle/Isar .thy files. wenzelm@16234: INCOMPATIBILITY: elements of CPure (such as the CPure.intro / wenzelm@16234: CPure.elim / CPure.dest attributes) now appear in the Pure name space; wenzelm@16234: use isatool fixcpure to adapt your theory and ML sources. wenzelm@16234: wenzelm@16234: * New syntax 'name(i-j, i-, i, ...)' for referring to specific wenzelm@16234: selections of theorems in named facts via index ranges. wenzelm@16234: wenzelm@17097: * 'print_theorems': in theory mode, really print the difference wenzelm@17097: wrt. the last state (works for interactive theory development only), wenzelm@17097: in proof mode print all local facts (cf. 'print_facts'); wenzelm@17097: wenzelm@17397: * 'hide': option '(open)' hides only base names. wenzelm@17397: wenzelm@17275: * More efficient treatment of intermediate checkpoints in interactive wenzelm@17275: theory development. wenzelm@17275: berghofe@17663: * Code generator is now invoked via code_module (incremental code wenzelm@17664: generation) and code_library (modular code generation, ML structures wenzelm@17664: for each theory). INCOMPATIBILITY: new keywords 'file' and 'contains' wenzelm@17664: must be quoted when used as identifiers. wenzelm@17664: wenzelm@17664: * New 'value' command for reading, evaluating and printing terms using wenzelm@17664: the code generator. INCOMPATIBILITY: command keyword 'value' must be wenzelm@17664: quoted when used as identifier. berghofe@17663: wenzelm@16234: wenzelm@16234: *** Locales *** ballarin@17095: wenzelm@17385: * New commands for the interpretation of locale expressions in wenzelm@17385: theories (1), locales (2) and proof contexts (3). These generate wenzelm@17385: proof obligations from the expression specification. After the wenzelm@17385: obligations have been discharged, theorems of the expression are added wenzelm@17385: to the theory, target locale or proof context. The synopsis of the wenzelm@17385: commands is a follows: wenzelm@17385: ballarin@17095: (1) interpretation expr inst ballarin@17095: (2) interpretation target < expr ballarin@17095: (3) interpret expr inst wenzelm@17385: ballarin@17095: Interpretation in theories and proof contexts require a parameter ballarin@17095: instantiation of terms from the current context. This is applied to wenzelm@17385: specifications and theorems of the interpreted expression. wenzelm@17385: Interpretation in locales only permits parameter renaming through the wenzelm@17385: locale expression. Interpretation is smart in that interpretations wenzelm@17385: that are active already do not occur in proof obligations, neither are wenzelm@17385: instantiated theorems stored in duplicate. Use 'print_interps' to wenzelm@17385: inspect active interpretations of a particular locale. For details, ballarin@17436: see the Isar Reference manual. Examples can be found in ballarin@17436: HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy. wenzelm@16234: wenzelm@16234: INCOMPATIBILITY: former 'instantiate' has been withdrawn, use wenzelm@16234: 'interpret' instead. wenzelm@16234: wenzelm@17385: * New context element 'constrains' for adding type constraints to wenzelm@17385: parameters. wenzelm@17385: wenzelm@17385: * Context expressions: renaming of parameters with syntax wenzelm@17385: redeclaration. ballarin@17095: ballarin@17095: * Locale declaration: 'includes' disallowed. ballarin@17095: wenzelm@16234: * Proper static binding of attribute syntax -- i.e. types / terms / wenzelm@16234: facts mentioned as arguments are always those of the locale definition wenzelm@16234: context, independently of the context of later invocations. Moreover, wenzelm@16234: locale operations (renaming and type / term instantiation) are applied wenzelm@16234: to attribute arguments as expected. wenzelm@16234: wenzelm@16234: INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of wenzelm@16234: actual attributes; rare situations may require Attrib.attribute to wenzelm@16234: embed those attributes into Attrib.src that lack concrete syntax. wenzelm@16234: Attribute implementations need to cooperate properly with the static wenzelm@16234: binding mechanism. Basic parsers Args.XXX_typ/term/prop and wenzelm@16234: Attrib.XXX_thm etc. already do the right thing without further wenzelm@16234: intervention. Only unusual applications -- such as "where" or "of" wenzelm@16234: (cf. src/Pure/Isar/attrib.ML), which process arguments depending both wenzelm@16234: on the context and the facts involved -- may have to assign parsed wenzelm@16234: values to argument tokens explicitly. wenzelm@16234: wenzelm@16234: * Changed parameter management in theorem generation for long goal wenzelm@16234: statements with 'includes'. INCOMPATIBILITY: produces a different wenzelm@16234: theorem statement in rare situations. wenzelm@16234: ballarin@17228: * Locale inspection command 'print_locale' omits notes elements. Use ballarin@17228: 'print_locale!' to have them included in the output. ballarin@17228: wenzelm@16234: wenzelm@16234: *** Provers *** wenzelm@16234: wenzelm@16234: * Provers/hypsubst.ML: improved version of the subst method, for wenzelm@16234: single-step rewriting: it now works in bound variable contexts. New is wenzelm@16234: 'subst (asm)', for rewriting an assumption. INCOMPATIBILITY: may wenzelm@16234: rewrite a different subterm than the original subst method, which is wenzelm@16234: still available as 'simplesubst'. wenzelm@16234: wenzelm@16234: * Provers/quasi.ML: new transitivity reasoners for transitivity only wenzelm@16234: and quasi orders. wenzelm@16234: wenzelm@16234: * Provers/trancl.ML: new transitivity reasoner for transitive and wenzelm@16234: reflexive-transitive closure of relations. wenzelm@16234: wenzelm@16234: * Provers/blast.ML: new reference depth_limit to make blast's depth wenzelm@16234: limit (previously hard-coded with a value of 20) user-definable. wenzelm@16234: wenzelm@16234: * Provers/simplifier.ML has been moved to Pure, where Simplifier.setup wenzelm@16234: is peformed already. Object-logics merely need to finish their wenzelm@16234: initial simpset configuration as before. INCOMPATIBILITY. wenzelm@15703: berghofe@15475: schirmer@14700: *** HOL *** schirmer@14700: wenzelm@16234: * Symbolic syntax of Hilbert Choice Operator is now as follows: wenzelm@14878: wenzelm@14878: syntax (epsilon) wenzelm@14878: "_Eps" :: "[pttrn, bool] => 'a" ("(3\_./ _)" [0, 10] 10) wenzelm@14878: wenzelm@16234: The symbol \ is displayed as the alternative epsilon of LaTeX wenzelm@16234: and x-symbol; use option '-m epsilon' to get it actually printed. wenzelm@16234: Moreover, the mathematically important symbolic identifier \ wenzelm@16234: becomes available as variable, constant etc. INCOMPATIBILITY, wenzelm@16234: wenzelm@16234: * "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x". wenzelm@16234: Similarly for all quantifiers: "ALL x > y" etc. The x-symbol for >= wenzelm@17371: is \. New transitivity rules have been added to HOL/Orderings.thy to avigad@17016: support corresponding Isar calculations. wenzelm@16234: wenzelm@16234: * "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\" wenzelm@16234: instead of ":". wenzelm@16234: wenzelm@16234: * theory SetInterval: changed the syntax for open intervals: wenzelm@16234: wenzelm@16234: Old New wenzelm@16234: {..n(} {.. {\1<\.\.} nipkow@15046: \.\.\([^(}]*\)(} -> \.\.<\1} nipkow@15046: wenzelm@17533: * Theory Commutative_Ring (in Library): method comm_ring for proving wenzelm@17533: equalities in commutative rings; method 'algebra' provides a generic wenzelm@17533: interface. wenzelm@17389: wenzelm@17389: * Theory Finite_Set: changed the syntax for 'setsum', summation over wenzelm@16234: finite sets: "setsum (%x. e) A", which used to be "\x:A. e", is wenzelm@17371: now either "SUM x:A. e" or "\x \ A. e". The bound variable can paulson@17189: be a tuple pattern. wenzelm@16234: wenzelm@16234: Some new syntax forms are available: wenzelm@16234: wenzelm@16234: "\x | P. e" for "setsum (%x. e) {x. P}" wenzelm@16234: "\x = a..b. e" for "setsum (%x. e) {a..b}" wenzelm@16234: "\x = a..x < k. e" for "setsum (%x. e) {..x < k. e" used to be based on a separate wenzelm@16234: function "Summation", which has been discontinued. wenzelm@16234: wenzelm@16234: * theory Finite_Set: in structured induction proofs, the insert case wenzelm@16234: is now 'case (insert x F)' instead of the old counterintuitive 'case wenzelm@16234: (insert F x)'. wenzelm@16234: wenzelm@16234: * The 'refute' command has been extended to support a much larger wenzelm@16234: fragment of HOL, including axiomatic type classes, constdefs and wenzelm@16234: typedefs, inductive datatypes and recursion. wenzelm@16234: webertj@17700: * New tactics 'sat' and 'satx' to prove propositional tautologies. webertj@17700: Requires zChaff with proof generation to be installed. See webertj@17700: HOL/ex/SAT_Examples.thy for examples. webertj@17619: wenzelm@16234: * Datatype induction via method 'induct' now preserves the name of the wenzelm@16234: induction variable. For example, when proving P(xs::'a list) by wenzelm@16234: induction on xs, the induction step is now P(xs) ==> P(a#xs) rather wenzelm@16234: than P(list) ==> P(a#list) as previously. Potential INCOMPATIBILITY wenzelm@16234: in unstructured proof scripts. wenzelm@16234: wenzelm@16234: * Reworked implementation of records. Improved scalability for wenzelm@16234: records with many fields, avoiding performance problems for type wenzelm@16234: inference. Records are no longer composed of nested field types, but wenzelm@16234: of nested extension types. Therefore the record type only grows linear wenzelm@16234: in the number of extensions and not in the number of fields. The wenzelm@16234: top-level (users) view on records is preserved. Potential wenzelm@16234: INCOMPATIBILITY only in strange cases, where the theory depends on the wenzelm@16234: old record representation. The type generated for a record is called wenzelm@16234: _ext_type. wenzelm@16234: wenzelm@16234: Flag record_quick_and_dirty_sensitive can be enabled to skip the wenzelm@16234: proofs triggered by a record definition or a simproc (if wenzelm@16234: quick_and_dirty is enabled). Definitions of large records can take wenzelm@16234: quite long. wenzelm@16234: wenzelm@16234: New simproc record_upd_simproc for simplification of multiple record wenzelm@16234: updates enabled by default. Moreover, trivial updates are also wenzelm@16234: removed: r(|x := x r|) = r. INCOMPATIBILITY: old proofs break wenzelm@16234: occasionally, since simplification is more powerful by default. wenzelm@16234: wenzelm@17275: * typedef: proper support for polymorphic sets, which contain extra wenzelm@17275: type-variables in the term. wenzelm@17275: wenzelm@16234: * Simplifier: automatically reasons about transitivity chains wenzelm@16234: involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics wenzelm@16234: provided by Provers/trancl.ML as additional solvers. INCOMPATIBILITY: wenzelm@16234: old proofs break occasionally as simplification may now solve more wenzelm@16234: goals than previously. wenzelm@16234: wenzelm@16234: * Simplifier: converts x <= y into x = y if assumption y <= x is wenzelm@16234: present. Works for all partial orders (class "order"), in particular wenzelm@16234: numbers and sets. For linear orders (e.g. numbers) it treats ~ x < y wenzelm@16234: just like y <= x. wenzelm@16234: wenzelm@16234: * Simplifier: new simproc for "let x = a in f x". If a is a free or wenzelm@16234: bound variable or a constant then the let is unfolded. Otherwise wenzelm@16234: first a is simplified to b, and then f b is simplified to g. If wenzelm@16234: possible we abstract b from g arriving at "let x = b in h x", wenzelm@16234: otherwise we unfold the let and arrive at g. The simproc can be wenzelm@16234: enabled/disabled by the reference use_let_simproc. Potential wenzelm@16234: INCOMPATIBILITY since simplification is more powerful by default. webertj@15776: paulson@16563: * Classical reasoning: the meson method now accepts theorems as arguments. paulson@16563: paulson@17595: * Prover support: pre-release of the Isabelle-ATP linkup, which runs background paulson@17595: jobs to provide advice on the provability of subgoals. paulson@17595: wenzelm@16891: * Theory OrderedGroup and Ring_and_Field: various additions and wenzelm@16891: improvements to faciliate calculations involving equalities and wenzelm@16891: inequalities. wenzelm@16891: wenzelm@16891: The following theorems have been eliminated or modified wenzelm@16891: (INCOMPATIBILITY): avigad@16888: avigad@16888: abs_eq now named abs_of_nonneg wenzelm@17371: abs_of_ge_0 now named abs_of_nonneg wenzelm@17371: abs_minus_eq now named abs_of_nonpos avigad@16888: imp_abs_id now named abs_of_nonneg avigad@16888: imp_abs_neg_id now named abs_of_nonpos avigad@16888: mult_pos now named mult_pos_pos avigad@16888: mult_pos_le now named mult_nonneg_nonneg avigad@16888: mult_pos_neg_le now named mult_nonneg_nonpos avigad@16888: mult_pos_neg2_le now named mult_nonneg_nonpos2 avigad@16888: mult_neg now named mult_neg_neg avigad@16888: mult_neg_le now named mult_nonpos_nonpos avigad@16888: obua@23495: * The following lemmas in Ring_and_Field have been added to the simplifier: lp15@57253: obua@23495: zero_le_square lp15@57253: not_square_less_zero obua@23495: obua@23495: The following lemmas have been deleted from Real/RealPow: lp15@57253: obua@23495: realpow_zero_zero obua@23495: realpow_two obua@23495: realpow_less obua@23495: zero_le_power obua@23495: realpow_two_le obua@23495: abs_realpow_two lp15@57253: realpow_two_abs obua@23495: wenzelm@16891: * Theory Parity: added rules for simplifying exponents. wenzelm@16891: nipkow@17092: * Theory List: nipkow@17092: nipkow@17092: The following theorems have been eliminated or modified nipkow@17092: (INCOMPATIBILITY): nipkow@17092: nipkow@17092: list_all_Nil now named list_all.simps(1) nipkow@17092: list_all_Cons now named list_all.simps(2) nipkow@17092: list_all_conv now named list_all_iff nipkow@17092: set_mem_eq now named mem_iff nipkow@17092: wenzelm@16929: * Theories SetsAndFunctions and BigO (see HOL/Library) support wenzelm@16929: asymptotic "big O" calculations. See the notes in BigO.thy. wenzelm@16929: avigad@16888: avigad@16888: *** HOL-Complex *** avigad@16888: wenzelm@16891: * Theory RealDef: better support for embedding natural numbers and wenzelm@16891: integers in the reals. wenzelm@16891: wenzelm@16891: The following theorems have been eliminated or modified wenzelm@16891: (INCOMPATIBILITY): wenzelm@16891: avigad@17016: exp_ge_add_one_self now requires no hypotheses avigad@17016: real_of_int_add reversed direction of equality (use [symmetric]) avigad@17016: real_of_int_minus reversed direction of equality (use [symmetric]) avigad@17016: real_of_int_diff reversed direction of equality (use [symmetric]) avigad@17016: real_of_int_mult reversed direction of equality (use [symmetric]) wenzelm@16891: wenzelm@16891: * Theory RComplete: expanded support for floor and ceiling functions. avigad@16888: avigad@16962: * Theory Ln is new, with properties of the natural logarithm avigad@16962: wenzelm@17423: * Hyperreal: There is a new type constructor "star" for making wenzelm@17423: nonstandard types. The old type names are now type synonyms: wenzelm@17423: wenzelm@17423: hypreal = real star wenzelm@17423: hypnat = nat star wenzelm@17423: hcomplex = complex star wenzelm@17423: wenzelm@17423: * Hyperreal: Many groups of similarly-defined constants have been huffman@17442: replaced by polymorphic versions (INCOMPATIBILITY): wenzelm@17423: wenzelm@17423: star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex wenzelm@17423: wenzelm@17423: starset <-- starsetNat, starsetC wenzelm@17423: *s* <-- *sNat*, *sc* wenzelm@17423: starset_n <-- starsetNat_n, starsetC_n wenzelm@17423: *sn* <-- *sNatn*, *scn* wenzelm@17423: InternalSets <-- InternalNatSets, InternalCSets wenzelm@17423: huffman@17442: starfun <-- starfun{Nat,Nat2,C,RC,CR} wenzelm@17423: *f* <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR* huffman@17442: starfun_n <-- starfun{Nat,Nat2,C,RC,CR}_n wenzelm@17423: *fn* <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn* huffman@17442: InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs wenzelm@17423: wenzelm@17423: * Hyperreal: Many type-specific theorems have been removed in favor of huffman@17442: theorems specific to various axiomatic type classes (INCOMPATIBILITY): huffman@17442: huffman@17442: add_commute <-- {hypreal,hypnat,hcomplex}_add_commute huffman@17442: add_assoc <-- {hypreal,hypnat,hcomplex}_add_assocs huffman@17442: OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left huffman@17442: OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right wenzelm@17423: right_minus <-- hypreal_add_minus huffman@17442: left_minus <-- {hypreal,hcomplex}_add_minus_left huffman@17442: mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute huffman@17442: mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc huffman@17442: mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left wenzelm@17423: mult_1_right <-- hcomplex_mult_one_right wenzelm@17423: mult_zero_left <-- hcomplex_mult_zero_left huffman@17442: left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib wenzelm@17423: right_distrib <-- hypnat_add_mult_distrib2 huffman@17442: zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one wenzelm@17423: right_inverse <-- hypreal_mult_inverse wenzelm@17423: left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left huffman@17442: order_refl <-- {hypreal,hypnat}_le_refl huffman@17442: order_trans <-- {hypreal,hypnat}_le_trans huffman@17442: order_antisym <-- {hypreal,hypnat}_le_anti_sym huffman@17442: order_less_le <-- {hypreal,hypnat}_less_le huffman@17442: linorder_linear <-- {hypreal,hypnat}_le_linear huffman@17442: add_left_mono <-- {hypreal,hypnat}_add_left_mono huffman@17442: mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2 wenzelm@17423: add_nonneg_nonneg <-- hypreal_le_add_order wenzelm@17423: wenzelm@17423: * Hyperreal: Separate theorems having to do with type-specific wenzelm@17423: versions of constants have been merged into theorems that apply to the huffman@17442: new polymorphic constants (INCOMPATIBILITY): huffman@17442: huffman@17442: STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set huffman@17442: STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set huffman@17442: STAR_Un <-- {STAR,NatStar,STARC}_Un huffman@17442: STAR_Int <-- {STAR,NatStar,STARC}_Int huffman@17442: STAR_Compl <-- {STAR,NatStar,STARC}_Compl huffman@17442: STAR_subset <-- {STAR,NatStar,STARC}_subset huffman@17442: STAR_mem <-- {STAR,NatStar,STARC}_mem huffman@17442: STAR_mem_Compl <-- {STAR,STARC}_mem_Compl huffman@17442: STAR_diff <-- {STAR,STARC}_diff huffman@17442: STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real, huffman@17442: STARC_hcomplex_of_complex}_image_subset huffman@17442: starset_n_Un <-- starset{Nat,C}_n_Un huffman@17442: starset_n_Int <-- starset{Nat,C}_n_Int huffman@17442: starset_n_Compl <-- starset{Nat,C}_n_Compl huffman@17442: starset_n_diff <-- starset{Nat,C}_n_diff huffman@17442: InternalSets_Un <-- Internal{Nat,C}Sets_Un huffman@17442: InternalSets_Int <-- Internal{Nat,C}Sets_Int huffman@17442: InternalSets_Compl <-- Internal{Nat,C}Sets_Compl huffman@17442: InternalSets_diff <-- Internal{Nat,C}Sets_diff huffman@17442: InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff huffman@17442: InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n huffman@17442: starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq huffman@17442: starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C} huffman@17442: starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR} huffman@17442: starfun <-- starfun{Nat,Nat2,C,RC,CR} huffman@17442: starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult huffman@17442: starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add huffman@17442: starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus huffman@17442: starfun_diff <-- starfun{C,RC,CR}_diff huffman@17442: starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o huffman@17442: starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2 huffman@17442: starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun huffman@17442: starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse huffman@17442: starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq huffman@17442: starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff wenzelm@17423: starfun_Id <-- starfunC_Id huffman@17442: starfun_approx <-- starfun{Nat,CR}_approx huffman@17442: starfun_capprox <-- starfun{C,RC}_capprox wenzelm@17423: starfun_abs <-- starfunNat_rabs huffman@17442: starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel huffman@17442: starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2 wenzelm@17423: starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox huffman@17442: starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox huffman@17442: starfun_add_capprox <-- starfun{C,RC}_add_capprox wenzelm@17423: starfun_add_approx <-- starfunCR_add_approx wenzelm@17423: starfun_inverse_inverse <-- starfunC_inverse_inverse huffman@17442: starfun_divide <-- starfun{C,CR,RC}_divide huffman@17442: starfun_n <-- starfun{Nat,C}_n huffman@17442: starfun_n_mult <-- starfun{Nat,C}_n_mult huffman@17442: starfun_n_add <-- starfun{Nat,C}_n_add wenzelm@17423: starfun_n_add_minus <-- starfunNat_n_add_minus huffman@17442: starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun huffman@17442: starfun_n_minus <-- starfun{Nat,C}_n_minus huffman@17442: starfun_n_eq <-- starfun{Nat,C}_n_eq huffman@17442: huffman@17442: star_n_add <-- {hypreal,hypnat,hcomplex}_add huffman@17442: star_n_minus <-- {hypreal,hcomplex}_minus huffman@17442: star_n_diff <-- {hypreal,hcomplex}_diff huffman@17442: star_n_mult <-- {hypreal,hcomplex}_mult huffman@17442: star_n_inverse <-- {hypreal,hcomplex}_inverse huffman@17442: star_n_le <-- {hypreal,hypnat}_le huffman@17442: star_n_less <-- {hypreal,hypnat}_less huffman@17442: star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num huffman@17442: star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num wenzelm@17423: star_n_abs <-- hypreal_hrabs wenzelm@17423: star_n_divide <-- hcomplex_divide wenzelm@17423: huffman@17442: star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add huffman@17442: star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus wenzelm@17423: star_of_diff <-- hypreal_of_real_diff huffman@17442: star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult huffman@17442: star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one huffman@17442: star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero huffman@17442: star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff huffman@17442: star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff huffman@17442: star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff huffman@17442: star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse huffman@17442: star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide huffman@17442: star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat huffman@17442: star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int huffman@17442: star_of_number_of <-- {hypreal,hcomplex}_number_of wenzelm@17423: star_of_number_less <-- number_of_less_hypreal_of_real_iff wenzelm@17423: star_of_number_le <-- number_of_le_hypreal_of_real_iff wenzelm@17423: star_of_eq_number <-- hypreal_of_real_eq_number_of_iff wenzelm@17423: star_of_less_number <-- hypreal_of_real_less_number_of_iff wenzelm@17423: star_of_le_number <-- hypreal_of_real_le_number_of_iff wenzelm@17423: star_of_power <-- hypreal_of_real_power wenzelm@17423: star_of_eq_0 <-- hcomplex_of_complex_zero_iff wenzelm@17423: huffman@17442: * Hyperreal: new method "transfer" that implements the transfer huffman@17442: principle of nonstandard analysis. With a subgoal that mentions huffman@17442: nonstandard types like "'a star", the command "apply transfer" huffman@17442: replaces it with an equivalent one that mentions only standard types. huffman@17442: To be successful, all free variables must have standard types; non- huffman@17442: standard variables must have explicit universal quantifiers. huffman@17442: wenzelm@17641: * Hyperreal: A theory of Taylor series. wenzelm@17641: wenzelm@14655: wenzelm@14682: *** HOLCF *** wenzelm@14682: wenzelm@17533: * Discontinued special version of 'constdefs' (which used to support wenzelm@17533: continuous functions) in favor of the general Pure one with full wenzelm@17533: type-inference. wenzelm@17533: wenzelm@17533: * New simplification procedure for solving continuity conditions; it wenzelm@17533: is much faster on terms with many nested lambda abstractions (cubic huffman@17442: instead of exponential time). huffman@17442: wenzelm@17533: * New syntax for domain package: selector names are now optional. huffman@17442: Parentheses should be omitted unless argument is lazy, for example: huffman@17442: huffman@17442: domain 'a stream = cons "'a" (lazy "'a stream") huffman@17442: wenzelm@17533: * New command 'fixrec' for defining recursive functions with pattern wenzelm@17533: matching; defining multiple functions with mutual recursion is also wenzelm@17533: supported. Patterns may include the constants cpair, spair, up, sinl, wenzelm@17533: sinr, or any data constructor defined by the domain package. The given wenzelm@17533: equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for wenzelm@17533: syntax and examples. wenzelm@17533: wenzelm@17533: * New commands 'cpodef' and 'pcpodef' for defining predicate subtypes wenzelm@17533: of cpo and pcpo types. Syntax is exactly like the 'typedef' command, wenzelm@17533: but the proof obligation additionally includes an admissibility wenzelm@17533: requirement. The packages generate instances of class cpo or pcpo, wenzelm@17533: with continuity and strictness theorems for Rep and Abs. huffman@17442: huffman@17584: * HOLCF: Many theorems have been renamed according to a more standard naming huffman@17584: scheme (INCOMPATIBILITY): huffman@17584: huffman@17584: foo_inject: "foo$x = foo$y ==> x = y" huffman@17584: foo_eq: "(foo$x = foo$y) = (x = y)" huffman@17584: foo_less: "(foo$x << foo$y) = (x << y)" huffman@17584: foo_strict: "foo$UU = UU" huffman@17584: foo_defined: "... ==> foo$x ~= UU" huffman@17584: foo_defined_iff: "(foo$x = UU) = (x = UU)" huffman@17584: wenzelm@14682: paulson@14885: *** ZF *** paulson@14885: wenzelm@16234: * ZF/ex: theories Group and Ring provide examples in abstract algebra, wenzelm@16234: including the First Isomorphism Theorem (on quotienting by the kernel wenzelm@16234: of a homomorphism). wenzelm@15089: wenzelm@15089: * ZF/Simplifier: install second copy of type solver that actually wenzelm@16234: makes use of TC rules declared to Isar proof contexts (or locales); wenzelm@16234: the old version is still required for ML proof scripts. wenzelm@15703: wenzelm@15703: wenzelm@17445: *** Cube *** wenzelm@17445: wenzelm@17445: * Converted to Isar theory format; use locales instead of axiomatic wenzelm@17445: theories. wenzelm@17445: wenzelm@17445: wenzelm@15703: *** ML *** wenzelm@15703: haftmann@21339: * Pure/library.ML: added ##>, ##>>, #>> -- higher-order counterparts haftmann@21339: for ||>, ||>>, |>>, haftmann@21339: wenzelm@15973: * Pure/library.ML no longer defines its own option datatype, but uses wenzelm@16234: that of the SML basis, which has constructors NONE and SOME instead of wenzelm@16234: None and Some, as well as exception Option.Option instead of OPTION. wenzelm@16234: The functions the, if_none, is_some, is_none have been adapted wenzelm@16234: accordingly, while Option.map replaces apsome. wenzelm@15973: wenzelm@16860: * Pure/library.ML: the exception LIST has been given up in favour of wenzelm@16860: the standard exceptions Empty and Subscript, as well as wenzelm@16860: Library.UnequalLengths. Function like Library.hd and Library.tl are wenzelm@36856: superceded by the standard hd and tl functions etc. wenzelm@16860: wenzelm@16860: A number of basic list functions are no longer exported to the ML wenzelm@16860: toplevel, as they are variants of predefined functions. The following wenzelm@16234: suggests how one can translate existing code: wenzelm@15973: wenzelm@15973: rev_append xs ys = List.revAppend (xs, ys) wenzelm@15973: nth_elem (i, xs) = List.nth (xs, i) wenzelm@15973: last_elem xs = List.last xs wenzelm@15973: flat xss = List.concat xss wenzelm@16234: seq fs = List.app fs wenzelm@15973: partition P xs = List.partition P xs wenzelm@15973: mapfilter f xs = List.mapPartial f xs wenzelm@15973: wenzelm@16860: * Pure/library.ML: several combinators for linear functional wenzelm@16860: transformations, notably reverse application and composition: wenzelm@16860: wenzelm@17371: x |> f f #> g wenzelm@17371: (x, y) |-> f f #-> g wenzelm@16860: haftmann@17495: * Pure/library.ML: introduced/changed precedence of infix operators: haftmann@17495: haftmann@17495: infix 1 |> |-> ||> ||>> |>> |>>> #> #->; haftmann@17495: infix 2 ?; haftmann@17495: infix 3 o oo ooo oooo; haftmann@17495: infix 4 ~~ upto downto; haftmann@17495: haftmann@17495: Maybe INCOMPATIBILITY when any of those is used in conjunction with other haftmann@17495: infix operators. haftmann@17495: wenzelm@17408: * Pure/library.ML: natural list combinators fold, fold_rev, and haftmann@16869: fold_map support linear functional transformations and nesting. For wenzelm@16860: example: wenzelm@16860: wenzelm@16860: fold f [x1, ..., xN] y = wenzelm@16860: y |> f x1 |> ... |> f xN wenzelm@16860: wenzelm@16860: (fold o fold) f [xs1, ..., xsN] y = wenzelm@16860: y |> fold f xs1 |> ... |> fold f xsN wenzelm@16860: wenzelm@16860: fold f [x1, ..., xN] = wenzelm@16860: f x1 #> ... #> f xN wenzelm@16860: wenzelm@16860: (fold o fold) f [xs1, ..., xsN] = wenzelm@16860: fold f xs1 #> ... #> fold f xsN wenzelm@16860: wenzelm@17408: * Pure/library.ML: the following selectors on type 'a option are wenzelm@17408: available: wenzelm@17408: wenzelm@17408: the: 'a option -> 'a (*partial*) wenzelm@17408: these: 'a option -> 'a where 'a = 'b list haftmann@17402: the_default: 'a -> 'a option -> 'a haftmann@17402: the_list: 'a option -> 'a list haftmann@17402: wenzelm@17408: * Pure/General: structure AList (cf. Pure/General/alist.ML) provides wenzelm@17408: basic operations for association lists, following natural argument haftmann@17564: order; moreover the explicit equality predicate passed here avoids haftmann@17495: potentially expensive polymorphic runtime equality checks. haftmann@17495: The old functions may be expressed as follows: wenzelm@17408: wenzelm@17408: assoc = uncurry (AList.lookup (op =)) wenzelm@17408: assocs = these oo AList.lookup (op =) wenzelm@17408: overwrite = uncurry (AList.update (op =)) o swap wenzelm@17408: haftmann@17564: * Pure/General: structure AList (cf. Pure/General/alist.ML) provides haftmann@17564: haftmann@17564: val make: ('a -> 'b) -> 'a list -> ('a * 'b) list haftmann@17564: val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list haftmann@17564: haftmann@17564: replacing make_keylist and keyfilter (occassionally used) haftmann@17564: Naive rewrites: haftmann@17564: haftmann@17564: make_keylist = AList.make haftmann@17564: keyfilter = AList.find (op =) haftmann@17564: haftmann@17564: * eq_fst and eq_snd now take explicit equality parameter, thus haftmann@17564: avoiding eqtypes. Naive rewrites: haftmann@17564: haftmann@17564: eq_fst = eq_fst (op =) haftmann@17564: eq_snd = eq_snd (op =) haftmann@17564: haftmann@17564: * Removed deprecated apl and apr (rarely used). haftmann@17564: Naive rewrites: haftmann@17564: haftmann@17564: apl (n, op) =>>= curry op n haftmann@17564: apr (op, m) =>>= fn n => op (n, m) haftmann@17564: wenzelm@17408: * Pure/General: structure OrdList (cf. Pure/General/ord_list.ML) wenzelm@17408: provides a reasonably efficient light-weight implementation of sets as wenzelm@17408: lists. wenzelm@17408: wenzelm@17408: * Pure/General: generic tables (cf. Pure/General/table.ML) provide a wenzelm@17408: few new operations; existing lookup and update are now curried to wenzelm@17408: follow natural argument order (for use with fold etc.); wenzelm@17408: INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort. wenzelm@17408: wenzelm@17408: * Pure/General: output via the Isabelle channels of wenzelm@17408: writeln/warning/error etc. is now passed through Output.output, with a wenzelm@17408: hook for arbitrary transformations depending on the print_mode wenzelm@17408: (cf. Output.add_mode -- the first active mode that provides a output wenzelm@17408: function wins). Already formatted output may be embedded into further wenzelm@17408: text via Output.raw; the result of Pretty.string_of/str_of and derived wenzelm@17408: functions (string_of_term/cterm/thm etc.) is already marked raw to wenzelm@17408: accommodate easy composition of diagnostic messages etc. Programmers wenzelm@17408: rarely need to care about Output.output or Output.raw at all, with wenzelm@17408: some notable exceptions: Output.output is required when bypassing the wenzelm@17408: standard channels (writeln etc.), or in token translations to produce wenzelm@17408: properly formatted results; Output.raw is required when capturing wenzelm@17408: already output material that will eventually be presented to the user wenzelm@17408: a second time. For the default print mode, both Output.output and wenzelm@17408: Output.raw have no effect. wenzelm@17408: wenzelm@17408: * Pure/General: Output.time_accumulator NAME creates an operator ('a wenzelm@17408: -> 'b) -> 'a -> 'b to measure runtime and count invocations; the wenzelm@17408: cumulative results are displayed at the end of a batch session. wenzelm@17408: wenzelm@17408: * Pure/General: File.sysify_path and File.quote_sysify path have been wenzelm@17408: replaced by File.platform_path and File.shell_path (with appropriate wenzelm@17408: hooks). This provides a clean interface for unusual systems where the wenzelm@17408: internal and external process view of file names are different. wenzelm@17408: wenzelm@16689: * Pure: more efficient orders for basic syntactic entities: added wenzelm@16689: fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord wenzelm@16689: and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is wenzelm@16689: NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast wenzelm@16689: orders now -- potential INCOMPATIBILITY for code that depends on a wenzelm@16689: particular order for Symtab.keys, Symtab.dest, etc. (consider using wenzelm@16689: Library.sort_strings on result). wenzelm@16689: wenzelm@17408: * Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types, wenzelm@17408: fold_types traverse types/terms from left to right, observing natural wenzelm@36856: argument order. Supercedes previous foldl_XXX versions, add_frees, wenzelm@17408: add_vars etc. have been adapted as well: INCOMPATIBILITY. wenzelm@17408: wenzelm@16151: * Pure: name spaces have been refined, with significant changes of the wenzelm@16234: internal interfaces -- INCOMPATIBILITY. Renamed cond_extern(_table) wenzelm@36856: to extern(_table). The plain name entry path is superceded by a wenzelm@16234: general 'naming' context, which also includes the 'policy' to produce wenzelm@16234: a fully qualified name and external accesses of a fully qualified wenzelm@36856: name; NameSpace.extend is superceded by context dependent wenzelm@16234: Sign.declare_name. Several theory and proof context operations modify wenzelm@16234: the naming context. Especially note Theory.restore_naming and wenzelm@16234: ProofContext.restore_naming to get back to a sane state; note that wenzelm@16234: Theory.add_path is no longer sufficient to recover from wenzelm@16234: Theory.absolute_path in particular. wenzelm@16234: wenzelm@16234: * Pure: new flags short_names (default false) and unique_names wenzelm@16234: (default true) for controlling output of qualified names. If wenzelm@16234: short_names is set, names are printed unqualified. If unique_names is wenzelm@16234: reset, the name prefix is reduced to the minimum required to achieve wenzelm@16234: the original result when interning again, even if there is an overlap wenzelm@16234: with earlier declarations. wenzelm@16151: wenzelm@16456: * Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is wenzelm@16456: now 'extend', and 'merge' gets an additional Pretty.pp argument wenzelm@16456: (useful for printing error messages). INCOMPATIBILITY. wenzelm@16456: wenzelm@16456: * Pure: major reorganization of the theory context. Type Sign.sg and wenzelm@16456: Theory.theory are now identified, referring to the universal wenzelm@16456: Context.theory (see Pure/context.ML). Actual signature and theory wenzelm@16456: content is managed as theory data. The old code and interfaces were wenzelm@16456: spread over many files and structures; the new arrangement introduces wenzelm@16456: considerable INCOMPATIBILITY to gain more clarity: wenzelm@16456: wenzelm@16456: Context -- theory management operations (name, identity, inclusion, wenzelm@16456: parents, ancestors, merge, etc.), plus generic theory data; wenzelm@16456: wenzelm@16456: Sign -- logical signature and syntax operations (declaring consts, wenzelm@16456: types, etc.), plus certify/read for common entities; wenzelm@16456: wenzelm@16456: Theory -- logical theory operations (stating axioms, definitions, wenzelm@16456: oracles), plus a copy of logical signature operations (consts, wenzelm@16456: types, etc.); also a few basic management operations (Theory.copy, wenzelm@16456: Theory.merge, etc.) wenzelm@16456: wenzelm@16456: The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm wenzelm@16456: etc.) as well as the sign field in Thm.rep_thm etc. have been retained wenzelm@16456: for convenience -- they merely return the theory. wenzelm@16456: wenzelm@36856: * Pure: type Type.tsig is superceded by theory in most interfaces. wenzelm@17193: wenzelm@16547: * Pure: the Isar proof context type is already defined early in Pure wenzelm@16547: as Context.proof (note that ProofContext.context and Proof.context are wenzelm@16547: aliases, where the latter is the preferred name). This enables other wenzelm@16547: Isabelle components to refer to that type even before Isar is present. wenzelm@16547: wenzelm@16373: * Pure/sign/theory: discontinued named name spaces (i.e. classK, wenzelm@16373: typeK, constK, axiomK, oracleK), but provide explicit operations for wenzelm@16373: any of these kinds. For example, Sign.intern typeK is now wenzelm@16373: Sign.intern_type, Theory.hide_space Sign.typeK is now wenzelm@16373: Theory.hide_types. Also note that former wenzelm@16373: Theory.hide_classes/types/consts are now wenzelm@16373: Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions wenzelm@16373: internalize their arguments! INCOMPATIBILITY. wenzelm@16373: wenzelm@16506: * Pure: get_thm interface (of PureThy and ProofContext) expects wenzelm@16506: datatype thmref (with constructors Name and NameSelection) instead of wenzelm@16506: plain string -- INCOMPATIBILITY; wenzelm@16506: wenzelm@16151: * Pure: cases produced by proof methods specify options, where NONE wenzelm@16234: means to remove case bindings -- INCOMPATIBILITY in wenzelm@16234: (RAW_)METHOD_CASES. wenzelm@16151: wenzelm@16373: * Pure: the following operations retrieve axioms or theorems from a wenzelm@16373: theory node or theory hierarchy, respectively: wenzelm@16373: wenzelm@16373: Theory.axioms_of: theory -> (string * term) list wenzelm@16373: Theory.all_axioms_of: theory -> (string * term) list wenzelm@16373: PureThy.thms_of: theory -> (string * thm) list wenzelm@16373: PureThy.all_thms_of: theory -> (string * thm) list wenzelm@16373: wenzelm@16718: * Pure: print_tac now outputs the goal through the trace channel. wenzelm@16718: wenzelm@17408: * Isar toplevel: improved diagnostics, mostly for Poly/ML only. wenzelm@17408: Reference Toplevel.debug (default false) controls detailed printing wenzelm@17408: and tracing of low-level exceptions; Toplevel.profiling (default 0) wenzelm@17408: controls execution profiling -- set to 1 for time and 2 for space wenzelm@17408: (both increase the runtime). wenzelm@17408: wenzelm@17408: * Isar session: The initial use of ROOT.ML is now always timed, wenzelm@17408: i.e. the log will show the actual process times, in contrast to the wenzelm@17408: elapsed wall-clock time that the outer shell wrapper produces. wenzelm@17408: wenzelm@17408: * Simplifier: improved handling of bound variables (nameless wenzelm@16997: representation, avoid allocating new strings). Simprocs that invoke wenzelm@16997: the Simplifier recursively should use Simplifier.inherit_bounds to wenzelm@17720: avoid local name clashes. Failure to do so produces warnings wenzelm@17720: "Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds wenzelm@17720: for further details. wenzelm@16234: wenzelm@17166: * ML functions legacy_bindings and use_legacy_bindings produce ML fact wenzelm@17166: bindings for all theorems stored within a given theory; this may help wenzelm@17166: in porting non-Isar theories to Isar ones, while keeping ML proof wenzelm@17166: scripts for the time being. wenzelm@17166: wenzelm@17457: * ML operator HTML.with_charset specifies the charset begin used for wenzelm@17457: generated HTML files. For example: wenzelm@17457: wenzelm@17457: HTML.with_charset "utf-8" use_thy "Hebrew"; wenzelm@17538: HTML.with_charset "utf-8" use_thy "Chinese"; wenzelm@17457: wenzelm@16234: wenzelm@16234: *** System *** wenzelm@16234: wenzelm@16234: * Allow symlinks to all proper Isabelle executables (Isabelle, wenzelm@16234: isabelle, isatool etc.). wenzelm@16234: wenzelm@16234: * ISABELLE_DOC_FORMAT setting specifies preferred document format (for wenzelm@16234: isatool doc, isatool mkdir, display_drafts etc.). wenzelm@16234: wenzelm@16234: * isatool usedir: option -f allows specification of the ML file to be wenzelm@16234: used by Isabelle; default is ROOT.ML. wenzelm@16234: wenzelm@16251: * New isatool version outputs the version identifier of the Isabelle wenzelm@16251: distribution being used. wenzelm@16251: wenzelm@16251: * HOL: new isatool dimacs2hol converts files in DIMACS CNF format wenzelm@16234: (containing Boolean satisfiability problems) into Isabelle/HOL wenzelm@16234: theories. wenzelm@15703: wenzelm@15703: wenzelm@14655: wenzelm@14606: New in Isabelle2004 (April 2004) wenzelm@14606: -------------------------------- wenzelm@13280: skalberg@14171: *** General *** skalberg@14171: ballarin@14398: * Provers/order.ML: new efficient reasoner for partial and linear orders. ballarin@14398: Replaces linorder.ML. ballarin@14398: wenzelm@14606: * Pure: Greek letters (except small lambda, \), as well as Gothic wenzelm@14606: (\...\\...\), calligraphic (\...\), and Euler skalberg@14173: (\...\), are now considered normal letters, and can therefore skalberg@14173: be used anywhere where an ASCII letter (a...zA...Z) has until skalberg@14173: now. COMPATIBILITY: This obviously changes the parsing of some skalberg@14173: terms, especially where a symbol has been used as a binder, say skalberg@14173: '\x. ...', which is now a type error since \x will be parsed skalberg@14173: as an identifier. Fix it by inserting a space around former skalberg@14173: symbols. Call 'isatool fixgreek' to try to fix parsing errors in skalberg@14173: existing theory and ML files. skalberg@14171: paulson@14237: * Pure: Macintosh and Windows line-breaks are now allowed in theory files. paulson@14237: wenzelm@14731: * Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now wenzelm@14731: allowed in identifiers. Similar to Greek letters \<^isub> is now considered wenzelm@14731: a normal (but invisible) letter. For multiple letter subscripts repeat wenzelm@14731: \<^isub> like this: x\<^isub>1\<^isub>2. kleing@14233: kleing@14333: * Pure: There are now sub-/superscripts that can span more than one kleing@14333: character. Text between \<^bsub> and \<^esub> is set in subscript in wenzelm@14606: ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in wenzelm@14606: superscript. The new control characters are not identifier parts. kleing@14333: schirmer@14561: * Pure: Control-symbols of the form \<^raw:...> will literally print the wenzelm@14606: content of "..." to the latex file instead of \isacntrl... . The "..." wenzelm@14606: may consist of any printable characters excluding the end bracket >. schirmer@14361: paulson@14237: * Pure: Using new Isar command "finalconsts" (or the ML functions paulson@14237: Theory.add_finals or Theory.add_finals_i) it is now possible to paulson@14237: declare constants "final", which prevents their being given a definition paulson@14237: later. It is useful for constants whose behaviour is fixed axiomatically skalberg@14224: rather than definitionally, such as the meta-logic connectives. skalberg@14224: wenzelm@14606: * Pure: 'instance' now handles general arities with general sorts wenzelm@14606: (i.e. intersections of classes), skalberg@14503: kleing@14547: * Presentation: generated HTML now uses a CSS style sheet to make layout wenzelm@14731: (somewhat) independent of content. It is copied from lib/html/isabelle.css. kleing@14547: It can be changed to alter the colors/layout of generated pages. kleing@14547: wenzelm@14556: ballarin@14175: *** Isar *** ballarin@14175: ballarin@14508: * Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac, ballarin@14508: cut_tac, subgoal_tac and thin_tac: ballarin@14175: - Now understand static (Isar) contexts. As a consequence, users of Isar ballarin@14175: locales are no longer forced to write Isar proof scripts. ballarin@14175: For details see Isar Reference Manual, paragraph 4.3.2: Further tactic ballarin@14175: emulations. ballarin@14175: - INCOMPATIBILITY: names of variables to be instantiated may no ballarin@14211: longer be enclosed in quotes. Instead, precede variable name with `?'. ballarin@14211: This is consistent with the instantiation attribute "where". ballarin@14211: ballarin@14257: * Attributes "where" and "of": ballarin@14285: - Now take type variables of instantiated theorem into account when reading ballarin@14285: the instantiation string. This fixes a bug that caused instantiated ballarin@14285: theorems to have too special types in some circumstances. ballarin@14285: - "where" permits explicit instantiations of type variables. ballarin@14257: wenzelm@14556: * Calculation commands "moreover" and "also" no longer interfere with wenzelm@14556: current facts ("this"), admitting arbitrary combinations with "then" wenzelm@14556: and derived forms. kleing@14283: ballarin@14211: * Locales: ballarin@14211: - Goal statements involving the context element "includes" no longer ballarin@14211: generate theorems with internal delta predicates (those ending on ballarin@14211: "_axioms") in the premise. ballarin@14211: Resolve particular premise with .intro to obtain old form. ballarin@14211: - Fixed bug in type inference ("unify_frozen") that prevented mix of target ballarin@14211: specification and "includes" elements in goal statement. ballarin@14254: - Rule sets .intro and .axioms no longer declared as ballarin@14254: [intro?] and [elim?] (respectively) by default. ballarin@14508: - Experimental command for instantiation of locales in proof contexts: ballarin@14551: instantiate