2010-08-26 re-added accidental omission
haftmann [Thu, 26 Aug 2010 13:25:14 +0200] rev 38775
re-added accidental omission
2010-08-26 tuned includes
haftmann [Thu, 26 Aug 2010 12:19:50 +0200] rev 38774
tuned includes
2010-08-26 prevent line breaks after Scala symbolic operators
haftmann [Thu, 26 Aug 2010 12:19:49 +0200] rev 38773
prevent line breaks after Scala symbolic operators
2010-08-26 corrected semantics of presentation_stmt_names; do not print includes on presentation selection
haftmann [Thu, 26 Aug 2010 10:23:25 +0200] rev 38772
corrected semantics of presentation_stmt_names; do not print includes on presentation selection
2010-08-26 code_include Scala: qualify module nmae
haftmann [Thu, 26 Aug 2010 10:16:22 +0200] rev 38771
code_include Scala: qualify module nmae
2010-08-25 merged
haftmann [Wed, 25 Aug 2010 22:47:04 +0200] rev 38770
merged
2010-08-25 preliminary implementation of hierarchical module name space
haftmann [Wed, 25 Aug 2010 16:33:05 +0200] rev 38769
preliminary implementation of hierarchical module name space
2010-08-25 tuned
haftmann [Wed, 25 Aug 2010 16:33:05 +0200] rev 38768
tuned
2010-08-27 proper context for various Thy_Output options, via official configuration options in ML and Isar;
wenzelm [Fri, 27 Aug 2010 12:40:20 +0200] rev 38767
proper context for various Thy_Output options, via official configuration options in ML and Isar;
2010-08-26 Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
wenzelm [Fri, 27 Aug 2010 00:09:56 +0200] rev 38766
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
2010-08-26 eliminated old 'local' command;
wenzelm [Fri, 27 Aug 2010 00:02:32 +0200] rev 38765
eliminated old 'local' command;
2010-08-26 more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm [Thu, 26 Aug 2010 21:04:22 +0200] rev 38764
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
2010-08-26 Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
wenzelm [Thu, 26 Aug 2010 20:42:09 +0200] rev 38763
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
2010-08-26 slightly more abstract data handling in Fast_Lin_Arith;
wenzelm [Thu, 26 Aug 2010 17:37:26 +0200] rev 38762
slightly more abstract data handling in Fast_Lin_Arith;
2010-08-26 theory data merge: prefer left side uniformly;
wenzelm [Thu, 26 Aug 2010 17:01:12 +0200] rev 38761
theory data merge: prefer left side uniformly;
2010-08-26 tuned;
wenzelm [Thu, 26 Aug 2010 16:56:45 +0200] rev 38760
tuned;
2010-08-26 simplification/standardization of some theory data;
wenzelm [Thu, 26 Aug 2010 16:34:10 +0200] rev 38759
simplification/standardization of some theory data;
2010-08-26 misc tuning and simplification, notably theory data;
wenzelm [Thu, 26 Aug 2010 16:25:25 +0200] rev 38758
misc tuning and simplification, notably theory data;
2010-08-26 renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm [Thu, 26 Aug 2010 15:48:08 +0200] rev 38757
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
2010-08-26 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm [Thu, 26 Aug 2010 13:09:12 +0200] rev 38756
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
2010-08-26 standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
wenzelm [Thu, 26 Aug 2010 12:06:00 +0200] rev 38755
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
2010-08-26 merged
wenzelm [Thu, 26 Aug 2010 11:33:36 +0200] rev 38754
merged
2010-08-26 merged
blanchet [Thu, 26 Aug 2010 10:42:22 +0200] rev 38753
merged
2010-08-26 consider "locality" when assigning weights to facts
blanchet [Thu, 26 Aug 2010 10:42:06 +0200] rev 38752
consider "locality" when assigning weights to facts
2010-08-26 add a bonus for chained facts, since they are likely to be relevant;
blanchet [Thu, 26 Aug 2010 09:23:21 +0200] rev 38751
add a bonus for chained facts, since they are likely to be relevant; (especially in a Mirabelle run!) -- chained facts used to be included forcibly, then were treated as any other fact; the current approach seems more flexible
2010-08-26 merged
blanchet [Thu, 26 Aug 2010 09:03:18 +0200] rev 38750
merged
2010-08-25 add a penalty for lambda-abstractions;
blanchet [Thu, 26 Aug 2010 01:03:08 +0200] rev 38749
add a penalty for lambda-abstractions; the penalty will kick in only when the goal contains no lambdas, in which case Sledgehammer previously totally disallowed any higher-order construct; this was too drastic; lambdas are dangerous because they rapidly lead to unsound proofs; e.g. COMBI_def COMBS_def not_Cons_self2 with explicit_apply
2010-08-25 renaming
blanchet [Thu, 26 Aug 2010 00:49:38 +0200] rev 38748
renaming
2010-08-25 fiddle with relevance filter
blanchet [Thu, 26 Aug 2010 00:49:04 +0200] rev 38747
fiddle with relevance filter
2010-08-25 update docs
blanchet [Wed, 25 Aug 2010 19:47:25 +0200] rev 38746
update docs
2010-08-25 reorganize options regarding to the relevance threshold and decay
blanchet [Wed, 25 Aug 2010 19:41:18 +0200] rev 38745
reorganize options regarding to the relevance threshold and decay
2010-08-25 make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet [Wed, 25 Aug 2010 17:49:52 +0200] rev 38744
make relevance filter work in term of a "max_relevant" option + use Vampire SOS; "max_relevant" is more reliable than "max_relevant_per_iter"; also made sure that the option is monotone -- larger values should lead to more axioms -- which wasn't always the case before; SOS for Vampire makes a difference of about 3% (i.e. 3% more proofs are found)
2010-08-25 simplify more code
blanchet [Wed, 25 Aug 2010 09:42:28 +0200] rev 38743
simplify more code
2010-08-25 cosmetics
blanchet [Wed, 25 Aug 2010 09:34:28 +0200] rev 38742
cosmetics
2010-08-25 get rid of "defs_relevant" feature;
blanchet [Wed, 25 Aug 2010 09:32:43 +0200] rev 38741
get rid of "defs_relevant" feature; nobody uses it and it works poorly
2010-08-25 make SML/NJ happy
blanchet [Wed, 25 Aug 2010 09:05:22 +0200] rev 38740
make SML/NJ happy
2010-08-25 renamed "relevance_convergence" to "relevance_decay"
blanchet [Wed, 25 Aug 2010 09:02:07 +0200] rev 38739
renamed "relevance_convergence" to "relevance_decay"
2010-08-24 make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet [Tue, 24 Aug 2010 22:57:22 +0200] rev 38738
make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
2010-08-24 better workaround for E's off-by-one-second issue
blanchet [Tue, 24 Aug 2010 21:40:03 +0200] rev 38737
better workaround for E's off-by-one-second issue
2010-08-26 merged
bulwahn [Thu, 26 Aug 2010 09:12:00 +0200] rev 38736
merged
2010-08-25 renaming variables to conform to prolog names
bulwahn [Wed, 25 Aug 2010 16:59:55 +0200] rev 38735
renaming variables to conform to prolog names
2010-08-25 changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn [Wed, 25 Aug 2010 16:59:53 +0200] rev 38734
changing hotel trace definition; adding simple handling of numerals on natural numbers
2010-08-25 added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn [Wed, 25 Aug 2010 16:59:51 +0200] rev 38733
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
2010-08-25 moving preprocessing to values in prolog generation
bulwahn [Wed, 25 Aug 2010 16:59:51 +0200] rev 38732
moving preprocessing to values in prolog generation
2010-08-25 invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn [Wed, 25 Aug 2010 16:59:50 +0200] rev 38731
invocation of values for prolog execution does not require invocation of code_pred anymore
2010-08-25 adding hotel keycard example for prolog generation
bulwahn [Wed, 25 Aug 2010 16:59:49 +0200] rev 38730
adding hotel keycard example for prolog generation
2010-08-25 improving output of set comprehensions; adding style_check flags
bulwahn [Wed, 25 Aug 2010 16:59:48 +0200] rev 38729
improving output of set comprehensions; adding style_check flags
2010-08-25 improving ensure_groundness in prolog generation; added further example
bulwahn [Wed, 25 Aug 2010 16:59:48 +0200] rev 38728
improving ensure_groundness in prolog generation; added further example
2010-08-25 adding very basic transformation to ensure groundness before negations
bulwahn [Wed, 25 Aug 2010 16:59:46 +0200] rev 38727
adding very basic transformation to ensure groundness before negations
2010-08-26 Markup_Tree.select: uniform treatment of root_range wrt. singularities, yielding empty result stream;
wenzelm [Thu, 26 Aug 2010 11:31:21 +0200] rev 38726
Markup_Tree.select: uniform treatment of root_range wrt. singularities, yielding empty result stream;
2010-08-26 tuned signature;
wenzelm [Thu, 26 Aug 2010 11:29:43 +0200] rev 38725
tuned signature;
2010-08-25 Pretty: tuned markup objects;
wenzelm [Wed, 25 Aug 2010 22:57:40 +0200] rev 38724
Pretty: tuned markup objects;
2010-08-25 tuned;
wenzelm [Wed, 25 Aug 2010 22:45:24 +0200] rev 38723
tuned;
2010-08-25 organized markup properties via apply/unapply patterns;
wenzelm [Wed, 25 Aug 2010 22:37:53 +0200] rev 38722
organized markup properties via apply/unapply patterns;
2010-08-25 added some proof state markup, notably number of subgoals (e.g. for indentation);
wenzelm [Wed, 25 Aug 2010 21:31:22 +0200] rev 38721
added some proof state markup, notably number of subgoals (e.g. for indentation); tuned;
2010-08-25 ML_Context.eval: produce antiquotation environment preferably in invisible context, to avoid displaced report messages from ML_Compiler;
wenzelm [Wed, 25 Aug 2010 20:43:03 +0200] rev 38720
ML_Context.eval: produce antiquotation environment preferably in invisible context, to avoid displaced report messages from ML_Compiler;
2010-08-25 merged
wenzelm [Wed, 25 Aug 2010 18:46:22 +0200] rev 38719
merged
2010-08-25 tuned code
Christian Urban <urbanc@in.tum.de> [Wed, 25 Aug 2010 20:04:49 +0800] rev 38718
tuned code
2010-08-25 quotient package: deal correctly with frees in lifted theorems
Christian Urban <urbanc@in.tum.de> [Wed, 25 Aug 2010 18:26:58 +0800] rev 38717
quotient package: deal correctly with frees in lifted theorems
2010-08-25 approximation_oracle: actually match true/false in ML, not arbitrary values;
wenzelm [Wed, 25 Aug 2010 18:38:49 +0200] rev 38716
approximation_oracle: actually match true/false in ML, not arbitrary values;
2010-08-25 renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm [Wed, 25 Aug 2010 18:36:22 +0200] rev 38715
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
2010-08-25 more precise Command.State accumulation;
wenzelm [Wed, 25 Aug 2010 18:19:04 +0200] rev 38714
more precise Command.State accumulation;
2010-08-25 eliminated some old camel case stuff;
wenzelm [Wed, 25 Aug 2010 17:45:35 +0200] rev 38713
eliminated some old camel case stuff;
2010-08-25 some attempts to recover Isabelle/ML style from the depths of time;
wenzelm [Wed, 25 Aug 2010 17:34:10 +0200] rev 38712
some attempts to recover Isabelle/ML style from the depths of time;
2010-08-25 keep persistent production tables as immutable vectors -- potential performance improvement on modern hardware;
wenzelm [Wed, 25 Aug 2010 16:13:55 +0200] rev 38711
keep persistent production tables as immutable vectors -- potential performance improvement on modern hardware;
2010-08-25 select tangible text ranges here -- more robust against ongoing changes of the markup query model
wenzelm [Wed, 25 Aug 2010 15:41:14 +0200] rev 38710
select tangible text ranges here -- more robust against ongoing changes of the markup query model
2010-08-25 structure Thm: less pervasive names;
wenzelm [Wed, 25 Aug 2010 15:12:49 +0200] rev 38709
structure Thm: less pervasive names;
2010-08-25 discontinued obsolete 'global' and 'local' commands;
wenzelm [Wed, 25 Aug 2010 14:18:09 +0200] rev 38708
discontinued obsolete 'global' and 'local' commands;
2010-08-25 merged
wenzelm [Wed, 25 Aug 2010 11:30:45 +0200] rev 38707
merged
2010-08-25 merged
hoelzl [Wed, 25 Aug 2010 11:10:08 +0200] rev 38706
merged
2010-08-24 moved generic lemmas in Probability to HOL
hoelzl [Tue, 24 Aug 2010 14:41:37 +0200] rev 38705
moved generic lemmas in Probability to HOL
2010-08-25 merged
wenzelm [Wed, 25 Aug 2010 11:13:58 +0200] rev 38704
merged
2010-08-25 traling newline on standard output
haftmann [Wed, 25 Aug 2010 09:44:54 +0200] rev 38703
traling newline on standard output
2010-08-25 Quotient Package / lemma for regularization of bex1_rel for equivalence relations
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 25 Aug 2010 11:17:33 +0900] rev 38702
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
2010-08-24 merged
blanchet [Tue, 24 Aug 2010 20:09:30 +0200] rev 38701
merged
2010-08-24 make Mirabelle happy
blanchet [Tue, 24 Aug 2010 19:55:34 +0200] rev 38700
make Mirabelle happy
2010-08-24 compute names lazily;
blanchet [Tue, 24 Aug 2010 19:19:28 +0200] rev 38699
compute names lazily; there's no point to compute the names of 10000 facts when only 500 are used
2010-08-24 clean handling of whether a fact is chained or not;
blanchet [Tue, 24 Aug 2010 18:03:43 +0200] rev 38698
clean handling of whether a fact is chained or not; more elegant and reliable than encoding it in the name
2010-08-24 don't backtick facts that contain schematic variables, since this doesn't work (for some reason)
blanchet [Tue, 24 Aug 2010 16:43:52 +0200] rev 38697
don't backtick facts that contain schematic variables, since this doesn't work (for some reason)
2010-08-24 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet [Tue, 24 Aug 2010 16:24:11 +0200] rev 38696
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
2010-08-24 revert this idea of automatically invoking "metisFT" when "metis" fails;
blanchet [Tue, 24 Aug 2010 15:29:13 +0200] rev 38695
revert this idea of automatically invoking "metisFT" when "metis" fails; there are very few good reasons why "metisFT" should succeed when "metis" fails, and "metisFT" tends to "diverge" more often than "metis -- furthermore the exception handling code wasn't working properly
2010-08-24 use matching of types than just equality - this is needed in nominal to cope with type variables
Christian Urban <urbanc@in.tum.de> [Tue, 24 Aug 2010 22:38:45 +0800] rev 38694
use matching of types than just equality - this is needed in nominal to cope with type variables
2010-08-24 merge
blanchet [Tue, 24 Aug 2010 15:08:05 +0200] rev 38693
merge
2010-08-24 cosmetics
blanchet [Tue, 24 Aug 2010 15:07:54 +0200] rev 38692
cosmetics
2010-08-24 use a soft time limit for E
blanchet [Tue, 24 Aug 2010 15:06:47 +0200] rev 38691
use a soft time limit for E with a hard limit, "eproof"/"eproof.pl" doesn't show the proof is there's less than 1 second left, yielding "the ATP output is malformed"
2010-08-24 make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet [Tue, 24 Aug 2010 14:36:29 +0200] rev 38690
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
2010-08-23 cosmetics
blanchet [Mon, 23 Aug 2010 23:32:11 +0200] rev 38689
cosmetics
2010-08-23 optimize fact filter some more
blanchet [Mon, 23 Aug 2010 23:24:24 +0200] rev 38688
optimize fact filter some more
2010-08-23 cache previous iteration's weights, and keep track of what's dirty and what's clean;
blanchet [Mon, 23 Aug 2010 23:00:57 +0200] rev 38687
cache previous iteration's weights, and keep track of what's dirty and what's clean; this speeds up the relevance filter significantly
2010-08-23 modified relevance filter
blanchet [Mon, 23 Aug 2010 21:55:52 +0200] rev 38686
modified relevance filter
2010-08-23 adjust fudge factors in the light of experiments
blanchet [Mon, 23 Aug 2010 21:11:30 +0200] rev 38685
adjust fudge factors in the light of experiments
2010-08-23 invert semantics of "relevance_convergence", to make it more intuitive
blanchet [Mon, 23 Aug 2010 18:53:11 +0200] rev 38684
invert semantics of "relevance_convergence", to make it more intuitive
2010-08-23 if no facts were selected on first iteration, try again with a lower threshold
blanchet [Mon, 23 Aug 2010 18:39:12 +0200] rev 38683
if no facts were selected on first iteration, try again with a lower threshold
2010-08-23 weed out junk in relevance filter
blanchet [Mon, 23 Aug 2010 18:25:49 +0200] rev 38682
weed out junk in relevance filter
2010-08-23 no need for "eq" facts -- good old "=" is good enough for us
blanchet [Mon, 23 Aug 2010 18:04:31 +0200] rev 38681
no need for "eq" facts -- good old "=" is good enough for us
2010-08-23 revert unintended change
blanchet [Mon, 23 Aug 2010 17:53:49 +0200] rev 38680
revert unintended change
2010-08-23 destroy elim rules before checking for finite exhaustive facts
blanchet [Mon, 23 Aug 2010 17:49:18 +0200] rev 38679
destroy elim rules before checking for finite exhaustive facts
2010-08-23 prevent double inclusion of the fact "True_or_False" in TPTP problems
blanchet [Mon, 23 Aug 2010 16:12:41 +0200] rev 38678
prevent double inclusion of the fact "True_or_False" in TPTP problems
2010-08-24 merged
haftmann [Tue, 24 Aug 2010 10:07:14 +0200] rev 38677
merged
2010-08-24 tuned
haftmann [Tue, 24 Aug 2010 09:06:17 +0200] rev 38676
tuned
2010-08-23 merged
haftmann [Mon, 23 Aug 2010 11:57:32 +0200] rev 38675
merged
2010-08-23 dropped pre_post_conv
haftmann [Mon, 23 Aug 2010 11:57:22 +0200] rev 38674
dropped pre_post_conv
2010-08-23 use conv alias
haftmann [Mon, 23 Aug 2010 11:51:33 +0200] rev 38673
use conv alias
2010-08-23 preliminary versions of static_eval_conv(_simple)
haftmann [Mon, 23 Aug 2010 11:51:32 +0200] rev 38672
preliminary versions of static_eval_conv(_simple)
2010-08-23 use Code_Thingol.static_eval_conv_simple
haftmann [Mon, 23 Aug 2010 11:51:32 +0200] rev 38671
use Code_Thingol.static_eval_conv_simple
2010-08-23 added static_eval_conv
haftmann [Mon, 23 Aug 2010 11:51:32 +0200] rev 38670
added static_eval_conv
2010-08-23 refined and unified naming convention for dynamic code evaluation techniques
haftmann [Mon, 23 Aug 2010 11:09:49 +0200] rev 38669
refined and unified naming convention for dynamic code evaluation techniques
2010-08-23 tap_thy conversional
haftmann [Mon, 23 Aug 2010 11:09:48 +0200] rev 38668
tap_thy conversional
2010-08-23 dropped now obsolete purge_data -- happens implicitly on change of theory identity
haftmann [Mon, 23 Aug 2010 11:09:48 +0200] rev 38667
dropped now obsolete purge_data -- happens implicitly on change of theory identity
2010-08-24 merged
bulwahn [Tue, 24 Aug 2010 08:22:17 +0200] rev 38666
merged
2010-08-23 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn [Mon, 23 Aug 2010 16:47:57 +0200] rev 38665
introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
2010-08-23 added support for xsymbol syntax for mode annotations in code_pred command
bulwahn [Mon, 23 Aug 2010 16:47:55 +0200] rev 38664
added support for xsymbol syntax for mode annotations in code_pred command
2010-08-25 tuned raw Sidekick output;
wenzelm [Wed, 25 Aug 2010 11:13:27 +0200] rev 38663
tuned raw Sidekick output;
2010-08-24 Text.Range.is_singleton;
wenzelm [Tue, 24 Aug 2010 23:49:07 +0200] rev 38662
Text.Range.is_singleton;
2010-08-24 Markup_Tree.+: new info tends to sink to bottom, where it is prefered by select;
wenzelm [Tue, 24 Aug 2010 21:34:38 +0200] rev 38661
Markup_Tree.+: new info tends to sink to bottom, where it is prefered by select;
2010-08-24 tuned;
wenzelm [Tue, 24 Aug 2010 21:22:01 +0200] rev 38660
tuned;
2010-08-24 Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
wenzelm [Tue, 24 Aug 2010 21:20:08 +0200] rev 38659
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree; tuned;
2010-08-24 tuned root markup;
wenzelm [Tue, 24 Aug 2010 20:36:48 +0200] rev 38658
tuned root markup;
2010-08-23 misc tuning of important special cases;
wenzelm [Mon, 23 Aug 2010 20:50:00 +0200] rev 38657
misc tuning of important special cases;
2010-08-23 Rewrite the Probability theory.
hoelzl [Mon, 23 Aug 2010 19:35:57 +0200] rev 38656
Rewrite the Probability theory. Introduced pinfreal as real numbers with infinity. Use pinfreal as value for measures. Introduces Lebesgue Measure based on the integral in Multivariate Analysis. Proved Radon Nikodym for arbitrary sigma finite measure spaces.
(0) -30000 -10000 -3000 -1000 -120 +120 +1000 +3000 +10000 +30000 tip