Thu, 01 Dec 2011 22:14:35 +0100 bulwahn making catch_match polymorphic
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn compilations return genuine flag to quickcheck framework
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn extending quickcheck's result by the genuine flag
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn reporting random compilation also catches match exceptions internally
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn the narrowing also indicates if counterexample is potentially spurious
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn the simple exhaustive compilation also indicates if counterexample is potentially spurious;
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn quickcheck-random compilation also indicates if the counterexample is potentially spurious or not
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn changing the exhaustive generator signatures;
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn adding examples for quickcheck-random
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn removing exception handling now that is caught at some other point;
Thu, 01 Dec 2011 22:14:35 +0100 bulwahn quickcheck random can also find potential counterexamples;
Thu, 01 Dec 2011 20:54:48 +0100 wenzelm merged
Thu, 01 Dec 2011 20:52:16 +0100 nipkow merged IMP/Util into IMP/Vars
Thu, 01 Dec 2011 15:41:58 +0100 hoelzl use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
Thu, 01 Dec 2011 15:41:58 +0100 hoelzl cardinality of sets of lists
Thu, 01 Dec 2011 15:41:48 +0100 hoelzl do not import examples Probability theory
Thu, 01 Dec 2011 14:03:57 +0100 hoelzl moved theorems about distribution to the definition; removed oopsed-lemma
Thu, 01 Dec 2011 14:03:57 +0100 hoelzl rename finite_prob_space.setsum_distribution, it collides with prob_space.setsum_distribution
Thu, 01 Dec 2011 14:03:57 +0100 hoelzl remove duplicate theorem setsum_real_distribution
Thu, 01 Dec 2011 14:29:14 +0100 wenzelm clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
Thu, 01 Dec 2011 13:34:16 +0100 blanchet updated Sledgehammer docs with new/renamed options
Thu, 01 Dec 2011 13:34:14 +0100 blanchet added "minimize" option for more control over automatic minimization
Thu, 01 Dec 2011 13:34:13 +0100 blanchet renamed "slicing" to "slice"
Thu, 01 Dec 2011 13:34:12 +0100 blanchet tuning
Thu, 01 Dec 2011 13:34:12 +0100 blanchet minor example tweak
Thu, 01 Dec 2011 12:25:27 +0100 wenzelm renamed inner syntax categories "num" to "num_token" and "xnum" to "xnum_token";
Thu, 01 Dec 2011 11:54:39 +0100 wenzelm updated markup conforming to ML side;
Wed, 30 Nov 2011 23:30:08 +0100 wenzelm discontinued obsolete datatype "alt_names";
Wed, 30 Nov 2011 21:14:01 +0100 wenzelm misc tuning;
Wed, 30 Nov 2011 19:18:17 +0100 wenzelm merged
Wed, 30 Nov 2011 18:50:46 +0100 kuncar removed outdated comment moved back and updated (at the direct request of Christian Urban)
Wed, 30 Nov 2011 15:07:10 +0100 bulwahn more stable introduction of the internally used unknown term
Wed, 30 Nov 2011 18:07:14 +0100 wenzelm prefer typedef without alternative name;
Wed, 30 Nov 2011 17:30:01 +0100 wenzelm prefer cpodef without extra definition;
Wed, 30 Nov 2011 16:27:10 +0100 wenzelm prefer typedef without extra definition and alternative name;
Wed, 30 Nov 2011 16:05:15 +0100 wenzelm tuned layout;
Wed, 30 Nov 2011 16:03:18 +0100 wenzelm tuned header;
Wed, 30 Nov 2011 12:09:29 +0100 wenzelm updated version information;
Wed, 30 Nov 2011 11:36:46 +0100 kuncar removed outdated comment
Wed, 30 Nov 2011 10:07:32 +0100 bulwahn adding examples for the potential counterexamples in the simple scheme
Wed, 30 Nov 2011 09:35:58 +0100 bulwahn also potential counterexamples in the simple exhaustive testing in quickcheck
Wed, 30 Nov 2011 09:21:18 +0100 bulwahn quickcheck does not show evaluation terms of equations if they are simply free variables to avoid duplicated output; tuned
Wed, 30 Nov 2011 09:21:15 +0100 bulwahn adding more verbose messages to exhaustive quickcheck
Wed, 30 Nov 2011 09:21:11 +0100 bulwahn quickcheck narrowing also shows potential counterexamples
Wed, 30 Nov 2011 09:21:09 +0100 bulwahn adding a exception-safe term reification step in quickcheck; adding examples
Wed, 30 Nov 2011 09:21:07 +0100 bulwahn quickcheck returns counterexamples that are potentially spurious due to underspecified code equations and match exceptions
Wed, 30 Nov 2011 09:21:04 +0100 bulwahn adding parsing of potential configuration to quickcheck command
Wed, 30 Nov 2011 09:21:02 +0100 bulwahn adding quickcheck's potential configuration
Tue, 29 Nov 2011 22:45:21 +0100 wenzelm more conventional file name;
Tue, 29 Nov 2011 21:50:00 +0100 wenzelm merged
Tue, 29 Nov 2011 15:52:51 +0100 kuncar updated documentation for the quotient package
Tue, 29 Nov 2011 18:22:31 +0100 kuncar merged
Tue, 29 Nov 2011 14:16:06 +0100 kuncar alternative names of morphisms in the definition of a quotient type can be specified
Tue, 29 Nov 2011 14:33:18 +0100 bulwahn adjusting antiquote_setup (cf. d83797ef0d2d)
Tue, 29 Nov 2011 21:50:00 +0100 wenzelm clarified Time vs. Timing;
Tue, 29 Nov 2011 21:29:53 +0100 wenzelm separate compilation of PIDE vs. Pure sources, which enables independent Scala library;
Tue, 29 Nov 2011 20:18:02 +0100 wenzelm clarified modules;
Tue, 29 Nov 2011 20:17:11 +0100 wenzelm tuned proofs;
Tue, 29 Nov 2011 19:49:36 +0100 wenzelm rearranged files;
Tue, 29 Nov 2011 06:09:41 +0100 huffman merged
Mon, 28 Nov 2011 21:28:01 +0100 huffman use HOL_basic_ss instead of HOL_ss for internal simproc proofs (cf. b36eedcd1633)
Mon, 28 Nov 2011 22:18:19 +0100 wenzelm explicit indication of modules for independent Scala library;
Mon, 28 Nov 2011 22:05:32 +0100 wenzelm separate module for concrete Isabelle markup;
Mon, 28 Nov 2011 20:39:08 +0100 wenzelm renamed Isabelle_Markup to Isabelle_Rendering to emphasize its meaning and make room for Pure Isabelle_Markup module;
Mon, 28 Nov 2011 20:31:53 +0100 wenzelm tuned signature (according to ML version);
Mon, 28 Nov 2011 18:08:07 +0100 nipkow merged
Mon, 28 Nov 2011 11:22:36 +0100 nipkow Hide Product_Type.Times - too precious an identifier
Mon, 28 Nov 2011 17:06:29 +0100 wenzelm more antiquotations;
Mon, 28 Nov 2011 17:06:20 +0100 wenzelm tuned messages;
Mon, 28 Nov 2011 17:05:41 +0100 wenzelm avoid stepping outside of context -- plain zero_var_indexes should be sufficient;
Mon, 28 Nov 2011 12:13:27 +0100 bulwahn increasing timeout to avoid test failures
Sun, 27 Nov 2011 23:12:03 +0100 wenzelm merged
Sun, 27 Nov 2011 13:32:20 +0100 nipkow merged
Sun, 27 Nov 2011 13:31:52 +0100 nipkow simplified Collecting1 and renamed: step -> step', step_cs -> step
Sun, 27 Nov 2011 23:10:19 +0100 wenzelm more antiquotations;
Sun, 27 Nov 2011 23:06:59 +0100 wenzelm tuned proof;
Sun, 27 Nov 2011 22:20:07 +0100 wenzelm permissive update for improved "tool compliance";
Sun, 27 Nov 2011 22:03:22 +0100 wenzelm just one data slot per module;
Sun, 27 Nov 2011 21:53:38 +0100 wenzelm tuned;
Sun, 27 Nov 2011 14:40:08 +0100 wenzelm tuned;
Sun, 27 Nov 2011 14:26:57 +0100 wenzelm more antiquotations;
Sun, 27 Nov 2011 14:20:31 +0100 wenzelm misc tuning;
Sun, 27 Nov 2011 13:12:42 +0100 wenzelm refined "literal" document style, with some correspondence to actual text source;
Sun, 27 Nov 2011 12:52:52 +0100 wenzelm modernized section about congruence rules;
Sat, 26 Nov 2011 17:10:03 +0100 wenzelm sharing of token source with span source;
Sat, 26 Nov 2011 14:14:51 +0100 wenzelm memoing of forked proofs;
Sat, 26 Nov 2011 13:10:12 +0100 wenzelm tuned;
Fri, 25 Nov 2011 23:04:12 +0100 wenzelm removed obsolete argument (cf. 954e9d6782ea);
Fri, 25 Nov 2011 22:21:37 +0100 wenzelm merged
Fri, 25 Nov 2011 19:07:26 +0100 krauss removed obsolete uses of Local_Theory.restore -- package composition P1 #> P2 no longer requires them due to 57def0b39696: P2 finds the results of P1 in the auxiliary context
Fri, 25 Nov 2011 12:18:39 +0100 nipkow merged
Fri, 25 Nov 2011 12:18:33 +0100 nipkow tuned
Fri, 25 Nov 2011 22:10:51 +0100 wenzelm increased stack limits (again, cf. d9cf3520083c and 77c3e74bd954);
Fri, 25 Nov 2011 21:29:11 +0100 wenzelm explicit change_parser thread, which avoids undirected Future.fork with its tendency towards hundreds of worker threads;
Fri, 25 Nov 2011 21:27:16 +0100 wenzelm tuned proofs;
Fri, 25 Nov 2011 18:37:14 +0100 wenzelm retain stderr and include it in syslog, which is buffered and thus increases the chance that users see remains from crashes etc.;
Fri, 25 Nov 2011 16:32:29 +0100 wenzelm prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
Fri, 25 Nov 2011 14:23:36 +0100 wenzelm recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
Fri, 25 Nov 2011 13:14:58 +0100 wenzelm proper intro? declarations -- NB: elim? indexes on major premise, which is too flexible here and easily leads to inefficiences (cf. a0336f8b6558);
Fri, 25 Nov 2011 10:52:30 +0100 bulwahn improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
Fri, 25 Nov 2011 00:18:59 +0000 Christian Urban in a local context, also the free variable case needs to be analysed default tip
Thu, 24 Nov 2011 21:41:58 +0100 wenzelm speed-up proof;
Thu, 24 Nov 2011 21:15:20 +0100 wenzelm more abstract datatype stamp, which avoids pointless allocation of mutable cells;
Thu, 24 Nov 2011 21:01:06 +0100 wenzelm modernized some old-style infix operations, which were left over from the time of ML proof scripts;
Thu, 24 Nov 2011 20:45:34 +0100 wenzelm simplified -- empty_ss already contains minimal mksimps;
Thu, 24 Nov 2011 19:58:37 +0100 nipkow Abstract interpretation is now based uniformly on annotated programs,
Wed, 23 Nov 2011 23:31:32 +0100 wenzelm tuned;
Wed, 23 Nov 2011 23:07:59 +0100 wenzelm tuned;
Wed, 23 Nov 2011 22:59:39 +0100 wenzelm modernized some old-style infix operations, which were left over from the time of ML proof scripts;
Wed, 23 Nov 2011 22:07:55 +0100 wenzelm tuned;
Wed, 23 Nov 2011 13:46:46 +0100 wenzelm updated according to bdcaa3f3a2f4;
Wed, 23 Nov 2011 13:41:42 +0100 wenzelm updated proof;
Wed, 23 Nov 2011 07:44:56 +0100 nipkow tuned
Wed, 23 Nov 2011 07:00:01 +0100 huffman remove outdated comment
Mon, 21 Nov 2011 23:29:53 +0100 wenzelm NEWS;
Mon, 21 Nov 2011 23:04:45 +0100 wenzelm simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
Mon, 21 Nov 2011 23:03:31 +0100 wenzelm drop vacuous decls;
Mon, 21 Nov 2011 21:38:08 +0100 wenzelm tuned;
Mon, 21 Nov 2011 19:52:50 +0100 wenzelm tuned header;
Mon, 21 Nov 2011 18:07:13 +0100 kuncar misspelled name
Sun, 20 Nov 2011 21:28:07 +0100 wenzelm eliminated obsolete "standard";
Sun, 20 Nov 2011 21:07:10 +0100 wenzelm eliminated obsolete "standard";
Sun, 20 Nov 2011 21:07:06 +0100 wenzelm eliminated obsolete "standard";
Sun, 20 Nov 2011 21:05:23 +0100 wenzelm eliminated obsolete "standard";
Sun, 20 Nov 2011 20:59:30 +0100 wenzelm eliminated obsolete "standard";
Sun, 20 Nov 2011 20:26:13 +0100 wenzelm explicit is better than implicit;
Sun, 20 Nov 2011 20:15:02 +0100 wenzelm eliminated obsolete "standard";
Sun, 20 Nov 2011 17:57:09 +0100 wenzelm tuned signature;
Sun, 20 Nov 2011 17:44:41 +0100 wenzelm 'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
Sun, 20 Nov 2011 17:32:27 +0100 wenzelm updated comment;
Sun, 20 Nov 2011 17:04:59 +0100 wenzelm more uniform signature;
Sun, 20 Nov 2011 16:59:37 +0100 wenzelm uniform cert_vars/read_vars;
Sun, 20 Nov 2011 16:58:12 +0100 wenzelm eliminated dead code;
Sun, 20 Nov 2011 15:21:22 +0100 wenzelm clarified certify vs. sharing;
Sun, 20 Nov 2011 13:29:12 +0100 wenzelm tuned;
Sat, 19 Nov 2011 21:23:16 +0100 wenzelm NEWS;
Sat, 19 Nov 2011 21:18:38 +0100 wenzelm added ML antiquotation @{attributes};
Sat, 19 Nov 2011 17:20:17 +0100 wenzelm merged
Sat, 19 Nov 2011 12:42:21 +0100 blanchet made SML/NJ happy
Sat, 19 Nov 2011 16:16:33 +0100 wenzelm simplified Locale.add_thmss, after partial evaluation of attributes;
Sat, 19 Nov 2011 15:34:37 +0100 wenzelm misc tuning;
Sat, 19 Nov 2011 15:04:36 +0100 wenzelm tuned;
Sat, 19 Nov 2011 14:31:43 +0100 wenzelm refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
Sat, 19 Nov 2011 13:36:38 +0100 wenzelm discontinued slightly odd write_prv, which would still write .prv files the first time, and deprive add-on tools from date stamp change (cf. 157e74588c49);
Sat, 19 Nov 2011 13:02:50 +0100 wenzelm do not store vacuous theorem specifications -- relevant for frugal local theory content;
Sat, 19 Nov 2011 12:33:18 +0100 wenzelm put structural attributes of 'obtains' in outer declarations, to make it actually work in local theory;
Fri, 18 Nov 2011 22:32:57 +0100 wenzelm sequential lemmas to accomodate static rule attributes;
Fri, 18 Nov 2011 21:55:24 +0100 wenzelm partial evaluation of locale facts leads to static scoping of rule attributes;
Fri, 18 Nov 2011 21:50:50 +0100 wenzelm tuned message;
Fri, 18 Nov 2011 16:42:31 +0100 blanchet gave SML/NJ a chance
Fri, 18 Nov 2011 14:47:08 +0100 blanchet more robust options
Fri, 18 Nov 2011 13:50:01 +0100 bulwahn adding another example for lifting definitions
Fri, 18 Nov 2011 13:42:07 +0100 bulwahn improving header
Fri, 18 Nov 2011 11:47:12 +0100 blanchet more "metis" calls in example
Fri, 18 Nov 2011 11:47:12 +0100 blanchet be more silent when auto minimizing
Fri, 18 Nov 2011 11:47:12 +0100 blanchet less offensive terminology
Fri, 18 Nov 2011 11:47:12 +0100 blanchet more "metis" calls in example
Fri, 18 Nov 2011 11:47:12 +0100 blanchet minor textual improvement
Fri, 18 Nov 2011 11:47:12 +0100 blanchet pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
Fri, 18 Nov 2011 11:47:12 +0100 blanchet wrap lambdas earlier, to get more control over beta/eta
Fri, 18 Nov 2011 11:47:12 +0100 blanchet move eta-contraction to before translation to Metis, to ensure everything stays in sync
Fri, 18 Nov 2011 11:47:12 +0100 blanchet avoid that var names get changed by resolution in Metis lambda-lifting
Fri, 18 Nov 2011 11:47:12 +0100 blanchet better threading of type encodings between Sledgehammer and "metis"
Fri, 18 Nov 2011 11:47:12 +0100 blanchet fixed bugs in lambda-lifting code -- ensure distinct names for variables
Fri, 18 Nov 2011 11:47:12 +0100 blanchet protect prefix against variant mutations
Fri, 18 Nov 2011 11:47:12 +0100 blanchet example cleanup
Fri, 18 Nov 2011 11:47:12 +0100 blanchet example cleanup
Fri, 18 Nov 2011 11:47:12 +0100 blanchet don't propagate user-set "type_enc" or "lam_trans" to Metis calls
Fri, 18 Nov 2011 11:47:12 +0100 blanchet don't needlessly pass "lam_lifted" option to "metis" call for SMT proof
Fri, 18 Nov 2011 11:47:12 +0100 blanchet avoid spurious messages in "lam_lifting" mode
Fri, 18 Nov 2011 11:47:12 +0100 blanchet eta-contract to avoid needless "lambda" wrappers
Fri, 18 Nov 2011 11:47:12 +0100 blanchet quiet down SMT
Fri, 18 Nov 2011 11:47:12 +0100 blanchet more aggressive lambda hiding (if we anyway need to pass an option to Metis)
Fri, 18 Nov 2011 11:47:12 +0100 blanchet updated Sledgehammer docs
Fri, 18 Nov 2011 11:47:12 +0100 blanchet don't pass "lam_lifted" option to "metis" unless there's a good reason
Fri, 18 Nov 2011 11:47:12 +0100 blanchet no needless reconstructors
Fri, 18 Nov 2011 11:47:12 +0100 blanchet removed more clutter
Fri, 18 Nov 2011 11:47:12 +0100 blanchet removed needless baggage
Fri, 18 Nov 2011 06:50:05 +0100 huffman Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
Fri, 18 Nov 2011 04:56:35 +0100 huffman merged
Thu, 17 Nov 2011 18:31:00 +0100 huffman Groups.thy: generalize several lemmas from class ab_group_add to class group_add
Thu, 17 Nov 2011 15:07:46 +0100 huffman HOL-Word: removed more duplicate theorems
Thu, 17 Nov 2011 14:52:05 +0100 huffman HOL-Word: removed many duplicate theorems (see NEWS)
Thu, 17 Nov 2011 14:24:10 +0100 huffman Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
Thu, 17 Nov 2011 12:38:03 +0100 huffman move definitions of bitwise operators into appropriate document section
Thu, 17 Nov 2011 12:29:48 +0100 huffman HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
Thu, 17 Nov 2011 08:07:54 +0100 huffman HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
Thu, 17 Nov 2011 07:55:09 +0100 huffman simplify implementation of approx_reorient_simproc
Thu, 17 Nov 2011 07:31:37 +0100 huffman name simp theorems st_0 and st_1
Thu, 17 Nov 2011 07:15:30 +0100 huffman remove redundant simp rules plus_enat_0
Thu, 17 Nov 2011 21:58:10 +0100 wenzelm eliminated slightly odd Rep' with dynamically-scoped [simplified];
(0) -30000 -10000 -3000 -1000 -192 +192 +1000 +3000 +10000 +30000 tip