src/HOL/Tools/SMT/smt_solver.ML
Thu, 06 Feb 2025 12:07:47 +0100 wenzelm avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
Wed, 05 Feb 2025 21:29:13 +0100 wenzelm clarified signature: re-use existing Process_Result.T (NB: err_lines are de-facto program startup errors, anything else is redirected to out_lines);
Wed, 18 Dec 2024 10:21:58 +0100 desharna added option "cache_dir" to Sledgehammer
Fri, 20 Sep 2024 14:28:13 +0200 wenzelm clarified signature: more explicit operations;
Wed, 26 Jul 2023 12:04:25 +0200 Mathias Fleury support for let in Alethe name bindings;
Wed, 15 Feb 2023 10:56:23 +0100 blanchet added refute mode to Sledgehammer to find 'counterexamples'
Fri, 25 Mar 2022 13:52:23 +0100 blanchet first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
Fri, 11 Mar 2022 09:22:13 +0100 desharna used more descriptive assert names in SMT-Lib output
Mon, 07 Feb 2022 15:26:22 +0100 blanchet added possibility of extra options to SMT slices
Mon, 31 Jan 2022 16:09:23 +0100 blanchet crude implementation of centralized slicing
Wed, 17 Nov 2021 15:09:10 +0100 fleury generate problems with correct logic for veriT
Thu, 04 Nov 2021 16:50:03 +0100 Mathias Fleury proper support of verit's return code for timeout
Wed, 20 Oct 2021 18:13:17 +0200 wenzelm discontinued obsolete "val extend = I" for data slots;
Thu, 15 Oct 2020 14:42:05 +0200 wenzelm more standard Value.print_time;
Mon, 12 Oct 2020 18:59:44 +0200 Mathias Fleury add reconstruction for the SMT solver veriT
Fri, 02 Oct 2020 10:18:50 +0200 desharna Add more tacing to sledgehammer_isar_trace
Sun, 20 Sep 2020 20:47:59 +0200 wenzelm clarified signature;
Fri, 07 Jun 2019 11:08:29 +0200 blanchet handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
Mon, 27 May 2019 15:08:51 +0200 wenzelm more direct invocation of Windows exe: avoid extra bash, cygpath, exec;
Fri, 24 May 2019 20:16:35 +0200 wenzelm avoid extra subprocess -- potentially more robust on Cygwin;
Fri, 04 Jan 2019 23:22:53 +0100 wenzelm isabelle update -u control_cartouches;
Tue, 30 Oct 2018 16:24:04 +0100 fleury add reconstruction by veriT in method smt
Sun, 28 Jan 2018 19:28:52 +0100 wenzelm clarified take/drop/chop prefix/suffix;
Fri, 22 Sep 2017 13:46:11 -0300 blanchet real oracle
Tue, 18 Oct 2016 16:03:30 +0200 wenzelm clarified modules;
Fri, 20 May 2016 07:54:54 +0200 fleury better handling of veriT's 'unknown' status
Mon, 07 Mar 2016 21:09:28 +0100 wenzelm File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
Fri, 25 Sep 2015 20:37:59 +0200 wenzelm moved remaining display.ML to more_thm.ML;
Thu, 27 Aug 2015 22:36:09 +0200 blanchet generate proper error instead of exception if goal cannot be atomized
Sat, 18 Jul 2015 20:47:08 +0200 wenzelm prefer tactics with explicit context;
Wed, 08 Jul 2015 21:33:00 +0200 wenzelm clarified context;
Wed, 08 Jul 2015 19:28:43 +0200 wenzelm Variable.focus etc.: optional bindings provided by user;
Sat, 25 Apr 2015 09:48:06 +0200 blanchet made CVC4 support work also without unsat cores
Fri, 06 Mar 2015 23:44:57 +0100 wenzelm clarified context;
Fri, 06 Mar 2015 15:58:56 +0100 wenzelm Thm.cterm_of and Thm.ctyp_of operate on local context;
Thu, 28 Aug 2014 00:40:38 +0200 blanchet renamed new SMT module from 'SMT2' to 'SMT'
Wed, 02 Oct 2013 22:54:42 +0200 blanchet make SMT integration slacker w.r.t. bad apples (facts)
Sat, 27 Jul 2013 16:35:51 +0200 wenzelm standardized aliases;
Wed, 12 Dec 2012 03:47:02 +0100 blanchet made MaSh evaluation driver work with SMT solvers
Thu, 23 Aug 2012 13:03:29 +0200 wenzelm prefer classic take_prefix/take_suffix over chop_while (cf. 0659e84bdc5f);
Thu, 26 Jul 2012 11:08:16 +0200 blanchet Z3 prints so many warnings that the very informative abnormal termination exception hardly ever gets raised -- better be more aggressive here
Thu, 26 Jul 2012 10:48:03 +0200 blanchet Sledgehammer already has its own ways of reporting and recovering from crashes in external provers -- no need to additionally print scores of warnings (cf. 4b0daca2bf88)
Wed, 30 May 2012 23:10:42 +0200 boehmes introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
Mon, 23 Apr 2012 21:44:36 +0200 wenzelm more standard method setup;
Tue, 27 Mar 2012 16:59:13 +0300 blanchet renamed "smt_fixed" to "smt_read_only_certificates"
Thu, 01 Mar 2012 10:12:58 +0100 boehmes fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
Wed, 29 Feb 2012 17:43:41 +0100 boehmes SMT fails if the chosen SMT solver is not installed
Wed, 15 Feb 2012 23:19:30 +0100 wenzelm renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
Tue, 14 Feb 2012 19:18:57 +0100 wenzelm normalized aliases;
Mon, 02 May 2011 16:33:21 +0200 wenzelm added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
Fri, 08 Apr 2011 19:04:08 +0200 boehmes check if negating the goal is possible (avoids CTERM exception for Pure propositions)
Mon, 14 Feb 2011 10:40:43 +0100 boehmes more explicit errors to inform users about problems related to SMT solvers;
Mon, 17 Jan 2011 17:45:52 +0100 boehmes made Z3 the default SMT solver again
Mon, 10 Jan 2011 17:39:54 +0100 boehmes dropped superfluous error message
Sat, 08 Jan 2011 17:14:48 +0100 wenzelm misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
Thu, 06 Jan 2011 17:51:56 +0100 boehmes differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
Mon, 20 Dec 2010 22:02:57 +0100 boehmes avoid ML structure aliases (especially single-letter abbreviations)
Mon, 20 Dec 2010 08:17:23 +0100 boehmes perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
Fri, 17 Dec 2010 15:30:43 +0100 blanchet run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
Fri, 17 Dec 2010 12:10:08 +0100 blanchet make timeout part of the SMT filter's tail
less more (0) -100 -60 tip