src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML
Wed, 15 Feb 2023 10:56:23 +0100 blanchet added refute mode to Sledgehammer to find 'counterexamples'
Sat, 22 Jan 2022 11:33:31 +0100 desharna added spying of fact filtering timing
Mon, 10 Jan 2022 14:13:23 +0100 desharna proper abstraction of function variables when instantiating induction rules in Sledgehammer
Sun, 19 Dec 2021 11:15:21 +0100 desharna merged
Fri, 17 Dec 2021 09:57:22 +0100 desharna added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
Fri, 17 Dec 2021 16:36:42 +0100 blanchet tuned comment
Tue, 28 Sep 2021 22:08:51 +0200 wenzelm clarified antiquotations;
Wed, 14 Jul 2021 16:09:57 +0200 blanchet prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
Wed, 14 Jul 2021 15:18:11 +0200 blanchet tuning
Thu, 08 Jul 2021 17:23:01 +0200 desharna refactored Sledgehammer option "induction_rules"
Thu, 08 Jul 2021 15:25:30 +0200 desharna promoted "sledgehammer_instantiate_inducts" to proper option "induction_rules"
Tue, 08 Sep 2020 11:32:57 +0200 desharna [mirabelle] add initial documentation in Sledgehammer's doc
Fri, 29 Nov 2019 20:57:04 +0100 wenzelm more informative spec rules: optional name;
Tue, 26 Mar 2019 14:23:18 +0100 wenzelm clarified signature: avoid direct comparison on type rough_classification;
Sat, 05 Jan 2019 17:24:33 +0100 wenzelm isabelle update -u control_cartouches;
Fri, 04 Jan 2019 23:22:53 +0100 wenzelm isabelle update -u control_cartouches;
Thu, 01 Feb 2018 15:12:57 +0100 wenzelm tuned signature: more operations;
Thu, 07 Sep 2017 23:13:15 +0200 blanchet better duplicate detection
Mon, 10 Apr 2017 21:05:31 +0200 wenzelm tuned signature;
Wed, 26 Oct 2016 22:40:28 +0200 blanchet tuning
Tue, 21 Jun 2016 14:42:47 +0200 wenzelm position information for literal facts;
Sat, 05 Mar 2016 17:01:45 +0100 wenzelm tuned signature -- clarified modules;
Thu, 07 Jan 2016 16:29:41 +0100 wenzelm tuned signature;
Thu, 07 Jan 2016 15:53:39 +0100 wenzelm more uniform treatment of package internals;
Sun, 20 Dec 2015 13:06:26 +0100 wenzelm renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
Thu, 12 Nov 2015 21:12:09 +0100 blanchet use cartouches instead of backquotes
Sun, 04 Oct 2015 17:48:34 +0200 blanchet speed up MaSh duplicate check
Sun, 30 Aug 2015 17:34:29 +0200 wenzelm trim context for persistent storage;
Sun, 16 Aug 2015 18:19:30 +0200 wenzelm prefer theory_id operations;
Sun, 16 Aug 2015 14:48:37 +0200 wenzelm clarified context;
Sun, 16 Aug 2015 11:46:08 +0200 wenzelm tuned signature;
Thu, 13 Aug 2015 11:05:19 +0200 wenzelm tuned signature, in accordance to sortBy in Scala;
Thu, 11 Jun 2015 22:47:53 +0200 wenzelm support for 'consider' command;
Sun, 03 May 2015 22:15:29 +0200 blanchet improved one-line preplaying (don't rely on 'using x by simp' to mean 'by (simp add: x)' and beware of inaccessible '(local.)this')
Wed, 08 Apr 2015 19:39:08 +0200 wenzelm proper context for Object_Logic operations;
Fri, 03 Apr 2015 18:36:19 +0200 wenzelm tuned;
Wed, 01 Apr 2015 10:35:43 +0200 wenzelm tuned signature;
Thu, 19 Mar 2015 22:30:57 +0100 wenzelm more position information;
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};
Fri, 07 Nov 2014 16:36:55 +0100 wenzelm plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
Thu, 06 Nov 2014 13:36:19 +0100 wenzelm prefer explicit Keyword.keywords;
Mon, 08 Sep 2014 14:04:03 +0200 blanchet never include hidden names -- these cannot be referenced afterward
Thu, 28 Aug 2014 20:05:39 +0200 blanchet gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
Tue, 19 Aug 2014 23:17:51 +0200 wenzelm tuned signature -- moved type src to Token, without aliases;
Mon, 18 Aug 2014 17:19:58 +0200 blanchet reordered some (co)datatype property names for more consistency
Sat, 16 Aug 2014 21:11:08 +0200 wenzelm updated to named_theorems;
Fri, 01 Aug 2014 14:43:57 +0200 blanchet whitespace tuning
Fri, 01 Aug 2014 14:43:57 +0200 blanchet remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
Fri, 01 Aug 2014 14:43:57 +0200 blanchet generate backquotes without markup, since this confuses preplay; bump up spying version identifier;
Tue, 01 Jul 2014 16:47:10 +0200 blanchet added hidden check to Sledgehammer fact filters, to avoid picking up facts like 'Nat.nat_induct0'
Tue, 01 Jul 2014 16:47:10 +0200 blanchet whitespace tuning
Wed, 28 May 2014 17:42:36 +0200 blanchet more generous max number of suggestions, for more safety
Fri, 21 Mar 2014 20:33:56 +0100 wenzelm more qualified names;
Sat, 15 Mar 2014 11:22:25 +0100 wenzelm more explicit treatment of verbose mode, which includes concealed entries;
Fri, 14 Mar 2014 12:23:59 +0100 wenzelm back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
Fri, 14 Mar 2014 10:08:36 +0100 wenzelm just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
Mon, 10 Mar 2014 18:06:23 +0100 wenzelm clarified Args.check_src: retain name space information for error output;
Mon, 10 Mar 2014 15:20:41 +0100 wenzelm prefer Name_Space.pretty with its builtin markup;
Thu, 06 Mar 2014 10:12:47 +0100 wenzelm tuned signature;
Mon, 03 Feb 2014 15:33:18 +0100 blanchet tuning
Fri, 31 Jan 2014 16:10:39 +0100 blanchet tuning
Sun, 26 Jan 2014 13:45:40 +0100 wenzelm tuned signature;
Sat, 25 Jan 2014 22:06:07 +0100 wenzelm explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
Thu, 02 Jan 2014 19:07:36 +0100 blanchet removed pointless warning (cf. http://stackoverflow.com/questions/20233463/isabelle-metis-proof-state-contains-the-universal-sort/20785045#20785045)
Tue, 19 Nov 2013 18:34:04 +0100 blanchet tuning
Fri, 18 Oct 2013 00:05:31 +0200 blanchet make sure add: doesn't add duplicates, and works for [no_atp] facts
Thu, 17 Oct 2013 23:41:00 +0200 blanchet no fact subsumption -- this only confuses later code, e.g. 'add:'
Tue, 15 Oct 2013 15:26:58 +0200 blanchet drop only real duplicates, not subsumed facts -- this confuses MaSh
Wed, 09 Oct 2013 16:38:48 +0200 blanchet normalize more equalities
Wed, 09 Oct 2013 08:28:36 +0200 blanchet added TODO
Wed, 09 Oct 2013 08:12:53 +0200 blanchet crank up limit a bit -- truly huge background theories are still nearly 3 times larger
Tue, 08 Oct 2013 21:19:46 +0200 blanchet minor fact filter speedups
Tue, 08 Oct 2013 20:56:35 +0200 blanchet more gracefully handle huge theories in relevance filters
Tue, 08 Oct 2013 16:40:03 +0200 blanchet further optimization in relevance filter
Tue, 08 Oct 2013 14:53:33 +0200 blanchet further speed up duplicate elimination
Tue, 08 Oct 2013 14:41:25 +0200 blanchet more efficient theorem variable normalization
Wed, 02 Oct 2013 22:54:42 +0200 blanchet strengthen top sort check
Tue, 24 Sep 2013 11:02:42 +0200 blanchet encode goal digest in spying log (to detect duplicates)
Wed, 11 Sep 2013 18:37:47 +0200 blanchet reintroduced 8d8f72aa5c0b, which does make a small difference in practice, but implemented more efficiently
Wed, 11 Sep 2013 14:07:24 +0200 blanchet tuning
Wed, 11 Sep 2013 14:07:24 +0200 blanchet disable some checks for huge background theories
Wed, 11 Sep 2013 14:07:24 +0200 blanchet tuning
Wed, 11 Sep 2013 14:07:24 +0200 blanchet reintroduced half of f99ee3adb81d -- that part definitely looks useless (and is inefficient)
Wed, 11 Sep 2013 14:07:24 +0200 blanchet reverted f99ee3adb81d -- that old logic seems to make a difference still today
Tue, 10 Sep 2013 15:56:52 +0200 blanchet faster detection of tautologies
Tue, 10 Sep 2013 15:56:51 +0200 blanchet slight speed optimization
Tue, 10 Sep 2013 15:56:51 +0200 blanchet got rid of another slowdown factor in relevance filter
Tue, 10 Sep 2013 15:56:51 +0200 blanchet removed completely needless, inefficient code
Tue, 10 Sep 2013 15:56:51 +0200 blanchet minor speed optimization
Tue, 10 Sep 2013 15:56:51 +0200 blanchet got rid of another taboo that appears to make no difference in practice (and that slows down the relevance filter)
Tue, 10 Sep 2013 15:56:51 +0200 blanchet avoid double traversal of term
Tue, 10 Sep 2013 15:56:51 +0200 blanchet got rid of old, needless logic
Tue, 10 Sep 2013 15:56:51 +0200 blanchet faster uniquification
Tue, 10 Sep 2013 15:56:51 +0200 blanchet stronger fact normalization
Tue, 10 Sep 2013 15:56:51 +0200 blanchet gracefully handle huge thys
Tue, 10 Sep 2013 15:56:51 +0200 blanchet speed up detection of simp rules
Thu, 16 May 2013 13:34:13 +0200 blanchet tuning -- renamed '_from_' to '_of_' in Sledgehammer
Wed, 15 May 2013 17:43:42 +0200 blanchet renamed Sledgehammer functions with 'for' in their names to 'of'
Fri, 15 Feb 2013 11:47:33 +0100 haftmann systematic conversions between nat and nibble/char;
Thu, 14 Feb 2013 15:27:10 +0100 haftmann reform of predicate compiler / quickcheck theories:
Thu, 31 Jan 2013 17:54:05 +0100 blanchet distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
Thu, 10 Jan 2013 23:34:20 +0100 blanchet make name table work the way it was intended to
Sun, 06 Jan 2013 17:38:29 +0100 blanchet put single-theorem names before multi-theorem ones (broken since 5d147d492792)
Sat, 05 Jan 2013 11:24:09 +0100 blanchet tuned blacklisting in relevance filter
Fri, 04 Jan 2013 21:56:20 +0100 blanchet refined class handling, to prevent cycles in fact graph
Fri, 04 Jan 2013 21:56:19 +0100 blanchet learn from low-level, inside-class facts
Fri, 04 Jan 2013 21:56:19 +0100 blanchet tuning
Mon, 17 Dec 2012 22:06:28 +0100 blanchet add a timeout in induction rule instantiation
Thu, 13 Dec 2012 23:47:01 +0100 blanchet get rid of some junk facts in the MaSh evaluation driver
Wed, 12 Dec 2012 21:48:29 +0100 blanchet tweaked which facts are included for MaSh evaluations
Wed, 12 Dec 2012 21:48:29 +0100 blanchet don't query blacklisted theorems in evaluation driver
Wed, 12 Dec 2012 21:48:29 +0100 blanchet export a pair of ML functions
Wed, 12 Dec 2012 15:38:47 +0100 blanchet further fix related to bd9a0028b063 -- that change was per se right, but it exposed a bug in the pattern for "all"
Wed, 12 Dec 2012 15:25:17 +0100 blanchet better tautology check -- don't reject "prod_cases3" for example
Wed, 12 Dec 2012 13:28:23 +0100 blanchet really all facts means really all facts (well, almost)
Wed, 12 Dec 2012 02:47:45 +0100 blanchet merge aliased theorems in MaSh dependencies, modulo symmetry of equality
Wed, 12 Dec 2012 00:14:58 +0100 blanchet push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
Sat, 08 Dec 2012 21:54:28 +0100 blanchet don't blacklist "case" theorems -- this causes problems in MaSh later
less more (0) -120 tip