bulwahn [Thu, 21 Jul 2011 08:31:35 +0200] rev 43937
adapting two examples in Predicate_Compile_Examples
blanchet [Wed, 20 Jul 2011 23:47:27 +0200] rev 43936
use a more robust naming convention for "polymorphic" frees -- the check is an overapproximation but that's fine as far as soundness is concerned
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 20 Jul 2011 16:15:33 +0200] rev 43935
merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 20 Jul 2011 16:14:49 +0200] rev 43934
Quotient Package: handle Bound variables in rep_abs_rsp_tac not only at top-level of the goal
hoelzl [Wed, 20 Jul 2011 15:42:23 +0200] rev 43933
add code generator setup and tests for ereal
boehmes [Wed, 20 Jul 2011 15:09:53 +0200] rev 43932
removed debugging facilities accidentally left in the committed code
boehmes [Wed, 20 Jul 2011 13:29:54 +0200] rev 43931
more precise dependencies
boehmes [Wed, 20 Jul 2011 13:27:01 +0200] rev 43930
merged
boehmes [Wed, 20 Jul 2011 12:23:20 +0200] rev 43929
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
boehmes [Wed, 20 Jul 2011 09:23:12 +0200] rev 43928
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
boehmes [Wed, 20 Jul 2011 09:23:09 +0200] rev 43927
removed old (unused) SMT monomorphizer
krauss [Wed, 20 Jul 2011 13:24:49 +0200] rev 43926
added UNION
hoelzl [Wed, 20 Jul 2011 10:48:00 +0200] rev 43925
merged
hoelzl [Tue, 19 Jul 2011 14:38:48 +0200] rev 43924
rename Fin to enat
hoelzl [Tue, 19 Jul 2011 14:38:29 +0200] rev 43923
add ereal to typeclass infinity
hoelzl [Tue, 19 Jul 2011 14:37:49 +0200] rev 43922
add nat => enat coercion
hoelzl [Tue, 19 Jul 2011 14:37:09 +0200] rev 43921
Introduce infinity type class
hoelzl [Tue, 19 Jul 2011 14:36:12 +0200] rev 43920
Rename extreal => ereal
hoelzl [Tue, 19 Jul 2011 14:35:44 +0200] rev 43919
rename Nat_Infinity (inat) to Extended_Nat (enat)
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 20 Jul 2011 10:11:08 +0200] rev 43918
HOL/Import reorganization/cleaning. Factor 9 speedup. Remove Import XML parser in favor of much faster of Isabelle's XML parser. Remove ImportRecording since we can use Isabelle images.
kleing [Wed, 20 Jul 2011 08:46:17 +0200] rev 43917
build an image for HOL-IMP
bulwahn [Wed, 20 Jul 2011 08:16:42 +0200] rev 43916
deactivating quickcheck invocation in this example until the Interrupt issue is understood
bulwahn [Wed, 20 Jul 2011 08:16:41 +0200] rev 43915
removing inner time limits in quickcheck
bulwahn [Wed, 20 Jul 2011 08:16:39 +0200] rev 43914
updating documentation about quickcheck; adding information about try
bulwahn [Wed, 20 Jul 2011 08:16:38 +0200] rev 43913
adapting example in Predicate_Compile_Examples
bulwahn [Wed, 20 Jul 2011 08:16:36 +0200] rev 43912
exporting function in quickcheck; adapting mutabelle script
bulwahn [Wed, 20 Jul 2011 08:16:35 +0200] rev 43911
more information for the user how to deactivate quickcheck_narrowing if he does not want to use it
bulwahn [Wed, 20 Jul 2011 08:16:33 +0200] rev 43910
making messages more informative
bulwahn [Wed, 20 Jul 2011 08:16:32 +0200] rev 43909
only use exhaustive testing in this quickcheck example
blanchet [Wed, 20 Jul 2011 00:37:42 +0200] rev 43908
parse equalities correctly in Nitrox parser
blanchet [Wed, 20 Jul 2011 00:37:42 +0200] rev 43907
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet [Wed, 20 Jul 2011 00:37:42 +0200] rev 43906
generate slightly less type information -- this should be sound since type arguments should keep things cleanly apart
blanchet [Wed, 20 Jul 2011 00:37:42 +0200] rev 43905
avoid calling "Term.is_first_order" (indirectly) on a term with loose de Bruijns -- this is not necessary anyway because of the Abs check in "simple_translate_lambdas"
blanchet [Wed, 20 Jul 2011 00:37:42 +0200] rev 43904
remove offset from Mirabelle output
krauss [Tue, 19 Jul 2011 11:15:38 +0200] rev 43903
the HOL4PROOFS setting is actually HOL4_PROOFS
haftmann [Tue, 19 Jul 2011 07:14:14 +0200] rev 43902
merged
haftmann [Mon, 18 Jul 2011 21:52:34 +0200] rev 43901
proof tuning
haftmann [Mon, 18 Jul 2011 21:49:39 +0200] rev 43900
generalization; various notation and proof tuning
haftmann [Mon, 18 Jul 2011 21:34:01 +0200] rev 43899
avoid misunderstandable names
haftmann [Mon, 18 Jul 2011 21:15:51 +0200] rev 43898
moved lemmas to appropriate theory
krauss [Tue, 19 Jul 2011 00:16:18 +0200] rev 43897
forgotten qualifier
krauss [Tue, 19 Jul 2011 00:07:21 +0200] rev 43896
values_timeout defaults to 600.0 on SML/NJ -- saves us from cluttering all theories equivalent declarations
krauss [Mon, 18 Jul 2011 23:48:28 +0200] rev 43895
killed use of PolyML.makestring
krauss [Mon, 18 Jul 2011 23:35:50 +0200] rev 43894
added experimental mira configuration for HOL Light importer
boehmes [Mon, 18 Jul 2011 18:52:52 +0200] rev 43893
allow rules with premises to be declared as z3_rule (to circumvent incompleteness of Z3 proof reconstruction)
bulwahn [Mon, 18 Jul 2011 13:49:26 +0200] rev 43892
unactivating narrowing-based quickcheck by default
bulwahn [Mon, 18 Jul 2011 13:48:35 +0200] rev 43891
making active configuration public in narrowing-based quickcheck
bulwahn [Mon, 18 Jul 2011 11:38:14 +0200] rev 43890
declare tester in this quickcheck example
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43889
adding code equations for partial_term_of for rational numbers
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43888
adapting an experimental setup to changes in quickcheck's infrastructure
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43887
adding narrowing instances for real and rational
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43886
adapting quickcheck based on the analysis of the predicate compiler
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43885
adapting prolog-based tester
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43884
quickcheck does not deactivate testers if none are given
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43883
adapting mutabelle to latest changes in quickcheck; removing unused code in mutabelle
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43882
renaming quickcheck_tester to quickcheck_batch_tester; tuned
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43881
changing parser in quickcheck to activate and deactivate the testers
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43880
adapting SML_Quickcheck to new quickcheck infrastructure
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43879
enabling parallel execution of testers but removing more informative quickcheck output
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43878
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43877
removing generator registration
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43876
parametrized test_term functions in quickcheck
bulwahn [Mon, 18 Jul 2011 10:34:21 +0200] rev 43875
adding random, exhaustive and SML quickcheck as testers
haftmann [Sun, 17 Jul 2011 22:25:14 +0200] rev 43874
more on complement
haftmann [Sun, 17 Jul 2011 22:24:08 +0200] rev 43873
more on complement
haftmann [Sun, 17 Jul 2011 20:57:56 +0200] rev 43872
more consistent theorem names
haftmann [Sun, 17 Jul 2011 20:46:51 +0200] rev 43871
more lemmas about SUP
haftmann [Sun, 17 Jul 2011 20:29:54 +0200] rev 43870
structuring duals together
haftmann [Sun, 17 Jul 2011 20:23:39 +0200] rev 43869
merged
haftmann [Sun, 17 Jul 2011 20:23:33 +0200] rev 43868
more lemmas about Sup
haftmann [Sun, 17 Jul 2011 19:55:17 +0200] rev 43867
generalized INT_anti_mono
haftmann [Sun, 17 Jul 2011 19:48:02 +0200] rev 43866
moving UNIV = ... equations to their proper theories
haftmann [Sun, 17 Jul 2011 15:15:58 +0200] rev 43865
further generalization from sets to complete lattices
blanchet [Sun, 17 Jul 2011 14:21:19 +0200] rev 43864
fixed lambda-liftg: must ensure the formulas are in close form
blanchet [Sun, 17 Jul 2011 14:12:45 +0200] rev 43863
ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet [Sun, 17 Jul 2011 14:11:35 +0200] rev 43862
pass kind to lambda-translation function
blanchet [Sun, 17 Jul 2011 14:11:35 +0200] rev 43861
more refactoring of preprocessing
blanchet [Sun, 17 Jul 2011 14:11:35 +0200] rev 43860
more refactoring of preprocessing, so as to be able to centralize it
blanchet [Sun, 17 Jul 2011 14:11:35 +0200] rev 43859
renamed internal data structure
blanchet [Sun, 17 Jul 2011 14:11:35 +0200] rev 43858
simplify code -- there are no lambdas in helpers anyway
blanchet [Sun, 17 Jul 2011 14:11:35 +0200] rev 43857
added lambda-lifting to Sledgehammer (rough)
blanchet [Sun, 17 Jul 2011 14:11:34 +0200] rev 43856
move more lambda-handling logic to Sledgehammer, from ATP module, for formal dependency reasons
haftmann [Sun, 17 Jul 2011 08:45:06 +0200] rev 43855
merged
haftmann [Sat, 16 Jul 2011 22:28:35 +0200] rev 43854
generalized some lemmas
haftmann [Sat, 16 Jul 2011 22:04:02 +0200] rev 43853
consolidated bot and top classes, tuned notation
haftmann [Sat, 16 Jul 2011 21:53:50 +0200] rev 43852
tuned notation
wenzelm [Sat, 16 Jul 2011 22:17:27 +0200] rev 43851
clarified bash_output_fifo;
wenzelm [Sat, 16 Jul 2011 20:52:41 +0200] rev 43850
moved bash operations to Isabelle_System (cf. Scala version);
wenzelm [Sat, 16 Jul 2011 20:14:58 +0200] rev 43849
access to process output stream via auxiliary fifo;
wenzelm [Sat, 16 Jul 2011 18:41:35 +0200] rev 43848
some file and directory operations;
wenzelm [Sat, 16 Jul 2011 18:20:02 +0200] rev 43847
more general bash_process, which allows to terminate background processes as well;
wenzelm [Sat, 16 Jul 2011 18:11:14 +0200] rev 43846
updated to Poly/ML SVN 1328, which is considered 5.4.2;
wenzelm [Sat, 16 Jul 2011 17:11:49 +0200] rev 43845
added File.fold_pages for streaming of large files;
prefer \f notation;
wenzelm [Sat, 16 Jul 2011 16:51:12 +0200] rev 43844
tuned;
Cezary Kaliszyk <kaliszyk@in.tum.de> [Sat, 16 Jul 2011 00:01:17 +0200] rev 43843
HOL/Import: Fix errors with _mk_list
wenzelm [Fri, 15 Jul 2011 16:51:01 +0200] rev 43842
Element.activate: leave check of binding where actually applied to the context -- allow internal qualifications, or non-identifier fact names like "assumes *: A" (see also 1183951365de);
wenzelm [Fri, 15 Jul 2011 14:07:12 +0200] rev 43841
simplified malformed YXML markup -- special controls are visible in IsabelleText font;
wenzelm [Fri, 15 Jul 2011 13:29:00 +0200] rev 43840
less ambitious ProofGeneral markup, which occasionally breaks plain-old regexps in elisp;
wenzelm [Fri, 15 Jul 2011 13:28:16 +0200] rev 43839
more robust Binding.pretty/print in typical error sitations with spaces etc. (NB: markup can only provide *additional* emphasis and is occasionally suppressed in TTY mode or tooltips);
wenzelm [Fri, 15 Jul 2011 00:49:38 +0200] rev 43838
more visible printing of empty binding;
wenzelm [Fri, 15 Jul 2011 00:03:47 +0200] rev 43837
do not check vacous bindings, which routinely occur in locale expressions and long theorem statements etc.;
wenzelm [Thu, 14 Jul 2011 23:05:25 +0200] rev 43836
more quotes;
wenzelm [Thu, 14 Jul 2011 22:53:43 +0200] rev 43835
merged
krauss [Thu, 14 Jul 2011 22:08:11 +0200] rev 43834
added missing dependencies;
simplified (by overapproximating a bit);
alphabetical order
haftmann [Thu, 14 Jul 2011 19:43:45 +0200] rev 43833
merged
haftmann [Thu, 14 Jul 2011 17:15:24 +0200] rev 43832
merged
haftmann [Thu, 14 Jul 2011 17:14:54 +0200] rev 43831
tuned notation and proofs
blanchet [Thu, 14 Jul 2011 17:29:30 +0200] rev 43830
move error logic closer to user
blanchet [Thu, 14 Jul 2011 17:29:30 +0200] rev 43829
allow lambda-lifting without triggers
blanchet [Thu, 14 Jul 2011 16:50:05 +0200] rev 43828
move lambda translation option from ATP to Sledgehammer, to avoid accidentally breaking Metis (its reconstruction code can only deal with combinators)
blanchet [Thu, 14 Jul 2011 16:50:05 +0200] rev 43827
added option to control which lambda translation to use (for experiments)
blanchet [Thu, 14 Jul 2011 15:14:38 +0200] rev 43826
fix subtle type inference bug in Metis -- different occurrences of the same Skolem might need to be typed differently, using paramify_vars overconstraints the typing problem
blanchet [Thu, 14 Jul 2011 15:14:38 +0200] rev 43825
use monomorphic encoding as fallback, since they tend to produce fewer type errors
blanchet [Thu, 14 Jul 2011 15:14:38 +0200] rev 43824
don't generate Waldmeister problems with only a conjecture, since it makes it crash sometimes
blanchet [Thu, 14 Jul 2011 15:14:37 +0200] rev 43823
clearer unsound message
blanchet [Thu, 14 Jul 2011 15:14:37 +0200] rev 43822
clarify fine soundness point
blanchet [Thu, 14 Jul 2011 15:14:37 +0200] rev 43821
always unfold "Let"s is Sledgehammer, Metis, and MESON
blanchet [Thu, 14 Jul 2011 15:14:37 +0200] rev 43820
unbreak Nitrox's parsing
haftmann [Thu, 14 Jul 2011 00:21:56 +0200] rev 43819
merged
haftmann [Thu, 14 Jul 2011 00:20:43 +0200] rev 43818
tuned lemma positions and proofs
haftmann [Thu, 14 Jul 2011 00:16:41 +0200] rev 43817
tuned notation
haftmann [Wed, 13 Jul 2011 23:49:56 +0200] rev 43816
uniqueness lemmas for bot and top
haftmann [Wed, 13 Jul 2011 23:41:13 +0200] rev 43815
adjusted to tightened specification of classes bot and top
haftmann [Wed, 13 Jul 2011 19:43:12 +0200] rev 43814
moved lemmas bot_less and less_top to classes bot and top respectively
haftmann [Wed, 13 Jul 2011 19:40:18 +0200] rev 43813
tightened specification of classes bot and top: uniqueness of top and bot elements
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43812
honor the TPTP environment variable as the root of include relative paths -- that's a weird convention but without it Nitrox will fail at CASC
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43811
better temp name creation for Nitrox -- still very hackish though, but should get us through CASC-23 and CASC-J6
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43810
more exhaustive testing in Nitrox
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43809
no timeout for Nitrox
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43808
avoid relying on piping to "isabelle tty", because this gives errors on some Linuxes about the standard input not being a tty
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43807
added arithmetic decision procedure to CASC setup
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43806
added some arithmetic functions, for THF with arithmetic
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43805
pull in arithmetic theories
blanchet [Wed, 13 Jul 2011 22:16:19 +0200] rev 43804
cleanly separate TPTP related files from other examples
bulwahn [Wed, 13 Jul 2011 21:59:54 +0200] rev 43803
increasing timeout to avoid spurious failures
haftmann [Wed, 13 Jul 2011 18:36:11 +0200] rev 43802
merged
haftmann [Wed, 13 Jul 2011 07:26:31 +0200] rev 43801
more generalization towards complete lattices
krauss [Wed, 13 Jul 2011 15:50:45 +0200] rev 43800
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 13 Jul 2011 04:00:32 +0900] rev 43799
merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 13 Jul 2011 11:31:36 +0900] rev 43798
Tuned
wenzelm [Thu, 14 Jul 2011 22:30:31 +0200] rev 43797
more precise integer Markup.properties/XML.attributes: disallow ML-style ~ minus;
wenzelm [Wed, 13 Jul 2011 22:05:55 +0200] rev 43796
added term_sharing.ML;
wenzelm [Wed, 13 Jul 2011 21:44:15 +0200] rev 43795
recovered some runtime sharing from d6b6c74a8bcf, without the global memory bottleneck;
wenzelm [Wed, 13 Jul 2011 20:36:18 +0200] rev 43794
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
minor tuning;
wenzelm [Wed, 13 Jul 2011 20:13:27 +0200] rev 43793
low-level tuning;
wenzelm [Wed, 13 Jul 2011 16:42:14 +0200] rev 43792
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
Sorts.certify_class: prefer the persistent name;
wenzelm [Wed, 13 Jul 2011 10:57:09 +0200] rev 43791
XML.pretty with depth limit;
wenzelm [Tue, 12 Jul 2011 23:22:22 +0200] rev 43790
more thorough Variable.check_name: Binding.check for logical entities within the term language;
wenzelm [Tue, 12 Jul 2011 23:20:34 +0200] rev 43789
tuned;
wenzelm [Tue, 12 Jul 2011 20:53:14 +0200] rev 43788
merged
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 13 Jul 2011 00:43:07 +0900] rev 43787
Update HOLLightCompat
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 13 Jul 2011 00:29:33 +0900] rev 43786
Update files generated in HOL/Import/HOLLight
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 13 Jul 2011 00:23:24 +0900] rev 43785
HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
wenzelm [Tue, 12 Jul 2011 20:11:11 +0200] rev 43784
ML pp for XML.tree;
wenzelm [Tue, 12 Jul 2011 20:11:00 +0200] rev 43783
made SML/NJ happy;
wenzelm [Tue, 12 Jul 2011 19:49:35 +0200] rev 43782
clarified YXML.detect;
wenzelm [Tue, 12 Jul 2011 19:47:40 +0200] rev 43781
retain some terminology of "XML attributes";
wenzelm [Tue, 12 Jul 2011 19:36:46 +0200] rev 43780
more uniform Properties in ML and Scala;
wenzelm [Tue, 12 Jul 2011 18:00:05 +0200] rev 43779
more uniform Term and Term_XML modules;
wenzelm [Tue, 12 Jul 2011 17:53:06 +0200] rev 43778
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm [Tue, 12 Jul 2011 15:32:16 +0200] rev 43777
tuned signature -- less cryptic ASCII names;
wenzelm [Tue, 12 Jul 2011 15:17:37 +0200] rev 43776
discontinued obsolete Isabelle_Syntax and Parse_Value -- superseded by Outer_Syntax.quote_string and XML.Encode, Term_XML.Encode etc.;
wenzelm [Tue, 12 Jul 2011 15:12:50 +0200] rev 43775
added Parse.properties (again) -- allow empty list like Parse_Value.properties but unlike Parse.properties of ef86de9c98aa;
wenzelm [Tue, 12 Jul 2011 14:54:29 +0200] rev 43774
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
wenzelm [Tue, 12 Jul 2011 14:33:08 +0200] rev 43773
more precise Symbol_Pos.quote_string;
wenzelm [Tue, 12 Jul 2011 13:45:05 +0200] rev 43772
clarified YXML.embed_controls -- this is idempotent and cannot be nested;
wenzelm [Tue, 12 Jul 2011 13:39:29 +0200] rev 43771
allow empty body for raw_message -- important for Invoke_Scala;
wenzelm [Tue, 12 Jul 2011 11:45:13 +0200] rev 43770
Isabelle string syntax allows literal control characters;
wenzelm [Tue, 12 Jul 2011 11:19:42 +0200] rev 43769
glyphs from DejaVu for ASCII control characters 5, 6, 7, 127, which have a special meaning in Isabelle or Poly/ML;
wenzelm [Tue, 12 Jul 2011 11:16:56 +0200] rev 43768
more precise exceptions;
wenzelm [Tue, 12 Jul 2011 10:44:30 +0200] rev 43767
tuned XML modules;
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 12 Jul 2011 16:00:05 +0900] rev 43766
Quotient example: Lists with distinct elements
wenzelm [Mon, 11 Jul 2011 23:20:40 +0200] rev 43765
merged
haftmann [Mon, 11 Jul 2011 18:44:58 +0200] rev 43764
explicit code equation for equality
wenzelm [Mon, 11 Jul 2011 23:15:27 +0200] rev 43763
tuned error messages;
wenzelm [Mon, 11 Jul 2011 23:15:04 +0200] rev 43762
tuned;
wenzelm [Mon, 11 Jul 2011 22:55:47 +0200] rev 43761
tuned signature -- corresponding to Scala version;
wenzelm [Mon, 11 Jul 2011 22:50:29 +0200] rev 43760
made SML/NJ happy;
tuned error;
wenzelm [Mon, 11 Jul 2011 22:19:11 +0200] rev 43759
more uniform padded_markup, which is important for caret visibility despite absence of markup;
wenzelm [Mon, 11 Jul 2011 17:22:31 +0200] rev 43758
merged
haftmann [Mon, 11 Jul 2011 07:04:30 +0200] rev 43757
merged
haftmann [Sun, 10 Jul 2011 22:42:53 +0200] rev 43756
tuned proofs
haftmann [Sun, 10 Jul 2011 22:17:33 +0200] rev 43755
tuned notation
haftmann [Sun, 10 Jul 2011 22:11:32 +0200] rev 43754
tuned notation
haftmann [Sun, 10 Jul 2011 21:56:39 +0200] rev 43753
tuned notation
wenzelm [Mon, 11 Jul 2011 17:22:15 +0200] rev 43752
NEWS;
wenzelm [Mon, 11 Jul 2011 17:14:30 +0200] rev 43751
proper InvocationTargetException.getCause for indirect exceptions;
capture hard errors to ensure protocol integrity;
tuned error messages;
wenzelm [Mon, 11 Jul 2011 17:11:54 +0200] rev 43750
tuned error message;
wenzelm [Mon, 11 Jul 2011 17:10:32 +0200] rev 43749
tuned signature;
wenzelm [Mon, 11 Jul 2011 16:48:02 +0200] rev 43748
JVM method invocation service via Scala layer;
wenzelm [Mon, 11 Jul 2011 15:56:30 +0200] rev 43747
tuned signature;
wenzelm [Mon, 11 Jul 2011 11:13:33 +0200] rev 43746
some support for raw messages, which bypass standard Symbol/YXML decoding;
tuned signature;
wenzelm [Mon, 11 Jul 2011 10:27:50 +0200] rev 43745
tuned XML.Cache parameters;
wenzelm [Sun, 10 Jul 2011 23:46:05 +0200] rev 43744
some support to invoke Scala methods under program control;
wenzelm [Sun, 10 Jul 2011 21:46:41 +0200] rev 43743
merged;
haftmann [Sun, 10 Jul 2011 21:39:03 +0200] rev 43742
merged
haftmann [Sun, 10 Jul 2011 15:45:35 +0200] rev 43741
tuned proofs and notation
haftmann [Sun, 10 Jul 2011 14:26:07 +0200] rev 43740
more succinct proofs
haftmann [Sun, 10 Jul 2011 14:14:19 +0200] rev 43739
more succinct proofs
bulwahn [Sun, 10 Jul 2011 19:33:27 +0200] rev 43738
adding a very liberal timeout for values after a test case failed due to the restricted timeout
bulwahn [Sun, 10 Jul 2011 14:02:27 +0200] rev 43737
improved NEWS
bulwahn [Sat, 09 Jul 2011 21:18:20 +0200] rev 43736
NEWS
bulwahn [Sat, 09 Jul 2011 21:09:09 +0200] rev 43735
standardized String.concat towards implode (cf. c37a1f29bbc0)
bulwahn [Sat, 09 Jul 2011 19:29:25 +0200] rev 43734
adding quickcheck examples for evaluating floor and ceiling functions
bulwahn [Sat, 09 Jul 2011 19:28:33 +0200] rev 43733
adding code equations to execute floor and ceiling on rational and real numbers
bulwahn [Sat, 09 Jul 2011 13:41:58 +0200] rev 43732
adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
wenzelm [Sun, 10 Jul 2011 20:59:04 +0200] rev 43731
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
tuned signature;
wenzelm [Sun, 10 Jul 2011 17:58:11 +0200] rev 43730
lambda terms with XML data representation in Scala;
avoid `class` in signature;
wenzelm [Sun, 10 Jul 2011 16:34:17 +0200] rev 43729
XML data representation of lambda terms;
wenzelm [Sun, 10 Jul 2011 16:31:04 +0200] rev 43728
YXML.string_of_body convenience;
wenzelm [Sun, 10 Jul 2011 16:13:37 +0200] rev 43727
made SML/NJ happy;
wenzelm [Sun, 10 Jul 2011 16:09:08 +0200] rev 43726
tuned signature;
wenzelm [Sun, 10 Jul 2011 15:48:15 +0200] rev 43725
more abstract signature;
tuned;
wenzelm [Sun, 10 Jul 2011 13:51:21 +0200] rev 43724
simplified XML_Data;
wenzelm [Sun, 10 Jul 2011 13:00:22 +0200] rev 43723
less currying in Scala;
wenzelm [Sun, 10 Jul 2011 00:21:19 +0200] rev 43722
propagate header changes to prover process;
simplified Document case classes;
Document.State.assignments: indexed by Version_ID;
wenzelm [Sat, 09 Jul 2011 21:53:27 +0200] rev 43721
echo prover input via raw_messages, for improved protocol tracing;
wenzelm [Sat, 09 Jul 2011 18:54:50 +0200] rev 43720
tuned;
wenzelm [Sat, 09 Jul 2011 18:35:00 +0200] rev 43719
tuned signature;
wenzelm [Sat, 09 Jul 2011 18:15:23 +0200] rev 43718
clarified propagation of node name and header;
wenzelm [Sat, 09 Jul 2011 17:14:08 +0200] rev 43717
more precise treatment of prover definedness;
wenzelm [Sat, 09 Jul 2011 16:53:19 +0200] rev 43716
tuned source structure;
wenzelm [Sat, 09 Jul 2011 13:29:33 +0200] rev 43715
some support for blobs (arbitrary text files) within document nodes;
wenzelm [Sat, 09 Jul 2011 12:56:51 +0200] rev 43714
tuned signature;
wenzelm [Fri, 08 Jul 2011 22:00:53 +0200] rev 43713
moved global state to structure Document (again);
wenzelm [Fri, 08 Jul 2011 21:44:47 +0200] rev 43712
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
tuned signatures;
tuned module dependencies;
wenzelm [Fri, 08 Jul 2011 20:27:09 +0200] rev 43711
less stateful outer_syntax;
wenzelm [Fri, 08 Jul 2011 17:04:38 +0200] rev 43710
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
Position.offset discriminates postions precisely, now also available for Position.line/line_file;
wenzelm [Fri, 08 Jul 2011 16:13:34 +0200] rev 43709
discontinued special treatment of hard tabulators;
wenzelm [Fri, 08 Jul 2011 16:01:14 +0200] rev 43708
eliminated hard tabs;
wenzelm [Fri, 08 Jul 2011 15:18:28 +0200] rev 43707
merged
nipkow [Fri, 08 Jul 2011 12:18:46 +0200] rev 43706
merged
nipkow [Thu, 07 Jul 2011 21:53:53 +0200] rev 43705
added translation to fix critical pair between abbreviations for surj and ~=
bulwahn [Thu, 07 Jul 2011 23:33:14 +0200] rev 43704
floor and ceiling definitions are not code equations -- this enables trivial evaluation of floor and ceiling
wenzelm [Fri, 08 Jul 2011 15:17:40 +0200] rev 43703
standardized String.concat towards implode;
wenzelm [Fri, 08 Jul 2011 14:37:19 +0200] rev 43702
more abstract Thy_Load.load_file/use_file for external theory resources;
prefer Boogie_Loader.parse_b2i on already loaded text, bypassing former File.fold_lines optimization;
wenzelm [Fri, 08 Jul 2011 13:59:54 +0200] rev 43701
comment;
wenzelm [Fri, 08 Jul 2011 11:50:58 +0200] rev 43700
clarified Thy_Load.digest_file -- read ML files only once;
wenzelm [Fri, 08 Jul 2011 11:13:21 +0200] rev 43699
tuned signature;
wenzelm [Thu, 07 Jul 2011 23:55:15 +0200] rev 43698
simplified make_option/dest_option;
added make_variant/dest_variant -- usual representation of datatypes;