src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
Thu, 09 Jun 2011 00:16:28 +0200 blanchet added a really fully typed translation as a fallback for Metis, in rare cases where Metis correctly proves a theorem but has type-unsound steps in it (which is likelier to happen with some of the lighter translations)
Wed, 08 Jun 2011 08:47:43 +0200 blanchet exploit new semantics of "max_new_instances"
Wed, 08 Jun 2011 08:47:43 +0200 blanchet make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
Wed, 08 Jun 2011 08:47:43 +0200 blanchet killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
Tue, 07 Jun 2011 14:38:42 +0200 blanchet prioritize more relevant facts for monomorphization
Tue, 07 Jun 2011 14:17:35 +0200 blanchet more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
Tue, 07 Jun 2011 14:17:35 +0200 blanchet workaround current "max_new_instances" semantics
Tue, 07 Jun 2011 11:13:52 +0200 blanchet move away from old SMT monomorphizer
Tue, 07 Jun 2011 11:05:09 +0200 blanchet merged
Tue, 07 Jun 2011 11:04:17 +0200 blanchet use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
Tue, 07 Jun 2011 10:46:09 +0200 blanchet removed "verbose" option -- there won't be any verbosity in there, according to Sascha (yeah)
Tue, 07 Jun 2011 10:43:18 +0200 blanchet nicely thread monomorphism verbosity in Metis and Sledgehammer
Tue, 07 Jun 2011 10:24:16 +0200 boehmes clarified meaning of monomorphization configuration option by renaming it
Tue, 07 Jun 2011 08:52:35 +0200 blanchet obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
Tue, 07 Jun 2011 07:45:12 +0200 blanchet use new monomorphization code
Tue, 07 Jun 2011 07:06:24 +0200 blanchet renamed ML function
Tue, 07 Jun 2011 06:58:52 +0200 blanchet pass props not thms to ATP translator
Mon, 06 Jun 2011 23:43:28 +0200 blanchet removed confusing slicing logic
Mon, 06 Jun 2011 23:26:40 +0200 blanchet suggest first reconstructor that timed out, not last (i.e. metis not metisFT in most cases)
Mon, 06 Jun 2011 23:11:14 +0200 blanchet effectively reenable slices for SPASS and Vampire -- they were disabled by mistake
Mon, 06 Jun 2011 20:56:06 +0200 blanchet removed old optimization that isn't one anyone
Mon, 06 Jun 2011 20:56:06 +0200 blanchet Metis code cleanup
Mon, 06 Jun 2011 20:36:36 +0200 blanchet enable new Metis
Mon, 06 Jun 2011 20:36:35 +0200 blanchet more preparations towards hijacking Metis
Mon, 06 Jun 2011 20:36:35 +0200 blanchet remove more occurrences of "metisX", preparing for the D Day when it will silently hijack "metis" and "metisFT"
Mon, 06 Jun 2011 20:36:35 +0200 blanchet make "metisX"'s default more like old "metis"
Mon, 06 Jun 2011 20:36:34 +0200 blanchet temporarily added "MetisX" as reconstructor and minimizer
Mon, 06 Jun 2011 20:36:34 +0200 blanchet show what failed to play
Wed, 01 Jun 2011 10:29:43 +0200 blanchet tuned names
Tue, 31 May 2011 16:38:36 +0200 blanchet tuned name
Tue, 31 May 2011 16:38:36 +0200 blanchet make "prepare_atp_problem" more robust w.r.t. choice of type system
Tue, 31 May 2011 16:38:36 +0200 blanchet don't preprocess twice
Tue, 31 May 2011 16:38:36 +0200 blanchet more work on new metis that exploits the powerful new type encodings
Tue, 31 May 2011 16:38:36 +0200 blanchet first step in sharing more code between ATP and Metis translation
Mon, 30 May 2011 17:00:38 +0200 blanchet don't slice if there are too few facts
Mon, 30 May 2011 17:00:38 +0200 blanchet automatically minimize with Metis when this can be done within a few seconds
Mon, 30 May 2011 17:00:38 +0200 blanchet minimize with Metis if possible
Mon, 30 May 2011 17:00:38 +0200 blanchet support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
Sun, 29 May 2011 19:40:56 +0200 blanchet always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
Fri, 27 May 2011 10:30:08 +0200 blanchet cleanup proof text generation code
Fri, 27 May 2011 10:30:08 +0200 blanchet try both "metis" and (on failure) "metisFT" in replay
Fri, 27 May 2011 10:30:08 +0200 blanchet show time taken for reconstruction
Fri, 27 May 2011 10:30:08 +0200 blanchet more concise output
Fri, 27 May 2011 10:30:08 +0200 blanchet handle non-auto try case of Sledgehammer better
Fri, 27 May 2011 10:30:08 +0200 blanchet renamed "metis_timeout" to "preplay_timeout" and continued implementation
Fri, 27 May 2011 10:30:07 +0200 blanchet added syntax for specifying Metis timeout (currently used only by SMT solvers)
Fri, 27 May 2011 10:30:07 +0200 blanchet make output more concise
Fri, 27 May 2011 10:30:07 +0200 blanchet merge timeout messages from several ATPs into one message to avoid clutter
Fri, 27 May 2011 10:30:07 +0200 blanchet fix soundness bug in Sledgehammer: distinguish params in goals from fixed variables in context
Fri, 27 May 2011 10:30:07 +0200 blanchet fully support all type system encodings in typed formats (TFF, THF)
Fri, 27 May 2011 10:30:07 +0200 blanchet towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
Tue, 24 May 2011 17:04:35 +0200 blanchet respect user's intention better when setting the CNF_UEQ type system (esp. w.r.t. "erased")
Tue, 24 May 2011 10:01:03 +0200 blanchet more work on parsing LEO-II proofs and extracting uses of extensionality
Tue, 24 May 2011 00:01:33 +0200 blanchet slightly gracefuller handling of LEO-II and Satallax output
Tue, 24 May 2011 00:01:33 +0200 blanchet identify HOL functions with THF functions
Tue, 24 May 2011 00:01:33 +0200 blanchet started adding support for THF output (but no lambdas)
Tue, 24 May 2011 00:01:33 +0200 blanchet filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
Tue, 24 May 2011 00:01:33 +0200 blanchet tuning -- the "appropriate" terminology is inspired from TPTP
Sun, 22 May 2011 14:51:42 +0200 blanchet improved Waldmeister support -- even run it by default on unit equational goals
Sun, 22 May 2011 14:51:41 +0200 blanchet fish out axioms in Waldmeister output
less more (0) -100 -60 tip