Mon, 12 Nov 2012 11:32:22 +0100 blanchet thread context correctly when printing backquoted facts
Sun, 11 Nov 2012 19:56:02 +0100 haftmann dropped dead code;
Sun, 11 Nov 2012 19:24:01 +0100 haftmann modernized, simplified and compacted oracle and proof method glue code;
Fri, 09 Nov 2012 19:21:47 +0100 nipkow merged
Fri, 09 Nov 2012 19:16:31 +0100 nipkow fixed underscores
Fri, 09 Nov 2012 14:31:26 +0100 immler moved lemmas into projective_family; added header for theory Projective_Family
Fri, 09 Nov 2012 14:14:45 +0100 immler removed redundant/unnecessary assumptions from projective_family
Wed, 07 Nov 2012 14:41:49 +0100 immler assume probability spaces; allow empty index set
Wed, 07 Nov 2012 11:33:27 +0100 immler added projective_family; generalized generator in product_prob_space to projective_family
Tue, 06 Nov 2012 11:03:28 +0100 immler moved lemmas further up
Thu, 08 Nov 2012 20:02:41 +0100 bulwahn tuned proofs
Thu, 08 Nov 2012 19:55:37 +0100 bulwahn using hyp_subst_tac that allows to pass the current simpset to avoid the renamed bound variable warning in the simplifier
Thu, 08 Nov 2012 19:55:35 +0100 bulwahn hyp_subst_tac allows to pass an optional simpset to the internal simplifier call to avoid renamed bound variable warnings in the simplifier call
Thu, 08 Nov 2012 19:55:19 +0100 bulwahn NEWS
Thu, 08 Nov 2012 19:55:17 +0100 bulwahn rewriting with the simpset that is passed to the simproc
Thu, 08 Nov 2012 17:11:04 +0100 bulwahn handling x : S y pattern with the default mechanism instead of raising an exception in the set_comprehension_pointfree simproc
Thu, 08 Nov 2012 17:06:59 +0100 bulwahn tuned
Thu, 08 Nov 2012 16:25:26 +0100 bulwahn syntactic tuning and restructuring of set_comprehension_pointfree simproc
Thu, 08 Nov 2012 11:59:50 +0100 bulwahn using more proper simpset in tactic of set_comprehension_pointfree simproc to avoid renamed bound variable warnings in recursive simplifier calls
Thu, 08 Nov 2012 11:59:49 +0100 bulwahn improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
Thu, 08 Nov 2012 11:59:48 +0100 bulwahn adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
Thu, 08 Nov 2012 11:59:47 +0100 bulwahn importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
Thu, 08 Nov 2012 11:59:47 +0100 bulwahn handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
Thu, 08 Nov 2012 11:59:46 +0100 bulwahn simplified structure of patterns in set_comprehension_simproc
Thu, 08 Nov 2012 10:02:38 +0100 haftmann refined stack of library theories implementing int and/or nat by target language numerals
Wed, 07 Nov 2012 20:48:04 +0100 haftmann restored SML code check which got unintentionally broken: must explicitly check for error during compilation;
Tue, 06 Nov 2012 19:18:35 +0100 hoelzl add support for function application to measurability prover
Tue, 06 Nov 2012 15:15:33 +0100 blanchet renamed Sledgehammer option
Tue, 06 Nov 2012 15:12:31 +0100 blanchet always show timing for structured proofs
Tue, 06 Nov 2012 14:46:21 +0100 blanchet use implications rather than disjunctions to improve readability
Tue, 06 Nov 2012 13:47:51 +0100 blanchet avoid name clashes
Tue, 06 Nov 2012 13:09:02 +0100 blanchet fixed more "Trueprop" issues
Tue, 06 Nov 2012 12:38:45 +0100 blanchet removed needless sort
Tue, 06 Nov 2012 11:24:48 +0100 blanchet avoid double "Trueprop"s
Tue, 06 Nov 2012 11:20:56 +0100 blanchet use original formulas for hypotheses and conclusion to avoid mismatches
Tue, 06 Nov 2012 11:20:56 +0100 blanchet track formula roles in proofs and use that to determine whether the conjecture should be negated or not
Tue, 06 Nov 2012 11:20:56 +0100 blanchet correct parsing of E dependencies
Tue, 06 Nov 2012 11:20:56 +0100 blanchet proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
Mon, 05 Nov 2012 11:40:51 +0100 nipkow tuned
Sun, 04 Nov 2012 18:41:27 +0100 nipkow code for while directly, not via while_option
Sun, 04 Nov 2012 18:38:18 +0100 nipkow executable true liveness analysis incl an approximating version
Sun, 04 Nov 2012 17:36:26 +0100 nipkow now that sets are executable again, no more special treatment of variable sets
Fri, 02 Nov 2012 16:16:48 +0100 blanchet handle non-unit clauses gracefully
Fri, 02 Nov 2012 16:16:48 +0100 blanchet several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
Fri, 02 Nov 2012 14:23:54 +0100 hoelzl use measurability prover
Fri, 02 Nov 2012 14:23:40 +0100 hoelzl add measurability prover; add support for Borel sets
Fri, 02 Nov 2012 14:00:39 +0100 hoelzl add syntax and a.e.-rules for (conditional) probability on predicates
Fri, 02 Nov 2012 14:00:39 +0100 hoelzl infinite product measure is invariant under adding prefixes
Fri, 02 Nov 2012 14:00:39 +0100 hoelzl for the product measure it is enough if only one measure is sigma-finite
Fri, 02 Nov 2012 12:00:51 +0100 berghofe Allow parentheses around left-hand sides of array associations
Thu, 01 Nov 2012 15:00:48 +0100 blanchet made MaSh more robust in the face of duplicate "nicknames" (which can happen e.g. if you have a lemma called foo(1) and another called foo_1 in the same theory)
Thu, 01 Nov 2012 13:32:57 +0100 blanchet regenerated SMT certificates
Thu, 01 Nov 2012 11:34:00 +0100 blanchet regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
Wed, 31 Oct 2012 11:23:21 +0100 blanchet fixed bool vs. prop mismatch
Wed, 31 Oct 2012 11:23:21 +0100 blanchet removed "refute" command from Isar manual, now that it has been moved outside "Main"
Wed, 31 Oct 2012 11:23:21 +0100 blanchet repaired "Mutabelle" after Refute move
Wed, 31 Oct 2012 11:23:21 +0100 blanchet less verbose -- the warning will reach the users anyway by other means
Wed, 31 Oct 2012 11:23:21 +0100 blanchet tuned messages
Wed, 31 Oct 2012 11:23:21 +0100 blanchet moved "SAT" before "FunDef" and moved back all SAT-related ML files to where they belong
Wed, 31 Oct 2012 11:23:21 +0100 blanchet fixes related to Refute's move
Wed, 31 Oct 2012 11:23:21 +0100 blanchet added a timeout around script that relies on the network
Wed, 31 Oct 2012 11:23:21 +0100 blanchet took out "using only ..." comments in Sledgehammer generated metis/smt calls, until these can be generated soundly
Wed, 31 Oct 2012 11:23:21 +0100 blanchet moved Refute to "HOL/Library" to speed up building "Main" even more
Wed, 31 Oct 2012 11:23:21 +0100 blanchet tuning
Wed, 31 Oct 2012 11:23:21 +0100 blanchet use metaquantification when possible in Isar proofs
Wed, 31 Oct 2012 11:23:21 +0100 blanchet tuned code
Wed, 31 Oct 2012 11:23:21 +0100 blanchet tuning
Wed, 31 Oct 2012 11:23:21 +0100 blanchet soft SMT timeout
Sun, 28 Oct 2012 02:22:39 +0000 Christian Urban added function store_termination_rule to the signature, as it is used in Nominal2
Sat, 27 Oct 2012 20:59:50 +0200 wenzelm longer log, to accomodate final status line of isabelle build;
Wed, 24 Oct 2012 18:43:25 +0200 huffman transfer package: error message if preprocessing goal to object-logic formula fails
Wed, 24 Oct 2012 18:43:25 +0200 huffman transfer package: add test to prevent trying to make cterms from open terms
Wed, 24 Oct 2012 18:43:25 +0200 huffman transfer package: more flexible handling of equality relations using is_equality predicate
Wed, 24 Oct 2012 17:40:56 +0200 nipkow ensured that rewr_conv rule t = "t == u" literally not just modulo beta-eta
Mon, 22 Oct 2012 22:47:14 +0200 kuncar new theorems
Mon, 22 Oct 2012 22:24:34 +0200 haftmann incorporated constant chars into instantiation proof for enum;
Mon, 22 Oct 2012 19:02:36 +0200 haftmann close code theorems explicitly after preprocessing
Mon, 22 Oct 2012 17:09:49 +0200 wenzelm tuned proofs;
Mon, 22 Oct 2012 16:27:55 +0200 wenzelm further attempts to cope with large files via option jedit_text_overview_limit;
Mon, 22 Oct 2012 14:52:38 +0200 wenzelm more detailed Prover IDE NEWS;
Sun, 21 Oct 2012 22:32:22 +0200 wenzelm recovered explicit error message, which was lost in b8570ea1ce25;
Sun, 21 Oct 2012 22:31:39 +0200 wenzelm removed dead code;
Sun, 21 Oct 2012 22:12:22 +0200 wenzelm proper signatures;
Sun, 21 Oct 2012 22:11:38 +0200 wenzelm tuned;
Sun, 21 Oct 2012 17:04:13 +0200 webertj merged
Fri, 19 Oct 2012 15:12:52 +0200 webertj Renamed {left,right}_distrib to distrib_{right,left}.
Fri, 19 Oct 2012 10:46:42 +0200 webertj Tuned.
Sun, 21 Oct 2012 16:43:08 +0200 haftmann more conventional argument order;
Sun, 21 Oct 2012 08:39:41 +0200 bulwahn another refinement in the comprehension conversion
Sun, 21 Oct 2012 05:24:59 +0200 bulwahn refined simplifier call in comprehension_conv
Sun, 21 Oct 2012 05:24:56 +0200 bulwahn passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
Sat, 20 Oct 2012 17:40:51 +0200 wenzelm avoid STIX font, which tends to render badly;
Sat, 20 Oct 2012 17:15:40 +0200 wenzelm extra jar for scala-2.10.0-RC1;
Sat, 20 Oct 2012 15:46:48 +0200 wenzelm more explicit auxiliary classes to avoid warning "reflective access of structural type member method" of scala-2.10.0-RC1;
Sat, 20 Oct 2012 15:45:40 +0200 wenzelm avoid duplicate build of jars_fresh;
Sat, 20 Oct 2012 12:01:20 +0200 wenzelm obsolete, cf. README_REPOSITORY;
Sat, 20 Oct 2012 12:00:48 +0200 wenzelm accomodate scala-2.10.0-RC1;
Sat, 20 Oct 2012 10:00:21 +0200 haftmann tailored enum specification towards simple instantiation;
Sat, 20 Oct 2012 10:00:21 +0200 haftmann refined internal structure of Enum.thy
Sat, 20 Oct 2012 09:12:16 +0200 haftmann moved quite generic material from theory Enum to more appropriate places
Sat, 20 Oct 2012 09:09:37 +0200 bulwahn adding another test case for the set_comprehension_simproc to the theory in HOL/ex
Sat, 20 Oct 2012 09:09:35 +0200 bulwahn improving tactic in setcomprehension_simproc
Sat, 20 Oct 2012 09:09:34 +0200 bulwahn adjusting proofs
Sat, 20 Oct 2012 09:09:33 +0200 bulwahn tuned tactic in set_comprehension_pointfree simproc to handle composition of negation and vimage
Sat, 20 Oct 2012 09:09:32 +0200 bulwahn passing names and types of all bounds around in the simproc
Thu, 18 Oct 2012 10:06:27 +0200 bulwahn locally inverting previously applied simplifications with ex_simps in set_comprehension_pointfree
Fri, 19 Oct 2012 21:52:45 +0200 wenzelm more precise pixel_range: avoid popup when pointing into empty space after actual end-of-line;
Fri, 19 Oct 2012 21:19:06 +0200 wenzelm merged
Fri, 19 Oct 2012 17:54:16 +0200 kuncar don't include Quotient_Option - workaround to a transfer bug
Fri, 19 Oct 2012 21:18:34 +0200 wenzelm ignore old stuff and thus speed up the script greatly;
Fri, 19 Oct 2012 20:15:14 +0200 wenzelm proper find -mtime (file data) instead of -ctime (meta data);
Fri, 19 Oct 2012 17:52:21 +0200 wenzelm made SML/NJ happy;
Fri, 19 Oct 2012 12:08:13 +0200 wenzelm clarified Future.map (again): finished value is mapped in-place, which saves task structures and changes error behaviour slightly (tolerance against canceled group of old value etc.);
Thu, 18 Oct 2012 20:59:46 +0200 wenzelm back to polyml-5.4.1 (cf. b3110dec1a32) -- no cause of spurious interrupts;
Thu, 18 Oct 2012 20:45:15 +0200 wenzelm merged
Thu, 18 Oct 2012 20:00:45 +0200 wenzelm back to parallel HOL-BNF-Examples, which seems to have suffered from Future.map on canceled persistent futures;
Thu, 18 Oct 2012 19:58:30 +0200 wenzelm more basic Goal.reset_futures as snapshot of implicit state;
Thu, 18 Oct 2012 19:12:58 +0200 wenzelm tuned proof;
Thu, 18 Oct 2012 15:52:33 +0200 kuncar update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
Thu, 18 Oct 2012 15:52:32 +0200 kuncar tuned proofs
Thu, 18 Oct 2012 15:52:31 +0200 kuncar new theorem
Thu, 18 Oct 2012 15:47:01 +0200 wenzelm merged
Thu, 18 Oct 2012 15:44:14 +0200 wenzelm merged
Thu, 18 Oct 2012 15:28:49 +0200 wenzelm merged
Thu, 18 Oct 2012 15:16:39 +0200 wenzelm merged
Thu, 18 Oct 2012 15:15:08 +0200 wenzelm fixed proof (cf. a81f95693c68);
Thu, 18 Oct 2012 15:41:15 +0200 blanchet tuned Isar output
Thu, 18 Oct 2012 15:40:02 +0200 nipkow tuned
Thu, 18 Oct 2012 15:10:49 +0200 blanchet updated docs
Thu, 18 Oct 2012 15:05:17 +0200 blanchet renamed Isar-proof related options + changed semantics of Isar shrinking
Thu, 18 Oct 2012 14:26:45 +0200 blanchet tuning
Thu, 18 Oct 2012 13:46:24 +0200 blanchet fixed theorem lookup code in Isar proof reconstruction
Thu, 18 Oct 2012 13:37:53 +0200 blanchet tuning
Thu, 18 Oct 2012 13:19:44 +0200 blanchet refactor code
Thu, 18 Oct 2012 11:59:45 +0200 blanchet tuning
Thu, 18 Oct 2012 14:15:46 +0200 wenzelm more robust future_proof result with specific error message (e.g. relevant for incomplete proof of non-registered theorem);
Thu, 18 Oct 2012 13:57:27 +0200 wenzelm collective errors from use_thys and Session.finish/Goal.finish_futures -- avoid uninformative interrupts stemming from failure of goal forks that are not registered in the theory (e.g. unnamed theorems);
Thu, 18 Oct 2012 13:53:02 +0200 wenzelm more uniform group for map_future, which is relevant for cancel in worker_task vs. future_job -- prefer peer group despite 81d03a29980c;
Thu, 18 Oct 2012 13:26:49 +0200 wenzelm tuned message;
Thu, 18 Oct 2012 12:47:30 +0200 wenzelm tuned comment;
Thu, 18 Oct 2012 12:26:30 +0200 wenzelm avoid spurious "bad" markup for show/test_proof;
Thu, 18 Oct 2012 12:00:27 +0200 wenzelm more official Future.terminate;
Thu, 18 Oct 2012 09:19:37 +0200 haftmann simp results for simplification results of Inf/Sup expressions on bool;
Thu, 18 Oct 2012 09:17:00 +0200 haftmann no sort constraints on datatype constructors in internal bookkeeping
Wed, 17 Oct 2012 22:57:28 +0200 wenzelm HOL-BNF-Examples is sequential for now, due to spurious interrupts (again);
Wed, 17 Oct 2012 22:45:40 +0200 wenzelm merged
Wed, 17 Oct 2012 15:25:52 +0200 bulwahn comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
Wed, 17 Oct 2012 14:13:57 +0200 bulwahn checking for bound variables in the set expression; handling negation more generally
Wed, 17 Oct 2012 14:13:57 +0200 bulwahn set_comprehension_pointfree simproc now handles the complicated test case; tuned
Wed, 17 Oct 2012 14:13:57 +0200 bulwahn refined conversion to only react on proper set comprehensions; tuned
Wed, 17 Oct 2012 14:13:57 +0200 bulwahn moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
Wed, 17 Oct 2012 14:13:57 +0200 bulwahn employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
Wed, 17 Oct 2012 22:11:12 +0200 wenzelm another Future.shutdown after Future.cancel_groups (cf. 0d4106850eb2);
Wed, 17 Oct 2012 21:18:32 +0200 wenzelm more robust cancel_now: avoid shooting yourself in the foot;
Wed, 17 Oct 2012 21:04:51 +0200 wenzelm more robust Session.finish (batch mode): use Goal.finish_futures to exhibit remaining failures of disconnected goal forks (e.g. from unnamed theorems) and Goal.cancel_futures the purge the persistent state;
Wed, 17 Oct 2012 14:58:04 +0200 wenzelm proper 'oops' to force sequential checking here, and avoid spurious *** Interrupt stemming from crash of forked outer syntax element;
Wed, 17 Oct 2012 14:39:00 +0200 wenzelm added Output "Detach" button;
Wed, 17 Oct 2012 14:20:54 +0200 wenzelm skipped proofs appear as "bad" without counting as error;
Wed, 17 Oct 2012 13:20:08 +0200 wenzelm more method position information, notably finished_pos after end of previous text;
Wed, 17 Oct 2012 10:46:14 +0200 wenzelm more formal markup;
Wed, 17 Oct 2012 10:45:43 +0200 wenzelm tuned signature;
Wed, 17 Oct 2012 10:26:27 +0200 wenzelm more formal markup;
Wed, 17 Oct 2012 00:16:31 +0200 kuncar don't be so aggressive when expanding a transfer rule relation; rewrite only the relational part of the rule
Tue, 16 Oct 2012 22:38:34 +0200 wenzelm merged
Tue, 16 Oct 2012 20:31:08 +0200 blanchet added missing file
Tue, 16 Oct 2012 20:11:15 +0200 traytel tuned for document output
Tue, 16 Oct 2012 18:50:53 +0200 blanchet added proof minimization code from Steffen Smolka
Tue, 16 Oct 2012 18:07:59 +0200 traytel tuned blank lines
Tue, 16 Oct 2012 18:05:28 +0200 traytel tuned whitespace
Tue, 16 Oct 2012 17:33:08 +0200 popescua a few notations changed in HOL/BNF/Examples/Derivation_Trees
Tue, 16 Oct 2012 17:08:20 +0200 popescua ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
Tue, 16 Oct 2012 13:57:08 +0200 bulwahn adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
Tue, 16 Oct 2012 13:18:13 +0200 bulwahn tactic of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
Tue, 16 Oct 2012 13:18:12 +0200 bulwahn term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
Tue, 16 Oct 2012 13:18:10 +0200 bulwahn extending preprocessing of simproc to rewrite subset inequality into membership of powerset
Tue, 16 Oct 2012 13:15:58 +0200 popescua update ROOT with teh directory change in BNF
Tue, 16 Oct 2012 13:09:46 +0200 popescua changed name of BNF/Example directory from Infinite_Derivation_Trees to Derivation_Trees
Tue, 16 Oct 2012 22:13:46 +0200 wenzelm retain info dockable state via educated guess on window focus;
Tue, 16 Oct 2012 21:30:52 +0200 wenzelm support for more informative errors in lazy enumerations;
Tue, 16 Oct 2012 21:26:36 +0200 wenzelm more informative errors for 'also' and 'finally';
Tue, 16 Oct 2012 20:35:24 +0200 wenzelm tuned messages;
Tue, 16 Oct 2012 20:23:00 +0200 wenzelm more proof method text position information;
Tue, 16 Oct 2012 17:47:23 +0200 wenzelm clarified defer/prefer: more specific errors;
Tue, 16 Oct 2012 16:50:03 +0200 wenzelm updated Toplevel.proofs;
Tue, 16 Oct 2012 15:14:12 +0200 wenzelm more informative errors for 'proof' and 'apply' steps;
Tue, 16 Oct 2012 15:02:49 +0200 wenzelm more friendly handling of Pure.thy bootstrap errors;
Tue, 16 Oct 2012 14:14:37 +0200 wenzelm more informative error for stand-alone 'qed';
Tue, 16 Oct 2012 14:02:02 +0200 wenzelm further attempts to unify/simplify goal output;
Tue, 16 Oct 2012 13:06:40 +0200 wenzelm more informative error messages of initial/terminal proof methods;
Mon, 15 Oct 2012 19:03:02 +0200 wenzelm merged
Mon, 15 Oct 2012 16:18:48 +0200 bulwahn setcomprehension_pointfree simproc also works for set comprehension without an equation
Mon, 15 Oct 2012 15:43:12 +0200 wenzelm tuned message -- avoid extra blank lines;
Mon, 15 Oct 2012 15:28:56 +0200 wenzelm updated to polyml-5.5.0 which reduces chance of HOL-IMP failure (although it is hard to reproduce anyway);
Sun, 14 Oct 2012 21:02:14 +0200 Markus Kaiser store colors after build
Sun, 14 Oct 2012 19:16:39 +0200 bulwahn adding further test cases for the set_comprehension_pointfree simproc
Sun, 14 Oct 2012 19:16:35 +0200 bulwahn refined tactic in set_comprehension_pointfree simproc
Sun, 14 Oct 2012 19:16:33 +0200 bulwahn adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
Sun, 14 Oct 2012 19:16:32 +0200 bulwahn adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
Sun, 14 Oct 2012 19:16:32 +0200 bulwahn extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
Sat, 13 Oct 2012 21:09:20 +0200 wenzelm more informative error of initial/terminal proof steps;
Sat, 13 Oct 2012 19:53:04 +0200 wenzelm some attempts to unify/simplify pretty_goal;
Sat, 13 Oct 2012 18:04:11 +0200 wenzelm refined Proof.the_finished_goal with more informative error;
Sat, 13 Oct 2012 16:19:16 +0200 wenzelm tuned signature;
Sat, 13 Oct 2012 00:08:36 +0200 wenzelm improved adhoc height for small fonts;
Fri, 12 Oct 2012 23:38:48 +0200 wenzelm further refinement of jEdit line range, avoiding lack of final \n;
Fri, 12 Oct 2012 22:53:20 +0200 wenzelm more uniform tooltip color;
Fri, 12 Oct 2012 22:10:45 +0200 wenzelm more NEWS;
Fri, 12 Oct 2012 21:51:25 +0200 wenzelm merged
Fri, 12 Oct 2012 15:52:55 +0200 traytel disambiguated grammar
Fri, 12 Oct 2012 15:52:45 +0200 traytel tuned proofs
Fri, 12 Oct 2012 14:57:56 +0200 nipkow tuned
Fri, 12 Oct 2012 21:39:58 +0200 wenzelm simplified 'typedef' specifications: discontinued implicit set definition and alternative name;
Fri, 12 Oct 2012 21:22:35 +0200 wenzelm discontinued typedef with alternative name;
Fri, 12 Oct 2012 18:58:20 +0200 wenzelm discontinued obsolete typedef (open) syntax;
Fri, 12 Oct 2012 15:08:29 +0200 wenzelm discontinued typedef with implicit set_def;
Fri, 12 Oct 2012 14:05:30 +0200 wenzelm merged
Fri, 12 Oct 2012 12:21:01 +0200 bulwahn increading indexes to avoid clashes in the set_comprehension_pointfree simproc
Fri, 12 Oct 2012 13:55:13 +0200 wenzelm no special treatment of errors inside goal forks without transaction id, to avoid duplication in plain build with sequential log, for example;
Fri, 12 Oct 2012 13:46:41 +0200 wenzelm do not treat interrupt as error here, to avoid confusion in log etc.;
Fri, 12 Oct 2012 11:03:23 +0200 wenzelm more basic ML compiler messages -- avoid conflict of 638cefe3ee99 and cb7264721c91 concerning Protocol.message_positions;
Thu, 11 Oct 2012 23:10:49 +0200 wenzelm refined separator: FBreak needs to be free for proper breaking, extra space at end helps to work around last-line oddity in jEdit;
Thu, 11 Oct 2012 21:02:19 +0200 wenzelm merged
Thu, 11 Oct 2012 14:38:58 +0200 hoelzl cleanup borel_measurable_positive_integral_(fst|snd)
Thu, 11 Oct 2012 11:56:43 +0200 haftmann msetprod based directly on Multiset.fold;
Thu, 11 Oct 2012 11:56:43 +0200 haftmann avoid global interpretation
Thu, 11 Oct 2012 11:56:42 +0200 haftmann simplified construction of fold combinator on multisets;
Thu, 11 Oct 2012 20:38:02 +0200 wenzelm clarified output token markup (see also bc22daeed49e);
Thu, 11 Oct 2012 19:25:36 +0200 wenzelm refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
Thu, 11 Oct 2012 16:09:44 +0200 wenzelm refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
Thu, 11 Oct 2012 15:26:33 +0200 wenzelm tuned;
Thu, 11 Oct 2012 15:06:27 +0200 wenzelm tuned;
Thu, 11 Oct 2012 12:38:18 +0200 wenzelm more position information for hyperlink and placement of message;
Thu, 11 Oct 2012 12:37:38 +0200 wenzelm tuned;
Thu, 11 Oct 2012 00:13:21 +0200 krauss mira: discontinued special settings for lxbroy10, which are probably made obsolete by newer polyml
Wed, 10 Oct 2012 22:53:48 +0200 krauss removed unused legacy material from mira.py
Wed, 10 Oct 2012 17:43:23 +0200 wenzelm eliminated some remaining uses of typedef with implicit set definition;
Wed, 10 Oct 2012 16:41:19 +0200 Andreas Lochbihler merged
Wed, 10 Oct 2012 16:18:27 +0200 Andreas Lochbihler fix code equation for RBT_Impl.fold
Wed, 10 Oct 2012 15:17:18 +0200 Andreas Lochbihler merged
Wed, 10 Oct 2012 15:16:44 +0200 Andreas Lochbihler tail-recursive implementation for length
(0) -30000 -10000 -3000 -1000 -240 +240 +1000 +3000 +10000 +30000 tip