src/HOL/Tools/Nitpick/nitpick_hol.ML
Thu, 24 May 2012 17:46:35 +0200 blanchet fixed soundness bug in Nitpick related to unfolding -- the unfolding criterion must at least as strict when looking at a definitional axiom as elsewhere, otherwise we end up unfolding a constant's definition in its own definition, yielding a trivial equality
Fri, 11 May 2012 00:45:24 +0200 blanchet fixed "real" after they were redefined as a 'quotient_type'
Thu, 26 Apr 2012 20:09:38 +0200 blanchet fixed bug in handling of new numerals -- the left-hand side of "Numeral1 = 1" should be left alone and not translated to a built-in Kodkod numeral in the specification of the "numeral" function
Sun, 22 Apr 2012 14:16:46 +0200 blanchet more meaningful default value
Sun, 22 Apr 2012 14:16:45 +0200 blanchet handle exception (needed to solve TPTP problem SEU880^5)
Tue, 03 Apr 2012 17:26:30 +0900 griff renamed "rel_comp" to "relcomp" (to be consistent with, e.g., "relpow")
Mon, 26 Mar 2012 10:42:06 +0200 blanchet fixed Nitpick after numeral representation change (2a1953f0d20d)
Sun, 25 Mar 2012 20:15:39 +0200 huffman merged fork with new numeral representation (see NEWS)
Sun, 04 Mar 2012 23:20:43 +0100 blanchet addressed a quotient-type-related issue that arose with the port to "set"
Thu, 01 Mar 2012 11:28:56 +0100 blanchet improved handling of polymorphic quotient types, by avoiding comparing types that will generally differ in the polymorphic case
Thu, 01 Mar 2012 11:28:56 +0100 blanchet fixed handling of "Rep_" function of quotient types -- side-effect of "set" constructor reintroduction
Mon, 23 Jan 2012 17:40:32 +0100 blanchet renamed two files to make room for a new file
Tue, 17 Jan 2012 18:25:36 +0100 blanchet fixed a bug introduced when porting functions to set -- extensionality on sets break the form of equations expected elsewhere by Nitpick
Sat, 14 Jan 2012 21:16:15 +0100 wenzelm discontinued old-style Term.list_abs in favour of plain Term.abs;
Sat, 14 Jan 2012 19:06:05 +0100 wenzelm renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
Wed, 04 Jan 2012 12:09:53 +0100 blanchet improved "set" support by code inspection
Wed, 04 Jan 2012 12:09:53 +0100 blanchet tuning
Tue, 03 Jan 2012 23:41:59 +0100 blanchet fixed bisimilarity axiom -- avoid "insert" with wrong type
Tue, 03 Jan 2012 18:33:18 +0100 blanchet more robust destruction of "set Collect" idiom
Tue, 03 Jan 2012 18:33:18 +0100 blanchet handle starred predicates correctly w.r.t. "set"
Tue, 03 Jan 2012 18:33:18 +0100 blanchet simplify mem Collect
Tue, 03 Jan 2012 18:33:18 +0100 blanchet fixed set extensionality code
Tue, 03 Jan 2012 18:33:18 +0100 blanchet construct correct "set" type for wf goal
Tue, 03 Jan 2012 18:33:18 +0100 blanchet fixed Nitpick's typedef handling w.r.t. "set"
Tue, 03 Jan 2012 18:33:18 +0100 blanchet rationalized output (a bit)
Tue, 03 Jan 2012 18:33:17 +0100 blanchet port part of Nitpick to "set" type constructor
Tue, 03 Jan 2012 18:33:17 +0100 blanchet ported mono calculus to handle "set" type constructors
Sat, 24 Dec 2011 16:14:58 +0100 haftmann adjusted to set/pred distinction by means of type constructor `set`
Fri, 09 Dec 2011 14:14:37 +0100 kuncar make ctxt the first parameter
Thu, 27 Oct 2011 21:02:10 +0200 wenzelm localized quotient data;
Thu, 27 Oct 2011 20:26:38 +0200 wenzelm simplified/standardized signatures;
Thu, 27 Oct 2011 13:50:54 +0200 bulwahn respecting isabelle's programming style in the quotient package by simplifying quotdata_lookup function for data access
Wed, 17 Aug 2011 18:05:31 +0200 wenzelm modernized signature of Term.absfree/absdummy;
Tue, 02 Aug 2011 13:07:00 +0200 krauss updated unchecked forward reference
Tue, 02 Aug 2011 12:27:24 +0200 krauss replaced Nitpick's hardwired basic_ersatz_table by context data
Tue, 02 Aug 2011 10:36:50 +0200 krauss moved recdef package to HOL/Library/Old_Recdef.thy
Tue, 02 Aug 2011 10:03:14 +0200 krauss added dynamic ersatz_table to Nitpick's data slot
Thu, 14 Jul 2011 16:50:05 +0200 blanchet added option to control which lambda translation to use (for experiments)
Tue, 31 May 2011 16:38:36 +0200 blanchet first step in sharing more code between ATP and Metis translation
Tue, 24 May 2011 00:01:33 +0200 blanchet fixed de Bruijn index bug
Fri, 13 May 2011 22:55:00 +0200 wenzelm proper Proof.context for classical tactics;
Thu, 05 May 2011 12:40:48 +0200 blanchet query typedefs as well for monotonicity
Wed, 04 May 2011 19:35:48 +0200 blanchet exploit inferred monotonicity
Wed, 04 May 2011 18:48:25 +0200 blanchet [mq]: nitpick_tuning
Wed, 04 May 2011 18:43:42 +0200 blanchet fixed cardinality computation for function types such as "'a -> unit"
Wed, 20 Apr 2011 13:54:07 +0200 wenzelm added Theory.nodes_of convenience;
Tue, 19 Apr 2011 14:38:38 +0200 blanchet avoid relying on "Thm.definitionK" to pick up definitions -- this was an old hack related to locales (to avoid expanding locale constants to their low-level definition) that is no longer necessary
Tue, 19 Apr 2011 14:04:58 +0200 blanchet use "Spec_Rules" for finding axioms -- more reliable and cleaner
Tue, 19 Apr 2011 12:22:59 +0200 blanchet optimize trivial equalities early in Nitpick -- it shouldn't be the job of the peephole optimizer
Sat, 16 Apr 2011 16:15:37 +0200 wenzelm modernized structure Proof_Context;
Fri, 18 Mar 2011 10:17:37 +0100 blanchet always destroy constructor patterns, since this seems to be always useful
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
less more (0) -120 tip