Thu, 12 Apr 2012 10:13:33 +0200 bulwahn multiset operations are defined with lift_definitions;
Thu, 12 Apr 2012 07:33:14 +0200 huffman remove outdated comment
Wed, 11 Apr 2012 21:40:46 +0200 wenzelm rule composition via attribute "OF" (or ML functions OF/MRS) is more tolerant against multiple unifiers;
Wed, 11 Apr 2012 20:42:28 +0200 wenzelm standardized ML aliases;
Wed, 11 Apr 2012 15:02:48 +0200 wenzelm clarified proof_result: finish proof formally via head tr, not end_tr;
Wed, 11 Apr 2012 14:20:51 +0200 wenzelm tuned;
Wed, 11 Apr 2012 13:49:09 +0200 wenzelm more robust Future.fulfill wrt. duplicate assignment and interrupt;
Wed, 11 Apr 2012 13:37:46 +0200 wenzelm tuned message;
Wed, 11 Apr 2012 12:15:56 +0200 wenzelm always signal after cancel_group: passive tasks may have become active;
Wed, 11 Apr 2012 11:44:21 +0200 wenzelm just one dedicated execution per document version -- NB: non-monotonicity of cancel always requires fresh update;
Tue, 10 Apr 2012 23:05:24 +0200 wenzelm merged
Tue, 10 Apr 2012 16:10:50 +0100 Christian Urban moved lift_raw_const wrapper out of the Quotient-package into Nominal2
Tue, 10 Apr 2012 22:53:41 +0200 wenzelm misc tuning and simplification;
Tue, 10 Apr 2012 21:31:05 +0200 wenzelm static relevance of proof via syntax keywords;
Tue, 10 Apr 2012 20:42:17 +0200 wenzelm tuned future priorities: print 0, goal ~1, execute ~2;
Tue, 10 Apr 2012 16:50:30 +0200 wenzelm updated for Poly/ML SVN 1476;
Tue, 10 Apr 2012 11:42:15 +0200 wenzelm some coverage of HOL/TPTP;
Tue, 10 Apr 2012 06:45:15 +0100 sultana added graph-conversion utility for TPTP files
Tue, 10 Apr 2012 06:45:15 +0100 sultana moved non-interpret-specific code to different module
Mon, 09 Apr 2012 23:06:14 +0200 wenzelm disable parallel proofs (again) -- still suffering from instabilites wrt. interrupts;
Mon, 09 Apr 2012 21:29:47 +0200 wenzelm tuned proofs;
Mon, 09 Apr 2012 20:57:23 +0200 wenzelm slightly faster default compilation of Isabelle/Scala;
Mon, 09 Apr 2012 20:42:05 +0200 wenzelm more explicit last exec result;
Mon, 09 Apr 2012 19:50:04 +0200 wenzelm dynamic propagation of node "updated" status, which is required to propagate edits and re-assigments and allow direct memo_result;
Mon, 09 Apr 2012 17:38:39 +0200 wenzelm tuned;
Mon, 09 Apr 2012 17:22:23 +0200 wenzelm simplified Future.cancel/cancel_group (again) -- running threads only;
Mon, 09 Apr 2012 15:10:52 +0200 wenzelm added ML pretty-printing;
Sat, 07 Apr 2012 20:35:01 +0200 wenzelm merged
Sat, 07 Apr 2012 20:10:13 +0200 wenzelm merged
Sat, 07 Apr 2012 20:24:39 +0200 haftmann explicit constructor Nat leaves nat_of as conversion
Fri, 06 Apr 2012 19:23:51 +0200 haftmann abandoned almost redundant *_foldr lemmas
Fri, 06 Apr 2012 19:18:00 +0200 haftmann tuned
Fri, 06 Apr 2012 18:17:16 +0200 haftmann no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
Sat, 07 Apr 2012 19:59:09 +0200 wenzelm enable parallel proofs (cf. e8552cba702d), only affects packages so far;
Sat, 07 Apr 2012 19:28:44 +0200 wenzelm added static command status markup, to emphasize accepted but unassigned/unparsed commands (notably in overview panel);
Sat, 07 Apr 2012 18:08:29 +0200 wenzelm tuned proofs;
Sat, 07 Apr 2012 17:55:35 +0200 wenzelm more robust update_perspective, e.g. required after reload of buffer that is not at start position;
Sat, 07 Apr 2012 17:49:20 +0200 wenzelm tuned imports;
Sat, 07 Apr 2012 17:48:47 +0200 wenzelm updated header keywords;
Sat, 07 Apr 2012 16:59:27 +0200 wenzelm init message not bad;
Sat, 07 Apr 2012 16:41:59 +0200 wenzelm explicit checks stable_finished_theory/stable_command allow parallel asynchronous command transactions;
Fri, 06 Apr 2012 23:34:38 +0200 wenzelm discontinued obsolete last_execs (cf. cd3ab7625519);
Fri, 06 Apr 2012 14:40:00 +0200 huffman remove now-unnecessary type annotations from lift_definition commands
Fri, 06 Apr 2012 14:39:27 +0200 huffman more robust generation of quotient rules using tactics
Fri, 06 Apr 2012 13:50:07 +0200 huffman merged
Fri, 06 Apr 2012 10:37:46 +0200 huffman add function dest_Quotient
Fri, 06 Apr 2012 13:10:45 +0200 wenzelm standardized alias;
Fri, 06 Apr 2012 12:45:56 +0200 wenzelm fixed document;
Fri, 06 Apr 2012 12:10:50 +0200 wenzelm merged
Fri, 06 Apr 2012 09:35:47 +0200 huffman correct plumbing of proof contexts, so that force_rty_type won't generalize more type variables than it should
Thu, 05 Apr 2012 23:22:54 +0200 kuncar detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
Thu, 05 Apr 2012 22:00:27 +0200 kuncar make Quotient_Def.lift_raw_const working again
Thu, 05 Apr 2012 16:25:59 +0200 huffman use standard quotient lemmas to generate transfer rules
Thu, 05 Apr 2012 15:59:25 +0200 huffman add transfer lemmas for quotients
Thu, 05 Apr 2012 15:23:26 +0200 huffman define reflp directly, in the manner of symp and transp
Thu, 05 Apr 2012 14:14:16 +0200 huffman set up and use lift_definition for word operations
Thu, 05 Apr 2012 12:18:06 +0200 huffman lift_definition declares transfer_rule attribute
Thu, 05 Apr 2012 10:48:46 +0200 huffman configure transfer method for 'a word -> int
Thu, 05 Apr 2012 08:43:31 +0200 krauss added timestamps to logging of named thms
Thu, 05 Apr 2012 08:13:40 +0200 huffman merged
Wed, 04 Apr 2012 20:45:19 +0200 huffman merged
Wed, 04 Apr 2012 16:48:39 +0200 huffman add lemmas for generating transfer rules for typedefs
Thu, 05 Apr 2012 00:37:22 +0100 sultana tuned;
Wed, 04 Apr 2012 21:57:39 +0100 sultana improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
Wed, 04 Apr 2012 20:40:39 +0200 Cezary Kaliszyk merge
Wed, 04 Apr 2012 20:10:21 +0200 Cezary Kaliszyk HOL/Import more precise map types
Wed, 04 Apr 2012 20:09:56 +0200 Cezary Kaliszyk HOL/Import typed matches against Isabelle typedef result
Wed, 04 Apr 2012 19:20:52 +0200 kuncar connect the Quotient package to the Lifting package
Wed, 04 Apr 2012 17:51:12 +0200 kuncar support non-open typedefs; define cr_rel in terms of a rep function for typedefs
Wed, 04 Apr 2012 16:29:17 +0100 sultana tuned;
Wed, 04 Apr 2012 16:29:16 +0100 sultana added interpretation for formula conditional;
Wed, 04 Apr 2012 16:29:16 +0100 sultana refactored tptp lex;
Wed, 04 Apr 2012 16:29:16 +0100 sultana refactored tptp yacc to bring close to official spec;
Wed, 04 Apr 2012 16:05:52 +0200 huffman transfer method generalizes over free variables in goal
Wed, 04 Apr 2012 16:03:01 +0200 huffman add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
Wed, 04 Apr 2012 12:25:58 +0200 huffman update keywords file
Wed, 04 Apr 2012 14:27:20 +0200 huffman merged
Wed, 04 Apr 2012 11:50:08 +0200 huffman fix typos
Wed, 04 Apr 2012 13:42:01 +0200 huffman lift_definition command generates transfer rule
Wed, 04 Apr 2012 13:41:38 +0200 huffman prove_quot_theorem fixes types
Wed, 04 Apr 2012 14:08:24 +0200 bulwahn documenting options quickcheck_locale; adjusting IsarRef documentation of Quotient predicate; NEWS
Wed, 04 Apr 2012 12:22:51 +0200 bulwahn added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
Fri, 06 Apr 2012 12:02:24 +0200 wenzelm stop node execution at first unassigned exec;
Fri, 06 Apr 2012 11:49:08 +0200 wenzelm discontinued Document.update_perspective side-entry (cf. 546adfa8a6fc) -- NB: re-assignment is always necessary due to non-monotonic cancel_execution;
Thu, 05 Apr 2012 22:33:52 +0200 wenzelm more scalable execute/update: avoid redundant traversal of invisible/finished nodes;
Thu, 05 Apr 2012 14:34:54 +0200 wenzelm less ambitious memo_eval, since memo_result is still not robust here;
Thu, 05 Apr 2012 14:14:51 +0200 wenzelm less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
Thu, 05 Apr 2012 13:01:54 +0200 wenzelm more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
Thu, 05 Apr 2012 11:58:46 +0200 wenzelm Command.memo including physical interrupts (unlike Lazy.lazy);
Thu, 05 Apr 2012 10:45:53 +0200 wenzelm tuned;
Wed, 04 Apr 2012 21:03:30 +0200 wenzelm tuned -- NB: get_theory still needs to apply Lazy.force due to interrupt instabilities;
Wed, 04 Apr 2012 17:21:39 +0200 wenzelm proper signature constraint;
Wed, 04 Apr 2012 17:14:19 +0200 wenzelm tuned comments;
Wed, 04 Apr 2012 14:19:47 +0200 wenzelm separate module for prover command execution;
Wed, 04 Apr 2012 14:00:47 +0200 wenzelm tuned;
Wed, 04 Apr 2012 10:38:04 +0200 huffman fix typo in ML structure name
Wed, 04 Apr 2012 11:15:54 +0200 wenzelm removed obsolete isar-overview manual;
Wed, 04 Apr 2012 10:04:25 +0100 sultana dealing with SMLNJ errors 'value type in structure doesn't match signature spec' (which originated from lack of type generalisation) and 'unresolved flex record' wrt tptp_interpret
Wed, 04 Apr 2012 10:49:42 +0200 bulwahn merged
Wed, 04 Apr 2012 10:17:54 +0200 bulwahn rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
Wed, 04 Apr 2012 10:17:08 +0200 bulwahn improved equality check for modes in predicate compiler
Wed, 04 Apr 2012 09:00:10 +0200 huffman rename ML structure to avoid shadowing earlier name
Wed, 04 Apr 2012 07:47:42 +0200 huffman add type annotations to make SML happy (cf. ec6187036495)
Tue, 03 Apr 2012 23:49:24 +0200 huffman merged
Tue, 03 Apr 2012 22:31:00 +0200 huffman new transfer proof method
Tue, 03 Apr 2012 22:04:50 +0200 huffman renamed Tools/transfer.ML to Tools/legacy_transfer.ML
Tue, 03 Apr 2012 21:39:28 +0200 wenzelm prefer prog-prove, suppress isar-overview;
Tue, 03 Apr 2012 21:09:09 +0200 wenzelm merged
Tue, 03 Apr 2012 20:41:13 +0200 wenzelm merged
Tue, 03 Apr 2012 20:42:00 +0200 wenzelm formal integration of "prog-prove" manual;
Tue, 03 Apr 2012 20:37:52 +0200 wenzelm prefer static dependencies;
Tue, 03 Apr 2012 20:56:32 +0200 huffman merged
Tue, 03 Apr 2012 15:15:00 +0200 huffman modernized obsolete old-style theory name with proper new-style underscore
Tue, 03 Apr 2012 17:33:22 +0100 sultana removed use of CharVector in generated parser, to make SMLNJ happy
Tue, 03 Apr 2012 20:24:00 +0200 wenzelm avoid duplicate PIDE markup;
Tue, 03 Apr 2012 20:08:08 +0200 wenzelm less intrusive visibility;
Tue, 03 Apr 2012 19:49:14 +0200 wenzelm more robust re-import wrt. non-HHF assumptions;
Tue, 03 Apr 2012 19:33:46 +0200 wenzelm consider polyml-5.3.0 as "experimental" since it chokes on HOL-Codegenerator_Test, while 5.2.1 happens to work;
Tue, 03 Apr 2012 18:22:14 +0200 wenzelm close context elements via Expression.cert/read_declaration;
Tue, 03 Apr 2012 17:48:16 +0200 wenzelm merged
Tue, 03 Apr 2012 16:45:44 +0100 sultana added me to isatest email list
Tue, 03 Apr 2012 16:26:48 +0200 kuncar new package Lifting - initial commit
Tue, 03 Apr 2012 14:09:37 +0200 huffman add floor/ceiling lemmas suggested by René Thiemann
Tue, 03 Apr 2012 08:55:06 +0200 nipkow made sure that " is shown in tutorial text
Mon, 02 Apr 2012 21:26:46 +0100 Christian Urban merged
Mon, 02 Apr 2012 21:26:07 +0100 Christian Urban tuned proofs
Mon, 02 Apr 2012 20:12:19 +0200 nipkow merged
Mon, 02 Apr 2012 20:12:10 +0200 nipkow towards showing " in the tutorial
Mon, 02 Apr 2012 18:12:53 +0100 Christian Urban tuned proof in order to avoid warning message
Mon, 02 Apr 2012 16:07:24 +0200 huffman remove unnecessary qualifiers on names
Mon, 02 Apr 2012 16:06:24 +0200 huffman add lemma Suc_1
Mon, 02 Apr 2012 14:09:27 +0200 berghofe merged
Mon, 02 Apr 2012 13:58:59 +0200 berghofe Require "For" keyword to be followed by at least one whitespace, to avoid that
Tue, 03 Apr 2012 17:21:33 +0200 wenzelm normalize defs (again, cf. 008b7858f3c0);
Tue, 03 Apr 2012 16:53:32 +0200 wenzelm some context examples;
Tue, 03 Apr 2012 16:51:01 +0200 wenzelm retain literal non-HHF assumptions, to facilitate re-import in Generic_Target.import_export_proof;
Tue, 03 Apr 2012 16:49:05 +0200 wenzelm another attempt to avoid duplication of global vs. local syntax (reminiscent of old fork_mixfix);
Tue, 03 Apr 2012 16:27:32 +0200 wenzelm export into target context (again), to retain its 'defines' (e.g. abbreviation lcoeff in theory HOL/Algebra/UnivPoly);
Tue, 03 Apr 2012 16:10:34 +0200 wenzelm better drop background syntax if entity depends on parameters;
Tue, 03 Apr 2012 14:37:16 +0200 wenzelm more uniform abbrev vs. define;
Tue, 03 Apr 2012 13:47:15 +0200 wenzelm tuned;
Tue, 03 Apr 2012 11:28:21 +0200 wenzelm clarified generic_const vs. close_schematic_term;
Tue, 03 Apr 2012 11:21:17 +0200 wenzelm tuned;
Tue, 03 Apr 2012 10:59:20 +0200 wenzelm more uniform theory_abbrev with const_declaration;
Tue, 03 Apr 2012 10:04:41 +0200 wenzelm avoid const_declaration in aux. context (cf. locale_foundation);
Tue, 03 Apr 2012 09:47:20 +0200 wenzelm clarified background_foundation vs. theory_foundation (with const_declaration);
Tue, 03 Apr 2012 09:41:16 +0200 wenzelm tuned;
Mon, 02 Apr 2012 23:55:25 +0200 wenzelm more general standard_declaration;
Mon, 02 Apr 2012 23:27:24 +0200 wenzelm better restore after close_target;
Mon, 02 Apr 2012 21:52:03 +0200 wenzelm tuned;
Mon, 02 Apr 2012 21:49:27 +0200 wenzelm clarified standard_declaration vs. theory_declaration;
Mon, 02 Apr 2012 20:50:41 +0200 wenzelm smarter generic_const: plain alias for non-dependent case (e.g. prospective datatype or record syntax);
Mon, 02 Apr 2012 20:12:17 +0200 wenzelm tuned;
Mon, 02 Apr 2012 19:54:25 +0200 wenzelm tuned;
Mon, 02 Apr 2012 19:47:21 +0200 wenzelm tuned signature;
Mon, 02 Apr 2012 19:10:52 +0200 wenzelm misc tuning and simplification;
Mon, 02 Apr 2012 17:00:32 +0200 wenzelm better restore to first target, not last target;
Mon, 02 Apr 2012 16:35:09 +0200 wenzelm refined define/abbrev: allow extra fixes in aux. context vs. bottom target (NB: export_term expands defined variables, leaving fixed ones);
Mon, 02 Apr 2012 15:42:50 +0200 wenzelm more general Local_Theory.restore, allow any nesting level;
Mon, 02 Apr 2012 13:47:00 +0200 nipkow new tutorial
Mon, 02 Apr 2012 10:49:03 +0200 nipkow New manual Programming and Proving in Isabelle/HOL
Mon, 02 Apr 2012 09:18:16 +0200 huffman add simp rules for dvd on negative numerals
Sun, 01 Apr 2012 23:21:54 +0200 krauss merged, manually resolving conflicts due to session renaming (cf. 6488c5efec49)
Sun, 01 Apr 2012 23:09:36 +0200 krauss clarified terminology; added reference to bundle component
Sun, 01 Apr 2012 22:55:06 +0200 krauss less modest NEWS; CONTRIBUTORS
Sun, 01 Apr 2012 22:41:56 +0200 krauss renamed import session back to Import, conforming to directory name; NEWS
Sun, 01 Apr 2012 23:07:15 +0200 wenzelm more precise IsaMakefile (eg. see HOL-Algebra);
Sun, 01 Apr 2012 22:58:05 +0200 wenzelm more keywords;
Sun, 01 Apr 2012 22:40:15 +0200 wenzelm merged
Sun, 01 Apr 2012 22:14:59 +0200 krauss merged
Sun, 01 Apr 2012 22:03:45 +0200 krauss removed old HOL4 import -- corresponding exporter is lost, code is broken, no users known, maintenance nightmare
Sun, 01 Apr 2012 14:50:47 +0200 Cezary Kaliszyk Modernized HOL-Import for HOL Light
Sun, 01 Apr 2012 22:26:28 +0200 wenzelm merged
Sun, 01 Apr 2012 21:12:04 +0200 krauss adapted Mira configuration to dd04c8173bb2.
Sun, 01 Apr 2012 16:09:58 +0200 huffman removed Nat_Numeral.thy, moving all theorems elsewhere
Sun, 01 Apr 2012 22:02:14 +0200 wenzelm less brutal return from function, to allow caller to report error;
Sun, 01 Apr 2012 21:46:45 +0200 wenzelm more general context command with auxiliary fixes/assumes etc.;
Sun, 01 Apr 2012 21:45:25 +0200 wenzelm more precise type annotation (cf. 6523a21076a8);
Sun, 01 Apr 2012 20:42:19 +0200 wenzelm nothing specific about named target;
Sun, 01 Apr 2012 20:36:33 +0200 wenzelm clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
Sun, 01 Apr 2012 19:07:32 +0200 wenzelm added Attrib.global_notes/local_notes/generic_notes convenience;
Sun, 01 Apr 2012 19:04:52 +0200 wenzelm simplified;
Sun, 01 Apr 2012 18:01:19 +0200 wenzelm tuned signature;
Sun, 01 Apr 2012 15:23:43 +0200 wenzelm clarified Named_Target.target_declaration: propagate through other levels as well;
Sun, 01 Apr 2012 14:29:22 +0200 wenzelm Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
Sun, 01 Apr 2012 09:12:03 +0200 huffman tuned proofs
Sat, 31 Mar 2012 22:45:46 +0200 huffman merged
Sat, 31 Mar 2012 20:09:24 +0200 huffman tuned proof
Sat, 31 Mar 2012 19:10:58 +0200 huffman add lemma power_le_one
Sat, 31 Mar 2012 19:38:41 +0200 wenzelm tuned;
Sat, 31 Mar 2012 19:26:23 +0200 wenzelm tuned signature;
Sat, 31 Mar 2012 19:09:59 +0200 wenzelm more direct Local_Defs.contract;
Sat, 31 Mar 2012 15:29:49 +0200 wenzelm more precise Local_Defs.expand wrt. *local* prems only;
Sat, 31 Mar 2012 15:21:35 +0200 wenzelm tuned comment;
Fri, 30 Mar 2012 21:08:00 +0200 wenzelm more robust Scala 2.9.x interpreter invocation -- avoid separate interpreter thread and thus deadlock of Swing_Thread.now;
Fri, 30 Mar 2012 19:36:41 +0200 wenzelm tuned;
Fri, 30 Mar 2012 18:56:46 +0200 haftmann dropped empty files
Fri, 30 Mar 2012 18:56:02 +0200 haftmann dropped now obsolete Cset theories
Fri, 30 Mar 2012 17:25:34 +0200 wenzelm merged
Fri, 30 Mar 2012 17:22:17 +0200 wenzelm tuned proofs, less guesswork;
Fri, 30 Mar 2012 17:21:36 +0200 huffman merged
Fri, 30 Mar 2012 16:44:23 +0200 huffman load Tools/numeral.ML in Num.thy
Fri, 30 Mar 2012 16:43:07 +0200 huffman tuned proof
Fri, 30 Mar 2012 15:56:12 +0200 huffman set up numeral reorient simproc in Num.thy
Fri, 30 Mar 2012 15:43:30 +0200 huffman remove redundant simp rule
Fri, 30 Mar 2012 15:24:24 +0200 huffman add simp rules for eve/odd on numerals
Fri, 30 Mar 2012 14:27:29 +0200 huffman remove content-free theory ex/Arithmetic_Series_Complex.thy
Fri, 30 Mar 2012 14:25:32 +0200 huffman rephrase lemmas about arithmetic series using numeral '2'
Fri, 30 Mar 2012 14:00:18 +0200 huffman rephrase lemma card_Pow using '2' instead of 'Suc (Suc 0)'
Fri, 30 Mar 2012 12:32:35 +0200 huffman replace lemmas eval_nat_numeral with a simpler reformulation
Fri, 30 Mar 2012 12:02:23 +0200 huffman restate various simp rules for word operations using pred_numeral
Fri, 30 Mar 2012 11:52:26 +0200 huffman new lemmas for simplifying subtraction on nat numerals
Fri, 30 Mar 2012 11:16:35 +0200 huffman removed redundant nat-specific copies of theorems
Fri, 30 Mar 2012 10:41:27 +0200 huffman move more theorems from Nat_Numeral.thy to Num.thy
Fri, 30 Mar 2012 15:25:47 +0200 wenzelm "invariant" is free in main HOL (cf. 56adbf5bcc82, e64ffc96a49f);
Fri, 30 Mar 2012 13:12:02 +0200 wenzelm more robust ISABELLE_JDK_HOME settings, based on exisiting JAVA_HOME provided by isatest shell environment (which depends a lot on the host);
Fri, 30 Mar 2012 11:48:24 +0200 wenzelm more explicit isatest environment settings (from private .bashrc);
Fri, 30 Mar 2012 09:55:03 +0200 huffman merged
Fri, 30 Mar 2012 08:15:29 +0200 huffman fix search-and-replace errors
Fri, 30 Mar 2012 08:04:28 +0200 huffman move lemma card_UNIV_bool from Nat_Numeral.thy to Finite_Set.thy
Fri, 30 Mar 2012 09:08:29 +0200 huffman add constant pred_numeral k = numeral k - (1::nat);
Fri, 30 Mar 2012 08:11:52 +0200 huffman move lemmas from Nat_Numeral.thy to Nat.thy
Fri, 30 Mar 2012 08:10:28 +0200 huffman move lemmas from Nat_Numeral to Int.thy and Num.thy
Fri, 30 Mar 2012 09:32:18 +0200 bulwahn merged
Fri, 30 Mar 2012 08:44:01 +0200 bulwahn adding theory to prove completeness of the exhaustive generators
Fri, 30 Mar 2012 08:19:31 +0200 bulwahn refine bindings in quickcheck_common: do not conceal and do not declare as simps
Fri, 30 Mar 2012 08:19:29 +0200 bulwahn hiding fact not so aggressively
Fri, 30 Mar 2012 09:04:29 +0200 haftmann power on predicate relations
Fri, 30 Mar 2012 00:01:30 +0100 sultana made Mirabelle-SH's 'trivial' check optional;
Thu, 29 Mar 2012 22:52:24 +0200 wenzelm merged
Thu, 29 Mar 2012 22:43:50 +0200 wenzelm more specific notion of partiality (cf. Scala version);
Thu, 29 Mar 2012 21:13:48 +0200 kuncar use qualified names for rsp and rep_eq theorems in quotient_def
Thu, 29 Mar 2012 17:40:44 +0200 bulwahn announcing NEWS (cf. 446cfc760ccf)
Thu, 29 Mar 2012 14:47:31 +0200 huffman remove obsolete simp rule for powers
Thu, 29 Mar 2012 14:42:55 +0200 huffman remove duplicate lemmas power_m1_{even,odd} in favor of power_minus1_{even,odd}
Thu, 29 Mar 2012 14:39:05 +0200 huffman remove unneeded rewrite rules for powers of numerals
Thu, 29 Mar 2012 14:29:36 +0200 huffman remove duplicate lemma Suc_numeral
Thu, 29 Mar 2012 14:09:10 +0200 huffman move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
Thu, 29 Mar 2012 11:47:30 +0200 huffman bootstrap Num.thy before Power.thy;
Thu, 29 Mar 2012 08:59:56 +0200 haftmann educated guess to include jdk
(0) -30000 -10000 -3000 -1000 -240 +240 +1000 +3000 +10000 +30000 tip