Mon, 20 Dec 2010 23:30:32 +0100 updated to polyml-5.4.0;
wenzelm [Mon, 20 Dec 2010 23:30:32 +0100] rev 41332
updated to polyml-5.4.0;
Mon, 20 Dec 2010 23:26:17 +0100 tuned;
wenzelm [Mon, 20 Dec 2010 23:26:17 +0100] rev 41331
tuned;
Mon, 20 Dec 2010 23:19:16 +0100 updated for Poly/ML 5.4.0;
wenzelm [Mon, 20 Dec 2010 23:19:16 +0100] rev 41330
updated for Poly/ML 5.4.0;
Mon, 20 Dec 2010 22:27:26 +0100 merged
boehmes [Mon, 20 Dec 2010 22:27:26 +0100] rev 41329
merged
Mon, 20 Dec 2010 22:02:57 +0100 avoid ML structure aliases (especially single-letter abbreviations)
boehmes [Mon, 20 Dec 2010 22:02:57 +0100] rev 41328
avoid ML structure aliases (especially single-letter abbreviations)
Mon, 20 Dec 2010 21:04:45 +0100 added an additional beta reduction: unfolding of special quantifiers might leave terms unnormalized wrt to beta reductions
boehmes [Mon, 20 Dec 2010 21:04:45 +0100] rev 41327
added an additional beta reduction: unfolding of special quantifiers might leave terms unnormalized wrt to beta reductions
Mon, 20 Dec 2010 11:11:51 -0800 merged
huffman [Mon, 20 Dec 2010 11:11:51 -0800] rev 41326
merged
Mon, 20 Dec 2010 10:57:01 -0800 configure domain package to work with HOL option type
huffman [Mon, 20 Dec 2010 10:57:01 -0800] rev 41325
configure domain package to work with HOL option type
Mon, 20 Dec 2010 10:48:01 -0800 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman [Mon, 20 Dec 2010 10:48:01 -0800] rev 41324
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
Mon, 20 Dec 2010 09:48:16 -0800 simplify proofs
huffman [Mon, 20 Dec 2010 09:48:16 -0800] rev 41323
simplify proofs
Mon, 20 Dec 2010 09:41:55 -0800 beta-reduction simproc uses lemma Abs_cfun_inverse2 instead of beta_cfun, to be more robust against eta-contraction
huffman [Mon, 20 Dec 2010 09:41:55 -0800] rev 41322
beta-reduction simproc uses lemma Abs_cfun_inverse2 instead of beta_cfun, to be more robust against eta-contraction
Mon, 20 Dec 2010 08:26:47 -0800 configure domain package to work with HOL sum type
huffman [Mon, 20 Dec 2010 08:26:47 -0800] rev 41321
configure domain package to work with HOL sum type
Mon, 20 Dec 2010 07:50:47 -0800 replace list_map function with an abbreviation
huffman [Mon, 20 Dec 2010 07:50:47 -0800] rev 41320
replace list_map function with an abbreviation
Mon, 20 Dec 2010 18:48:13 +0100 merged
blanchet [Mon, 20 Dec 2010 18:48:13 +0100] rev 41319
merged
Mon, 20 Dec 2010 14:55:24 +0100 run SPASS and Vampire in SOS mode only if >= 50 facts are passed -- otherwise we are probably minimizing and then it's better if the prover is run only once with a full strategy
blanchet [Mon, 20 Dec 2010 14:55:24 +0100] rev 41318
run SPASS and Vampire in SOS mode only if >= 50 facts are passed -- otherwise we are probably minimizing and then it's better if the prover is run only once with a full strategy
Mon, 20 Dec 2010 14:17:49 +0100 disable feature that was enabled by mistake
blanchet [Mon, 20 Dec 2010 14:17:49 +0100] rev 41317
disable feature that was enabled by mistake
Mon, 20 Dec 2010 14:10:40 +0100 make exceptions more transparent in "debug" mode
blanchet [Mon, 20 Dec 2010 14:10:40 +0100] rev 41316
make exceptions more transparent in "debug" mode
Mon, 20 Dec 2010 14:09:45 +0100 remove spurious line
blanchet [Mon, 20 Dec 2010 14:09:45 +0100] rev 41315
remove spurious line
Mon, 20 Dec 2010 13:52:44 +0100 use the options provided by Stephan Schulz -- much better
blanchet [Mon, 20 Dec 2010 13:52:44 +0100] rev 41314
use the options provided by Stephan Schulz -- much better
Mon, 20 Dec 2010 12:12:35 +0100 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet [Mon, 20 Dec 2010 12:12:35 +0100] rev 41313
optionally supply constant weights to E -- turned off by default until properly parameterized
Mon, 20 Dec 2010 17:31:58 +0100 merged
haftmann [Mon, 20 Dec 2010 17:31:58 +0100] rev 41312
merged
Mon, 20 Dec 2010 15:37:25 +0100 type_lifting for predicates
haftmann [Mon, 20 Dec 2010 15:37:25 +0100] rev 41311
type_lifting for predicates
Mon, 20 Dec 2010 16:44:33 +0100 proper identifiers for consts and types;
wenzelm [Mon, 20 Dec 2010 16:44:33 +0100] rev 41310
proper identifiers for consts and types;
Mon, 20 Dec 2010 15:24:25 +0100 some attempts to fit diagnostic output into regular TTY (75-80 characters per line);
wenzelm [Mon, 20 Dec 2010 15:24:25 +0100] rev 41309
some attempts to fit diagnostic output into regular TTY (75-80 characters per line);
Mon, 20 Dec 2010 15:19:15 +0100 tuned/clarified some component settings;
wenzelm [Mon, 20 Dec 2010 15:19:15 +0100] rev 41308
tuned/clarified some component settings; explicit comments about common mistakes;
Mon, 20 Dec 2010 14:44:00 +0100 slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
wenzelm [Mon, 20 Dec 2010 14:44:00 +0100] rev 41307
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version); more robust rm_tree -- somewhat dangerous and not exported; tuned;
Mon, 20 Dec 2010 13:36:25 +0100 purged some comments (Locale_Test is already clean thanks to configuration options);
wenzelm [Mon, 20 Dec 2010 13:36:25 +0100] rev 41306
purged some comments (Locale_Test is already clean thanks to configuration options);
Mon, 20 Dec 2010 13:24:04 +0100 actually enable show_hyps option, unlike local_setup in 6da953d30f48 which merely affects the (temporary) auxiliary context;
wenzelm [Mon, 20 Dec 2010 13:24:04 +0100] rev 41305
actually enable show_hyps option, unlike local_setup in 6da953d30f48 which merely affects the (temporary) auxiliary context;
Mon, 20 Dec 2010 09:45:26 +0100 merged
boehmes [Mon, 20 Dec 2010 09:45:26 +0100] rev 41304
merged
Mon, 20 Dec 2010 09:31:47 +0100 updated SMT certificates
boehmes [Mon, 20 Dec 2010 09:31:47 +0100] rev 41303
updated SMT certificates
Mon, 20 Dec 2010 09:06:37 +0100 re-introduced support for nonlinear multiplication in Z3 (overriding the built-in linear multiplication of the SMT-LIB class of solvers)
boehmes [Mon, 20 Dec 2010 09:06:37 +0100] rev 41302
re-introduced support for nonlinear multiplication in Z3 (overriding the built-in linear multiplication of the SMT-LIB class of solvers)
Mon, 20 Dec 2010 08:45:27 +0100 derived SMT solver classes override inherited properties (properties of derived classes have a higher priority than properties of base classes)
boehmes [Mon, 20 Dec 2010 08:45:27 +0100] rev 41301
derived SMT solver classes override inherited properties (properties of derived classes have a higher priority than properties of base classes)
Mon, 20 Dec 2010 08:17:23 +0100 perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
boehmes [Mon, 20 Dec 2010 08:17:23 +0100] rev 41300
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
Mon, 20 Dec 2010 08:55:36 +0100 merge
haftmann [Mon, 20 Dec 2010 08:55:36 +0100] rev 41299
merge
Fri, 17 Dec 2010 22:00:54 +0100 more convenient order of type variables
haftmann [Fri, 17 Dec 2010 22:00:54 +0100] rev 41298
more convenient order of type variables
Sun, 19 Dec 2010 18:38:50 -0800 rename function cprod_map to prod_map
huffman [Sun, 19 Dec 2010 18:38:50 -0800] rev 41297
rename function cprod_map to prod_map
Sun, 19 Dec 2010 18:15:21 -0800 switch to transparent ascription, to avoid warning messages
huffman [Sun, 19 Dec 2010 18:15:21 -0800] rev 41296
switch to transparent ascription, to avoid warning messages
Sun, 19 Dec 2010 18:11:20 -0800 types -> type_synonym
huffman [Sun, 19 Dec 2010 18:11:20 -0800] rev 41295
types -> type_synonym
Sun, 19 Dec 2010 18:10:54 -0800 fix typo
huffman [Sun, 19 Dec 2010 18:10:54 -0800] rev 41294
fix typo
Sun, 19 Dec 2010 17:39:20 -0800 merged
huffman [Sun, 19 Dec 2010 17:39:20 -0800] rev 41293
merged
Sun, 19 Dec 2010 17:37:19 -0800 use deflations over type 'udom u' to represent predomains;
huffman [Sun, 19 Dec 2010 17:37:19 -0800] rev 41292
use deflations over type 'udom u' to represent predomains; removed now-unnecessary class liftdomain;
Sun, 19 Dec 2010 10:33:46 -0800 add lemma u_map_oo
huffman [Sun, 19 Dec 2010 10:33:46 -0800] rev 41291
add lemma u_map_oo
Sun, 19 Dec 2010 09:52:33 -0800 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman [Sun, 19 Dec 2010 09:52:33 -0800] rev 41290
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
Sun, 19 Dec 2010 06:59:01 -0800 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman [Sun, 19 Dec 2010 06:59:01 -0800] rev 41289
minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
Sun, 19 Dec 2010 06:39:19 -0800 powerdomain theories require class 'bifinite' instead of 'domain'
huffman [Sun, 19 Dec 2010 06:39:19 -0800] rev 41288
powerdomain theories require class 'bifinite' instead of 'domain'
Sun, 19 Dec 2010 06:34:41 -0800 type 'defl' takes a type parameter again (cf. b525988432e9)
huffman [Sun, 19 Dec 2010 06:34:41 -0800] rev 41287
type 'defl' takes a type parameter again (cf. b525988432e9)
Sun, 19 Dec 2010 05:15:31 -0800 reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman [Sun, 19 Dec 2010 05:15:31 -0800] rev 41286
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
Sun, 19 Dec 2010 04:06:02 -0800 renamed Bifinite.thy to Representable.thy
huffman [Sun, 19 Dec 2010 04:06:02 -0800] rev 41285
renamed Bifinite.thy to Representable.thy
Fri, 17 Dec 2010 16:43:45 -0800 renamed CompactBasis.thy to Compact_Basis.thy
huffman [Fri, 17 Dec 2010 16:43:45 -0800] rev 41284
renamed CompactBasis.thy to Compact_Basis.thy
Sun, 19 Dec 2010 19:03:49 +0100 merged
boehmes [Sun, 19 Dec 2010 19:03:49 +0100] rev 41283
merged
Sun, 19 Dec 2010 18:55:21 +0100 updated SMT certificates
boehmes [Sun, 19 Dec 2010 18:55:21 +0100] rev 41282
updated SMT certificates
Sun, 19 Dec 2010 18:54:29 +0100 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes [Sun, 19 Dec 2010 18:54:29 +0100] rev 41281
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions); removed odd retyping during folify (instead, keep all terms well-typed)
Sun, 19 Dec 2010 17:55:56 +0100 only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
boehmes [Sun, 19 Dec 2010 17:55:56 +0100] rev 41280
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general); hide internal constants z3div and z3mod; rewrite div/mod to z3div/z3mod instead of adding extra rules characterizing div/mod in terms of z3div/z3mod
Sun, 19 Dec 2010 13:25:18 +0100 escape backticks in altstrings
blanchet [Sun, 19 Dec 2010 13:25:18 +0100] rev 41279
escape backticks in altstrings
Sun, 19 Dec 2010 11:48:42 +0100 added a timestamp to Nitpick in verbose mode for debugging purposes;
blanchet [Sun, 19 Dec 2010 11:48:42 +0100] rev 41278
added a timestamp to Nitpick in verbose mode for debugging purposes; turn on verbose mode for the examples
Sun, 19 Dec 2010 00:13:25 +0100 reduce the minimizer slack and add verbose information
blanchet [Sun, 19 Dec 2010 00:13:25 +0100] rev 41277
reduce the minimizer slack and add verbose information
Sat, 18 Dec 2010 23:31:22 +0100 two layers of timeouts seem to be less reliable than a single layer
blanchet [Sat, 18 Dec 2010 23:31:22 +0100] rev 41276
two layers of timeouts seem to be less reliable than a single layer
Sat, 18 Dec 2010 22:15:39 +0100 move relevance filter into hard timeout
blanchet [Sat, 18 Dec 2010 22:15:39 +0100] rev 41275
move relevance filter into hard timeout
Sat, 18 Dec 2010 21:24:34 +0100 handle timeouts in Mirabelle more gracefully
blanchet [Sat, 18 Dec 2010 21:24:34 +0100] rev 41274
handle timeouts in Mirabelle more gracefully
Sat, 18 Dec 2010 21:07:34 +0100 made the relevance filter treat unatomizable facts like "atomize_all" properly (these result in problems that get E spinning seemingly forever);
blanchet [Sat, 18 Dec 2010 21:07:34 +0100] rev 41273
made the relevance filter treat unatomizable facts like "atomize_all" properly (these result in problems that get E spinning seemingly forever); disabled inductions (E and SPASS don't like them), but kept "ext" (seems more harmless)
Sat, 18 Dec 2010 18:43:16 +0100 Add mixins to locale dependencies.
ballarin [Sat, 18 Dec 2010 18:43:16 +0100] rev 41272
Add mixins to locale dependencies.
Sat, 18 Dec 2010 18:43:14 +0100 Enable show_hyps, which appears to be set in batch mode but in an interactive session.
ballarin [Sat, 18 Dec 2010 18:43:14 +0100] rev 41271
Enable show_hyps, which appears to be set in batch mode but in an interactive session.
Sat, 18 Dec 2010 18:43:13 +0100 Add mixins to sublocale command.
ballarin [Sat, 18 Dec 2010 18:43:13 +0100] rev 41270
Add mixins to sublocale command.
Sat, 18 Dec 2010 14:02:14 +0100 tuning
blanchet [Sat, 18 Dec 2010 14:02:14 +0100] rev 41269
tuning
Sat, 18 Dec 2010 13:48:24 +0100 higher hard timeout
blanchet [Sat, 18 Dec 2010 13:48:24 +0100] rev 41268
higher hard timeout
Sat, 18 Dec 2010 13:43:46 +0100 lower threshold where the binary algorithm kick in and use the same value for automatic minimization
blanchet [Sat, 18 Dec 2010 13:43:46 +0100] rev 41267
lower threshold where the binary algorithm kick in and use the same value for automatic minimization
Sat, 18 Dec 2010 13:38:14 +0100 compile
blanchet [Sat, 18 Dec 2010 13:38:14 +0100] rev 41266
compile
Sat, 18 Dec 2010 13:34:32 +0100 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet [Sat, 18 Dec 2010 13:34:32 +0100] rev 41265
let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
Sat, 18 Dec 2010 12:55:33 +0100 use minimizing prover in Mirabelle
blanchet [Sat, 18 Dec 2010 12:55:33 +0100] rev 41264
use minimizing prover in Mirabelle
Sat, 18 Dec 2010 12:53:56 +0100 renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet [Sat, 18 Dec 2010 12:53:56 +0100] rev 41263
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
Sat, 18 Dec 2010 12:46:58 +0100 factored out running a prover with (optionally) an implicit minimizer phrase
blanchet [Sat, 18 Dec 2010 12:46:58 +0100] rev 41262
factored out running a prover with (optionally) an implicit minimizer phrase
Fri, 17 Dec 2010 23:18:39 +0100 merged;
wenzelm [Fri, 17 Dec 2010 23:18:39 +0100] rev 41261
merged;
Fri, 17 Dec 2010 23:09:53 +0100 remove option that doesn't work in Mirabelle anyway (Mirabelle uses Sledgehammer_Provers, not Sledgehammer_Run)
blanchet [Fri, 17 Dec 2010 23:09:53 +0100] rev 41260
remove option that doesn't work in Mirabelle anyway (Mirabelle uses Sledgehammer_Provers, not Sledgehammer_Run)
Fri, 17 Dec 2010 22:15:08 +0100 more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet [Fri, 17 Dec 2010 22:15:08 +0100] rev 41259
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
Fri, 17 Dec 2010 21:47:13 +0100 convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
blanchet [Fri, 17 Dec 2010 21:47:13 +0100] rev 41258
convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
Fri, 17 Dec 2010 21:32:06 +0100 merged
blanchet [Fri, 17 Dec 2010 21:32:06 +0100] rev 41257
merged
Fri, 17 Dec 2010 21:31:19 +0100 put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet [Fri, 17 Dec 2010 21:31:19 +0100] rev 41256
put the SMT weights back where they belong, so that they're also used by Mirabelle
Fri, 17 Dec 2010 18:23:56 +0100 added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet [Fri, 17 Dec 2010 18:23:56 +0100] rev 41255
added debugging option to find out how good the relevance filter was at identifying relevant facts
Fri, 17 Dec 2010 22:23:56 +0100 extra checking of name bindings for classes, types, consts;
wenzelm [Fri, 17 Dec 2010 22:23:56 +0100] rev 41254
extra checking of name bindings for classes, types, consts; tuned;
Fri, 17 Dec 2010 20:21:35 +0100 more explicit references to structure Raw_Simplifier;
wenzelm [Fri, 17 Dec 2010 20:21:35 +0100] rev 41253
more explicit references to structure Raw_Simplifier;
Fri, 17 Dec 2010 18:38:33 +0100 merged
wenzelm [Fri, 17 Dec 2010 18:38:33 +0100] rev 41252
merged
Fri, 17 Dec 2010 18:33:35 +0100 merged
wenzelm [Fri, 17 Dec 2010 18:33:35 +0100] rev 41251
merged
Fri, 17 Dec 2010 18:15:56 +0100 merged
wenzelm [Fri, 17 Dec 2010 18:15:56 +0100] rev 41250
merged
Fri, 17 Dec 2010 18:10:37 +0100 Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
wenzelm [Fri, 17 Dec 2010 18:10:37 +0100] rev 41249
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
Fri, 17 Dec 2010 18:32:40 +0100 dropped slightly odd Conv.tap_thy
haftmann [Fri, 17 Dec 2010 18:32:40 +0100] rev 41248
dropped slightly odd Conv.tap_thy
Fri, 17 Dec 2010 18:24:44 +0100 avoid slightly odd Conv.tap_thy
haftmann [Fri, 17 Dec 2010 18:24:44 +0100] rev 41247
avoid slightly odd Conv.tap_thy
Fri, 17 Dec 2010 18:24:44 +0100 allocate intermediate directories in module hierarchy
haftmann [Fri, 17 Dec 2010 18:24:44 +0100] rev 41246
allocate intermediate directories in module hierarchy
Fri, 17 Dec 2010 16:55:27 +0100 export experimental options
blanchet [Fri, 17 Dec 2010 16:55:27 +0100] rev 41245
export experimental options
Fri, 17 Dec 2010 16:45:31 +0100 merged
blanchet [Fri, 17 Dec 2010 16:45:31 +0100] rev 41244
merged
Fri, 17 Dec 2010 16:20:02 +0100 compile
blanchet [Fri, 17 Dec 2010 16:20:02 +0100] rev 41243
compile
Fri, 17 Dec 2010 15:30:43 +0100 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet [Fri, 17 Dec 2010 15:30:43 +0100] rev 41242
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 make timeout part of the SMT filter's tail
blanchet [Fri, 17 Dec 2010 12:10:08 +0100] rev 41241
make timeout part of the SMT filter's tail
Fri, 17 Dec 2010 12:02:57 +0100 merge
blanchet [Fri, 17 Dec 2010 12:02:57 +0100] rev 41240
merge
Fri, 17 Dec 2010 12:02:46 +0100 split "smt_filter" into head and tail
blanchet [Fri, 17 Dec 2010 12:02:46 +0100] rev 41239
split "smt_filter" into head and tail
Fri, 17 Dec 2010 12:01:49 +0100 fewer facts to SInE-E
blanchet [Fri, 17 Dec 2010 12:01:49 +0100] rev 41238
fewer facts to SInE-E
Fri, 17 Dec 2010 11:12:37 +0100 Z3 sometimes reports two errors, with the first one referring to line 1 for some strange reason -- but it makes no sense to kill line 1, so we traverse the errors in reverse and consider only the last error
blanchet [Fri, 17 Dec 2010 11:12:37 +0100] rev 41237
Z3 sometimes reports two errors, with the first one referring to line 1 for some strange reason -- but it makes no sense to kill line 1, so we traverse the errors in reverse and consider only the last error
Fri, 17 Dec 2010 09:56:04 +0100 trap one more Z3 error
blanchet [Fri, 17 Dec 2010 09:56:04 +0100] rev 41236
trap one more Z3 error
Fri, 17 Dec 2010 15:30:00 +0100 fixed the command-line syntax for setting Yices' random seed
boehmes [Fri, 17 Dec 2010 15:30:00 +0100] rev 41235
fixed the command-line syntax for setting Yices' random seed
Fri, 17 Dec 2010 15:07:32 +0100 merged
boehmes [Fri, 17 Dec 2010 15:07:32 +0100] rev 41234
merged
Fri, 17 Dec 2010 14:59:06 +0100 updated SMT certificates
boehmes [Fri, 17 Dec 2010 14:59:06 +0100] rev 41233
updated SMT certificates
Fri, 17 Dec 2010 14:36:33 +0100 fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes [Fri, 17 Dec 2010 14:36:33 +0100] rev 41232
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order; fixed introduction of explicit application: use explicit application for every additional argument (grouping of arguments caused confusion when translating into the intermediate format)
Fri, 17 Dec 2010 12:14:18 +0100 adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn [Fri, 17 Dec 2010 12:14:18 +0100] rev 41231
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
Fri, 17 Dec 2010 17:48:05 +0100 updated generated file;
wenzelm [Fri, 17 Dec 2010 17:48:05 +0100] rev 41230
updated generated file;
Fri, 17 Dec 2010 17:43:54 +0100 replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm [Fri, 17 Dec 2010 17:43:54 +0100] rev 41229
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
Fri, 17 Dec 2010 17:08:56 +0100 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm [Fri, 17 Dec 2010 17:08:56 +0100] rev 41228
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
Fri, 17 Dec 2010 16:25:21 +0100 tuned signature;
wenzelm [Fri, 17 Dec 2010 16:25:21 +0100] rev 41227
tuned signature;
Fri, 17 Dec 2010 14:09:37 +0100 clarified exports of structure Simplifier;
wenzelm [Fri, 17 Dec 2010 14:09:37 +0100] rev 41226
clarified exports of structure Simplifier;
Fri, 17 Dec 2010 13:45:43 +0100 refer to regular structure Simplifier;
wenzelm [Fri, 17 Dec 2010 13:45:43 +0100] rev 41225
refer to regular structure Simplifier;
Fri, 17 Dec 2010 13:12:58 +0100 tuned;
wenzelm [Fri, 17 Dec 2010 13:12:58 +0100] rev 41224
tuned;
Fri, 17 Dec 2010 08:37:35 +0100 updated SMT certificates
boehmes [Fri, 17 Dec 2010 08:37:35 +0100] rev 41223
updated SMT certificates
Fri, 17 Dec 2010 00:27:40 +0100 more precise/correct SMT error handling
blanchet [Fri, 17 Dec 2010 00:27:40 +0100] rev 41222
more precise/correct SMT error handling
Fri, 17 Dec 2010 00:11:06 +0100 fixed off-by-one and return proper error code -- never underestimate the number of oddities in Perl
blanchet [Fri, 17 Dec 2010 00:11:06 +0100] rev 41221
fixed off-by-one and return proper error code -- never underestimate the number of oddities in Perl
Thu, 16 Dec 2010 22:45:02 +0100 discriminate SMT errors a bit better
blanchet [Thu, 16 Dec 2010 22:45:02 +0100] rev 41220
discriminate SMT errors a bit better
Thu, 16 Dec 2010 22:43:22 +0100 keep track of errors in Z3 input file for debugging purposes
blanchet [Thu, 16 Dec 2010 22:43:22 +0100] rev 41219
keep track of errors in Z3 input file for debugging purposes
Thu, 16 Dec 2010 21:53:31 +0100 better propagation of stdout in case of failure + comply with strict/warnings
blanchet [Thu, 16 Dec 2010 21:53:31 +0100] rev 41218
better propagation of stdout in case of failure + comply with strict/warnings
Thu, 16 Dec 2010 21:21:52 +0100 merge
blanchet [Thu, 16 Dec 2010 21:21:52 +0100] rev 41217
merge
Thu, 16 Dec 2010 21:21:13 +0100 cleaner handling of temporary files
blanchet [Thu, 16 Dec 2010 21:21:13 +0100] rev 41216
cleaner handling of temporary files
Thu, 16 Dec 2010 20:14:21 +0000 merged
paulson [Thu, 16 Dec 2010 20:14:21 +0000] rev 41215
merged
Thu, 16 Dec 2010 20:14:04 +0000 made sml/nj happy
paulson [Thu, 16 Dec 2010 20:14:04 +0000] rev 41214
made sml/nj happy
Thu, 16 Dec 2010 21:02:08 +0100 added self-correcting wrapper for Z3 -- see comment in the file for details
blanchet [Thu, 16 Dec 2010 21:02:08 +0100] rev 41213
added self-correcting wrapper for Z3 -- see comment in the file for details
(0) -30000 -10000 -3000 -1000 -120 +120 +1000 +3000 +10000 +30000 tip