src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
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
Sun, 22 May 2011 14:51:01 +0200 blanchet added support for remote Waldmeister
Sun, 22 May 2011 14:49:35 +0200 blanchet reorganized ATP formats a little bit
Fri, 20 May 2011 12:47:59 +0200 blanchet make sure the Vampire incomplete proof detection code kicks in
Fri, 20 May 2011 12:47:59 +0200 blanchet automatically use "metisFT" when typed helpers are necessary
Fri, 20 May 2011 12:47:58 +0200 blanchet more informative message when Sledgehammer finds an unsound proof
Thu, 19 May 2011 10:24:13 +0200 blanchet tweaked ATP type systems further based on Judgment Day
Thu, 19 May 2011 10:24:13 +0200 blanchet reintroduce TFF workaround for limitations of actual TFF implementations (ToFoF, SNARK)
Tue, 17 May 2011 15:11:36 +0200 blanchet renamed thin to light, fat to heavy
Tue, 17 May 2011 15:11:36 +0200 blanchet run blacklist algorithm only if slicing is on
Tue, 17 May 2011 15:11:36 +0200 blanchet renamed "shallow" to "thin" and make it the default
Tue, 17 May 2011 15:11:36 +0200 blanchet added syntax for "shallow" encodings
Fri, 13 May 2011 10:10:43 +0200 blanchet tuned comment
Fri, 13 May 2011 10:10:43 +0200 blanchet fixed off-by-one bug
Thu, 12 May 2011 15:29:19 +0200 blanchet robustly detect how many type args were passed to the ATP, even if some of them were omitted
Thu, 12 May 2011 15:29:19 +0200 blanchet further lower penalty associated with existentials in Sledgehammer's relevance filter, so that "exhaust" facts are picked up
Thu, 12 May 2011 15:29:19 +0200 blanchet renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
Thu, 12 May 2011 15:29:19 +0200 blanchet remove unused parameter
Thu, 12 May 2011 15:29:19 +0200 blanchet reduced penalty associated with existential quantifiers
Thu, 12 May 2011 15:29:19 +0200 blanchet fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
Thu, 12 May 2011 15:29:19 +0200 blanchet improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
Thu, 12 May 2011 15:29:19 +0200 blanchet tune whitespace
Thu, 12 May 2011 15:29:18 +0200 blanchet added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
Thu, 12 May 2011 15:29:18 +0200 blanchet allow each slice to have its own type system
Thu, 12 May 2011 15:29:18 +0200 blanchet renamed type systems for more consistency
Fri, 06 May 2011 13:34:59 +0200 blanchet allow each prover to specify its own formula kind for symbols occurring in the conjecture
Thu, 05 May 2011 12:40:48 +0200 blanchet no lies in debug output (e.g. "slice 2 of 1")
Wed, 04 May 2011 22:56:33 +0200 blanchet renamed "many_typed" to "simple" (as in simple types)
Wed, 04 May 2011 22:47:13 +0200 blanchet added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
Tue, 03 May 2011 21:46:49 +0200 blanchet fixed per-ATP dangerous axiom detection -- embarrassing bugs introduced in change a7a30721767a
Tue, 03 May 2011 01:04:03 +0200 blanchet no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
Tue, 03 May 2011 00:10:22 +0200 blanchet replaced some Unsynchronized.refs with Config.Ts
Mon, 02 May 2011 22:52:15 +0200 blanchet tuning
less more (0) -100 -60 tip