blanchet [Tue, 23 Mar 2010 11:39:21 +0100] rev 35963
added options to Sledgehammer;
syntax: sledgehammer [options] goal_no, where "[options]" and "goal_no" are optional
blanchet [Mon, 22 Mar 2010 15:23:18 +0100] rev 35962
make "sledgehammer" and "atp_minimize" improper commands
wenzelm [Thu, 25 Mar 2010 21:27:04 +0100] rev 35961
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed (SUBTLE CHANGE IN SEMANTICS);
officially export weaken as Sorts.classrel_derivation;
tuned comments;
wenzelm [Thu, 25 Mar 2010 21:14:15 +0100] rev 35960
removed unused AxClass.of_sort derivation;
wenzelm [Wed, 24 Mar 2010 22:30:33 +0100] rev 35959
more precise dependencies;
wenzelm [Wed, 24 Mar 2010 22:08:03 +0100] rev 35958
merged
bulwahn [Wed, 24 Mar 2010 17:41:25 +0100] rev 35957
merged
bulwahn [Wed, 24 Mar 2010 17:40:44 +0100] rev 35956
removed predicate_compile_core.ML from HOL-ex session
bulwahn [Wed, 24 Mar 2010 17:40:44 +0100] rev 35955
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
bulwahn [Wed, 24 Mar 2010 17:40:44 +0100] rev 35954
adopting examples to Library move
bulwahn [Wed, 24 Mar 2010 17:40:43 +0100] rev 35953
moved further predicate compile files to HOL-Library
bulwahn [Wed, 24 Mar 2010 17:40:43 +0100] rev 35952
added simple setup for arithmetic on natural numbers
bulwahn [Wed, 24 Mar 2010 17:40:43 +0100] rev 35951
removing uncommented examples in setup theory of predicate compile quickcheck
bulwahn [Wed, 24 Mar 2010 17:40:43 +0100] rev 35950
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
wenzelm [Wed, 24 Mar 2010 21:55:30 +0100] rev 35949
slightly more logical bootstrap order -- also helps to sort out proof terms extension;
huffman [Wed, 24 Mar 2010 07:50:21 -0700] rev 35948
remove ancient continuity tactic
boehmes [Wed, 24 Mar 2010 15:21:42 +0100] rev 35947
removed Cache_IO component
boehmes [Wed, 24 Mar 2010 14:08:07 +0100] rev 35946
updated SMT certificates
boehmes [Wed, 24 Mar 2010 14:03:52 +0100] rev 35945
inhibit invokation of external SMT solver
boehmes [Wed, 24 Mar 2010 12:30:21 +0100] rev 35944
more precise dependencies
boehmes [Wed, 24 Mar 2010 09:44:47 +0100] rev 35943
cache_io is now just a single ML file instead of a component
boehmes [Wed, 24 Mar 2010 09:43:34 +0100] rev 35942
use internal SHA1 digest implementation for generating hash keys
boehmes [Wed, 24 Mar 2010 08:22:43 +0100] rev 35941
remove component Cache_IO (external dependency on MD5 will be replaced by internal SHA1 digest implementation)
huffman [Tue, 23 Mar 2010 19:03:05 -0700] rev 35940
merged
huffman [Tue, 23 Mar 2010 13:42:12 -0700] rev 35939
minimize dependencies
huffman [Tue, 23 Mar 2010 12:20:27 -0700] rev 35938
sublocale comm_monoid_add < setprod --> sublocale comm_monoid_mult < setprod
boehmes [Tue, 23 Mar 2010 22:43:53 +0100] rev 35937
use ml_platform instead of ml_system to distinguish library names
boehmes [Tue, 23 Mar 2010 20:46:47 +0100] rev 35936
merged
boehmes [Tue, 23 Mar 2010 20:46:08 +0100] rev 35935
use LONG rather than INT to represent the C datatype size_t
wenzelm [Tue, 23 Mar 2010 19:35:33 +0100] rev 35934
merged
huffman [Tue, 23 Mar 2010 10:07:39 -0700] rev 35933
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman [Tue, 23 Mar 2010 09:39:21 -0700] rev 35932
move letrec stuff to new file HOLCF/ex/Letrec.thy
wenzelm [Tue, 23 Mar 2010 19:35:03 +0100] rev 35931
more accurate dependencies;
wenzelm [Tue, 23 Mar 2010 17:26:41 +0100] rev 35930
even less ambitious isatest for smlnj;
hoelzl [Tue, 23 Mar 2010 16:18:44 +0100] rev 35929
Unhide measure_space.positive defined in Caratheodory.
hoelzl [Tue, 23 Mar 2010 16:17:41 +0100] rev 35928
Generate image for HOL-Probability
wenzelm [Tue, 23 Mar 2010 12:29:41 +0100] rev 35927
updated Thm.add_axiom/add_def;
huffman [Mon, 22 Mar 2010 23:34:23 -0700] rev 35926
completely remove constants cpair, cfst, csnd
huffman [Mon, 22 Mar 2010 23:33:23 -0700] rev 35925
use Pair instead of cpair in Fixrec_ex.thy
huffman [Mon, 22 Mar 2010 23:33:02 -0700] rev 35924
use Pair instead of cpair
huffman [Mon, 22 Mar 2010 23:02:43 -0700] rev 35923
define CLetrec using Pair, fst, snd instead of cpair, cfst, csnd
huffman [Mon, 22 Mar 2010 22:55:26 -0700] rev 35922
define csplit using fst, snd
huffman [Mon, 22 Mar 2010 22:43:21 -0700] rev 35921
convert lemma fix_cprod to use Pair, fst, snd
huffman [Mon, 22 Mar 2010 22:42:34 -0700] rev 35920
remove unused syntax for as_pat, lazy_pat
huffman [Mon, 22 Mar 2010 22:41:41 -0700] rev 35919
add lemmas fst_monofun, snd_monofun
huffman [Mon, 22 Mar 2010 21:37:48 -0700] rev 35918
use Pair instead of cpair
huffman [Mon, 22 Mar 2010 21:33:31 -0700] rev 35917
use fixrec_simp instead of fixpat
huffman [Mon, 22 Mar 2010 21:31:32 -0700] rev 35916
use Pair, fst, snd instead of cpair, cfst, csnd
huffman [Mon, 22 Mar 2010 21:11:54 -0700] rev 35915
remove admw predicate
huffman [Mon, 22 Mar 2010 20:54:52 -0700] rev 35914
remove contlub predicate
huffman [Mon, 22 Mar 2010 16:02:51 -0700] rev 35913
merged
huffman [Mon, 22 Mar 2010 15:53:25 -0700] rev 35912
error -> raise Fail
wenzelm [Mon, 22 Mar 2010 23:48:27 +0100] rev 35911
merged
wenzelm [Mon, 22 Mar 2010 23:20:55 +0100] rev 35910
merged
wenzelm [Mon, 22 Mar 2010 22:56:46 +0100] rev 35909
merged
huffman [Mon, 22 Mar 2010 15:45:54 -0700] rev 35908
remove unused adm_tac.ML
huffman [Mon, 22 Mar 2010 15:42:07 -0700] rev 35907
avoid dependence on adm_tac solver
huffman [Mon, 22 Mar 2010 15:23:16 -0700] rev 35906
remove obsolete holcf_logic.ML
huffman [Mon, 22 Mar 2010 15:05:20 -0700] rev 35905
fix ML warning in domain_library.ML
huffman [Mon, 22 Mar 2010 14:58:21 -0700] rev 35904
fix ML warnings in repdef.ML
huffman [Mon, 22 Mar 2010 14:44:37 -0700] rev 35903
fix ML warnings in fixrec.ML
huffman [Mon, 22 Mar 2010 14:11:13 -0700] rev 35902
fix ML warnings in pcpodef.ML
huffman [Mon, 22 Mar 2010 13:27:35 -0700] rev 35901
fix LaTeX overfull hbox warnings in HOLCF document
huffman [Mon, 22 Mar 2010 12:52:51 -0700] rev 35900
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
wenzelm [Mon, 22 Mar 2010 20:59:22 +0100] rev 35899
explicit Simplifier.global_context;
wenzelm [Mon, 22 Mar 2010 20:58:52 +0100] rev 35898
recovered header;
wenzelm [Mon, 22 Mar 2010 19:29:11 +0100] rev 35897
eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
wenzelm [Mon, 22 Mar 2010 19:26:12 +0100] rev 35896
inlined version of old-style Drule.add_axioms -- Specification.axiom is not bootstrapped yet;
wenzelm [Mon, 22 Mar 2010 19:25:14 +0100] rev 35895
use Specification.axiom, together with Drule.export_without_context that was implicit in the former PureThy.add_axioms (cf. f81557a124d5);
wenzelm [Mon, 22 Mar 2010 19:23:03 +0100] rev 35894
added Specification.axiom convenience;
blanchet [Mon, 22 Mar 2010 15:07:07 +0100] rev 35893
detect OFCLASS() axioms in Nitpick;
this is necessary now that "Thy.all_axioms_of" returns such axiom (starting with change 46cfc4b8112e)
bulwahn [Mon, 22 Mar 2010 13:48:15 +0100] rev 35892
merged
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35891
contextifying the compilation of the predicate compiler
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35890
removed unused Predicate_Compile_Set
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35889
avoiding fishing for split_asm rule in the predicate compiler
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35888
contextifying the proof procedure in the predicate compiler
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35887
making flat triples to nested tuple to remove general triple functions
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35886
reduced the debug output functions from 2 to 1
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35885
some improvements thanks to Makarius source code review
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35884
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35883
enabling a previously broken example of the predicate compiler again
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35882
improving handling of case expressions in predicate rewriting
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35881
adding depth_limited_random compilation to predicate compiler
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35880
a new simpler random compilation for the predicate compiler
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35879
reviving the classical depth-limited computation in the predicate compiler
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35878
cleaning the function flattening
bulwahn [Mon, 22 Mar 2010 08:30:13 +0100] rev 35877
generalized split transformation in the function flattening
bulwahn [Mon, 22 Mar 2010 08:30:12 +0100] rev 35876
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn [Mon, 22 Mar 2010 08:30:12 +0100] rev 35875
restructuring function flattening
bulwahn [Mon, 22 Mar 2010 08:30:12 +0100] rev 35874
renaming mk_prems to flatten in the function flattening
bulwahn [Mon, 22 Mar 2010 08:30:12 +0100] rev 35873
simplifying function flattening
boehmes [Mon, 22 Mar 2010 11:45:09 +0100] rev 35872
removed warning_count (known causes for warnings have been resolved)
blanchet [Mon, 22 Mar 2010 10:38:28 +0100] rev 35871
remove the iteration counter from Sledgehammer's minimizer
blanchet [Mon, 22 Mar 2010 10:25:44 +0100] rev 35870
merged
blanchet [Mon, 22 Mar 2010 10:25:07 +0100] rev 35869
start work on direct proof reconstruction for Sledgehammer
blanchet [Fri, 19 Mar 2010 16:04:15 +0100] rev 35868
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
"full" sounds like "full types" or something, not like a structured Isar proof -- at some point I hope to make this an option that's orthogonal to the prover
blanchet [Fri, 19 Mar 2010 15:33:18 +0100] rev 35867
move all ATP setup code into ATP_Wrapper
blanchet [Fri, 19 Mar 2010 15:07:44 +0100] rev 35866
move the Sledgehammer Isar commands together into one file;
this will make easier to add options and reorganize them later
blanchet [Fri, 19 Mar 2010 13:02:18 +0100] rev 35865
more Sledgehammer refactoring
boehmes [Mon, 22 Mar 2010 09:54:22 +0100] rev 35864
use a proof context instead of a local theory
boehmes [Mon, 22 Mar 2010 09:46:04 +0100] rev 35863
provide a hook to safely manipulate verification conditions
boehmes [Mon, 22 Mar 2010 09:40:11 +0100] rev 35862
replaced old-style Drule.add_axiom by Specification.axiomatization
boehmes [Mon, 22 Mar 2010 09:39:10 +0100] rev 35861
removed e-mail address from error message
haftmann [Mon, 22 Mar 2010 09:32:28 +0100] rev 35860
merged
haftmann [Sun, 21 Mar 2010 08:46:50 +0100] rev 35859
tuned whitespace
haftmann [Sun, 21 Mar 2010 08:46:49 +0100] rev 35858
handle hidden polymorphism in class target (without class target syntax, though)
wenzelm [Mon, 22 Mar 2010 00:51:18 +0100] rev 35857
replaced Theory.add_axioms(_i) by more primitive Theory.add_axiom;
wenzelm [Mon, 22 Mar 2010 00:48:56 +0100] rev 35856
replaced PureThy.add_axioms by more basic Drule.add_axiom, which is old-style nonetheless;
wenzelm [Sun, 21 Mar 2010 22:24:04 +0100] rev 35855
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
more uniform add_axiom/add_def;
wenzelm [Sun, 21 Mar 2010 22:13:31 +0100] rev 35854
Logic.mk_of_sort convenience;
wenzelm [Sun, 21 Mar 2010 19:30:19 +0100] rev 35853
more explicit invented name;
wenzelm [Sun, 21 Mar 2010 19:28:25 +0100] rev 35852
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
wenzelm [Sun, 21 Mar 2010 19:04:46 +0100] rev 35851
do not open ML structures;
wenzelm [Sun, 21 Mar 2010 17:28:35 +0100] rev 35850
modernized overloaded definitions;
wenzelm [Sun, 21 Mar 2010 17:12:31 +0100] rev 35849
standard headers;
wenzelm [Sun, 21 Mar 2010 16:51:37 +0100] rev 35848
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm [Sun, 21 Mar 2010 15:57:40 +0100] rev 35847
eliminated old constdefs;
haftmann [Sun, 21 Mar 2010 06:59:23 +0100] rev 35846
corrected setup for of_list
wenzelm [Sat, 20 Mar 2010 17:33:11 +0100] rev 35845
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
Christian Urban <urbanc@in.tum.de> [Sat, 20 Mar 2010 02:23:41 +0100] rev 35844
added lemma infinite_Un