src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
Fri, 18 Nov 2011 11:47:12 +0100 blanchet removed more clutter
Fri, 18 Nov 2011 11:47:12 +0100 blanchet removed needless baggage
Wed, 16 Nov 2011 17:06:14 +0100 blanchet give each time slice its own lambda translation
Wed, 16 Nov 2011 16:35:19 +0100 blanchet thread in additional options to minimizer
Wed, 16 Nov 2011 13:22:36 +0100 blanchet make metis reconstruction handling more flexible
Wed, 16 Nov 2011 09:42:27 +0100 blanchet parse lambda translation option in Metis
Tue, 15 Nov 2011 22:15:51 +0100 blanchet rename configuration option to more reasonable length
Tue, 15 Nov 2011 22:13:39 +0100 blanchet started implementing lambda-lifting in Metis
Sun, 06 Nov 2011 22:18:54 +0100 blanchet more millisecond cleanup
Sun, 06 Nov 2011 22:18:54 +0100 blanchet try "smt" as a fallback for ATPs if "metis" fails/times out
Sun, 06 Nov 2011 22:18:54 +0100 blanchet more detailed preplay output
Sun, 06 Nov 2011 22:18:54 +0100 blanchet tuning
Sun, 06 Nov 2011 13:57:17 +0100 blanchet use "Time.time" rather than milliseconds internally
Sun, 06 Nov 2011 13:46:26 +0100 blanchet tune: no need for option
Sat, 29 Oct 2011 13:15:58 +0200 blanchet added DFG unsorted support (like in the old days)
Sat, 29 Oct 2011 13:15:58 +0200 blanchet added sorted DFG output for coming version of SPASS
Sat, 29 Oct 2011 13:15:58 +0200 blanchet check "sound" flag before doing something unsound...
Tue, 13 Sep 2011 11:24:58 +0200 blanchet simplified unsound proof detection by removing impossible case
Sat, 10 Sep 2011 00:44:25 +0200 blanchet continue with minimization in debug mode in spite of unsoundness
Fri, 02 Sep 2011 14:43:20 +0200 blanchet renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
Fri, 02 Sep 2011 14:43:20 +0200 blanchet fewer TPTP important messages
Thu, 01 Sep 2011 13:18:27 +0200 blanchet always measure time for ATPs -- auto minimization relies on it
Thu, 01 Sep 2011 13:18:27 +0200 blanchet make "sound" sound and "unsound" more sound, based on evaluation
Tue, 30 Aug 2011 16:25:10 +0200 blanchet fixed just introduced silly bug
Tue, 30 Aug 2011 16:11:42 +0200 blanchet tuning
Tue, 30 Aug 2011 16:07:45 +0200 blanchet flip logic of boolean option so it's off by default
Tue, 30 Aug 2011 14:12:55 +0200 nik improved handling of induction rules in Sledgehammer
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
less more (0) -100 -60 tip