src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
Sat, 18 Dec 2010 23:31:22 +0100 blanchet two layers of timeouts seem to be less reliable than a single layer
Sat, 18 Dec 2010 22:15:39 +0100 blanchet move relevance filter into hard timeout
Sat, 18 Dec 2010 21:24:34 +0100 blanchet handle timeouts in Mirabelle more gracefully
Sat, 18 Dec 2010 13:48:24 +0100 blanchet higher hard timeout
Sat, 18 Dec 2010 13:38:14 +0100 blanchet compile
Sat, 18 Dec 2010 12:55:33 +0100 blanchet use minimizing prover in Mirabelle
Fri, 17 Dec 2010 23:09:53 +0100 blanchet remove option that doesn't work in Mirabelle anyway (Mirabelle uses Sledgehammer_Provers, not Sledgehammer_Run)
Fri, 17 Dec 2010 18:23:56 +0100 blanchet added debugging option to find out how good the relevance filter was at identifying relevant facts
Fri, 17 Dec 2010 16:20:02 +0100 blanchet compile
Thu, 16 Dec 2010 15:44:32 +0100 blanchet removed unused variable
Thu, 16 Dec 2010 15:12:17 +0100 blanchet tuning
Wed, 15 Dec 2010 11:26:29 +0100 blanchet added support for "type_sys" option to Mirabelle
Wed, 15 Dec 2010 11:26:29 +0100 blanchet honor "metisFT" in Mirabelle
Wed, 15 Dec 2010 11:26:28 +0100 blanchet implemented partially-typed "tags" type encoding
Wed, 08 Dec 2010 22:17:53 +0100 blanchet implicitly call the minimizer for SMT solvers that don't return an unsat core
Wed, 08 Dec 2010 22:17:52 +0100 blanchet renamings
Wed, 08 Dec 2010 22:17:52 +0100 blanchet split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
Mon, 06 Dec 2010 11:41:24 +0100 blanchet handle "max_relevant" uniformly
Fri, 03 Dec 2010 18:29:14 +0100 blanchet replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
Thu, 25 Nov 2010 14:13:48 +0100 blanchet reverted c059d550afec -- the triviality check had apparently nothing to do with spontaneous Interrupt exceptions
Tue, 23 Nov 2010 23:10:13 +0100 blanchet disable triviality checking -- it might be the source of the spurious Interrupt exceptions that make it almost impossible to run Judgement Day
Tue, 23 Nov 2010 19:01:21 +0100 blanchet make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
Sat, 20 Nov 2010 00:53:26 +0100 wenzelm renamed raw "explode" function to "raw_explode" to emphasize its meaning;
Mon, 15 Nov 2010 18:56:30 +0100 blanchet turn on Sledgehammer verbosity so we can track down crashes
Sat, 13 Nov 2010 19:55:45 +0100 wenzelm total Symbol.source;
Sun, 07 Nov 2010 18:19:04 +0100 blanchet always use a hard timeout in Mirabelle
Sun, 07 Nov 2010 18:15:13 +0100 blanchet use "smt" (rather than "metis") to reconstruct SMT proofs
Fri, 05 Nov 2010 09:05:22 +0100 blanchet make Mirabelle work correctly if the prover (e.g. the SMT solver) returns no timing information
Thu, 04 Nov 2010 15:30:48 +0100 blanchet remove " s" suffix since seconds are now implicit
Thu, 04 Nov 2010 14:59:44 +0100 blanchet use the SMT integration's official list of built-ins
Tue, 02 Nov 2010 20:55:12 +0100 wenzelm simplified some time constants;
Tue, 26 Oct 2010 21:01:28 +0200 blanchet standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
Tue, 26 Oct 2010 16:56:54 +0200 blanchet remove needless context argument;
Tue, 26 Oct 2010 13:50:57 +0200 blanchet proper error handling for SMT solvers in Sledgehammer
Fri, 22 Oct 2010 18:31:45 +0200 blanchet tuning
Fri, 22 Oct 2010 15:02:27 +0200 blanchet generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
Fri, 22 Oct 2010 14:47:43 +0200 blanchet replaced references with proper record that's threaded through
Fri, 22 Oct 2010 14:10:32 +0200 blanchet fixed signature of "is_smt_solver_installed";
Fri, 22 Oct 2010 13:48:21 +0200 blanchet remove more needless code ("run_smt_solvers");
Fri, 22 Oct 2010 11:58:33 +0200 blanchet bring ATPs and SMT solvers more in line with each other
Fri, 22 Oct 2010 11:11:34 +0200 blanchet make Sledgehammer minimizer fully work with SMT
Fri, 22 Oct 2010 09:50:18 +0200 blanchet generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
Thu, 21 Oct 2010 16:25:40 +0200 blanchet first step in adding support for an SMT backend to Sledgehammer
Thu, 21 Oct 2010 14:55:09 +0200 blanchet use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
Thu, 16 Sep 2010 16:12:02 +0200 blanchet rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
Tue, 14 Sep 2010 23:38:36 +0200 blanchet tuning
Tue, 14 Sep 2010 16:34:26 +0200 blanchet handle relevance filter corner cases more gracefully;
Tue, 14 Sep 2010 14:47:53 +0200 blanchet added a timeout around "try" call in Mirabelle
Mon, 13 Sep 2010 15:11:10 +0200 blanchet indicate triviality in the list of proved things
Mon, 13 Sep 2010 15:01:31 +0200 blanchet indicate which goals are trivial
Mon, 13 Sep 2010 14:29:05 +0200 blanchet keep track of trivial vs. nontrivial calls using "try" for 30 seconds
Sat, 11 Sep 2010 10:24:57 +0200 blanchet make Mirabelle happy
Wed, 01 Sep 2010 23:10:01 +0200 blanchet minor refactoring
Wed, 01 Sep 2010 23:04:47 +0200 blanchet translate the axioms to FOF once and for all ATPs
Wed, 01 Sep 2010 22:33:31 +0200 blanchet run relevance filter in a thread, to avoid blocking
Wed, 01 Sep 2010 18:41:23 +0200 blanchet share the relevance filter among the provers
Tue, 31 Aug 2010 23:50:59 +0200 blanchet finished renaming
Fri, 27 Aug 2010 15:37:03 +0200 blanchet drop chained facts
Thu, 26 Aug 2010 10:42:06 +0200 blanchet consider "locality" when assigning weights to facts
Tue, 24 Aug 2010 19:55:34 +0200 blanchet make Mirabelle happy
Wed, 18 Aug 2010 17:16:37 +0200 blanchet get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
Mon, 09 Aug 2010 12:05:48 +0200 blanchet move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
Thu, 29 Jul 2010 23:37:10 +0200 blanchet fix Mirabelle timeout
Thu, 29 Jul 2010 23:24:07 +0200 blanchet make Mirabelle happy
Thu, 29 Jul 2010 14:42:09 +0200 blanchet "axiom_clauses" -> "axioms" (these are no longer clauses)
Tue, 27 Jul 2010 20:16:03 +0200 blanchet compile
Tue, 27 Jul 2010 17:04:09 +0200 blanchet implemented "sublinear" minimization algorithm
Mon, 26 Jul 2010 14:14:24 +0200 blanchet make TPTP generator accept full first-order formulas
Tue, 29 Jun 2010 11:20:05 +0200 blanchet compile
Tue, 29 Jun 2010 11:14:52 +0200 blanchet compile
Mon, 28 Jun 2010 11:04:02 +0200 blanchet compile
Fri, 25 Jun 2010 17:08:39 +0200 blanchet renamed "Sledgehammer_FOL_Clauses" to "Metis_Clauses", so that Metis doesn't depend on Sledgehammer
Fri, 25 Jun 2010 16:42:06 +0200 blanchet merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
Tue, 22 Jun 2010 19:08:25 +0200 blanchet turn on "natural form" filtering in the Mirabelle tests, to see how it performs
Mon, 17 May 2010 23:54:15 +0200 wenzelm prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
Mon, 17 May 2010 15:11:25 +0200 wenzelm renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
Sun, 16 May 2010 00:02:11 +0200 wenzelm prefer structure Parse_Spec;
Wed, 28 Apr 2010 16:56:03 +0200 blanchet make Mirabelle happy
Mon, 26 Apr 2010 21:50:36 +0200 blanchet compile
Fri, 23 Apr 2010 16:55:51 +0200 blanchet move some sledgehammer stuff out of "atp_manager.ML"
Fri, 23 Apr 2010 12:24:30 +0200 blanchet compile
Mon, 19 Apr 2010 15:15:21 +0200 blanchet make Sledgehammer's minimizer also minimize Isar proofs
Thu, 01 Apr 2010 11:12:08 +0200 blanchet add missing goal argument
Thu, 25 Mar 2010 17:55:55 +0100 blanchet make Mirabelle happy again
Mon, 22 Mar 2010 10:38:28 +0100 blanchet remove the iteration counter from Sledgehammer's minimizer
Fri, 19 Mar 2010 15:33:18 +0100 blanchet move all ATP setup code into ATP_Wrapper
Fri, 19 Mar 2010 15:07:44 +0100 blanchet move the Sledgehammer Isar commands together into one file;
Thu, 18 Mar 2010 13:43:50 +0100 blanchet fix Mirabelle after renaming Sledgehammer structures
Fri, 05 Mar 2010 23:51:32 +0100 wenzelm use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
Thu, 10 Dec 2009 18:10:59 +0100 boehmes only invoke metisFT if metis failed
Tue, 08 Dec 2009 23:05:23 +0100 boehmes also consider the fully-typed version of metis for Mirabelle measurements
Thu, 29 Oct 2009 16:59:12 +0100 wenzelm modernized some structure names;
Wed, 28 Oct 2009 22:18:00 +0100 wenzelm renamed raw Proof.get_goal to Proof.raw_goal;
Tue, 27 Oct 2009 18:00:50 +0100 boehmes measure runtime of ATPs only if requested
Tue, 20 Oct 2009 12:06:17 +0200 boehmes proper exceptions instead of unhandled partiality
Sun, 18 Oct 2009 16:25:59 +0200 nipkow merged
Sun, 18 Oct 2009 16:25:04 +0200 nipkow added sums of squares for standard deviation
Sat, 17 Oct 2009 23:47:27 +0200 wenzelm made SML/NJ happy;
Thu, 15 Oct 2009 17:49:30 +0200 wenzelm natural argument order for prover;
Thu, 15 Oct 2009 17:06:19 +0200 wenzelm ATP_Manager.get_prover: canonical argument order;
Thu, 15 Oct 2009 11:49:27 +0200 wenzelm eliminated extraneous wrapping of public records;
Thu, 15 Oct 2009 00:55:29 +0200 wenzelm structure ATP_Manager: eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
Wed, 14 Oct 2009 23:44:37 +0200 wenzelm modernized structure names;
Sun, 04 Oct 2009 11:45:41 +0200 boehmes avoid exception Option: only apply "the" if needed
Sat, 03 Oct 2009 12:05:40 +0200 boehmes re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
Thu, 01 Oct 2009 15:19:23 +0200 nipkow resolved conflict
Thu, 01 Oct 2009 11:35:13 +0200 nipkow record max lemmas used
Tue, 29 Sep 2009 16:24:36 +0200 wenzelm explicit indication of Unsynchronized.ref;
Thu, 24 Sep 2009 17:25:42 +0200 nipkow record how many "proof"s are solved by s/h
Sat, 19 Sep 2009 10:19:12 +0200 nipkow restructured code
Fri, 18 Sep 2009 23:08:53 +0200 nipkow modified minimization log
Fri, 18 Sep 2009 14:40:06 +0200 nipkow skip &&& goals
Thu, 17 Sep 2009 19:13:07 +0200 nipkow removed misleading log line
Wed, 16 Sep 2009 12:47:14 +0200 nipkow revised lemma counting
Tue, 15 Sep 2009 15:29:11 +0200 boehmes added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
Mon, 14 Sep 2009 19:30:48 +0200 nipkow count number of iterations required for minimization (and fixed bug: minimization was always called)
Sun, 13 Sep 2009 02:10:41 +0200 wenzelm explicitly export type abbreviations (as usual in SML97);
Sat, 12 Sep 2009 16:30:48 +0200 wenzelm standard headers and text sections;
Thu, 10 Sep 2009 15:57:55 +0200 nipkow position information is now passed to all actions;
Thu, 10 Sep 2009 12:53:49 +0200 nipkow logging number of metis lemmas
less more (0) -120 tip