src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
Wed, 09 Feb 2011 17:18:58 +0100 blanchet renamed field
Mon, 10 Jan 2011 15:45:46 +0100 wenzelm eliminated Int.toString;
Tue, 21 Dec 2010 10:18:56 +0100 blanchet added "sledgehammer_tac" as possible reconstructor in Mirabelle
Tue, 21 Dec 2010 06:53:20 +0100 blanchet avoid weird symbols in path
Tue, 21 Dec 2010 06:06:28 +0100 blanchet mechanism to keep SMT input and output files around in Mirabelle
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
less more (0) -100 -60 tip