Wed, 28 Apr 2010 16:56:18 +0200 haftmann avoid code_datatype antiquotation
Wed, 28 Apr 2010 19:46:09 +0200 bulwahn merged
Wed, 28 Apr 2010 16:45:51 +0200 bulwahn added an example with a free function variable to the Predicate Compile examples
Wed, 28 Apr 2010 16:45:50 +0200 bulwahn removed local clone in the predicate compiler
Wed, 28 Apr 2010 16:45:48 +0200 bulwahn improving proof procedure for transforming cases rule in the predicate compiler to handle free variables of function type
Thu, 29 Apr 2010 17:47:53 +0200 wenzelm allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
Thu, 29 Apr 2010 17:29:53 +0200 wenzelm 'write': actually observe the proof structure (like 'let' or 'fix');
Thu, 29 Apr 2010 17:15:23 +0200 wenzelm adapted ProofContext.infer_type;
Thu, 29 Apr 2010 16:55:22 +0200 wenzelm ProofContext.read_const: allow for type constraint (for fixed variable);
Thu, 29 Apr 2010 16:53:08 +0200 wenzelm avoid clash with keyword 'write';
Thu, 29 Apr 2010 11:05:13 +0200 wenzelm allow mixfix syntax for fixes within a proof body -- should now work thanks to fully authentic syntax;
Thu, 29 Apr 2010 11:00:32 +0200 wenzelm uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
Wed, 28 Apr 2010 19:43:45 +0200 wenzelm disabled spurious invocation of (interactive) sledgehammer;
Wed, 28 Apr 2010 17:29:58 +0200 wenzelm merged
Wed, 28 Apr 2010 16:56:03 +0200 blanchet make Mirabelle happy
Wed, 28 Apr 2010 16:47:56 +0200 blanchet remove removed option
Wed, 28 Apr 2010 16:15:45 +0200 blanchet merge
Wed, 28 Apr 2010 16:14:56 +0200 blanchet parentheses around nested cases
Wed, 28 Apr 2010 16:06:27 +0200 blanchet merged
Wed, 28 Apr 2010 16:05:38 +0200 blanchet add an Isar proof found with Sledgehammer that involves a Skolem constant (internally)
Wed, 28 Apr 2010 16:03:49 +0200 blanchet reintroduced short names for HOL->FOL constants; other parts of the code rely on these
Wed, 28 Apr 2010 15:53:17 +0200 blanchet save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
Wed, 28 Apr 2010 15:34:55 +0200 blanchet unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
Wed, 28 Apr 2010 14:19:26 +0200 blanchet redo Sledgehammer proofs (and get rid of "neg_clausify")
Wed, 28 Apr 2010 13:32:45 +0200 blanchet removed "sorts" option, continued
Wed, 28 Apr 2010 13:00:30 +0200 blanchet remove Sledgehammer's "sorts" option to annotate variables with sorts in proof;
Wed, 28 Apr 2010 12:49:52 +0200 blanchet insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
Wed, 28 Apr 2010 12:46:50 +0200 blanchet support Vampire definitions of constants as "let" constructs in Isar proofs
Tue, 27 Apr 2010 18:58:05 +0200 blanchet tuning
Tue, 27 Apr 2010 18:07:51 +0200 blanchet redid the proofs with the latest Sledgehammer;
Tue, 27 Apr 2010 18:02:46 +0200 blanchet remove Nitpick functions that are now implemented in Sledgehammer
Tue, 27 Apr 2010 18:01:41 +0200 blanchet added total goal count as argument + message when killing ATPs
Tue, 27 Apr 2010 17:44:33 +0200 blanchet make Sledgehammer more friendly if no subgoal is left
Tue, 27 Apr 2010 17:05:39 +0200 blanchet polish Isar proofs: don't mention facts twice, and don't show one-liner "structured" proofs
Tue, 27 Apr 2010 16:12:51 +0200 blanchet reintroduce missing "gen_all_vars" call
Tue, 27 Apr 2010 16:00:20 +0200 blanchet fix types of "fix" variables to help proof reconstruction and aid readability
Tue, 27 Apr 2010 14:55:10 +0200 blanchet allow schematic variables in types in terms that are reconstructed by Sledgehammer
Tue, 27 Apr 2010 14:27:47 +0200 blanchet in Sledgehammer "debug" mode, the names of most variables are already short and sweet, so most of the entries of the "const_trans_table" don't have a raison d'etre anymore
Tue, 27 Apr 2010 12:07:07 +0200 blanchet new Isar proof construction code: stringfy axiom names correctly
Tue, 27 Apr 2010 11:44:01 +0200 blanchet honor "shrink_proof" Sledgehammer option
Tue, 27 Apr 2010 11:24:47 +0200 blanchet remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
Wed, 28 Apr 2010 15:42:10 +0200 haftmann updated keywords
Wed, 28 Apr 2010 15:17:13 +0200 haftmann exported cert_tyco, read_tyco
Wed, 28 Apr 2010 15:17:09 +0200 haftmann added code_reflect command
Wed, 28 Apr 2010 14:54:17 +0200 haftmann merged
Wed, 28 Apr 2010 11:26:10 +0200 haftmann fix "fors" for proof of monotonicity
Wed, 28 Apr 2010 14:01:54 +0200 Cezary Kaliszyk merge
Wed, 28 Apr 2010 14:01:13 +0200 Cezary Kaliszyk merge
Wed, 28 Apr 2010 13:29:40 +0200 Cezary Kaliszyk Tuned FSet
Wed, 28 Apr 2010 13:30:52 +0200 haftmann merged
Wed, 28 Apr 2010 13:30:34 +0200 haftmann try to observe intended meaning of add_registration interface more closely
Wed, 28 Apr 2010 13:30:17 +0200 haftmann codified comment
Wed, 28 Apr 2010 13:29:57 +0200 haftmann merged
Wed, 28 Apr 2010 13:29:39 +0200 haftmann empty class specifcations observe default sort
Wed, 28 Apr 2010 16:56:51 +0200 wenzelm document some known problems with Mac OS;
Wed, 28 Apr 2010 16:12:21 +0200 wenzelm removed redundant/ignored sort constraint;
Wed, 28 Apr 2010 16:11:13 +0200 wenzelm tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
Wed, 28 Apr 2010 13:32:00 +0200 wenzelm made SML/NJ happy;
Wed, 28 Apr 2010 12:23:14 +0200 wenzelm updated keywords;
Wed, 28 Apr 2010 12:21:55 +0200 wenzelm command 'defaultsort' is renamed to 'default_sort', it works within a local theory context;
Wed, 28 Apr 2010 12:18:49 +0200 wenzelm removed material that is out of scope of this manual;
Wed, 28 Apr 2010 12:07:52 +0200 wenzelm renamed command 'defaultsort' to 'default_sort';
Wed, 28 Apr 2010 11:41:27 +0200 wenzelm localized default sort;
Wed, 28 Apr 2010 11:13:11 +0200 wenzelm more systematic naming of tsig operations;
Wed, 28 Apr 2010 11:09:19 +0200 wenzelm modernized/simplified Sign.set_defsort;
Wed, 28 Apr 2010 10:51:34 +0200 wenzelm get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
Wed, 28 Apr 2010 10:43:08 +0200 wenzelm export Type.minimize_sort;
Wed, 28 Apr 2010 08:25:02 +0200 haftmann term_typ: print styled term
Tue, 27 Apr 2010 22:23:12 +0200 wenzelm merged
Tue, 27 Apr 2010 11:17:50 -0700 huffman merged
Tue, 27 Apr 2010 11:03:04 -0700 huffman generalize types of path operations
Tue, 27 Apr 2010 10:54:24 -0700 huffman generalize more continuity lemmas
Tue, 27 Apr 2010 10:39:52 -0700 huffman generalized many lemmas about continuity
Mon, 26 Apr 2010 22:21:03 -0700 huffman simplify definition of continuous_on; generalize some lemmas
Mon, 26 Apr 2010 20:03:01 -0700 huffman move intervals section heading
Mon, 26 Apr 2010 19:58:51 -0700 huffman remove unused, redundant constant inv_on
Mon, 26 Apr 2010 19:55:50 -0700 huffman reorganize subsection headings
Mon, 26 Apr 2010 17:56:39 -0700 huffman remove redundant lemma
Mon, 26 Apr 2010 16:28:58 -0700 huffman more lemmas to Vec1.thy
Mon, 26 Apr 2010 15:51:10 -0700 huffman simplify proof
Mon, 26 Apr 2010 15:44:54 -0700 huffman move more lemmas into Vec1.thy
Mon, 26 Apr 2010 15:22:03 -0700 huffman move proof of Fashoda meet theorem into separate file
Mon, 26 Apr 2010 12:19:57 -0700 huffman move definitions and theorems for type real^1 to separate theory file
Tue, 27 Apr 2010 21:53:55 +0200 wenzelm removed obsolete sanity check -- Sign.certify_sort is stable;
Tue, 27 Apr 2010 21:46:10 +0200 wenzelm monotonic sort certification: sorts are no longer minimized at the kernel boundary, only when reading input from the end-user;
Tue, 27 Apr 2010 21:34:22 +0200 wenzelm really minimize sorts after certification -- looks like this is intended here;
Tue, 27 Apr 2010 19:44:04 +0200 wenzelm tuned signature;
Tue, 27 Apr 2010 16:24:57 +0200 wenzelm merged
Tue, 27 Apr 2010 12:20:17 +0200 haftmann tuned whitespace
Tue, 27 Apr 2010 12:20:09 +0200 haftmann got rid of [simplified]
Tue, 27 Apr 2010 11:52:41 +0200 haftmann got rid of [simplified]
Tue, 27 Apr 2010 10:51:39 +0200 blanchet fix SML/NJ compilation (I hope)
Tue, 27 Apr 2010 16:09:15 +0200 wenzelm tuned classrel completion -- bypass composition with reflexive edges;
Tue, 27 Apr 2010 15:23:05 +0200 wenzelm tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
Tue, 27 Apr 2010 15:03:19 +0200 wenzelm tuned aritiy completion -- slightly less intermediate data structures;
Tue, 27 Apr 2010 14:41:27 +0200 wenzelm clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
Tue, 27 Apr 2010 14:19:47 +0200 wenzelm misc tuning and simplification;
Tue, 27 Apr 2010 10:42:41 +0200 haftmann NEWS and CONTRIBUTORS
Tue, 27 Apr 2010 09:49:40 +0200 haftmann explicit is better than implicit
Tue, 27 Apr 2010 09:49:36 +0200 haftmann tuned class linordered_field_inverse_zero
Tue, 27 Apr 2010 08:18:25 +0200 haftmann merged
Tue, 27 Apr 2010 08:17:40 +0200 haftmann instances for *_inverse_zero classes
Tue, 27 Apr 2010 08:17:39 +0200 haftmann canonical import
Mon, 26 Apr 2010 15:38:14 +0200 haftmann merged
Mon, 26 Apr 2010 15:37:50 +0200 haftmann use new classes (linordered_)field_inverse_zero
Mon, 26 Apr 2010 23:46:45 +0200 blanchet merged
Mon, 26 Apr 2010 23:45:51 +0200 blanchet renamed option
Mon, 26 Apr 2010 23:45:32 +0200 blanchet fixes 2a5c6e7b55cb;
Mon, 26 Apr 2010 21:50:36 +0200 blanchet compile
Mon, 26 Apr 2010 21:41:54 +0200 blanchet make compile (and not just load dynamically)
Mon, 26 Apr 2010 21:25:32 +0200 blanchet merge
Mon, 26 Apr 2010 21:20:43 +0200 blanchet introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
Mon, 26 Apr 2010 21:18:20 +0200 blanchet adapt code to reflect new signature of "neg_clausify"
Mon, 26 Apr 2010 21:17:41 +0200 blanchet rename options and keep track of conjecture shape (to facilitate proof reconstruction)
Mon, 26 Apr 2010 21:17:04 +0200 blanchet rename options
Mon, 26 Apr 2010 21:16:35 +0200 blanchet make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
Mon, 26 Apr 2010 21:14:28 +0200 blanchet remove needless code that was copy-pasted from Quickcheck (where it made sense)
Sun, 25 Apr 2010 15:30:33 +0200 blanchet cosmetics
Sun, 25 Apr 2010 15:04:20 +0200 blanchet cosmetics
Sun, 25 Apr 2010 14:40:36 +0200 blanchet move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
Sun, 25 Apr 2010 11:38:46 +0200 blanchet support readable names even when Isar proof reconstruction is enabled -- useful for debugging
Sun, 25 Apr 2010 10:22:31 +0200 blanchet cosmetics: rename internal functions
Sun, 25 Apr 2010 00:33:26 +0200 blanchet cosmetics
Sun, 25 Apr 2010 00:25:44 +0200 blanchet remove "show_skolems" option and change style of record declarations
Sun, 25 Apr 2010 00:10:30 +0200 blanchet remove "skolemize" option from Nitpick, since Skolemization is always useful
Sat, 24 Apr 2010 17:48:21 +0200 blanchet removed Nitpick's "uncurry" option
Sat, 24 Apr 2010 16:44:45 +0200 blanchet fix typesetting
Sat, 24 Apr 2010 16:43:03 +0200 blanchet Fruhjahrsputz: remove three mostly useless Nitpick options
Sat, 24 Apr 2010 16:33:01 +0200 blanchet remove type annotations as comments;
Sat, 24 Apr 2010 16:17:30 +0200 blanchet cosmetics
Sat, 24 Apr 2010 16:05:42 +0200 blanchet better error reporting;
Fri, 23 Apr 2010 19:36:49 +0200 blanchet cosmetics
Fri, 23 Apr 2010 19:26:39 +0200 blanchet reuse timestamp function
Fri, 23 Apr 2010 19:18:39 +0200 blanchet stop referring to Sledgehammer_Util stuff all over Nitpick code; instead, redeclare any needed function in Nitpick_Util as synonym for the Sledgehammer_Util function of the same name
Fri, 23 Apr 2010 19:16:52 +0200 blanchet make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
Fri, 23 Apr 2010 19:12:49 +0200 blanchet remove some bloat
Fri, 23 Apr 2010 18:11:41 +0200 blanchet now rename the file "atp_wrapper.ML" to "atp_systems.ML" + fix typo in "SystemOnTPTP" script
Fri, 23 Apr 2010 18:06:41 +0200 blanchet renamed module "ATP_Wrapper" to "ATP_Systems"
Fri, 23 Apr 2010 17:38:25 +0200 blanchet move the minimizer to the Sledgehammer directory
Fri, 23 Apr 2010 16:59:48 +0200 blanchet remove debugging code
Fri, 23 Apr 2010 16:55:51 +0200 blanchet move some sledgehammer stuff out of "atp_manager.ML"
Fri, 23 Apr 2010 16:21:47 +0200 blanchet give an error if no ATP is set
Fri, 23 Apr 2010 16:15:35 +0200 blanchet move the Sledgehammer menu options to "sledgehammer_isar.ML"
Fri, 23 Apr 2010 15:48:34 +0200 blanchet centralized ATP-specific error handling in "atp_wrapper.ML"
Fri, 23 Apr 2010 13:16:50 +0200 blanchet handle ATP proof delimiters in a cleaner, more extensible fashion
Mon, 26 Apr 2010 22:59:28 +0200 wenzelm merged
Mon, 26 Apr 2010 11:08:49 -0700 huffman fix another if-then-else parse error
Mon, 26 Apr 2010 10:57:04 -0700 huffman fix if-then-else parse error
Mon, 26 Apr 2010 09:45:22 -0700 huffman merged
Mon, 26 Apr 2010 09:37:46 -0700 huffman fix syntax precedence declarations for UNION, INTER, SUP, INF
Mon, 26 Apr 2010 09:26:31 -0700 huffman syntax precedence for If and Let
Mon, 26 Apr 2010 09:21:25 -0700 huffman fix lots of looping simp calls and other warnings
Sun, 25 Apr 2010 23:22:29 -0700 huffman fix duplicate simp rule warnings
Sun, 25 Apr 2010 20:48:19 -0700 huffman define finer-than ordering on net type; move some theorems into Limits.thy
Sun, 25 Apr 2010 16:23:40 -0700 huffman generalize type of continuous_on
Sun, 25 Apr 2010 11:58:39 -0700 huffman define nets directly as filters, instead of as filter bases
Mon, 26 Apr 2010 21:50:28 +0200 wenzelm use 'example_proof' (invisible);
Mon, 26 Apr 2010 21:45:08 +0200 wenzelm command 'example_proof' opens an empty proof body;
Mon, 26 Apr 2010 21:36:44 +0200 wenzelm proofs that are discontinued via 'oops' are treated as relevant --- for improved robustness of the final join of all proofs, which is hooked to results that are missing here;
Mon, 26 Apr 2010 20:30:50 +0200 wenzelm eliminanated some unreferenced identifiers;
Mon, 26 Apr 2010 16:08:04 +0200 wenzelm merged
Mon, 26 Apr 2010 15:14:14 +0200 Cezary Kaliszyk add bounded_lattice_bot and bounded_lattice_top type classes
Mon, 26 Apr 2010 13:43:31 +0200 haftmann merged
Mon, 26 Apr 2010 11:34:19 +0200 haftmann dropped group_simps, ring_simps, field_eq_simps
Mon, 26 Apr 2010 11:34:17 +0200 haftmann class division_ring_inverse_zero
Mon, 26 Apr 2010 11:34:15 +0200 haftmann dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
Mon, 26 Apr 2010 11:34:15 +0200 haftmann line break
Mon, 26 Apr 2010 14:44:41 +0200 wenzelm removed unused AxClass.class_intros;
Mon, 26 Apr 2010 11:20:18 +0200 wenzelm updated Sign.add_type_abbrev;
Mon, 26 Apr 2010 07:47:18 +0200 haftmann merged
Sun, 25 Apr 2010 08:25:34 +0200 haftmann field_simps as named theorems
Sun, 25 Apr 2010 23:26:40 +0200 wenzelm merged
Sun, 25 Apr 2010 10:23:03 -0700 huffman generalize more constants and lemmas
Sun, 25 Apr 2010 09:01:03 -0700 huffman simplify types of path operations (use real instead of real^1)
Sun, 25 Apr 2010 07:41:57 -0700 huffman add lemmas convex_real_interval and convex_box
Sat, 24 Apr 2010 21:29:22 -0700 huffman generalize more constants and lemmas
Sat, 24 Apr 2010 19:32:20 -0700 huffman generalize constant closest_point
Sat, 24 Apr 2010 14:06:19 -0700 huffman minimize imports
Sat, 24 Apr 2010 13:34:11 -0700 huffman fix imports
Sat, 24 Apr 2010 13:31:52 -0700 huffman document generation for Multivariate_Analysis
Sat, 24 Apr 2010 11:11:09 -0700 huffman move l2-norm stuff into separate theory file
Sat, 24 Apr 2010 09:37:24 -0700 huffman convert proofs to Isar-style
Sat, 24 Apr 2010 09:34:36 -0700 huffman Library/Fraction_Field.thy: ordering relations for fractions
Sun, 25 Apr 2010 23:09:32 +0200 wenzelm renamed Drule.unconstrainTs to Thm.unconstrain_allTs to accomdate the version by krauss/schropp;
Sun, 25 Apr 2010 22:50:47 +0200 wenzelm more systematic treatment of data -- avoid slightly odd nested tuples here;
Sun, 25 Apr 2010 21:18:04 +0200 wenzelm replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
Sun, 25 Apr 2010 21:02:36 +0200 wenzelm misc tuning and simplification;
Sun, 25 Apr 2010 19:44:47 +0200 wenzelm simplified some private bindings;
Sun, 25 Apr 2010 19:09:37 +0200 wenzelm classrel and arity completion by krauss/schropp;
Sun, 25 Apr 2010 16:10:05 +0200 wenzelm removed obsolete/unused Proof.match_bind;
Sun, 25 Apr 2010 15:52:03 +0200 wenzelm modernized naming conventions of main Isar proof elements;
Sun, 25 Apr 2010 15:13:33 +0200 wenzelm goals: simplified handling of implicit variables -- removed obsolete warning;
Fri, 23 Apr 2010 23:42:46 +0200 wenzelm updated generated files;
Fri, 23 Apr 2010 23:38:01 +0200 wenzelm cover 'schematic_lemma' etc.;
Fri, 23 Apr 2010 23:35:43 +0200 wenzelm mark schematic statements explicitly;
Fri, 23 Apr 2010 23:33:48 +0200 wenzelm eliminated spurious schematic statements;
Fri, 23 Apr 2010 22:48:07 +0200 wenzelm explicit 'schematic_theorem' etc. for schematic theorem statements;
Fri, 23 Apr 2010 22:39:49 +0200 wenzelm collapse category "schematic goal" in keyword table -- Proof General does not know about this;
Fri, 23 Apr 2010 21:00:28 +0200 wenzelm added keyword category "schematic goal", which prevents any attempt to fork the proof;
Fri, 23 Apr 2010 19:36:23 +0200 wenzelm merged
Fri, 23 Apr 2010 19:32:37 +0200 haftmann merged
Fri, 23 Apr 2010 16:45:53 +0200 haftmann separated instantiation of division_by_zero
Fri, 23 Apr 2010 16:38:51 +0200 haftmann epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
Fri, 23 Apr 2010 16:17:25 +0200 haftmann adapted to new times_divide_eq simp situation
Fri, 23 Apr 2010 16:17:25 +0200 haftmann epheremal replacement of field_simps by field_eq_simps
Fri, 23 Apr 2010 16:17:24 +0200 haftmann dequalified fact name
Fri, 23 Apr 2010 15:18:00 +0200 haftmann sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
Fri, 23 Apr 2010 15:18:00 +0200 haftmann separated instantiation of division_by_zero
Fri, 23 Apr 2010 15:17:59 +0200 haftmann dequalified fact name
Fri, 23 Apr 2010 15:17:59 +0200 haftmann less special treatment of times_divide_eq [simp]
Fri, 23 Apr 2010 15:17:59 +0200 haftmann sharpened constraint (c.f. 4e7f5b22dd7d)
Fri, 23 Apr 2010 13:58:15 +0200 haftmann more localization; tuned proofs
Fri, 23 Apr 2010 13:58:14 +0200 haftmann more localization; factored out lemmas for division_ring
Fri, 23 Apr 2010 13:58:14 +0200 haftmann modernized abstract algebra theories
Fri, 23 Apr 2010 10:50:17 +0200 haftmann merged
Thu, 22 Apr 2010 12:07:00 +0200 haftmann swapped generic code generator and SML code generator
Fri, 23 Apr 2010 19:36:04 +0200 wenzelm removed obsolete Named_Thm_Set -- Named_Thms provides efficient member operation;
Fri, 23 Apr 2010 18:30:01 +0200 wenzelm Item_Net/Named_Thms: export efficient member operation;
Fri, 23 Apr 2010 16:12:57 +0200 bulwahn initialized atps reference with standard setup in the atp manager
Fri, 23 Apr 2010 12:24:30 +0200 blanchet compile
Fri, 23 Apr 2010 12:07:12 +0200 blanchet handle SPASS's equality predicate correctly in Isar proof reconstruction
Fri, 23 Apr 2010 11:34:49 +0200 blanchet merged
Fri, 23 Apr 2010 11:32:36 +0200 blanchet added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
Thu, 22 Apr 2010 16:30:54 +0200 blanchet remove hack that is no longer necessary now that "ATP_Wrapper" properly detects which ATPs are installed
Thu, 22 Apr 2010 16:30:04 +0200 blanchet set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
Thu, 22 Apr 2010 15:01:36 +0200 blanchet minor code cleanup
Thu, 22 Apr 2010 14:47:52 +0200 blanchet if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
Thu, 22 Apr 2010 13:50:58 +0200 blanchet "remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
Thu, 22 Apr 2010 10:54:56 +0200 blanchet Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
Thu, 22 Apr 2010 10:13:02 +0200 blanchet postprocess ATP output in "overlord" mode to make it more readable
Wed, 21 Apr 2010 17:06:26 +0200 blanchet failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
Wed, 21 Apr 2010 16:38:03 +0200 blanchet generate command-line in addition to timestamp in ATP output file, for debugging purposes
Wed, 21 Apr 2010 16:21:19 +0200 blanchet pass relevant options from "sledgehammer" to "sledgehammer minimize";
Fri, 23 Apr 2010 10:00:53 +0200 Cezary Kaliszyk Finite set theory
Thu, 22 Apr 2010 22:12:12 +0200 wenzelm merged
Thu, 22 Apr 2010 20:39:48 +0100 paulson Tidied up using s/l
Thu, 22 Apr 2010 22:01:06 +0200 wenzelm split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
Thu, 22 Apr 2010 11:55:19 +0200 Cezary Kaliszyk fun_rel introduction and list_rel elimination for quotient package
Thu, 22 Apr 2010 09:30:39 +0200 haftmann lemmas concerning remdups
Thu, 22 Apr 2010 09:30:36 +0200 haftmann lemma dlist_ext
(0) -30000 -10000 -3000 -1000 -240 +240 +1000 +3000 +10000 +30000 tip