Wed, 21 Oct 2009 16:54:04 +0200 blanchet fixed the "expect" mechanism of Refute in the face of timeouts
Wed, 21 Oct 2009 16:53:00 +0200 blanchet removed "nitpick_const_simp" attribute from Record's "simps";
Wed, 21 Oct 2009 15:54:31 +0200 haftmann merged
Wed, 21 Oct 2009 15:54:01 +0200 haftmann more accurate removal
Wed, 21 Oct 2009 12:12:21 +0200 haftmann merged
Wed, 21 Oct 2009 12:09:37 +0200 haftmann curried inter as canonical list operation (beware of argument order)
Wed, 21 Oct 2009 14:08:04 +0200 boehmes merged
Wed, 21 Oct 2009 12:19:46 +0200 boehmes proper handling of single literal case,
Wed, 21 Oct 2009 12:48:28 +0100 paulson Removed the hard-wired white list of theorems for sledgehammer
Wed, 21 Oct 2009 11:19:11 +0100 paulson merged
Tue, 20 Oct 2009 16:32:51 +0100 paulson Some new lemmas concerning sets
Wed, 21 Oct 2009 12:08:52 +0200 haftmann merged
Wed, 21 Oct 2009 12:02:56 +0200 haftmann curried union as canonical list operation
Wed, 21 Oct 2009 12:02:19 +0200 haftmann tuned ML import
Wed, 21 Oct 2009 10:15:31 +0200 haftmann removed old-style \ and \\ infixes
Wed, 21 Oct 2009 08:16:25 +0200 haftmann merged
Wed, 21 Oct 2009 08:14:38 +0200 haftmann dropped redundant gen_ prefix
Tue, 20 Oct 2009 16:13:01 +0200 haftmann replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
Wed, 21 Oct 2009 16:41:22 +1100 kleing find_theorems: better handling of abbreviations (by Timothy Bourke)
Wed, 21 Oct 2009 00:36:12 +0200 wenzelm standardized basic operations on type option;
Tue, 20 Oct 2009 23:25:04 +0200 wenzelm eliminated THENL -- use THEN RANGE;
Tue, 20 Oct 2009 22:46:24 +0200 wenzelm tuned;
Tue, 20 Oct 2009 21:37:06 +0200 wenzelm fixed SML/NJ toplevel pp;
Tue, 20 Oct 2009 21:26:45 +0200 wenzelm backpatching of structure Proof and ProofContext -- avoid odd aliases;
Tue, 20 Oct 2009 21:22:37 +0200 wenzelm tuned;
Tue, 20 Oct 2009 20:54:31 +0200 wenzelm uniform use of Integer.min/max;
Tue, 20 Oct 2009 20:03:23 +0200 wenzelm modernized session SET_Protocol;
Tue, 20 Oct 2009 19:52:04 +0200 wenzelm modernized session Metis_Examples;
Tue, 20 Oct 2009 19:37:09 +0200 wenzelm modernized session Isar_Examples;
Tue, 20 Oct 2009 19:36:52 +0200 wenzelm tuned header;
Tue, 20 Oct 2009 19:29:24 +0200 wenzelm more accurate dependencies for HOL-SMT, which is a session with image;
Tue, 20 Oct 2009 19:28:01 +0200 wenzelm removed unused map_force;
Tue, 20 Oct 2009 15:02:48 +0100 paulson Removal of the unused atpset concept, the atp attribute and some related code.
Tue, 20 Oct 2009 15:03:17 +0200 boehmes additional schematic rules for Z3's rewrite rule
Tue, 20 Oct 2009 14:44:19 +0200 nipkow merged
Tue, 20 Oct 2009 14:44:02 +0200 nipkow added Hilbert_Choice section
Tue, 20 Oct 2009 14:22:54 +0200 boehmes merged
Tue, 20 Oct 2009 14:22:02 +0200 boehmes eliminated extraneous wrapping of public records,
Tue, 20 Oct 2009 12:06:17 +0200 boehmes proper exceptions instead of unhandled partiality
Tue, 20 Oct 2009 13:37:56 +0200 nipkow footnote: inv via inv_onto
Tue, 20 Oct 2009 13:20:42 +0200 nipkow added inv_def for compatibility as a lemma
Tue, 20 Oct 2009 11:36:19 +0200 wenzelm slightly less context-sensitive settings;
Tue, 20 Oct 2009 11:08:50 +0200 wenzelm merged
Tue, 20 Oct 2009 10:29:47 +0200 boehmes corrected paths to certificates,
Tue, 20 Oct 2009 10:11:30 +0200 boehmes added proof reconstructon for Z3,
Tue, 20 Oct 2009 10:46:42 +0200 wenzelm merged
Tue, 20 Oct 2009 08:10:47 +0200 haftmann merged
Tue, 20 Oct 2009 08:10:31 +0200 haftmann more accurate checkpoints
Mon, 19 Oct 2009 16:34:12 +0200 haftmann dropped lazy code equations
Mon, 19 Oct 2009 16:32:03 +0200 haftmann CONTRIBUTORS
Mon, 19 Oct 2009 23:02:56 +0200 wenzelm always qualify NJ's old List.foldl/foldr in Isabelle/ML;
Mon, 19 Oct 2009 23:02:23 +0200 wenzelm eliminated duplicate fold1 -- beware of argument order!
Mon, 19 Oct 2009 21:54:57 +0200 wenzelm uniform use of Integer.add/mult/sum/prod;
Mon, 19 Oct 2009 16:47:21 +0200 berghofe Removed dead code in function mk_deftab.
Mon, 19 Oct 2009 16:45:52 +0200 berghofe Removed unneeded reference to inv_def.
Mon, 19 Oct 2009 16:45:00 +0200 berghofe Replaced inv by the_inv_onto.
Mon, 19 Oct 2009 16:43:45 +0200 berghofe Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
Sun, 18 Oct 2009 22:19:05 +0200 wenzelm fixed proof (cf. d1d4d7a08a66);
Sun, 18 Oct 2009 22:16:37 +0200 wenzelm removed disjunctive group cancellation -- provers run independently;
Sun, 18 Oct 2009 21:13:29 +0200 wenzelm tuned;
Sun, 18 Oct 2009 20:53:40 +0200 wenzelm removed some unreferenced material;
Sun, 18 Oct 2009 18:08:04 +0200 nipkow merged
Sun, 18 Oct 2009 18:07:44 +0200 nipkow certificates for sos
Sun, 18 Oct 2009 16:25:59 +0200 nipkow merged
Sun, 18 Oct 2009 16:25:04 +0200 nipkow added sums of squares for standard deviation
Sun, 18 Oct 2009 12:07:56 +0200 nipkow merged
Sun, 18 Oct 2009 12:07:25 +0200 nipkow Inv -> inv_onto, inv abbr. inv_onto UNIV.
Sun, 18 Oct 2009 00:10:20 +0200 wenzelm disable indent-tabs-mode in Proof General / Emacs;
Sat, 17 Oct 2009 23:48:09 +0200 wenzelm merged
Sat, 17 Oct 2009 23:47:27 +0200 wenzelm made SML/NJ happy;
Sat, 17 Oct 2009 23:07:53 +0200 ballarin Merged.
Sat, 17 Oct 2009 22:58:18 +0200 ballarin Finished revisions of locales tutorial.
Thu, 15 Oct 2009 22:22:08 +0200 ballarin Changed part of the examples to int.
Thu, 15 Oct 2009 22:07:18 +0200 ballarin Save current state of locales tutorial.
Thu, 15 Oct 2009 22:06:43 +0200 ballarin Observe order of declaration when printing registrations.
Sat, 17 Oct 2009 22:35:28 +0200 wenzelm disable sunbroy2 for now;
Sat, 17 Oct 2009 22:34:19 +0200 wenzelm tuned/moved divide_and_conquer';
Sat, 17 Oct 2009 21:14:08 +0200 wenzelm tuned;
Sat, 17 Oct 2009 20:37:38 +0200 wenzelm removed separate record_quick_and_dirty_sensitive;
Sat, 17 Oct 2009 20:15:59 +0200 wenzelm simplified tactics;
Sat, 17 Oct 2009 19:04:35 +0200 wenzelm eliminated old List.foldr and OldTerm operations;
Sat, 17 Oct 2009 18:14:47 +0200 wenzelm removed unused names;
Sat, 17 Oct 2009 18:01:24 +0200 wenzelm misc tuning and simplification;
Sat, 17 Oct 2009 17:18:59 +0200 wenzelm less pervasive names;
Sat, 17 Oct 2009 16:58:03 +0200 wenzelm operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
Sat, 17 Oct 2009 16:40:41 +0200 wenzelm Method.cheating: check quick_and_dirty here;
Sat, 17 Oct 2009 16:34:39 +0200 wenzelm tuned;
Sat, 17 Oct 2009 16:33:14 +0200 wenzelm Unsynchronized.set etc.;
Sat, 17 Oct 2009 15:57:51 +0200 wenzelm indicate CRITICAL nature of various setmp combinators;
Sat, 17 Oct 2009 15:55:57 +0200 wenzelm ISABELLE_TOOL;
Sat, 17 Oct 2009 15:42:36 +0200 wenzelm tuned signature;
Sat, 17 Oct 2009 14:51:30 +0200 wenzelm merged
Sat, 17 Oct 2009 13:46:55 +0200 nipkow merged
Sat, 17 Oct 2009 13:46:39 +0200 nipkow added the_inv_onto
Sat, 17 Oct 2009 14:43:18 +0200 wenzelm eliminated hard tabulators, guessing at each author's individual tab-width;
Sat, 17 Oct 2009 01:05:59 +0200 wenzelm removed obsolete old goal command;
Sat, 17 Oct 2009 00:53:18 +0200 wenzelm legacy Drule.standard is no longer pervasive;
Sat, 17 Oct 2009 00:52:37 +0200 wenzelm explicitly qualify Drule.standard;
Fri, 16 Oct 2009 10:55:07 +0200 wenzelm tuned white space;
Fri, 16 Oct 2009 10:45:10 +0200 wenzelm local channels for tracing/debugging;
Fri, 16 Oct 2009 00:26:19 +0200 wenzelm made SML/NJ happy;
Thu, 15 Oct 2009 23:51:22 +0200 wenzelm sunbroy2: back to single-threaded mode for now -- deadlock in Poly/ML 5.3-SVN-900;
Thu, 15 Oct 2009 23:28:10 +0200 wenzelm replaced String.concat by implode;
Thu, 15 Oct 2009 23:10:35 +0200 wenzelm space_implode;
Thu, 15 Oct 2009 21:28:39 +0200 wenzelm normalized aliases of Output operations;
Thu, 15 Oct 2009 21:08:03 +0200 wenzelm eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
Thu, 15 Oct 2009 17:49:30 +0200 wenzelm natural argument order for prover;
Thu, 15 Oct 2009 17:06:19 +0200 wenzelm ATP_Manager.get_prover: canonical argument order;
Thu, 15 Oct 2009 17:04:45 +0200 wenzelm tuned proof (via atp_minimized);
Thu, 15 Oct 2009 16:15:22 +0200 wenzelm sort_strings (cf. Pure/library.ML);
Thu, 15 Oct 2009 15:53:33 +0200 wenzelm clarified File.platform_path vs. File.shell_path;
Thu, 15 Oct 2009 15:45:50 +0200 wenzelm exported File.shell_quote;
Thu, 15 Oct 2009 12:23:24 +0200 wenzelm misc tuning and recovery of Isabelle coding style;
Thu, 15 Oct 2009 11:49:27 +0200 wenzelm eliminated extraneous wrapping of public records;
Thu, 15 Oct 2009 11:23:03 +0200 wenzelm tuned signature;
Thu, 15 Oct 2009 11:12:09 +0200 wenzelm renamed functor HeapFun to Heap;
Thu, 15 Oct 2009 10:59:10 +0200 wenzelm misc tuning and clarification;
Thu, 15 Oct 2009 00:55:29 +0200 wenzelm structure ATP_Manager: eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
Wed, 14 Oct 2009 23:44:37 +0200 wenzelm modernized structure names;
Wed, 14 Oct 2009 23:13:38 +0200 wenzelm show direct GC time (which is included in CPU time);
Wed, 14 Oct 2009 22:57:44 +0200 wenzelm eliminated obsolete C/flip combinator;
Wed, 14 Oct 2009 21:14:53 +0200 nipkow added List.nth
Wed, 14 Oct 2009 16:45:26 +0200 wenzelm tuned make parameters for sunbroy2;
Wed, 14 Oct 2009 16:39:35 +0200 wenzelm accomodate old version of "tail", e.g. on sunbroy2;
Wed, 14 Oct 2009 16:08:51 +0200 wenzelm settings for parallel experimental Poly/ML 5.3;
Wed, 14 Oct 2009 13:56:56 +0200 haftmann sharpened name
Wed, 14 Oct 2009 12:20:01 +0200 haftmann more explicit notion of canonized code equations
Wed, 14 Oct 2009 12:19:17 +0200 haftmann more explicit notion of canonized code equations
Wed, 14 Oct 2009 12:04:16 +0200 haftmann tuned whitespace
Wed, 14 Oct 2009 12:03:16 +0200 haftmann tuned whitespace
Wed, 14 Oct 2009 11:56:44 +0200 haftmann dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
Tue, 13 Oct 2009 14:57:53 +0200 haftmann merged
Tue, 13 Oct 2009 14:08:01 +0200 haftmann deactivated Datatype.distinct_simproc
Tue, 13 Oct 2009 14:08:00 +0200 haftmann dropped Datatype.distinct_simproc; tuned
Tue, 13 Oct 2009 13:40:26 +0200 hoelzl order conjunctions to be printed without parentheses
Tue, 13 Oct 2009 12:02:55 +0200 hoelzl approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
Tue, 13 Oct 2009 09:13:24 +0200 haftmann merged
Mon, 12 Oct 2009 16:16:44 +0200 haftmann added add_tyconames; tuned
Tue, 13 Oct 2009 08:36:53 +0200 haftmann merged
Tue, 13 Oct 2009 08:36:39 +0200 haftmann more appropriate abstraction over distinctness rules
Mon, 12 Oct 2009 15:48:12 +0200 haftmann merged
Mon, 12 Oct 2009 15:46:38 +0200 haftmann intro_base_names combinator
Mon, 12 Oct 2009 15:26:50 +0200 haftmann merged
Mon, 12 Oct 2009 15:26:40 +0200 haftmann dropped dist_ss
Mon, 12 Oct 2009 14:22:54 +0200 haftmann merged
Mon, 12 Oct 2009 12:19:19 +0200 haftmann added is_IVar
Mon, 12 Oct 2009 12:19:19 +0200 haftmann factored out Code_Printer.aux_params
Mon, 12 Oct 2009 13:40:28 +0200 haftmann dropped rule duplicates
Mon, 12 Oct 2009 11:03:10 +0200 haftmann less non-standard combinators
Mon, 12 Oct 2009 10:24:08 +0200 haftmann nth replaces List.nth
Mon, 12 Oct 2009 10:24:07 +0200 haftmann dropped redundancy
Mon, 12 Oct 2009 09:25:27 +0200 haftmann dropped dead code
Mon, 12 Oct 2009 09:25:26 +0200 haftmann using distinct rules directly
Fri, 09 Oct 2009 13:40:34 +0200 haftmann simplified proof
Fri, 09 Oct 2009 13:34:40 +0200 haftmann dropped simproc_dist formally
Fri, 09 Oct 2009 13:34:34 +0200 haftmann tuned order of lemmas
Fri, 09 Oct 2009 10:00:47 +0200 haftmann term styles also cover antiquotations term_type and typeof
Fri, 09 Oct 2009 09:14:25 +0200 haftmann merged
Thu, 08 Oct 2009 19:33:03 +0200 haftmann lookup for datatype constructors considers type annotations to resolve overloading
Thu, 08 Oct 2009 15:59:17 +0200 haftmann added group_stmts
Thu, 08 Oct 2009 15:59:16 +0200 haftmann moved labelled_name to code_thingol
Thu, 08 Oct 2009 15:59:16 +0200 haftmann updated generated documentation
Thu, 08 Oct 2009 15:59:15 +0200 haftmann corrected syntax diagrams for styles
Thu, 08 Oct 2009 15:16:13 +0200 haftmann new generalized concept for term styles
Wed, 07 Oct 2009 16:57:56 +0200 haftmann generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
Thu, 08 Oct 2009 20:56:40 +0200 krauss isatest: store test identifiers
Wed, 07 Oct 2009 14:01:26 +0200 haftmann tuned proofs
Wed, 07 Oct 2009 14:01:26 +0200 haftmann added bot_boolE, top_boolI
Wed, 07 Oct 2009 12:06:04 +0200 haftmann do not use Locale.add_registration_eqs any longer
Wed, 07 Oct 2009 09:44:03 +0200 haftmann Inf/Sup now purely syntactic
Tue, 06 Oct 2009 20:19:54 +0200 haftmann merged
Tue, 06 Oct 2009 18:44:06 +0200 haftmann inf/sup1/2_iff are mere duplicates of underlying definitions: dropped
Tue, 06 Oct 2009 20:00:08 +0200 haftmann merged
Tue, 06 Oct 2009 18:27:00 +0200 haftmann added Coset as constructor
Tue, 06 Oct 2009 15:59:12 +0200 haftmann sets and cosets
Tue, 06 Oct 2009 15:51:34 +0200 haftmann added syntactic Inf and Sup
Mon, 05 Oct 2009 17:28:59 +0100 paulson merged
Mon, 05 Oct 2009 17:27:46 +0100 paulson New lemmas connected with the reals and infinite series
Mon, 05 Oct 2009 16:41:06 +0100 paulson New facts about domain and range in
Mon, 05 Oct 2009 16:55:56 +0200 haftmann merged
Mon, 05 Oct 2009 15:05:10 +0200 haftmann experimental de-facto abolishment of distinctness limit
Mon, 05 Oct 2009 15:04:45 +0200 haftmann tuned handling of type variable names further
Mon, 05 Oct 2009 08:36:33 +0200 haftmann variables in type schemes must be renamed simultaneously with variables in equations
Mon, 05 Oct 2009 11:48:06 +0200 haftmann explicitly unsynchronized
Mon, 05 Oct 2009 11:47:38 +0200 haftmann explicitly unsynchronized
Sun, 04 Oct 2009 12:59:22 +0200 boehmes recovered support for Spass: re-enabled writing problems in DFG format
Sun, 04 Oct 2009 11:45:41 +0200 boehmes avoid exception Option: only apply "the" if needed
Sun, 04 Oct 2009 07:01:22 +0200 nipkow merged
Wed, 30 Sep 2009 11:33:59 +0200 Philipp Meyer atp_minimal using chain_ths again
Sat, 03 Oct 2009 12:10:16 +0200 boehmes merged
Sat, 03 Oct 2009 12:05:40 +0200 boehmes re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
Fri, 02 Oct 2009 23:15:36 +0200 wenzelm eliminated dead code;
Fri, 02 Oct 2009 22:15:30 +0200 wenzelm eliminated dead code and redundant parameters;
Fri, 02 Oct 2009 22:15:08 +0200 wenzelm eliminated dead code;
Fri, 02 Oct 2009 22:02:54 +0200 wenzelm replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
Fri, 02 Oct 2009 22:02:11 +0200 wenzelm replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
Fri, 02 Oct 2009 21:42:31 +0200 wenzelm Refute.refute_goal: goal addressing from 1 as usual;
Fri, 02 Oct 2009 21:41:57 +0200 wenzelm Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
Fri, 02 Oct 2009 21:39:06 +0200 wenzelm clarified Proof.refine_insert -- always "refine" to apply standard method treatment (of conjunctions);
Fri, 02 Oct 2009 20:51:32 +0200 wenzelm misc tuning and simplification;
Fri, 02 Oct 2009 20:10:25 +0200 wenzelm macbroy2: bigger jobs first, to avoid danger of swapping during daytime;
Fri, 02 Oct 2009 10:35:13 +0200 wenzelm less ambitious heap settings;
Fri, 02 Oct 2009 04:44:56 +0200 haftmann merged
Thu, 01 Oct 2009 18:46:57 +0200 haftmann merged
Thu, 01 Oct 2009 17:11:48 +0200 haftmann proper merge of interpretation equations
Fri, 02 Oct 2009 00:10:08 +0200 wenzelm merged
Thu, 01 Oct 2009 23:03:59 +0200 ballarin Merged again.
Thu, 01 Oct 2009 20:52:18 +0200 ballarin Merged.
Thu, 01 Oct 2009 20:49:46 +0200 ballarin News entry: inheritance of mixins; print_interps.
Thu, 01 Oct 2009 20:37:33 +0200 ballarin Avoid administrative overhead for identity mixins.
Thu, 01 Oct 2009 23:49:05 +0200 wenzelm tuned;
Thu, 01 Oct 2009 23:27:05 +0200 wenzelm moved cache_conv to src/Pure/conv.ML, use Thm.cterm_cache;
Thu, 01 Oct 2009 22:40:29 +0200 wenzelm added Ctermtab, cterm_cache, thm_cache;
Thu, 01 Oct 2009 22:39:58 +0200 wenzelm added term_cache;
Thu, 01 Oct 2009 22:39:06 +0200 wenzelm Concurrently cached values.
Thu, 01 Oct 2009 20:47:26 +0200 wenzelm tuned header;
Thu, 01 Oct 2009 20:33:45 +0200 wenzelm core_sos_tac: SUBPROOF body operates on subgoal 1;
Thu, 01 Oct 2009 20:20:56 +0200 wenzelm merged
Thu, 01 Oct 2009 20:20:45 +0200 wenzelm updated generated files;
Thu, 01 Oct 2009 20:13:32 +0200 wenzelm enable slow-motion mode to accomodate unsynchronized refs within theory sources;
Thu, 01 Oct 2009 20:06:11 +0200 wenzelm avoid unsynchronized refs within theory sources;
Thu, 01 Oct 2009 20:04:44 +0200 wenzelm explicitly Unsynchronized;
Thu, 01 Oct 2009 13:32:03 +0200 Philipp Meyer additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Thu, 01 Oct 2009 11:54:01 +0200 Philipp Meyer changed core_sos_tac to use SUBPROOF
(0) -30000 -10000 -3000 -1000 -224 +224 +1000 +3000 +10000 +30000 tip