src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
Tue, 30 Aug 2011 14:12:55 +0200 nik added generation of induction rules
Fri, 26 Aug 2011 10:25:13 +0200 blanchet added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
Tue, 23 Aug 2011 16:07:01 +0200 blanchet exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
Tue, 23 Aug 2011 14:44:19 +0200 blanchet added formats to the slice and use TFF for remote Vampire
Mon, 22 Aug 2011 15:02:45 +0200 blanchet cleaner handling of polymorphic monotonicity inference
Mon, 22 Aug 2011 15:02:45 +0200 blanchet added option to control soundness of encodings more precisely, for evaluation purposes
Mon, 22 Aug 2011 15:02:45 +0200 blanchet make sound mode more sound (and clean up code)
Tue, 09 Aug 2011 09:05:22 +0200 blanchet move lambda-lifting code to ATP encoding, so it can be used by Metis
Thu, 28 Jul 2011 11:43:45 +0200 blanchet tuning
Tue, 26 Jul 2011 14:53:00 +0200 blanchet remove spurious message
Mon, 25 Jul 2011 14:10:12 +0200 blanchet introduced hybrid lambda translation
Mon, 25 Jul 2011 14:10:12 +0200 blanchet avoid needless type args for lifted-lambdas
Thu, 21 Jul 2011 21:29:10 +0200 blanchet make "concealed" lambda translation sound
Wed, 20 Jul 2011 23:47:27 +0200 blanchet use a more robust naming convention for "polymorphic" frees -- the check is an overapproximation but that's fine as far as soundness is concerned
Wed, 20 Jul 2011 12:23:20 +0200 boehmes generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
Wed, 20 Jul 2011 09:23:12 +0200 boehmes moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
Sun, 17 Jul 2011 14:21:19 +0200 blanchet fixed lambda-liftg: must ensure the formulas are in close form
Sun, 17 Jul 2011 14:12:45 +0200 blanchet ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
Sun, 17 Jul 2011 14:11:35 +0200 blanchet pass kind to lambda-translation function
Sun, 17 Jul 2011 14:11:35 +0200 blanchet added lambda-lifting to Sledgehammer (rough)
Sun, 17 Jul 2011 14:11:34 +0200 blanchet move more lambda-handling logic to Sledgehammer, from ATP module, for formal dependency reasons
Sat, 16 Jul 2011 20:52:41 +0200 wenzelm moved bash operations to Isabelle_System (cf. Scala version);
Thu, 14 Jul 2011 17:29:30 +0200 blanchet move error logic closer to user
Thu, 14 Jul 2011 16:50:05 +0200 blanchet move lambda translation option from ATP to Sledgehammer, to avoid accidentally breaking Metis (its reconstruction code can only deal with combinators)
Thu, 14 Jul 2011 15:14:37 +0200 blanchet clearer unsound message
Wed, 06 Jul 2011 17:19:34 +0100 blanchet better setup for experimental "z3_atp"
Sun, 03 Jul 2011 08:15:14 +0200 blanchet make SML/NJ happy
Fri, 01 Jul 2011 17:44:04 +0200 blanchet enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
Fri, 01 Jul 2011 15:53:38 +0200 blanchet renamed "type_sys" to "type_enc", which is more accurate
Thu, 30 Jun 2011 13:21:41 +0200 wenzelm standardized use of Path operations;
Mon, 27 Jun 2011 14:56:35 +0200 blanchet clarify minimizer output
Mon, 27 Jun 2011 14:56:28 +0200 blanchet added "sound" option to force Sledgehammer to be pedantically sound
Mon, 27 Jun 2011 13:52:47 +0200 blanchet removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
Tue, 21 Jun 2011 17:17:39 +0200 blanchet generate type predicates for existentials/skolems, otherwise some problems might not be provable
Mon, 20 Jun 2011 12:13:43 +0200 blanchet clean up SPASS FLOTTER hack
Mon, 20 Jun 2011 11:42:41 +0200 blanchet remove automatic recovery from (some) unsound proofs, now that we use sound encodings for all the interesting provers
Mon, 20 Jun 2011 10:41:02 +0200 blanchet deal with ATP time slices in a more flexible/robust fashion
Fri, 10 Jun 2011 17:52:09 +0200 blanchet name tuning
Fri, 10 Jun 2011 17:52:09 +0200 blanchet don't trim proofs in debug mode
Fri, 10 Jun 2011 12:01:15 +0200 blanchet pass --trim option to "eproof" script to speed up proof reconstruction
Thu, 09 Jun 2011 00:16:28 +0200 blanchet cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
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)
less more (0) -100 -60 tip