Fri, 26 Sep 2008 09:09:51 +0200 haftmann op = vs. eq
Thu, 25 Sep 2008 20:34:21 +0200 wenzelm moved future_scheduler flag to skip_proof.ML;
Thu, 25 Sep 2008 20:34:20 +0200 wenzelm added future_scheduler (from thy_info.ML);
Thu, 25 Sep 2008 20:34:19 +0200 wenzelm simplified promise;
Thu, 25 Sep 2008 20:34:18 +0200 wenzelm simplified Thm.promise;
Thu, 25 Sep 2008 20:34:17 +0200 wenzelm explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
Thu, 25 Sep 2008 20:34:15 +0200 wenzelm explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
Thu, 25 Sep 2008 19:15:50 +0200 haftmann circumvent problem with code redundancy
Thu, 25 Sep 2008 16:05:52 +0200 haftmann clarifed redundancy policy
Thu, 25 Sep 2008 14:37:32 +0200 wenzelm tuned comments;
Thu, 25 Sep 2008 14:35:03 +0200 wenzelm added release_results;
Thu, 25 Sep 2008 14:35:02 +0200 wenzelm abtract types: plain datatype with opaque signature matching;
Thu, 25 Sep 2008 14:35:01 +0200 wenzelm prove: error with original thread position;
Thu, 25 Sep 2008 13:21:13 +0200 wenzelm explicit type OrdList.T;
Thu, 25 Sep 2008 11:14:01 +0200 haftmann (temporary workaround)
Thu, 25 Sep 2008 10:17:23 +0200 haftmann (temporal deactivation)
Thu, 25 Sep 2008 10:17:22 +0200 haftmann non left-linear equations for nbe
Thu, 25 Sep 2008 09:28:08 +0200 haftmann non left-linear equations for nbe
Thu, 25 Sep 2008 09:28:07 +0200 haftmann map_force
Thu, 25 Sep 2008 09:28:06 +0200 haftmann matchess
Thu, 25 Sep 2008 09:28:05 +0200 haftmann burrow_fst
Thu, 25 Sep 2008 09:28:03 +0200 haftmann discontinued special treatment of op = vs. eq_class.eq
Wed, 24 Sep 2008 19:39:25 +0200 wenzelm report: produce individual status messages;
Wed, 24 Sep 2008 19:33:14 +0200 wenzelm protocol change: remapped message codes to make room for nested system messages (e.g. for protocol proxy);
Wed, 24 Sep 2008 19:33:13 +0200 wenzelm protocol change: remapped message codes to make room for nested system messages (e.g. for protocol proxy);
Wed, 24 Sep 2008 18:08:42 +0200 wenzelm init: OuterKeyword.report;
Tue, 23 Sep 2008 23:07:48 +0200 wenzelm prove_multi: immediate;
Tue, 23 Sep 2008 22:04:30 +0200 wenzelm added promise_result, prove_promise;
Tue, 23 Sep 2008 18:31:33 +0200 berghofe Corrected call of SUBPROOF in coherent_tac that used wrong context.
Tue, 23 Sep 2008 18:11:45 +0200 haftmann fixed outer syntax
Tue, 23 Sep 2008 18:11:44 +0200 haftmann case default fallback for NBE
Tue, 23 Sep 2008 18:11:43 +0200 haftmann fixed quickcheck parameter syntax
Tue, 23 Sep 2008 18:11:42 +0200 haftmann renamed rtype to typerep
Tue, 23 Sep 2008 17:28:58 +0200 wenzelm added fold_rev;
Tue, 23 Sep 2008 15:48:55 +0200 wenzelm added del_node, which is more efficient for sparse graphs;
Tue, 23 Sep 2008 15:48:54 +0200 wenzelm IntGraph.del_node;
Tue, 23 Sep 2008 15:48:53 +0200 wenzelm join_results: special case for empty list, works without multithreading;
Tue, 23 Sep 2008 15:48:52 +0200 wenzelm added dest_deriv, removed external type deriv;
Tue, 23 Sep 2008 15:48:51 +0200 wenzelm added conditional add_oracles, keep oracles_of_proof private;
Tue, 23 Sep 2008 15:48:50 +0200 wenzelm Thm.proof_of;
Mon, 22 Sep 2008 23:04:35 +0200 berghofe Added coherent logic prover.
Mon, 22 Sep 2008 23:01:54 +0200 berghofe New prover for coherent logic.
Mon, 22 Sep 2008 23:00:15 +0200 berghofe Added setup for coherent logic prover.
Mon, 22 Sep 2008 22:59:35 +0200 berghofe Added examples for coherent logic prover.
Mon, 22 Sep 2008 22:59:11 +0200 berghofe Examples for coherent logic prover.
Mon, 22 Sep 2008 19:46:24 +0200 urbanc made the perm_simp tactic to understand options such as (no_asm)
Mon, 22 Sep 2008 15:26:14 +0200 wenzelm type thm: fully internal derivation, no longer exported;
Mon, 22 Sep 2008 15:26:14 +0200 wenzelm added is_finished;
Mon, 22 Sep 2008 15:26:13 +0200 wenzelm added Promise constructor, which is similar to Oracle but may be replaced later;
Mon, 22 Sep 2008 15:26:11 +0200 wenzelm removed deriv.ML which is now incorporated into thm.ML;
Mon, 22 Sep 2008 15:26:11 +0200 wenzelm added reject_draft;
Mon, 22 Sep 2008 15:26:07 +0200 wenzelm type thm: fully internal derivation, no longer exported;
Mon, 22 Sep 2008 13:56:04 +0200 haftmann generic quickcheck framework
Mon, 22 Sep 2008 13:56:03 +0200 haftmann TEMPORARY: make batch run happy
Mon, 22 Sep 2008 13:56:01 +0200 haftmann absolute Library path
Mon, 22 Sep 2008 13:55:59 +0200 haftmann different session branches for HOL-Plain vs. Plain
Mon, 22 Sep 2008 08:00:28 +0200 haftmann temporary workaround for class constants
Mon, 22 Sep 2008 08:00:27 +0200 haftmann corrected sort intersection
Mon, 22 Sep 2008 08:00:26 +0200 haftmann some steps towards generic quickcheck framework
Mon, 22 Sep 2008 08:00:24 +0200 haftmann fixed headers
Mon, 22 Sep 2008 08:00:23 +0200 haftmann added some fragments from website
Sat, 20 Sep 2008 21:05:41 +0200 wenzelm made SML/NJ happy;
Fri, 19 Sep 2008 22:11:50 +0200 wenzelm Isar toplevel editor model.
Fri, 19 Sep 2008 21:22:31 +0200 wenzelm future tasks: support boolean priorities (true = high, false = low/irrelevant);
Fri, 19 Sep 2008 21:00:50 +0200 wenzelm output_sync is now public;
Fri, 19 Sep 2008 21:00:49 +0200 wenzelm added props_text (from isar_syn.ML);
Fri, 19 Sep 2008 21:00:48 +0200 wenzelm moved Isar editor commands from isar_syn.ML to isar.ML;
Fri, 19 Sep 2008 21:00:47 +0200 wenzelm moved Isar editor commands from isar_syn.ML to isar.ML;
Fri, 19 Sep 2008 21:00:46 +0200 wenzelm added Isar/isar.scala;
Fri, 19 Sep 2008 18:05:19 +0200 huffman avoid using implicit assumptions
Fri, 19 Sep 2008 17:54:04 +0200 huffman add theory graph to ZF document
Fri, 19 Sep 2008 09:41:17 +0200 haftmann made SMLNJ happy
Thu, 18 Sep 2008 22:30:17 +0200 wenzelm jar: include sources;
Thu, 18 Sep 2008 20:12:02 +0200 wenzelm tuned;
Thu, 18 Sep 2008 19:39:50 +0200 wenzelm eval_term: CRITICAL due to eval_result;
Thu, 18 Sep 2008 19:39:49 +0200 wenzelm begin_theory: Theory.checkpoint for immediate uses ensures that ML evaluation always starts with non-draft @{theory};
Thu, 18 Sep 2008 19:39:47 +0200 wenzelm updated generated file;
Thu, 18 Sep 2008 19:39:44 +0200 wenzelm simplified oracle interface;
Thu, 18 Sep 2008 14:06:58 +0200 wenzelm show: non-critical testing;
Thu, 18 Sep 2008 14:06:56 +0200 wenzelm added deriv.ML: Abstract derivations based on raw proof terms.
Thu, 18 Sep 2008 12:13:50 +0200 krauss termination prover for "fun" can be configured using context data.
Thu, 18 Sep 2008 10:57:30 +0200 wenzelm updated generated file;
Thu, 18 Sep 2008 10:57:23 +0200 wenzelm unchecked $ISABELLE_HOME_USER/etc/settings;
Wed, 17 Sep 2008 23:44:31 +0200 wenzelm use_text/use_file now depend on explicit ML name space;
Wed, 17 Sep 2008 23:23:49 +0200 wenzelm threads work only for Poly/ML 5.2 or later;
Wed, 17 Sep 2008 23:23:13 +0200 wenzelm * ML bindings produced via Isar commands are stored within the Isar context.
Wed, 17 Sep 2008 23:08:06 +0200 wenzelm added ML_prf;
Wed, 17 Sep 2008 23:04:27 +0200 wenzelm updated generated file;
Wed, 17 Sep 2008 22:06:59 +0200 wenzelm added inherit_env;
Wed, 17 Sep 2008 22:06:57 +0200 wenzelm added map_contexts;
Wed, 17 Sep 2008 22:06:54 +0200 wenzelm ML_prf: inherit env for all contexts within the proof;
Wed, 17 Sep 2008 22:06:52 +0200 wenzelm shutdown only if Multithreading.available;
Wed, 17 Sep 2008 21:27:44 +0200 wenzelm ML_Context.evaluate: proper context (for ML environment);
Wed, 17 Sep 2008 21:27:43 +0200 wenzelm ML_Context.evaluate: proper context (for ML environment);
Wed, 17 Sep 2008 21:27:38 +0200 wenzelm simplified ML_Context.eval_in -- expect immutable Proof.context value;
Wed, 17 Sep 2008 21:27:36 +0200 wenzelm proper thm antiquotations within ML solve obscure context problems (due to update of ML environment);
Wed, 17 Sep 2008 21:27:34 +0200 wenzelm simplified ML_Context.eval_in -- expect immutable Proof.context value;
Wed, 17 Sep 2008 21:27:32 +0200 wenzelm explicit handling of ML environment within generic context;
Wed, 17 Sep 2008 21:27:31 +0200 wenzelm added ML_prf;
Wed, 17 Sep 2008 21:27:24 +0200 wenzelm use_text/use_file now depend on explicit ML name space;
Wed, 17 Sep 2008 21:27:22 +0200 wenzelm ML name space -- dummy version of Poly/ML 5.2 facility.
Wed, 17 Sep 2008 21:27:20 +0200 wenzelm added ML-Systems/ml_name_space.ML;
Wed, 17 Sep 2008 21:27:18 +0200 wenzelm use ML_prf within proofs;
Wed, 17 Sep 2008 21:27:17 +0200 wenzelm local @{context};
Wed, 17 Sep 2008 21:27:14 +0200 wenzelm moved global ML bindings to global place;
Wed, 17 Sep 2008 21:27:08 +0200 wenzelm back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
Wed, 17 Sep 2008 21:27:03 +0200 wenzelm updated generated file;
Wed, 17 Sep 2008 15:59:23 +0200 krauss wf_finite_psubset[simp], in_finite_psubset[simp]
Wed, 17 Sep 2008 15:21:30 +0200 ballarin Public interface to interpretation morphism.
Wed, 17 Sep 2008 11:42:25 +0200 haftmann moved term_of syntax to separate theory
Wed, 17 Sep 2008 10:00:16 +0200 haftmann removed obsolete theory
Wed, 17 Sep 2008 07:32:04 +0200 haftmann added quickcheck.ML
Tue, 16 Sep 2008 18:01:25 +0200 wenzelm tuned comments;
Tue, 16 Sep 2008 18:01:24 +0200 wenzelm multithreading for Poly/ML 5.1 is no longer supported;
Tue, 16 Sep 2008 17:28:37 +0200 wenzelm tuned;
Tue, 16 Sep 2008 17:21:14 +0200 wenzelm updated system manual;
Tue, 16 Sep 2008 17:18:41 +0200 wenzelm Proof General / Emacs interface wrapper;
Tue, 16 Sep 2008 17:16:28 +0200 wenzelm Proof General: option -I is obsolete;
Tue, 16 Sep 2008 17:16:27 +0200 wenzelm added PROOFGENERAL_HOME;
Tue, 16 Sep 2008 17:16:25 +0200 wenzelm separate emacs tool for Proof General / Emacs;
Tue, 16 Sep 2008 16:13:31 +0200 haftmann added quickcheck stub
Tue, 16 Sep 2008 16:13:14 +0200 haftmann removed babel again
Tue, 16 Sep 2008 16:13:11 +0200 haftmann celver code lemma avoid type ambiguity problem with Haskell
Tue, 16 Sep 2008 16:13:09 +0200 haftmann a sophisticated char/nibble conversion combinator
Tue, 16 Sep 2008 16:13:06 +0200 haftmann moved term_of syntax to separate theory
Tue, 16 Sep 2008 15:37:33 +0200 wenzelm SimpleThread.fork;
Tue, 16 Sep 2008 15:37:32 +0200 wenzelm Simplified thread fork interface.
Tue, 16 Sep 2008 15:37:30 +0200 wenzelm added Concurrent/simple_thread.ML;
Tue, 16 Sep 2008 14:48:51 +0200 wenzelm updated generated file;
Tue, 16 Sep 2008 14:40:30 +0200 wenzelm misc tuning and modernization;
Tue, 16 Sep 2008 14:39:56 +0200 wenzelm check setting and tool;
Tue, 16 Sep 2008 12:27:05 +0200 ballarin Clearer separation of interpretation frontend and backend.
Tue, 16 Sep 2008 12:26:15 +0200 ballarin No interpretation of locale with dangling type frees.
Tue, 16 Sep 2008 12:25:26 +0200 ballarin Do not rely on locale assumption in interpretation.
Tue, 16 Sep 2008 12:25:04 +0200 paulson The metis method now fails in the usual manner, rather than raising an exception,
Tue, 16 Sep 2008 12:24:37 +0200 ballarin Fixed typo in locale declaration.
Tue, 16 Sep 2008 09:21:28 +0200 haftmann added babel
Tue, 16 Sep 2008 09:21:27 +0200 haftmann explicit size of characters
Tue, 16 Sep 2008 09:21:26 +0200 haftmann dropped superfluous code lemmas
Tue, 16 Sep 2008 09:21:24 +0200 haftmann evaluation using code generator
Tue, 16 Sep 2008 09:21:22 +0200 haftmann generic value command
Mon, 15 Sep 2008 20:51:58 +0200 wenzelm converted symbols.tex;
Mon, 15 Sep 2008 20:51:40 +0200 wenzelm tuned;
Mon, 15 Sep 2008 20:22:38 +0200 wenzelm converted misc.tex;
Mon, 15 Sep 2008 19:43:10 +0200 wenzelm tuned;
Mon, 15 Sep 2008 19:42:51 +0200 wenzelm generated files;
Mon, 15 Sep 2008 19:42:22 +0200 wenzelm converted present.tex;
Mon, 15 Sep 2008 17:32:12 +0200 wenzelm basic setup for generated document sources (cf. IsarRef/isar-ref.tex);
Mon, 15 Sep 2008 16:50:35 +0200 wenzelm load underscore package after iman etc.;
Mon, 15 Sep 2008 16:43:53 +0200 wenzelm tuned comment;
Mon, 15 Sep 2008 16:43:31 +0200 wenzelm added formal markup for setting, executable, tool;
Mon, 15 Sep 2008 16:42:09 +0200 wenzelm basic setup for generated document sources (cf. IsarRef/isar-ref.tex);
Mon, 15 Sep 2008 16:42:00 +0200 wenzelm converted basics.tex to theory file;
Mon, 15 Sep 2008 16:40:53 +0200 wenzelm added isatt markup;
Sun, 14 Sep 2008 21:50:35 +0200 haftmann New outline for codegen tutorial -- draft
Fri, 12 Sep 2008 12:04:20 +0200 wenzelm added extern_fact (local or global);
Fri, 12 Sep 2008 12:04:19 +0200 wenzelm print raw (internal) result names;
Fri, 12 Sep 2008 12:04:16 +0200 wenzelm more procise printing of fact names;
Fri, 12 Sep 2008 10:54:00 +0200 wenzelm pretty_fact: extern fact name wrt. the given context, assuming that is the proper one for presentation;
Thu, 11 Sep 2008 22:22:59 +0200 wenzelm cancel, shutdown: notify_all;
Thu, 11 Sep 2008 22:22:20 +0200 wenzelm finish: Future.shutdown last;
Thu, 11 Sep 2008 21:53:53 +0200 wenzelm eliminated requests, use global state variables uniformly;
Thu, 11 Sep 2008 21:04:09 +0200 wenzelm finish: Future.shutdown;
Thu, 11 Sep 2008 21:04:07 +0200 wenzelm added is_empty;
Thu, 11 Sep 2008 21:04:05 +0200 wenzelm shutdown: global join-and-shutdown operation;
Thu, 11 Sep 2008 18:07:58 +0200 wenzelm added focus, which indicates a particular collection of high-priority tasks;
Thu, 11 Sep 2008 13:43:42 +0200 wenzelm some general notes on future values;
Thu, 11 Sep 2008 13:24:19 +0200 wenzelm separate Concurrent/ROOT.ML;
Thu, 11 Sep 2008 13:24:14 +0200 wenzelm Parallel list combinators.
Thu, 11 Sep 2008 13:23:57 +0200 wenzelm added Concurrent/par_list.ML;
Wed, 10 Sep 2008 23:28:09 +0200 wenzelm added interrupt_task (external id);
Wed, 10 Sep 2008 23:19:36 +0200 wenzelm tuned;
Wed, 10 Sep 2008 22:29:36 +0200 wenzelm future_schedule: uninterruptible join;
Wed, 10 Sep 2008 21:50:32 +0200 wenzelm added future_scheduler (default false);
Wed, 10 Sep 2008 21:50:30 +0200 wenzelm replaced join_all by join_results, which returns Exn.results;
Wed, 10 Sep 2008 20:28:01 +0200 wenzelm workers: explicit activity flag;
Wed, 10 Sep 2008 19:44:29 +0200 wenzelm future: allow explicit group;
Wed, 10 Sep 2008 19:44:28 +0200 wenzelm cancel: invalidate group implicitly, via bool ref;
Wed, 10 Sep 2008 11:36:37 +0200 wenzelm auto_flush: uniform block buffering for all output streams;
Tue, 09 Sep 2008 23:48:38 +0200 wenzelm auto_flush stdout, stderr as well;
Tue, 09 Sep 2008 23:48:36 +0200 wenzelm proper values of no_interrupts, regular_interrupts;
Tue, 09 Sep 2008 23:30:05 +0200 wenzelm cancel: check_scheduler;
Tue, 09 Sep 2008 23:30:00 +0200 wenzelm simplified dequeue: provide Thread.self internally;
Tue, 09 Sep 2008 20:22:40 +0200 wenzelm eliminated cache, access queue efficiently via IntGraph.get_first;
Tue, 09 Sep 2008 20:22:30 +0200 wenzelm export get_first from underlying table;
Tue, 09 Sep 2008 19:57:54 +0200 wenzelm out_stream: block-buffered, with separate autoflush thread (every 50ms);
Tue, 09 Sep 2008 19:36:21 +0200 wenzelm babel: removed unnecessary "french" option, which actually enables french section names etc. on some LaTeX installations;
Tue, 09 Sep 2008 19:33:22 +0200 nipkow added comment
Tue, 09 Sep 2008 16:59:48 +0200 wenzelm human-readable printing of TaskQueue.task/group;
Tue, 09 Sep 2008 16:35:57 +0200 wenzelm * Changed defaults for unify configuration options;
Tue, 09 Sep 2008 16:29:34 +0200 wenzelm inherit group from running thread, or create a new one -- make it harder to re-use canceled groups;
Tue, 09 Sep 2008 16:29:32 +0200 wenzelm job: explicit 'ok' status -- false for canceled jobs;
Tue, 09 Sep 2008 16:17:08 +0200 paulson Overall exception handler in order to insulate our users from low-level bugs.
Tue, 09 Sep 2008 16:16:20 +0200 paulson more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
Tue, 09 Sep 2008 16:15:25 +0200 paulson Increasing the default limits in order to prevent unnecessary failures.
Mon, 08 Sep 2008 22:14:39 +0200 wenzelm send: broadcast condition while locked!
Mon, 08 Sep 2008 21:08:30 +0200 wenzelm proper signature constraint;
Mon, 08 Sep 2008 20:35:38 +0200 wenzelm tuned Mailbox.send;
Mon, 08 Sep 2008 20:33:29 +0200 wenzelm removed unused sync_interrupts;
Mon, 08 Sep 2008 20:33:27 +0200 wenzelm moved thread data to future.ML (again);
Mon, 08 Sep 2008 20:33:24 +0200 wenzelm more interrupt operations;
Mon, 08 Sep 2008 16:08:23 +0200 wenzelm moved task, thread_data, group, queue to task_queue.ML;
Mon, 08 Sep 2008 16:08:18 +0200 wenzelm Ordered queue of grouped tasks.
Mon, 08 Sep 2008 16:08:13 +0200 wenzelm added Concurrent/task_queue.ML;
Mon, 08 Sep 2008 00:25:34 +0200 wenzelm await: SYNCHRONIZED wait!
Mon, 08 Sep 2008 00:10:41 +0200 wenzelm tuned check_cache;
Sun, 07 Sep 2008 22:20:15 +0200 wenzelm added sync_interrupts, regular_interrupts;
Sun, 07 Sep 2008 22:20:11 +0200 wenzelm added sync_interrupts, regular_interrupts;
Sun, 07 Sep 2008 22:20:08 +0200 wenzelm opaque signature constraint abstracts local type abbrev;
Sun, 07 Sep 2008 22:19:58 +0200 wenzelm tuned;
Sun, 07 Sep 2008 22:19:46 +0200 wenzelm added change_result;
Sun, 07 Sep 2008 22:19:42 +0200 wenzelm Functional threads as future values.
Sun, 07 Sep 2008 22:19:31 +0200 wenzelm added Concurrent/future.ML;
Sun, 07 Sep 2008 17:48:50 +0200 wenzelm Default (mostly dummy) implementation of thread structures.
Sun, 07 Sep 2008 17:48:49 +0200 wenzelm *** MESSAGE REFERS TO PREVIOUS VERSION ***
Sun, 07 Sep 2008 17:46:44 +0200 wenzelm *** empty log message ***
Sun, 07 Sep 2008 17:46:43 +0200 wenzelm explicit use of universal.ML and dummy_thread.ML;
Sun, 07 Sep 2008 17:46:41 +0200 wenzelm added no_interrupts;
Sun, 07 Sep 2008 17:46:40 +0200 wenzelm added no_interrupts;
Sun, 07 Sep 2008 17:46:39 +0200 wenzelm tuned;
Sun, 07 Sep 2008 17:46:38 +0200 wenzelm send: broadcast to all waiting threads;
Sun, 07 Sep 2008 17:46:37 +0200 wenzelm added ML-Systems/thread_dummy.ML;
Sat, 06 Sep 2008 14:02:36 +0200 haftmann dropped "run" marker in monad syntax
Fri, 05 Sep 2008 11:50:35 +0200 wenzelm multithreading.ML provides dummy thread structures;
Fri, 05 Sep 2008 06:50:22 +0200 haftmann different bookkeeping for code equations
Fri, 05 Sep 2008 06:50:20 +0200 haftmann renamed structure CodeTarget to Code_Target
Fri, 05 Sep 2008 00:19:50 +0200 huffman instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
Thu, 04 Sep 2008 21:12:06 +0200 wenzelm proper header;
Thu, 04 Sep 2008 21:02:42 +0200 wenzelm added receive_timeout;
Thu, 04 Sep 2008 20:06:23 +0200 wenzelm check WRAPPER_OUTPUT node type;
Thu, 04 Sep 2008 20:05:48 +0200 wenzelm init: disallow "" as out stream;
Thu, 04 Sep 2008 19:47:13 +0200 wenzelm fixed deps: no Concurrent/receiver.ML yet;
Thu, 04 Sep 2008 19:45:13 +0200 wenzelm Concurrent message exchange via mailbox -- with unbounded queueing.
Thu, 04 Sep 2008 19:45:12 +0200 wenzelm added Concurrent/mailbox.ML;
Thu, 04 Sep 2008 17:24:18 +0200 huffman reorganize subsections
Thu, 04 Sep 2008 17:21:49 +0200 huffman rename INF_drop_prefix to INFM_drop_prefix
Thu, 04 Sep 2008 17:19:57 +0200 huffman add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
Thu, 04 Sep 2008 17:18:44 +0200 huffman move diff_add_cancel, add_diff_cancel from class ab_group_add to group_add
Thu, 04 Sep 2008 16:43:51 +0200 wenzelm tuned signature;
Thu, 04 Sep 2008 16:43:50 +0200 wenzelm added General/queue.ML;
(0) -10000 -3000 -1000 -240 +240 +1000 +3000 +10000 +30000 tip