Thu, 25 Aug 2011 14:06:34 +0200 lemma Compl_insert: "- insert x A = (-A) - {x}"
krauss [Thu, 25 Aug 2011 14:06:34 +0200] rev 44490
lemma Compl_insert: "- insert x A = (-A) - {x}"
Thu, 25 Aug 2011 11:15:31 +0200 avoid variable clashes by properly incrementing indices
boehmes [Thu, 25 Aug 2011 11:15:31 +0200] rev 44489
avoid variable clashes by properly incrementing indices
Thu, 25 Aug 2011 11:15:31 +0200 improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes [Thu, 25 Aug 2011 11:15:31 +0200] rev 44488
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
Thu, 25 Aug 2011 00:00:36 +0200 include chained facts for minimizer, otherwise it won't work
blanchet [Thu, 25 Aug 2011 00:00:36 +0200] rev 44487
include chained facts for minimizer, otherwise it won't work
Wed, 24 Aug 2011 22:12:30 +0200 remove Vampire imconplete proof detection -- the bug it was trying to work around has been fixed in version 1.8, and the check is too sensitive anyway
blanchet [Wed, 24 Aug 2011 22:12:30 +0200] rev 44486
remove Vampire imconplete proof detection -- the bug it was trying to work around has been fixed in version 1.8, and the check is too sensitive anyway
Fri, 26 Aug 2011 22:25:41 +0200 back to tradition Scratch.thy default -- execution wrt. perspective overcomes the main problems of 226563829580;
wenzelm [Fri, 26 Aug 2011 22:25:41 +0200] rev 44485
back to tradition Scratch.thy default -- execution wrt. perspective overcomes the main problems of 226563829580;
Fri, 26 Aug 2011 22:14:12 +0200 tuned Session.edit_node: update_perspective based on last_exec_offset;
wenzelm [Fri, 26 Aug 2011 22:14:12 +0200] rev 44484
tuned Session.edit_node: update_perspective based on last_exec_offset;
Fri, 26 Aug 2011 21:27:58 +0200 tuned signature -- iterate subsumes both fold and get_first;
wenzelm [Fri, 26 Aug 2011 21:27:58 +0200] rev 44483
tuned signature -- iterate subsumes both fold and get_first;
Fri, 26 Aug 2011 21:18:42 +0200 further clarification of Document.updated, based on last_common and after_entry;
wenzelm [Fri, 26 Aug 2011 21:18:42 +0200] rev 44482
further clarification of Document.updated, based on last_common and after_entry; tuned;
Fri, 26 Aug 2011 16:06:58 +0200 tuned signature;
wenzelm [Fri, 26 Aug 2011 16:06:58 +0200] rev 44481
tuned signature;
Fri, 26 Aug 2011 15:56:30 +0200 improved Document.edit: more accurate update_start and no_execs;
wenzelm [Fri, 26 Aug 2011 15:56:30 +0200] rev 44480
improved Document.edit: more accurate update_start and no_execs; tuned;
Fri, 26 Aug 2011 15:09:54 +0200 refined document state assignment: observe perspective, more explicit assignment message;
wenzelm [Fri, 26 Aug 2011 15:09:54 +0200] rev 44479
refined document state assignment: observe perspective, more explicit assignment message; misc tuning and clarification;
Thu, 25 Aug 2011 19:12:58 +0200 tuned signature -- emphasize traditional read/eval/print terminology, which is still relevant here;
wenzelm [Thu, 25 Aug 2011 19:12:58 +0200] rev 44478
tuned signature -- emphasize traditional read/eval/print terminology, which is still relevant here;
Thu, 25 Aug 2011 17:38:12 +0200 maintain last_execs assignment on Scala side;
wenzelm [Thu, 25 Aug 2011 17:38:12 +0200] rev 44477
maintain last_execs assignment on Scala side; prefer tables over IDs instead of objects;
Thu, 25 Aug 2011 16:44:06 +0200 propagate information about last command with exec state assignment through document model;
wenzelm [Thu, 25 Aug 2011 16:44:06 +0200] rev 44476
propagate information about last command with exec state assignment through document model;
Thu, 25 Aug 2011 13:24:41 +0200 tuned signature;
wenzelm [Thu, 25 Aug 2011 13:24:41 +0200] rev 44475
tuned signature;
Thu, 25 Aug 2011 11:41:48 +0200 slightly more abstract Command.Perspective;
wenzelm [Thu, 25 Aug 2011 11:41:48 +0200] rev 44474
slightly more abstract Command.Perspective;
Thu, 25 Aug 2011 11:27:37 +0200 slightly more abstract Text.Perspective;
wenzelm [Thu, 25 Aug 2011 11:27:37 +0200] rev 44473
slightly more abstract Text.Perspective;
Wed, 24 Aug 2011 23:20:05 +0200 tuned proofs;
wenzelm [Wed, 24 Aug 2011 23:20:05 +0200] rev 44472
tuned proofs;
Wed, 24 Aug 2011 23:19:40 +0200 tuned syntax -- avoid ambiguities;
wenzelm [Wed, 24 Aug 2011 23:19:40 +0200] rev 44471
tuned syntax -- avoid ambiguities;
Wed, 24 Aug 2011 23:00:53 +0200 more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm [Wed, 24 Aug 2011 23:00:53 +0200] rev 44470
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
Wed, 24 Aug 2011 09:23:26 -0700 delete commented-out dead code
huffman [Wed, 24 Aug 2011 09:23:26 -0700] rev 44469
delete commented-out dead code
Wed, 24 Aug 2011 09:08:07 -0700 merged
huffman [Wed, 24 Aug 2011 09:08:07 -0700] rev 44468
merged
Wed, 24 Aug 2011 09:08:00 -0700 change some subsection headings to subsubsection
huffman [Wed, 24 Aug 2011 09:08:00 -0700] rev 44467
change some subsection headings to subsubsection
Tue, 23 Aug 2011 16:47:48 -0700 remove unnecessary lemma card_ge1
huffman [Tue, 23 Aug 2011 16:47:48 -0700] rev 44466
remove unnecessary lemma card_ge1
Tue, 23 Aug 2011 16:17:22 -0700 move connected_real_lemma to the one place it is used
huffman [Tue, 23 Aug 2011 16:17:22 -0700] rev 44465
move connected_real_lemma to the one place it is used
Wed, 24 Aug 2011 17:30:25 +0200 merged
wenzelm [Wed, 24 Aug 2011 17:30:25 +0200] rev 44464
merged
Wed, 24 Aug 2011 15:25:39 +0200 make sure that all facts are passed to ATP from minimizer
blanchet [Wed, 24 Aug 2011 15:25:39 +0200] rev 44463
make sure that all facts are passed to ATP from minimizer
Wed, 24 Aug 2011 11:17:33 +0200 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet [Wed, 24 Aug 2011 11:17:33 +0200] rev 44462
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
Wed, 24 Aug 2011 11:17:33 +0200 specify timeout for "sledgehammer_tac"
blanchet [Wed, 24 Aug 2011 11:17:33 +0200] rev 44461
specify timeout for "sledgehammer_tac"
Wed, 24 Aug 2011 11:17:33 +0200 tuning
blanchet [Wed, 24 Aug 2011 11:17:33 +0200] rev 44460
tuning
Wed, 24 Aug 2011 10:59:22 +0900 Quotient Package: add mem_rsp, mem_prs, tune proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 24 Aug 2011 10:59:22 +0900] rev 44459
Quotient Package: add mem_rsp, mem_prs, tune proofs.
Tue, 23 Aug 2011 15:46:53 -0700 merged
huffman [Tue, 23 Aug 2011 15:46:53 -0700] rev 44458
merged
Tue, 23 Aug 2011 14:11:02 -0700 declare euclidean_simps [simp] at the point they are proved;
huffman [Tue, 23 Aug 2011 14:11:02 -0700] rev 44457
declare euclidean_simps [simp] at the point they are proved; avoid duplicate rule warnings;
Tue, 23 Aug 2011 07:12:05 -0700 merged
huffman [Tue, 23 Aug 2011 07:12:05 -0700] rev 44456
merged
Mon, 22 Aug 2011 18:15:33 -0700 merged
huffman [Mon, 22 Aug 2011 18:15:33 -0700] rev 44455
merged
Mon, 22 Aug 2011 17:22:49 -0700 avoid warnings
huffman [Mon, 22 Aug 2011 17:22:49 -0700] rev 44454
avoid warnings
Mon, 22 Aug 2011 16:49:45 -0700 comment out dead code to avoid compiler warnings
huffman [Mon, 22 Aug 2011 16:49:45 -0700] rev 44453
comment out dead code to avoid compiler warnings
Mon, 22 Aug 2011 10:43:10 -0700 legacy theorem names
huffman [Mon, 22 Aug 2011 10:43:10 -0700] rev 44452
legacy theorem names
Mon, 22 Aug 2011 10:19:39 -0700 remove duplicate lemma
huffman [Mon, 22 Aug 2011 10:19:39 -0700] rev 44451
remove duplicate lemma
Tue, 23 Aug 2011 23:18:13 +0200 fixed "hBOOL" of existential variables, and generate more helpers
blanchet [Tue, 23 Aug 2011 23:18:13 +0200] rev 44450
fixed "hBOOL" of existential variables, and generate more helpers
Tue, 23 Aug 2011 22:44:08 +0200 don't select facts when using sledgehammer_tac for reconstruction
blanchet [Tue, 23 Aug 2011 22:44:08 +0200] rev 44449
don't select facts when using sledgehammer_tac for reconstruction
Tue, 23 Aug 2011 20:35:41 +0200 don't perform a triviality check if the goal is skipped anyway
blanchet [Tue, 23 Aug 2011 20:35:41 +0200] rev 44448
don't perform a triviality check if the goal is skipped anyway
Tue, 23 Aug 2011 19:50:25 +0200 optional reconstructor
blanchet [Tue, 23 Aug 2011 19:50:25 +0200] rev 44447
optional reconstructor
Wed, 24 Aug 2011 17:25:45 +0200 misc tuning and simplification;
wenzelm [Wed, 24 Aug 2011 17:25:45 +0200] rev 44446
misc tuning and simplification;
Wed, 24 Aug 2011 17:16:48 +0200 tuned pri: prefer purging of canceled execution;
wenzelm [Wed, 24 Aug 2011 17:16:48 +0200] rev 44445
tuned pri: prefer purging of canceled execution;
Wed, 24 Aug 2011 17:14:31 +0200 tuned Document.node: maintain "touched" flag to indicate changes in entries etc.;
wenzelm [Wed, 24 Aug 2011 17:14:31 +0200] rev 44444
tuned Document.node: maintain "touched" flag to indicate changes in entries etc.;
Wed, 24 Aug 2011 16:49:48 +0200 clarified Document.Node.clear -- retain header (cf. ML version);
wenzelm [Wed, 24 Aug 2011 16:49:48 +0200] rev 44443
clarified Document.Node.clear -- retain header (cf. ML version);
Wed, 24 Aug 2011 16:27:27 +0200 clarified norm_header/header_edit -- disallow update of loaded theories;
wenzelm [Wed, 24 Aug 2011 16:27:27 +0200] rev 44442
clarified norm_header/header_edit -- disallow update of loaded theories;
Wed, 24 Aug 2011 15:55:43 +0200 misc tuning and simplification;
wenzelm [Wed, 24 Aug 2011 15:55:43 +0200] rev 44441
misc tuning and simplification;
Wed, 24 Aug 2011 15:30:43 +0200 ignore irrelevant timings;
wenzelm [Wed, 24 Aug 2011 15:30:43 +0200] rev 44440
ignore irrelevant timings;
Wed, 24 Aug 2011 13:40:10 +0200 print state only for visible command, to avoid wasting resources for the larger part of the text;
wenzelm [Wed, 24 Aug 2011 13:40:10 +0200] rev 44439
print state only for visible command, to avoid wasting resources for the larger part of the text;
Wed, 24 Aug 2011 13:38:07 +0200 early filtering of unchanged perspective;
wenzelm [Wed, 24 Aug 2011 13:38:07 +0200] rev 44438
early filtering of unchanged perspective;
Wed, 24 Aug 2011 13:37:43 +0200 more reliable update_perspective handler based on actual text visibility (e.g. on startup or when resizing without scrolling);
wenzelm [Wed, 24 Aug 2011 13:37:43 +0200] rev 44437
more reliable update_perspective handler based on actual text visibility (e.g. on startup or when resizing without scrolling);
Wed, 24 Aug 2011 13:03:39 +0200 update_perspective without actual edits, bypassing the full state assignment protocol;
wenzelm [Wed, 24 Aug 2011 13:03:39 +0200] rev 44436
update_perspective without actual edits, bypassing the full state assignment protocol; edit_nodes/Perspective: do not touch_descendants here; propagate editor scroll events via update_perspective; misc tuning;
Tue, 23 Aug 2011 21:19:24 +0200 tuned;
wenzelm [Tue, 23 Aug 2011 21:19:24 +0200] rev 44435
tuned;
Tue, 23 Aug 2011 21:14:59 +0200 handle potentially more approriate BufferUpdate.LOADED event;
wenzelm [Tue, 23 Aug 2011 21:14:59 +0200] rev 44434
handle potentially more approriate BufferUpdate.LOADED event;
Mon, 22 Aug 2011 23:39:05 +0200 special treatment of structure index 1 in Pure, including legacy warning;
wenzelm [Mon, 22 Aug 2011 23:39:05 +0200] rev 44433
special treatment of structure index 1 in Pure, including legacy warning;
Tue, 23 Aug 2011 19:49:21 +0200 compile
blanchet [Tue, 23 Aug 2011 19:49:21 +0200] rev 44432
compile
Tue, 23 Aug 2011 19:00:48 +0200 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet [Tue, 23 Aug 2011 19:00:48 +0200] rev 44431
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
Tue, 23 Aug 2011 18:42:05 +0200 beef up "sledgehammer_tac" reconstructor
blanchet [Tue, 23 Aug 2011 18:42:05 +0200] rev 44430
beef up "sledgehammer_tac" reconstructor
Tue, 23 Aug 2011 18:42:05 +0200 clean up Sledgehammer tactic
blanchet [Tue, 23 Aug 2011 18:42:05 +0200] rev 44429
clean up Sledgehammer tactic
Tue, 23 Aug 2011 17:44:31 +0200 fixed document;
wenzelm [Tue, 23 Aug 2011 17:44:31 +0200] rev 44428
fixed document;
Tue, 23 Aug 2011 17:43:06 +0200 tuned signature;
wenzelm [Tue, 23 Aug 2011 17:43:06 +0200] rev 44427
tuned signature;
Tue, 23 Aug 2011 17:12:54 +0200 merged
wenzelm [Tue, 23 Aug 2011 17:12:54 +0200] rev 44426
merged
Tue, 23 Aug 2011 16:37:23 +0200 always use TFF if possible
blanchet [Tue, 23 Aug 2011 16:37:23 +0200] rev 44425
always use TFF if possible
Tue, 23 Aug 2011 16:14:19 +0200 clearer separator in generated file names
blanchet [Tue, 23 Aug 2011 16:14:19 +0200] rev 44424
clearer separator in generated file names
Tue, 23 Aug 2011 16:07:01 +0200 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet [Tue, 23 Aug 2011 16:07:01 +0200] rev 44423
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
Tue, 23 Aug 2011 15:50:27 +0200 updated known failures for Z3 3.0 TPTP
blanchet [Tue, 23 Aug 2011 15:50:27 +0200] rev 44422
updated known failures for Z3 3.0 TPTP
Tue, 23 Aug 2011 15:18:46 +0200 updated Z3 docs
blanchet [Tue, 23 Aug 2011 15:18:46 +0200] rev 44421
updated Z3 docs
Tue, 23 Aug 2011 15:15:43 +0200 avoid TFF format with older Vampire versions
blanchet [Tue, 23 Aug 2011 15:15:43 +0200] rev 44420
avoid TFF format with older Vampire versions
Tue, 23 Aug 2011 14:57:16 +0200 update the Vampire related parts of the documentation
blanchet [Tue, 23 Aug 2011 14:57:16 +0200] rev 44419
update the Vampire related parts of the documentation
Tue, 23 Aug 2011 14:44:19 +0200 fixed TFF slicing
blanchet [Tue, 23 Aug 2011 14:44:19 +0200] rev 44418
fixed TFF slicing
Tue, 23 Aug 2011 14:44:19 +0200 kindly ask Vampire to output axiom names
blanchet [Tue, 23 Aug 2011 14:44:19 +0200] rev 44417
kindly ask Vampire to output axiom names
Tue, 23 Aug 2011 14:44:19 +0200 added formats to the slice and use TFF for remote Vampire
blanchet [Tue, 23 Aug 2011 14:44:19 +0200] rev 44416
added formats to the slice and use TFF for remote Vampire
Tue, 23 Aug 2011 07:14:09 +0200 tuned specifications, syntax and proofs
haftmann [Tue, 23 Aug 2011 07:14:09 +0200] rev 44415
tuned specifications, syntax and proofs
Mon, 22 Aug 2011 22:00:36 +0200 tuned specifications and syntax
haftmann [Mon, 22 Aug 2011 22:00:36 +0200] rev 44414
tuned specifications and syntax
Tue, 23 Aug 2011 03:34:17 +0900 Quotient Package: some infrastructure for lifting inside sets
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 23 Aug 2011 03:34:17 +0900] rev 44413
Quotient Package: some infrastructure for lifting inside sets
Mon, 22 Aug 2011 15:02:45 +0200 include all encodings in tests, now that the incompleteness of some encodings has been addressed
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44412
include all encodings in tests, now that the incompleteness of some encodings has been addressed
Mon, 22 Aug 2011 15:02:45 +0200 change Metis's default settings if type information axioms are generated
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44411
change Metis's default settings if type information axioms are generated
Mon, 22 Aug 2011 15:02:45 +0200 we must tag any type whose ground types intersect a nonmonotonic type
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44410
we must tag any type whose ground types intersect a nonmonotonic type
Mon, 22 Aug 2011 15:02:45 +0200 prefer the lighter, slightly unsound monotonicity-based encodings for Metis
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44409
prefer the lighter, slightly unsound monotonicity-based encodings for Metis
Mon, 22 Aug 2011 15:02:45 +0200 made reconstruction of type tag equalities "\?x = \?x" reliable
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44408
made reconstruction of type tag equalities "\?x = \?x" reliable
Mon, 22 Aug 2011 15:02:45 +0200 tuning ATP problem output
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44407
tuning ATP problem output
Mon, 22 Aug 2011 15:02:45 +0200 revert guard logic -- make sure that typing information is generated for existentials
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44406
revert guard logic -- make sure that typing information is generated for existentials
Mon, 22 Aug 2011 15:02:45 +0200 generate tag equations for existential variables
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44405
generate tag equations for existential variables
Mon, 22 Aug 2011 15:02:45 +0200 tuning, plus started implementing tag equation generation for existential variables
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44404
tuning, plus started implementing tag equation generation for existential variables
Mon, 22 Aug 2011 15:02:45 +0200 precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44403
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
Mon, 22 Aug 2011 15:02:45 +0200 clearer terminology
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44402
clearer terminology
Mon, 22 Aug 2011 15:02:45 +0200 renamed "heavy" to "uniform", based on discussion with Nick Smallbone
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44401
renamed "heavy" to "uniform", based on discussion with Nick Smallbone
Mon, 22 Aug 2011 15:02:45 +0200 removed unused configuration option
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44400
removed unused configuration option
Mon, 22 Aug 2011 15:02:45 +0200 added caching for (in)finiteness checks
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44399
added caching for (in)finiteness checks
Mon, 22 Aug 2011 15:02:45 +0200 remove needless typing information
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44398
remove needless typing information
Mon, 22 Aug 2011 15:02:45 +0200 cleaner handling of polymorphic monotonicity inference
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44397
cleaner handling of polymorphic monotonicity inference
Mon, 22 Aug 2011 15:02:45 +0200 started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44396
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
Mon, 22 Aug 2011 15:02:45 +0200 more precise warning
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44395
more precise warning
Mon, 22 Aug 2011 15:02:45 +0200 added option to control soundness of encodings more precisely, for evaluation purposes
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44394
added option to control soundness of encodings more precisely, for evaluation purposes
Mon, 22 Aug 2011 15:02:45 +0200 make sound mode more sound (and clean up code)
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44393
make sound mode more sound (and clean up code)
Mon, 22 Aug 2011 15:02:45 +0200 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44392
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
Mon, 22 Aug 2011 15:02:45 +0200 gracefully handle empty SPASS problems
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44391
gracefully handle empty SPASS problems
Mon, 22 Aug 2011 15:02:45 +0200 pass sound option to Sledgehammer tactic
blanchet [Mon, 22 Aug 2011 15:02:45 +0200] rev 44390
pass sound option to Sledgehammer tactic
Tue, 23 Aug 2011 16:53:05 +0200 tuned signature -- contrast physical output primitives versus Output.raw_message;
wenzelm [Tue, 23 Aug 2011 16:53:05 +0200] rev 44389
tuned signature -- contrast physical output primitives versus Output.raw_message;
Tue, 23 Aug 2011 16:41:16 +0200 tuned signature;
wenzelm [Tue, 23 Aug 2011 16:41:16 +0200] rev 44388
tuned signature;
Tue, 23 Aug 2011 16:39:21 +0200 discontinued slightly odd Future/Lazy.get_finished, which do not really fit into the execution model of Future.cancel/join_tasks (canceled tasks need to be dequeued and terminated explicitly);
wenzelm [Tue, 23 Aug 2011 16:39:21 +0200] rev 44387
discontinued slightly odd Future/Lazy.get_finished, which do not really fit into the execution model of Future.cancel/join_tasks (canceled tasks need to be dequeued and terminated explicitly);
Tue, 23 Aug 2011 15:48:41 +0200 some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm [Tue, 23 Aug 2011 15:48:41 +0200] rev 44386
some support for toplevel printing wrt. editor perspective (still inactive);
Tue, 23 Aug 2011 12:20:12 +0200 propagate editor perspective through document model;
wenzelm [Tue, 23 Aug 2011 12:20:12 +0200] rev 44385
propagate editor perspective through document model;
Mon, 22 Aug 2011 21:42:02 +0200 some support for editor perspective;
wenzelm [Mon, 22 Aug 2011 21:42:02 +0200] rev 44384
some support for editor perspective;
Mon, 22 Aug 2011 21:09:26 +0200 discontinued redundant Edit_Command_ID;
wenzelm [Mon, 22 Aug 2011 21:09:26 +0200] rev 44383
discontinued redundant Edit_Command_ID;
Mon, 22 Aug 2011 20:11:44 +0200 reduced warnings;
wenzelm [Mon, 22 Aug 2011 20:11:44 +0200] rev 44382
reduced warnings;
Mon, 22 Aug 2011 20:00:04 +0200 tuned message;
wenzelm [Mon, 22 Aug 2011 20:00:04 +0200] rev 44381
tuned message;
Mon, 22 Aug 2011 17:10:22 +0200 old-style numbered structure index is legacy feature (hardly ever used now);
wenzelm [Mon, 22 Aug 2011 17:10:22 +0200] rev 44380
old-style numbered structure index is legacy feature (hardly ever used now);
Mon, 22 Aug 2011 16:12:23 +0200 added official Text.Range.Ordering;
wenzelm [Mon, 22 Aug 2011 16:12:23 +0200] rev 44379
added official Text.Range.Ordering; some support for text perspective;
Mon, 22 Aug 2011 14:15:52 +0200 tuned signature;
wenzelm [Mon, 22 Aug 2011 14:15:52 +0200] rev 44378
tuned signature;
Mon, 22 Aug 2011 12:17:22 +0200 reverted some odd changes to HOL/Import (cf. 74c08021ab2e);
wenzelm [Mon, 22 Aug 2011 12:17:22 +0200] rev 44377
reverted some odd changes to HOL/Import (cf. 74c08021ab2e);
Mon, 22 Aug 2011 10:57:33 +0200 merged
krauss [Mon, 22 Aug 2011 10:57:33 +0200] rev 44376
merged
Sun, 21 Aug 2011 22:13:04 +0200 modernized specifications
krauss [Sun, 21 Aug 2011 22:13:04 +0200] rev 44375
modernized specifications
Sun, 21 Aug 2011 22:13:04 +0200 removed session HOL/Subst -- now subsumed my more modern HOL/ex/Unification.thy
krauss [Sun, 21 Aug 2011 22:13:04 +0200] rev 44374
removed session HOL/Subst -- now subsumed my more modern HOL/ex/Unification.thy
Sun, 21 Aug 2011 22:13:04 +0200 removed technical or trivial unused facts
krauss [Sun, 21 Aug 2011 22:13:04 +0200] rev 44373
removed technical or trivial unused facts
Sun, 21 Aug 2011 22:13:04 +0200 more precise authors and comments;
krauss [Sun, 21 Aug 2011 22:13:04 +0200] rev 44372
more precise authors and comments; tuned order and headers
Sun, 21 Aug 2011 22:13:04 +0200 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss [Sun, 21 Aug 2011 22:13:04 +0200] rev 44371
added proof of idempotence, roughly after HOL/Subst/Unify.thy
(0) -30000 -10000 -3000 -1000 -120 +120 +1000 +3000 +10000 +30000 tip