Sat, 15 Sep 2012 20:13:25 +0200 dropped some unused identifiers
haftmann [Sat, 15 Sep 2012 20:13:25 +0200] rev 49387
dropped some unused identifiers
Sat, 15 Sep 2012 16:09:53 +0200 export rel_mono theorem
traytel [Sat, 15 Sep 2012 16:09:53 +0200] rev 49386
export rel_mono theorem
Fri, 14 Sep 2012 22:23:11 +0200 merged two unfold steps
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49385
merged two unfold steps
Fri, 14 Sep 2012 22:23:11 +0200 took out one rotate_tac
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49384
took out one rotate_tac
Fri, 14 Sep 2012 22:23:11 +0200 killed spurious rotate_tac; use auto instead of blast
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49383
killed spurious rotate_tac; use auto instead of blast
Fri, 14 Sep 2012 22:23:11 +0200 moved blast tactic to where it is actually needed
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49382
moved blast tactic to where it is actually needed
Fri, 14 Sep 2012 22:23:11 +0200 fixed bug in "mk_map" for the "fun" case
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49381
fixed bug in "mk_map" for the "fun" case
Fri, 14 Sep 2012 22:23:11 +0200 correct generalization to 3 or more mutually recursive datatypes
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49380
correct generalization to 3 or more mutually recursive datatypes
Fri, 14 Sep 2012 22:23:11 +0200 provide more guidance, exploiting our knowledge of the goal
blanchet [Fri, 14 Sep 2012 22:23:11 +0200] rev 49379
provide more guidance, exploiting our knowledge of the goal
Fri, 14 Sep 2012 22:23:10 +0200 fixed issue with bound variables in prem prems + tuning
blanchet [Fri, 14 Sep 2012 22:23:10 +0200] rev 49378
fixed issue with bound variables in prem prems + tuning
Fri, 14 Sep 2012 22:23:10 +0200 use right version of "mk_UnIN"
blanchet [Fri, 14 Sep 2012 22:23:10 +0200] rev 49377
use right version of "mk_UnIN"
Fri, 14 Sep 2012 22:23:10 +0200 select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet [Fri, 14 Sep 2012 22:23:10 +0200] rev 49376
select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
Fri, 14 Sep 2012 22:23:10 +0200 tuned code before fixing "mk_induct_discharge_prem_prems_tac"
blanchet [Fri, 14 Sep 2012 22:23:10 +0200] rev 49375
tuned code before fixing "mk_induct_discharge_prem_prems_tac"
Fri, 14 Sep 2012 21:26:01 +0200 tuned proofs;
wenzelm [Fri, 14 Sep 2012 21:26:01 +0200] rev 49374
tuned proofs;
Fri, 14 Sep 2012 21:23:06 +0200 merged
wenzelm [Fri, 14 Sep 2012 21:23:06 +0200] rev 49373
merged
Fri, 14 Sep 2012 12:09:27 +0200 polished the induction
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49372
polished the induction
Fri, 14 Sep 2012 12:09:27 +0200 put the flat at the right place (to avoid exceptions)
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49371
put the flat at the right place (to avoid exceptions)
Fri, 14 Sep 2012 12:09:27 +0200 fixed variable exporting problem
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49370
fixed variable exporting problem
Fri, 14 Sep 2012 12:09:27 +0200 compile
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49369
compile
Fri, 14 Sep 2012 12:09:27 +0200 added induct tactic
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49368
added induct tactic
Fri, 14 Sep 2012 12:09:27 +0200 tuning
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49367
tuning
Fri, 14 Sep 2012 12:09:27 +0200 renamed "mk_UnN" to "mk_UnIN"
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49366
renamed "mk_UnN" to "mk_UnIN"
Fri, 14 Sep 2012 12:09:27 +0200 merged two commands
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49365
merged two commands
Fri, 14 Sep 2012 12:09:27 +0200 allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49364
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
Fri, 14 Sep 2012 12:09:27 +0200 distinguish between nested and nesting BNFs
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49363
distinguish between nested and nesting BNFs
Fri, 14 Sep 2012 12:09:27 +0200 make tactic more robust in the case where "asm_simp_tac" already finishes the job
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49362
make tactic more robust in the case where "asm_simp_tac" already finishes the job
Fri, 14 Sep 2012 12:09:27 +0200 derive induction via backward proof, to ensure that the premises are in the right order for constructors like "X x y x" where x and y are mutually recursive
blanchet [Fri, 14 Sep 2012 12:09:27 +0200] rev 49361
derive induction via backward proof, to ensure that the premises are in the right order for constructors like "X x y x" where x and y are mutually recursive
Fri, 14 Sep 2012 21:15:59 +0200 no longer react on global_settings (cf. 34ac36642a31);
wenzelm [Fri, 14 Sep 2012 21:15:59 +0200] rev 49360
no longer react on global_settings (cf. 34ac36642a31);
Fri, 14 Sep 2012 20:49:54 +0200 refined output panel: more value-oriented approach to update and caret focus;
wenzelm [Fri, 14 Sep 2012 20:49:54 +0200] rev 49359
refined output panel: more value-oriented approach to update and caret focus;
Fri, 14 Sep 2012 18:12:41 +0200 clarified markup names;
wenzelm [Fri, 14 Sep 2012 18:12:41 +0200] rev 49358
clarified markup names;
Fri, 14 Sep 2012 17:37:19 +0200 more general Document_Model.point_range;
wenzelm [Fri, 14 Sep 2012 17:37:19 +0200] rev 49357
more general Document_Model.point_range; more general Document_View.Active_Area; eliminated dead popup material;
Fri, 14 Sep 2012 13:52:16 +0200 more static handling of rendering options;
wenzelm [Fri, 14 Sep 2012 13:52:16 +0200] rev 49356
more static handling of rendering options;
Fri, 14 Sep 2012 12:46:33 +0200 tuned options (again);
wenzelm [Fri, 14 Sep 2012 12:46:33 +0200] rev 49355
tuned options (again);
Fri, 14 Sep 2012 12:29:02 +0200 more scalable option-group;
wenzelm [Fri, 14 Sep 2012 12:29:02 +0200] rev 49354
more scalable option-group;
Fri, 14 Sep 2012 10:01:42 +0200 tuned
nipkow [Fri, 14 Sep 2012 10:01:42 +0200] rev 49353
tuned
Thu, 13 Sep 2012 17:20:44 +0200 merged
wenzelm [Thu, 13 Sep 2012 17:20:44 +0200] rev 49352
merged
Thu, 13 Sep 2012 17:20:04 +0200 tuned proofs;
wenzelm [Thu, 13 Sep 2012 17:20:04 +0200] rev 49351
tuned proofs;
Thu, 13 Sep 2012 17:15:04 +0200 remove theory Real_Integration, not needed since 44e42d392c6e when Euclidean spaces where introduced
hoelzl [Thu, 13 Sep 2012 17:15:04 +0200] rev 49350
remove theory Real_Integration, not needed since 44e42d392c6e when Euclidean spaces where introduced
Thu, 13 Sep 2012 16:43:33 +0200 workaround for HOL-Mirabelle-ex oddities;
wenzelm [Thu, 13 Sep 2012 16:43:33 +0200] rev 49349
workaround for HOL-Mirabelle-ex oddities;
Thu, 13 Sep 2012 16:10:20 +0200 instructions for quick start in 20min;
wenzelm [Thu, 13 Sep 2012 16:10:20 +0200] rev 49348
instructions for quick start in 20min;
Thu, 13 Sep 2012 16:09:35 +0200 more liberal init_components: base dir may get created later when resolving missing components;
wenzelm [Thu, 13 Sep 2012 16:09:35 +0200] rev 49347
more liberal init_components: base dir may get created later when resolving missing components;
Thu, 13 Sep 2012 16:01:42 +0200 more efficient painting based on cached result;
wenzelm [Thu, 13 Sep 2012 16:01:42 +0200] rev 49346
more efficient painting based on cached result;
Thu, 13 Sep 2012 11:13:00 +0200 more standard init_components -- particularly important to pick up correct jdk/scala version;
wenzelm [Thu, 13 Sep 2012 11:13:00 +0200] rev 49345
more standard init_components -- particularly important to pick up correct jdk/scala version;
Thu, 13 Sep 2012 10:28:48 +0200 tuned
nipkow [Thu, 13 Sep 2012 10:28:48 +0200] rev 49344
tuned
Wed, 12 Sep 2012 23:38:12 +0200 merged
wenzelm [Wed, 12 Sep 2012 23:38:12 +0200] rev 49343
merged
Wed, 12 Sep 2012 23:06:39 +0200 rough and ready induction
blanchet [Wed, 12 Sep 2012 23:06:39 +0200] rev 49342
rough and ready induction
Wed, 12 Sep 2012 23:06:39 +0200 nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
blanchet [Wed, 12 Sep 2012 23:06:39 +0200] rev 49341
nicer error message, indicating which type is empty (relying on metatheoretical completeness proof here)
Wed, 12 Sep 2012 23:18:26 +0200 observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
wenzelm [Wed, 12 Sep 2012 23:18:26 +0200] rev 49340
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
Wed, 12 Sep 2012 22:00:29 +0200 eliminated some old material that is unused in the visible universe;
wenzelm [Wed, 12 Sep 2012 22:00:29 +0200] rev 49339
eliminated some old material that is unused in the visible universe;
Wed, 12 Sep 2012 17:26:05 +0200 tuning
blanchet [Wed, 12 Sep 2012 17:26:05 +0200] rev 49338
tuning
Wed, 12 Sep 2012 17:26:05 +0200 set up things for (co)induction sugar
blanchet [Wed, 12 Sep 2012 17:26:05 +0200] rev 49337
set up things for (co)induction sugar
Wed, 12 Sep 2012 17:26:05 +0200 tuning
blanchet [Wed, 12 Sep 2012 17:26:05 +0200] rev 49336
tuning
Wed, 12 Sep 2012 17:26:05 +0200 added sumEN_tupled_balanced
blanchet [Wed, 12 Sep 2012 17:26:05 +0200] rev 49335
added sumEN_tupled_balanced
Wed, 12 Sep 2012 16:54:24 +0200 load fonts into JavaFX as well;
wenzelm [Wed, 12 Sep 2012 16:54:24 +0200] rev 49334
load fonts into JavaFX as well;
Wed, 12 Sep 2012 16:27:44 +0200 some support for actual HTML rendering;
wenzelm [Wed, 12 Sep 2012 16:27:44 +0200] rev 49333
some support for actual HTML rendering;
Wed, 12 Sep 2012 15:01:25 +0200 merged
wenzelm [Wed, 12 Sep 2012 15:01:25 +0200] rev 49332
merged
Wed, 12 Sep 2012 12:43:34 +0200 free variable name tuning
blanchet [Wed, 12 Sep 2012 12:43:34 +0200] rev 49331
free variable name tuning
Wed, 12 Sep 2012 12:06:03 +0200 reuse generated names (they look better + slightly more efficient)
blanchet [Wed, 12 Sep 2012 12:06:03 +0200] rev 49330
reuse generated names (they look better + slightly more efficient)
Wed, 12 Sep 2012 11:47:51 +0200 desambiguate grammar (e.g. for Nil's mixfix ("[]"))
blanchet [Wed, 12 Sep 2012 11:47:51 +0200] rev 49329
desambiguate grammar (e.g. for Nil's mixfix ("[]"))
Wed, 12 Sep 2012 11:39:05 +0200 avoided duplicate lemma
blanchet [Wed, 12 Sep 2012 11:39:05 +0200] rev 49328
avoided duplicate lemma
Wed, 12 Sep 2012 11:38:22 +0200 put an underscore between names, for compatibility with old package (and also because it makes sense)
blanchet [Wed, 12 Sep 2012 11:38:22 +0200] rev 49327
put an underscore between names, for compatibility with old package (and also because it makes sense)
Wed, 12 Sep 2012 10:36:00 +0200 got rid of metis calls
blanchet [Wed, 12 Sep 2012 10:36:00 +0200] rev 49326
got rid of metis calls
Wed, 12 Sep 2012 10:35:56 +0200 tuning
blanchet [Wed, 12 Sep 2012 10:35:56 +0200] rev 49325
tuning
Wed, 12 Sep 2012 14:46:13 +0200 removed obsolete argument "int" and thus made SML/NJ happy (cf. 03bee3a6a1b7);
wenzelm [Wed, 12 Sep 2012 14:46:13 +0200] rev 49324
removed obsolete argument "int" and thus made SML/NJ happy (cf. 03bee3a6a1b7);
Wed, 12 Sep 2012 13:56:49 +0200 standardized ML aliases;
wenzelm [Wed, 12 Sep 2012 13:56:49 +0200] rev 49323
standardized ML aliases;
Wed, 12 Sep 2012 13:42:28 +0200 tuned headers;
wenzelm [Wed, 12 Sep 2012 13:42:28 +0200] rev 49322
tuned headers;
Wed, 12 Sep 2012 13:21:33 +0200 avoid spaces in markup names, which might cause problems in boundary situations (e.g. HTML class);
wenzelm [Wed, 12 Sep 2012 13:21:33 +0200] rev 49321
avoid spaces in markup names, which might cause problems in boundary situations (e.g. HTML class);
Wed, 12 Sep 2012 12:09:40 +0200 discontinued experiment with literal replacement text in PDF (cf. b646316f8b3c, 2ff10e613689);
wenzelm [Wed, 12 Sep 2012 12:09:40 +0200] rev 49320
discontinued experiment with literal replacement text in PDF (cf. b646316f8b3c, 2ff10e613689);
Wed, 12 Sep 2012 11:38:23 +0200 more robust interrupt handling;
wenzelm [Wed, 12 Sep 2012 11:38:23 +0200] rev 49319
more robust interrupt handling;
Wed, 12 Sep 2012 11:28:34 +0200 some attempts to synchronize ROOT/files and document/build;
wenzelm [Wed, 12 Sep 2012 11:28:34 +0200] rev 49318
some attempts to synchronize ROOT/files and document/build;
Wed, 12 Sep 2012 11:14:44 +0200 tuned error;
wenzelm [Wed, 12 Sep 2012 11:14:44 +0200] rev 49317
tuned error;
Wed, 12 Sep 2012 10:18:31 +0200 option_pred characterization
traytel [Wed, 12 Sep 2012 10:18:31 +0200] rev 49316
option_pred characterization
Wed, 12 Sep 2012 09:39:41 +0200 true vs. True in pattern matching
traytel [Wed, 12 Sep 2012 09:39:41 +0200] rev 49315
true vs. True in pattern matching
Wed, 12 Sep 2012 06:35:07 +0200 reduced theory dependencies
blanchet [Wed, 12 Sep 2012 06:35:07 +0200] rev 49314
reduced theory dependencies
Wed, 12 Sep 2012 06:30:35 +0200 tuning
blanchet [Wed, 12 Sep 2012 06:30:35 +0200] rev 49313
tuning
Wed, 12 Sep 2012 06:27:48 +0200 moved theorems closer to where they are used
blanchet [Wed, 12 Sep 2012 06:27:48 +0200] rev 49312
moved theorems closer to where they are used
Wed, 12 Sep 2012 06:27:36 +0200 tuning
blanchet [Wed, 12 Sep 2012 06:27:36 +0200] rev 49311
tuning
Wed, 12 Sep 2012 05:29:21 +0200 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet [Wed, 12 Sep 2012 05:29:21 +0200] rev 49310
renamed "Ordinals_and_Cardinals" to "Cardinals"
Wed, 12 Sep 2012 05:21:47 +0200 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet [Wed, 12 Sep 2012 05:21:47 +0200] rev 49309
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
Wed, 12 Sep 2012 05:03:18 +0200 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet [Wed, 12 Sep 2012 05:03:18 +0200] rev 49308
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
Wed, 12 Sep 2012 02:06:31 +0200 tuning
blanchet [Wed, 12 Sep 2012 02:06:31 +0200] rev 49307
tuning
Wed, 12 Sep 2012 02:05:06 +0200 tuning annotations
blanchet [Wed, 12 Sep 2012 02:05:06 +0200] rev 49306
tuning annotations
Wed, 12 Sep 2012 02:05:05 +0200 tuned antiquotations
blanchet [Wed, 12 Sep 2012 02:05:05 +0200] rev 49305
tuned antiquotations
Wed, 12 Sep 2012 02:05:04 +0200 tuning
blanchet [Wed, 12 Sep 2012 02:05:04 +0200] rev 49304
tuning
Wed, 12 Sep 2012 02:05:03 +0200 tuning
blanchet [Wed, 12 Sep 2012 02:05:03 +0200] rev 49303
tuning
Wed, 12 Sep 2012 00:55:11 +0200 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet [Wed, 12 Sep 2012 00:55:11 +0200] rev 49302
added optional qualifiers for constructors and destructors, similarly to the old package
Wed, 12 Sep 2012 00:20:37 +0200 adapted example
blanchet [Wed, 12 Sep 2012 00:20:37 +0200] rev 49301
adapted example
Wed, 12 Sep 2012 00:20:37 +0200 added attributes to theorems
blanchet [Wed, 12 Sep 2012 00:20:37 +0200] rev 49300
added attributes to theorems
Tue, 11 Sep 2012 23:27:19 +0200 merged
wenzelm [Tue, 11 Sep 2012 23:27:19 +0200] rev 49299
merged
Tue, 11 Sep 2012 22:31:43 +0200 support for sort constraints in new (co)data commands
blanchet [Tue, 11 Sep 2012 22:31:43 +0200] rev 49298
support for sort constraints in new (co)data commands
Tue, 11 Sep 2012 22:13:22 +0200 provide a programmatic interface for FP sugar
blanchet [Tue, 11 Sep 2012 22:13:22 +0200] rev 49297
provide a programmatic interface for FP sugar
Tue, 11 Sep 2012 23:26:03 +0200 some GUI support for color options;
wenzelm [Tue, 11 Sep 2012 23:26:03 +0200] rev 49296
some GUI support for color options;
Tue, 11 Sep 2012 22:59:25 +0200 more precise sections;
wenzelm [Tue, 11 Sep 2012 22:59:25 +0200] rev 49295
more precise sections;
Tue, 11 Sep 2012 22:54:12 +0200 provide color values via options;
wenzelm [Tue, 11 Sep 2012 22:54:12 +0200] rev 49294
provide color values via options;
Tue, 11 Sep 2012 20:22:03 +0200 prefer tuning parameters as public methods (again) -- to allow overriding in applications;
wenzelm [Tue, 11 Sep 2012 20:22:03 +0200] rev 49293
prefer tuning parameters as public methods (again) -- to allow overriding in applications;
Tue, 11 Sep 2012 19:49:17 +0200 updated keywords;
wenzelm [Tue, 11 Sep 2012 19:49:17 +0200] rev 49292
updated keywords;
(0) -30000 -10000 -3000 -1000 -300 -100 -96 +96 +100 +300 +1000 +3000 +10000 +30000 tip