src/HOL/Tools/Nitpick/nitpick_hol.ML
Sat, 19 Dec 2015 20:02:51 +0100 blanchet removed dead code
Tue, 01 Dec 2015 22:24:37 +0100 blanchet removed needless ML function
Tue, 13 Oct 2015 09:21:15 +0200 haftmann prod_case as canonical name for product type eliminator
Mon, 05 Oct 2015 15:57:25 +0200 blanchet avoid unsound simplification of (C (s x)) when s is a selector but not C's
Sun, 06 Sep 2015 22:14:51 +0200 haftmann prefer "uncurry" as canonical name for case distinction on products in combinatorial view
Sun, 16 Aug 2015 18:19:30 +0200 wenzelm prefer theory_id operations;
Mon, 01 Jun 2015 18:59:21 +0200 haftmann separate class for division operator, with particular syntax added in more specific classes
Wed, 08 Apr 2015 19:39:08 +0200 wenzelm proper context for Object_Logic operations;
Fri, 06 Mar 2015 15:58:56 +0100 wenzelm Thm.cterm_of and Thm.ctyp_of operate on local context;
Wed, 04 Mar 2015 19:53:18 +0100 wenzelm tuned signature -- prefer qualified names;
Sat, 24 Jan 2015 22:00:24 +0100 wenzelm tuned signature;
Wed, 26 Nov 2014 20:05:34 +0100 wenzelm renamed "pairself" to "apply2", in accordance to @{apply 2};
Mon, 24 Nov 2014 12:35:13 +0100 blanchet no need for subset operation as a primitive in Nitpick, esp. that its implementation was unsound (cf. Rene Thiemann's 16 Oct 2014 email on isabelle mailing list)
Fri, 31 Oct 2014 11:36:41 +0100 wenzelm discontinued obsolete Output.urgent_message;
Wed, 08 Oct 2014 17:09:07 +0200 wenzelm added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
Tue, 19 Aug 2014 09:34:41 +0200 blanchet tuning
Sat, 16 Aug 2014 22:14:57 +0200 wenzelm updated to named_theorems;
Thu, 12 Jun 2014 01:00:49 +0200 blanchet removed subsumed record-related code, now that records are registered as 'ctr_sugar'
Thu, 12 Jun 2014 01:00:49 +0200 blanchet made lookup more robust in the face of missing (dummy) case constant
Fri, 21 Mar 2014 20:33:56 +0100 wenzelm more qualified names;
Fri, 21 Mar 2014 12:34:50 +0100 wenzelm more qualified names;
Wed, 19 Mar 2014 22:10:33 +0100 wenzelm more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
Sat, 15 Mar 2014 11:59:18 +0100 wenzelm tuned signature;
Mon, 03 Mar 2014 23:05:30 +0100 blanchet fixed handling of 'case_prod' and other 'case' functions for interpreted types
Mon, 03 Mar 2014 22:33:22 +0100 blanchet support 'datatype_new'-defined datatypes in Nitpick + better support for 'codatatype's
Mon, 03 Mar 2014 22:33:22 +0100 blanchet tuned ML names
Mon, 03 Mar 2014 22:33:22 +0100 blanchet tuned code
Mon, 03 Mar 2014 22:33:22 +0100 blanchet removed nonstandard models from Nitpick
Mon, 03 Mar 2014 12:58:17 +0100 blanchet guard against unsound cases that arise when people peek into 'int' and similar types that are handled specially by Nitpick
Wed, 19 Feb 2014 08:34:34 +0100 blanchet adapted Nitpick to 'primrec' refactoring
Mon, 17 Feb 2014 22:54:38 +0100 blanchet simplified data structure by reducing the incidence of clumsy indices
Wed, 12 Feb 2014 08:35:57 +0100 blanchet renamed '{prod,sum,bool,unit}_case' to 'case_...'
Mon, 20 Jan 2014 19:51:56 +0100 blanchet have Nitpick lookup codatatypes
Thu, 16 Jan 2014 16:20:17 +0100 blanchet adapted to move of Wfrec
Thu, 19 Dec 2013 13:43:21 +0100 blanchet made timeouts in Sledgehammer not be 'option's -- simplified lots of code
Tue, 19 Nov 2013 10:05:53 +0100 haftmann eliminiated neg_numeral in favour of - (numeral _)
Sun, 10 Nov 2013 10:02:34 +0100 haftmann dropped obsolete check: dest_num always yields positive number
Tue, 24 Sep 2013 11:02:42 +0200 blanchet encode goal digest in spying log (to detect duplicates)
Mon, 23 Sep 2013 18:40:02 +0200 blanchet don't generalize w.r.t. wrong context -- better overgeneralize (since the instantiation phase will compensate for it)
Tue, 28 May 2013 19:59:54 +0200 blanchet don't create needless constant
Mon, 15 Apr 2013 10:41:03 +0200 blanchet not all Nitpick 'constructors' are injective -- careful
Fri, 15 Feb 2013 08:31:31 +0100 haftmann two target language numeral types: integer and natural, as replacement for code_numeral;
Wed, 31 Oct 2012 11:23:21 +0100 blanchet moved Refute to "HOL/Library" to speed up building "Main" even more
Fri, 12 Oct 2012 15:08:29 +0200 wenzelm discontinued typedef with implicit set_def;
Wed, 15 Aug 2012 11:04:56 +0200 blanchet fixed handling of "int" in the wake of its port to the quotient package
Wed, 15 Aug 2012 11:04:55 +0200 blanchet removed dead code
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
less more (0) -120 tip