src/HOL/Tools/Nitpick/nitpick_hol.ML
Wed, 09 Mar 2011 10:28:01 +0100 blanchet improved formula for specialization, to prevent needless specializations -- and removed dead code
Thu, 03 Mar 2011 11:20:48 +0100 blanchet simplify "need" option's syntax
Thu, 03 Mar 2011 11:20:48 +0100 blanchet renamed "preconstr" option "need"
Wed, 02 Mar 2011 14:50:16 +0100 blanchet lower threshold for implicitly using "nitpick_simp" for predicate definitions when "total_consts" is on
Mon, 28 Feb 2011 17:53:10 +0100 blanchet improved "is_special_eligible_arg" further, by approximating the cardinality of the argument (as is done quite successfully elsewhere)
Mon, 28 Feb 2011 17:53:10 +0100 blanchet made "fixed_is_special_eligible_arg" smarter w.r.t. pairs, and fixed previous unintended behavior because "andalso" ties more tightly than "orelse"
Mon, 21 Feb 2011 17:36:32 +0100 blanchet more work on "fix_datatype_vals" optimization (renamed "preconstruct")
Mon, 21 Feb 2011 11:50:38 +0100 blanchet improve optimization
Mon, 21 Feb 2011 11:50:31 +0100 blanchet tweaked Nitpick based on C++ memory model example
Mon, 21 Feb 2011 10:44:19 +0100 blanchet renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
Mon, 21 Feb 2011 10:42:29 +0100 blanchet always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
Sat, 08 Jan 2011 17:14:48 +0100 wenzelm misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
Tue, 07 Dec 2010 11:56:53 +0100 blanchet remove the "fin_fun" optimization in Nitpick -- it was always a hack and didn't help much
Tue, 07 Dec 2010 11:56:01 +0100 blanchet removed needless optimization for image -- there might be cases that benefit from it but there are others where it is clearly evil
Tue, 07 Dec 2010 11:56:01 +0100 blanchet simplified special handling of set products
Tue, 07 Dec 2010 11:56:01 +0100 blanchet fix special handling of set products
Tue, 07 Dec 2010 11:56:01 +0100 blanchet use heuristic to determine whether to keep or drop an existing "let" -- and drop all higher-order lets
Mon, 25 Oct 2010 21:06:56 +0200 wenzelm renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
Fri, 24 Sep 2010 15:53:13 +0200 wenzelm modernized structure Ord_List;
Mon, 20 Sep 2010 16:05:25 +0200 wenzelm renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
Tue, 14 Sep 2010 13:44:43 +0200 blanchet eliminate more clutter related to "fast_descrs" optimization
Tue, 14 Sep 2010 13:24:18 +0200 blanchet remove "fast_descs" option from Nitpick;
Mon, 13 Sep 2010 20:21:40 +0200 blanchet remove unreferenced identifiers
Sat, 11 Sep 2010 10:13:51 +0200 blanchet always handle type variables in typedefs as global
Sat, 28 Aug 2010 16:14:32 +0200 haftmann formerly unnamed infix equality now named HOL.eq
Fri, 27 Aug 2010 10:56:46 +0200 haftmann formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
Thu, 26 Aug 2010 20:51:17 +0200 haftmann formerly unnamed infix impliciation now named HOL.implies
Mon, 23 Aug 2010 14:54:17 +0200 blanchet perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
Mon, 09 Aug 2010 12:40:15 +0200 blanchet use "declaration" instead of "setup" to register Nitpick extensions
Mon, 09 Aug 2010 12:05:48 +0200 blanchet move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
Fri, 06 Aug 2010 18:11:30 +0200 blanchet added support for partial quotient types;
Fri, 06 Aug 2010 17:23:11 +0200 blanchet adapt occurrences of renamed Nitpick functions
Fri, 06 Aug 2010 17:05:29 +0200 blanchet local versions of Nitpick.register_xxx functions
Fri, 06 Aug 2010 11:35:10 +0200 blanchet quotient types registered as codatatypes are no longer quotient types
Fri, 06 Aug 2010 10:50:52 +0200 blanchet improved "merge_type_vars" option: map supersorts to subsorts, to avoid distinguishing, say, "{}", and "HOL.type"
Thu, 05 Aug 2010 20:17:50 +0200 blanchet added "whack"
Thu, 05 Aug 2010 18:33:07 +0200 blanchet handle "Rep_unit" & Co. gracefully
Thu, 05 Aug 2010 18:00:50 +0200 blanchet added support for "Abs_" and "Rep_" functions on quotient types
Thu, 05 Aug 2010 15:44:54 +0200 blanchet fiddle with specialization etc.
Thu, 05 Aug 2010 14:45:27 +0200 blanchet handle inductive predicates correctly after change in "def" semantics
Thu, 05 Aug 2010 14:32:24 +0200 blanchet don't specialize built-ins or constructors
Thu, 05 Aug 2010 14:10:18 +0200 blanchet prevent the expansion of too large definitions -- use equations for these instead
Thu, 05 Aug 2010 12:58:57 +0200 blanchet make nitpick accept "==" for "nitpick_(p)simp"s
Thu, 05 Aug 2010 12:40:12 +0200 blanchet fix bug in Nitpick's "equationalize" function (the prems were ignored) + make it do some basic extensionalization
Wed, 04 Aug 2010 10:39:35 +0200 blanchet get rid of all "optimizations" regarding "unit" and other cardinality-1 types
Tue, 03 Aug 2010 21:20:31 +0200 blanchet better "Pretty" handling
Tue, 03 Aug 2010 14:49:42 +0200 blanchet bump up the max cardinalities, to use up more of the time given to us by the user
Tue, 03 Aug 2010 13:29:26 +0200 blanchet fix newly introduced bug w.r.t. conditional equations
Tue, 03 Aug 2010 12:31:30 +0200 blanchet make Nitpick more flexible when parsing (p)simp rules
Tue, 03 Aug 2010 01:16:08 +0200 blanchet optimize local "def"s by treating them as definitions
Mon, 02 Aug 2010 19:15:15 +0200 blanchet careful about which linear inductive predicates should be starred
Mon, 02 Aug 2010 14:27:20 +0200 blanchet prefer implication to equality, to be more SAT solver friendly
Mon, 02 Aug 2010 12:36:50 +0200 blanchet fix bug with mutually recursive and nested codatatypes
Thu, 01 Jul 2010 16:54:44 +0200 haftmann "prod" and "sum" replace "*" and "+" respectively
Mon, 28 Jun 2010 15:03:07 +0200 haftmann merged constants "split" and "prod_case"
Mon, 21 Jun 2010 14:07:00 +0200 blanchet sort cases on the proper key
Mon, 21 Jun 2010 11:15:21 +0200 blanchet optimized code generated for datatype cases + more;
Tue, 01 Jun 2010 17:28:16 +0200 blanchet cosmetics
Tue, 01 Jun 2010 15:38:47 +0200 blanchet removed "nitpick_intro" attribute -- Nitpick noew uses Spec_Rules instead
Tue, 01 Jun 2010 14:14:02 +0200 blanchet honor xsymbols in Nitpick
Tue, 01 Jun 2010 12:20:08 +0200 blanchet added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
Tue, 01 Jun 2010 10:40:23 +0200 blanchet make Nitpick handle multiple typedef entries for same typedef
Tue, 01 Jun 2010 10:31:18 +0200 blanchet thread along context instead of theory for typedef lookup
Mon, 31 May 2010 18:00:28 +0200 blanchet don't include any axioms for "TYPE" in Nitpick
Mon, 31 May 2010 17:20:41 +0200 blanchet fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
Fri, 30 Apr 2010 14:58:21 +0200 blanchet catch the right exception
Thu, 29 Apr 2010 01:17:14 +0200 blanchet expand combinators in Isar proofs constructed by Sledgehammer;
Sun, 25 Apr 2010 00:33:26 +0200 blanchet cosmetics
Sun, 25 Apr 2010 00:25:44 +0200 blanchet remove "show_skolems" option and change style of record declarations
Sun, 25 Apr 2010 00:10:30 +0200 blanchet remove "skolemize" option from Nitpick, since Skolemization is always useful
Sat, 24 Apr 2010 17:48:21 +0200 blanchet removed Nitpick's "uncurry" option
Sat, 24 Apr 2010 16:33:01 +0200 blanchet remove type annotations as comments;
Fri, 23 Apr 2010 19:18:39 +0200 blanchet stop referring to Sledgehammer_Util stuff all over Nitpick code; instead, redeclare any needed function in Nitpick_Util as synonym for the Sledgehammer_Util function of the same name
Tue, 13 Apr 2010 11:43:11 +0200 blanchet cosmetics
Mon, 29 Mar 2010 15:50:18 +0200 blanchet get rid of Polyhash, since it's no longer used
Sat, 27 Mar 2010 21:38:38 +0100 wenzelm Typedef.info: separate global and local part, only the latter is transformed by morphisms;
Mon, 22 Mar 2010 15:07:07 +0100 blanchet detect OFCLASS() axioms in Nitpick;
Sat, 20 Mar 2010 17:33:11 +0100 wenzelm renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
Wed, 17 Mar 2010 09:14:43 +0100 blanchet added support for "specification" and "ax_specification" constructs to Nitpick
Sat, 13 Mar 2010 14:44:47 +0100 wenzelm adapted to localized typedef: handle single global interpretation only;
Thu, 11 Mar 2010 17:48:07 +0100 blanchet moved some Nitpick code around
Thu, 11 Mar 2010 12:22:11 +0100 blanchet added term postprocessor to Nitpick, to provide custom syntax for typedefs
Tue, 09 Mar 2010 14:18:21 +0100 blanchet improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
Tue, 09 Mar 2010 09:25:23 +0100 blanchet added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
Sun, 07 Mar 2010 12:19:47 +0100 wenzelm modernized structure Object_Logic;
Sat, 27 Feb 2010 23:13:01 +0100 wenzelm modernized structure Term_Ord;
Fri, 26 Feb 2010 18:38:23 +0100 blanchet use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
Fri, 26 Feb 2010 16:49:46 +0100 blanchet more work on the new monotonicity stuff in Nitpick
Thu, 25 Feb 2010 16:33:39 +0100 blanchet improved precision of infinite "shallow" datatypes in Nitpick;
Thu, 25 Feb 2010 10:08:44 +0100 blanchet cosmetics
Tue, 23 Feb 2010 19:10:25 +0100 blanchet support local definitions in Nitpick
Tue, 23 Feb 2010 14:50:44 +0100 blanchet optimized multisets in Nitpick by fishing "finite"
Tue, 23 Feb 2010 11:05:32 +0100 blanchet improved Nitpick's support for quotient types
Tue, 23 Feb 2010 10:02:14 +0100 blanchet catch IO errors in Nitpick's "kodkodi" invocation + shorten execution time of "Manual_Nits" example
Mon, 22 Feb 2010 19:31:00 +0100 blanchet enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
Mon, 22 Feb 2010 14:36:10 +0100 blanchet filter out trivial definitions in Nitpick (e.g. "Topology.topo" from AFP)
Mon, 22 Feb 2010 11:57:33 +0100 blanchet fixed a few bugs in Nitpick and removed unreferenced variables
Thu, 18 Feb 2010 18:48:07 +0100 blanchet added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
Wed, 17 Feb 2010 20:46:50 +0100 blanchet make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
Wed, 17 Feb 2010 10:28:04 +0100 blanchet synchronize Nitpick's wellfoundedness formulas caching
Sat, 13 Feb 2010 11:56:06 +0100 blanchet redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
Fri, 12 Feb 2010 19:44:37 +0100 blanchet various cosmetic changes to Nitpick
Tue, 09 Feb 2010 17:06:05 +0100 blanchet merged (manual for "nitpick_hol.ML" and "kodkod.ML")
Tue, 09 Feb 2010 16:07:51 +0100 blanchet optimization to quantifiers in Nitpick's handling of simp rules + renamed some SAT solvers
Fri, 05 Feb 2010 12:04:54 +0100 blanchet handle Nitpick's nonstandard model enumeration in a cleaner way;
Thu, 04 Feb 2010 16:03:15 +0100 blanchet split "nitpick_hol.ML" into two files to make it more manageable;
Fri, 05 Feb 2010 14:33:50 +0100 haftmann more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
Thu, 04 Feb 2010 13:36:52 +0100 blanchet four changes to Nitpick:
Tue, 02 Feb 2010 11:38:38 +0100 blanchet added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
Wed, 20 Jan 2010 10:38:06 +0100 blanchet some work on Nitpick's support for quotient types;
Fri, 18 Dec 2009 12:00:29 +0100 blanchet polished Nitpick's binary integer support etc.;
Thu, 17 Dec 2009 15:22:11 +0100 blanchet added support for binary nat/int representation to Nitpick
Mon, 14 Dec 2009 16:48:49 +0100 blanchet distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
Mon, 14 Dec 2009 12:30:26 +0100 blanchet get rid of polymorphic equality in Nitpick's code + a few minor cleanups
Fri, 04 Dec 2009 17:17:52 +0100 blanchet fix soundness bug in Nitpick's "destroy_constrs" optimization
Mon, 30 Nov 2009 11:42:49 +0100 haftmann modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
Tue, 24 Nov 2009 18:35:21 +0100 blanchet fix soundness bug in "uncurry" option of Nitpick
Tue, 24 Nov 2009 13:22:18 +0100 blanchet fix soundness bug in Nitpick's handling of negative literals (e.g., -1::rat)
Tue, 24 Nov 2009 10:33:02 +0100 blanchet fixed soundness bug / type error in handling of unpolarized (co)inductive predicates in Nitpick
Mon, 23 Nov 2009 18:29:00 +0100 blanchet fix Nitpick soundness bug related to "finite (UNIV::'a set)" where "'a" is constrained by a sort to be infinite
less more (0) -120 tip