src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
Sat, 02 Apr 2016 23:29:05 +0200 wenzelm prefer infix operations;
Mon, 01 Feb 2016 20:55:23 +0100 blanchet avoid error in Isar proof reconstruction if no ATP proof is available
Mon, 29 Jun 2015 23:44:53 +0200 blanchet removed chained facts from preplaying -- and careful about extra chained facts when removing 'proof -' and 'qed' from one-line Isar proofs
Tue, 03 Mar 2015 16:37:45 +0100 blanchet SPASS-Pirate is now called Pirate
Fri, 31 Oct 2014 11:36:41 +0100 wenzelm discontinued obsolete Output.urgent_message;
Sun, 12 Oct 2014 21:52:45 +0200 blanchet improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
Thu, 02 Oct 2014 18:10:09 +0200 blanchet more precise lemma insertion
Thu, 02 Oct 2014 17:40:46 +0200 blanchet insert lemmas closer to where they are needed, both for esthetics and (primarily) for correctness in case the lemma refers to a skolem
Thu, 02 Oct 2014 15:24:51 +0200 blanchet eliminate duplicate hypotheses (which can arise due to (un)clausification)
Thu, 02 Oct 2014 11:01:20 +0200 blanchet 'moura' method is also useful for reconstructing skolemization of lambda-lifting of formulas for other provers than Z3
Mon, 15 Sep 2014 14:31:32 +0200 blanchet tuning
Mon, 15 Sep 2014 12:30:06 +0200 blanchet removed accidental '@{print}'
Tue, 02 Sep 2014 16:38:26 +0200 steckerm Some work on the new waldmeister integration
Thu, 28 Aug 2014 23:57:26 +0200 blanchet renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
Thu, 28 Aug 2014 16:58:27 +0200 blanchet removed show stuttering
Thu, 28 Aug 2014 16:58:27 +0200 blanchet try 'skolem' method first for Z3
Thu, 28 Aug 2014 16:58:27 +0200 blanchet added 'skolem' method, esp. for 'obtain's generated from Z3 proofs
Thu, 28 Aug 2014 00:40:38 +0200 blanchet renamed new SMT module from 'SMT2' to 'SMT'
Tue, 05 Aug 2014 10:59:40 +0200 blanchet rationalize Skolem names
Tue, 05 Aug 2014 10:17:15 +0200 blanchet tuning
Tue, 05 Aug 2014 09:20:00 +0200 blanchet tuning
Mon, 04 Aug 2014 16:53:09 +0200 blanchet deal with E definitions
Mon, 04 Aug 2014 15:02:02 +0200 blanchet cleaner 'compress' option
Mon, 04 Aug 2014 13:16:18 +0200 blanchet tuned terminology (cf. 'isar_proofs' option)
Mon, 04 Aug 2014 12:28:42 +0200 blanchet rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
Sat, 02 Aug 2014 00:15:08 +0200 blanchet better duplicate detection
Fri, 01 Aug 2014 23:58:42 +0200 blanchet normalize conjectures vs. negated conjectures when comparing terms
Fri, 01 Aug 2014 23:33:43 +0200 blanchet tweaked 'clone' formula detection
Fri, 01 Aug 2014 23:29:50 +0200 blanchet fine-tuned Isar reconstruction, esp. boolean simplifications
Fri, 01 Aug 2014 23:29:49 +0200 blanchet centralized boolean simplification so that e.g. LEO-II benefits from it
Fri, 01 Aug 2014 20:44:29 +0200 blanchet better handling of variable names
Fri, 01 Aug 2014 20:08:50 +0200 blanchet nicer generated variable names
Fri, 01 Aug 2014 19:44:18 +0200 blanchet tuning
Fri, 01 Aug 2014 19:36:23 +0200 blanchet tuning
Fri, 01 Aug 2014 19:32:10 +0200 blanchet more precise handling of LEO-II skolemization
Fri, 01 Aug 2014 16:07:34 +0200 blanchet beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
Fri, 01 Aug 2014 14:43:57 +0200 blanchet added appropriate method for skolemization of Z3 steps to the mix
Fri, 01 Aug 2014 14:43:57 +0200 blanchet tentatively took out 'fastforce' from the set of tried methods -- it seems to be largely subsumed and is hard to silence
Fri, 01 Aug 2014 14:43:57 +0200 blanchet tuning
Fri, 01 Aug 2014 14:43:57 +0200 blanchet eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
Wed, 30 Jul 2014 14:03:13 +0200 fleury Skolemization for tmp_ite_elim rule in the SMT solver veriT.
Wed, 30 Jul 2014 14:03:13 +0200 fleury Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
Wed, 30 Jul 2014 14:03:13 +0200 fleury Whitespace and indentation correction.
Wed, 30 Jul 2014 14:03:12 +0200 fleury imported patch hilbert_choice_support
Wed, 30 Jul 2014 14:03:12 +0200 fleury veriT changes for lifted terms, and ite_elim rules.
Wed, 30 Jul 2014 14:03:12 +0200 fleury Subproofs for the SMT solver veriT.
Wed, 30 Jul 2014 14:03:12 +0200 fleury Basic support for the SMT prover veriT.
Wed, 30 Jul 2014 14:03:11 +0200 fleury Skolemization support for leo-II and Zipperposition.
Wed, 30 Jul 2014 00:50:41 +0200 blanchet also try 'metis' with 'full_types'
Thu, 24 Jul 2014 18:46:38 +0200 blanchet refined filter for ATP steps to avoid 'have True' steps in E proofs
Thu, 24 Jul 2014 18:46:38 +0200 blanchet 'shift_quantors' is not an E skolemization rule (cf. 3ab503b04bdb)
Tue, 24 Jun 2014 08:19:22 +0200 blanchet supports gradual skolemization (cf. Z3) by threading context through correctly
Tue, 24 Jun 2014 08:19:22 +0200 blanchet given two one-liners, only show the best of the two
Tue, 24 Jun 2014 08:19:22 +0200 blanchet don't generate Isar proof skeleton for what amounts to a one-line proof
Thu, 12 Jun 2014 17:10:12 +0200 blanchet renamed Sledgehammer options
Mon, 02 Jun 2014 15:10:18 +0200 fleury basic setup for zipperposition prover
Thu, 22 May 2014 05:23:50 +0200 blanchet properly reconstruct helpers in Z3 proofs
Thu, 22 May 2014 03:29:36 +0200 blanchet tuning
Fri, 16 May 2014 19:14:00 +0200 blanchet correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
Fri, 16 May 2014 19:13:50 +0200 blanchet use 'simp add:' syntax in Sledgehammer rather than 'using'
Sun, 04 May 2014 19:01:36 +0200 blanchet added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
Fri, 14 Mar 2014 11:44:11 +0100 blanchet consolidate consecutive steps that prove the same formula
Fri, 14 Mar 2014 11:15:46 +0100 blanchet undo rewrite rules (e.g. for 'fun_app') in Isar
Fri, 14 Mar 2014 11:05:45 +0100 blanchet debugging stuff
Fri, 14 Mar 2014 11:05:44 +0100 blanchet more simplification of trivial steps
Fri, 14 Mar 2014 11:05:37 +0100 blanchet tuning
Thu, 13 Mar 2014 13:18:14 +0100 blanchet tuning
Thu, 13 Mar 2014 13:18:14 +0100 blanchet simplified preplaying information
Thu, 13 Mar 2014 13:18:13 +0100 blanchet integrate SMT2 with Sledgehammer
Thu, 06 Mar 2014 10:12:47 +0100 wenzelm tuned signature;
Thu, 13 Feb 2014 13:16:17 +0100 blanchet avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
Tue, 04 Feb 2014 23:11:18 +0100 blanchet more generous Isar proof compression -- try to remove failing steps
Tue, 04 Feb 2014 23:11:18 +0100 blanchet do a second phase of proof compression after minimization
Tue, 04 Feb 2014 23:11:18 +0100 blanchet tuned code
Tue, 04 Feb 2014 23:11:18 +0100 blanchet split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
Mon, 03 Feb 2014 23:59:36 +0100 blanchet rationalized lists of methods
Mon, 03 Feb 2014 23:49:01 +0100 blanchet extended method list
Mon, 03 Feb 2014 19:32:02 +0100 blanchet generate comments in Isar proofs
Mon, 03 Feb 2014 19:32:02 +0100 blanchet renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
Mon, 03 Feb 2014 19:32:02 +0100 blanchet tuned behavior of 'smt' option
Mon, 03 Feb 2014 19:32:02 +0100 blanchet proper fresh name generation
Mon, 03 Feb 2014 17:13:31 +0100 blanchet added 'smt' option to control generation of 'by smt' proofs
Mon, 03 Feb 2014 16:53:58 +0100 blanchet renamed ML file
Mon, 03 Feb 2014 15:33:18 +0100 blanchet tuning
Mon, 03 Feb 2014 15:19:07 +0100 blanchet merged 'reconstructors' and 'proof methods'
Mon, 03 Feb 2014 13:37:23 +0100 blanchet tuning
Mon, 03 Feb 2014 11:58:38 +0100 blanchet tuned data structure
Mon, 03 Feb 2014 11:37:48 +0100 blanchet tuned data structure
Mon, 03 Feb 2014 10:14:18 +0100 blanchet less aggressive evaluation
Mon, 03 Feb 2014 10:14:18 +0100 blanchet added a new version of 'metis' to the mix
Mon, 03 Feb 2014 10:14:18 +0100 blanchet implemented new 'try0_isar' semantics
Mon, 03 Feb 2014 10:14:18 +0100 blanchet got rid of 'try0' step that is now redundant
Mon, 03 Feb 2014 10:14:18 +0100 blanchet centralize more preplaying
Mon, 03 Feb 2014 10:14:18 +0100 blanchet centralize preplaying
Mon, 03 Feb 2014 10:14:18 +0100 blanchet tuned
Sun, 02 Feb 2014 20:53:51 +0100 blanchet more data structure rationalization
Sun, 02 Feb 2014 20:53:51 +0100 blanchet more data structure rationalization
Sun, 02 Feb 2014 20:53:51 +0100 blanchet rationalized threading of 'metis' arguments
Sun, 02 Feb 2014 20:53:51 +0100 blanchet refactored data structure (step 3)
Sun, 02 Feb 2014 20:53:51 +0100 blanchet unform treatment of preplay_timeout = 0 and > 0
Sun, 02 Feb 2014 20:53:51 +0100 blanchet use Skolem proof methods appropriately
Sun, 02 Feb 2014 20:53:51 +0100 blanchet simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
Fri, 31 Jan 2014 19:16:41 +0100 blanchet generalized preplaying infrastructure to store various results for various methods
Fri, 31 Jan 2014 18:43:16 +0100 blanchet added a 'trace' option
Fri, 31 Jan 2014 18:43:16 +0100 blanchet moved code around
Fri, 31 Jan 2014 18:43:16 +0100 blanchet added 'algebra' to the mix
Fri, 31 Jan 2014 18:43:16 +0100 blanchet more informative trace
Fri, 31 Jan 2014 16:41:54 +0100 blanchet better tracing + syntactically correct 'metis' calls
Fri, 31 Jan 2014 16:26:43 +0100 blanchet tuned ML function names
Fri, 31 Jan 2014 16:10:39 +0100 blanchet tuning
Fri, 31 Jan 2014 16:07:20 +0100 blanchet moved ML code around
Fri, 31 Jan 2014 10:23:32 +0100 blanchet renamed many Sledgehammer ML files to clarify structure
Thu, 30 Jan 2014 14:37:53 +0100 blanchet renamed Sledgehammer options for symmetry between positive and negative versions
Thu, 19 Dec 2013 13:43:21 +0100 blanchet made timeouts in Sledgehammer not be 'option's -- simplified lots of code
Thu, 21 Nov 2013 12:29:29 +0100 blanchet fixed spying so that the envirnoment variables are queried at run-time not at build-time
Tue, 19 Nov 2013 18:34:04 +0100 blanchet tuning
Thu, 17 Oct 2013 01:34:34 +0200 blanchet added comment
Tue, 15 Oct 2013 15:31:18 +0200 blanchet use MePo with Auto Sledgehammer, because it's lighter than MaSh and always available
Fri, 04 Oct 2013 11:52:10 +0200 blanchet run fewer provers in "try" mode
Fri, 04 Oct 2013 11:28:28 +0200 blanchet count remote threads as well when balancing CPU usage -- otherwise jEdit users and other users of the "blocking" mode may have to wait for 2 * timeout if they e.g. have 4 cores and 5 provers (the typical situation)
less more (0) -120 tip