huffman [Fri, 22 Oct 2010 11:24:52 -0700] rev 40092
do proofs using Rep_Sprod_simps, Rep_Ssum_simps; remove unused lemmas
huffman [Fri, 22 Oct 2010 07:45:32 -0700] rev 40091
make discrete_cpo a subclass of chfin; remove chfin instances for fun, cfun
huffman [Fri, 22 Oct 2010 07:44:34 -0700] rev 40090
direct instantiation unit :: discrete_cpo
huffman [Fri, 22 Oct 2010 06:58:45 -0700] rev 40089
remove finite_po class
huffman [Fri, 22 Oct 2010 06:08:51 -0700] rev 40088
simplify proofs about flift; remove unneeded lemmas
huffman [Fri, 22 Oct 2010 05:54:54 -0700] rev 40087
simplify proof
huffman [Thu, 21 Oct 2010 15:21:39 -0700] rev 40086
minimize imports
huffman [Thu, 21 Oct 2010 15:19:07 -0700] rev 40085
add type annotation to avoid warning
huffman [Thu, 21 Oct 2010 12:51:36 -0700] rev 40084
simplify some proofs, convert to Isar style
huffman [Thu, 21 Oct 2010 12:03:49 -0700] rev 40083
rename lemma spair_lemma to spair_Sprod
huffman [Thu, 21 Oct 2010 06:03:18 -0700] rev 40082
pcpodef (open) 'a lift
huffman [Thu, 21 Oct 2010 05:44:38 -0700] rev 40081
remove intro! attribute from {sinl,sinr}_defined
huffman [Thu, 21 Oct 2010 05:35:32 -0700] rev 40080
simplify proofs of ssumE, sprodE
wenzelm [Sun, 24 Oct 2010 21:25:13 +0200] rev 40079
merged
nipkow [Sun, 24 Oct 2010 20:37:30 +0200] rev 40078
renamed nat_number
nipkow [Sun, 24 Oct 2010 20:19:00 +0200] rev 40077
nat_number -> eval_nat_numeral
krauss [Fri, 22 Oct 2010 23:45:20 +0200] rev 40076
some cleanup in Function_Lib
blanchet [Fri, 22 Oct 2010 17:15:46 +0200] rev 40075
merged
blanchet [Fri, 22 Oct 2010 16:45:55 +0200] rev 40074
compile
blanchet [Fri, 22 Oct 2010 16:37:11 +0200] rev 40073
added SMT solver to Sledgehammer docs
blanchet [Fri, 22 Oct 2010 16:11:43 +0200] rev 40072
more robust handling of "remote_" vs. non-"remote_" provers
blanchet [Fri, 22 Oct 2010 15:02:27 +0200] rev 40071
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet [Fri, 22 Oct 2010 14:47:43 +0200] rev 40070
replaced references with proper record that's threaded through
blanchet [Fri, 22 Oct 2010 14:10:32 +0200] rev 40069
fixed signature of "is_smt_solver_installed";
renaming
blanchet [Fri, 22 Oct 2010 13:57:54 +0200] rev 40068
renamed modules
blanchet [Fri, 22 Oct 2010 13:54:51 +0200] rev 40067
renamed files
blanchet [Fri, 22 Oct 2010 13:49:44 +0200] rev 40066
took out "smt"/"remote_smt" from default ATPs until they are properly implemented
blanchet [Fri, 22 Oct 2010 13:48:21 +0200] rev 40065
remove more needless code ("run_smt_solvers");
tuning
blanchet [Fri, 22 Oct 2010 12:15:31 +0200] rev 40064
got rid of duplicate functionality ("run_smt_solver_somehow");
added minimization command to SMT solver message
blanchet [Fri, 22 Oct 2010 11:58:33 +0200] rev 40063
bring ATPs and SMT solvers more in line with each other
blanchet [Fri, 22 Oct 2010 11:11:34 +0200] rev 40062
make Sledgehammer minimizer fully work with SMT
blanchet [Fri, 22 Oct 2010 09:50:18 +0200] rev 40061
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet [Thu, 21 Oct 2010 16:25:40 +0200] rev 40060
first step in adding support for an SMT backend to Sledgehammer
blanchet [Thu, 21 Oct 2010 14:55:09 +0200] rev 40059
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet [Thu, 21 Oct 2010 14:54:39 +0200] rev 40058
cosmetics
krauss [Fri, 22 Oct 2010 13:59:34 +0200] rev 40057
relation method: re-check given term with type constraints to avoid unspecific failure if ill-typed -- keep old behaviour for tactic version
hoelzl [Fri, 22 Oct 2010 12:01:12 +0200] rev 40056
Changed section title to please LaTeX.
bulwahn [Thu, 21 Oct 2010 20:26:35 +0200] rev 40055
temporary removed Predicate_Compile_Quickcheck_Examples from tests
bulwahn [Thu, 21 Oct 2010 19:13:11 +0200] rev 40054
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn [Thu, 21 Oct 2010 19:13:10 +0200] rev 40053
using a signature in core_data and moving some more functions to core_data
bulwahn [Thu, 21 Oct 2010 19:13:09 +0200] rev 40052
splitting large core file into core_data, mode_inference and predicate_compile_proof
bulwahn [Thu, 21 Oct 2010 19:13:09 +0200] rev 40051
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn [Thu, 21 Oct 2010 19:13:08 +0200] rev 40050
for now safely but unpractically assume no predicate is terminating
bulwahn [Thu, 21 Oct 2010 19:13:07 +0200] rev 40049
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn [Thu, 21 Oct 2010 19:13:06 +0200] rev 40048
adding option smart_depth_limiting to predicate compiler
huffman [Wed, 20 Oct 2010 21:26:51 -0700] rev 40047
merged
huffman [Wed, 20 Oct 2010 19:40:02 -0700] rev 40046
introduce function strict :: 'a -> 'b -> 'b, which works like Haskell's seq; use strict instead of strictify in various definitions
huffman [Wed, 20 Oct 2010 17:25:22 -0700] rev 40045
add lemma lub_eq_bottom_iff
huffman [Wed, 20 Oct 2010 16:19:25 -0700] rev 40044
combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman [Wed, 20 Oct 2010 13:22:30 -0700] rev 40043
constructor arguments with selectors must have pointed types
huffman [Wed, 20 Oct 2010 13:02:13 -0700] rev 40042
simplify check_and_sort_domain; more meaningful variable names
huffman [Tue, 19 Oct 2010 16:21:24 -0700] rev 40041
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman [Tue, 19 Oct 2010 15:01:51 -0700] rev 40040
rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
huffman [Tue, 19 Oct 2010 14:28:14 -0700] rev 40039
simplify some proofs; remove some unused lists of lemmas
huffman [Tue, 19 Oct 2010 11:07:42 -0700] rev 40038
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman [Tue, 19 Oct 2010 10:13:29 -0700] rev 40037
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman [Tue, 19 Oct 2010 07:05:04 -0700] rev 40036
simplify fixrec pattern match function
huffman [Sun, 17 Oct 2010 09:53:47 -0700] rev 40035
simplify some proofs
Christian Urban <urbanc@in.tum.de> [Tue, 19 Oct 2010 15:13:35 +0100] rev 40034
tuned
bulwahn [Tue, 19 Oct 2010 12:26:38 +0200] rev 40033
added some facts about factorial and dvd, div and mod
bulwahn [Tue, 19 Oct 2010 12:26:37 +0200] rev 40032
removing something that probably slipped into the Quotient_List theory
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 19 Oct 2010 11:44:42 +0900] rev 40031
Quotient package: partial equivalence introduction
Christian Urban <urbanc@in.tum.de> [Mon, 18 Oct 2010 14:25:15 +0100] rev 40030
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
huffman [Sat, 16 Oct 2010 17:10:23 -0700] rev 40029
remove dead code
huffman [Sat, 16 Oct 2010 17:09:57 -0700] rev 40028
remove old uses of 'simp_tac HOLCF_ss'
huffman [Sat, 16 Oct 2010 16:39:06 -0700] rev 40027
merged
huffman [Sat, 16 Oct 2010 16:22:42 -0700] rev 40026
remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman [Sat, 16 Oct 2010 15:26:30 -0700] rev 40025
reimplement proof automation for coinduct rules
huffman [Sat, 16 Oct 2010 14:41:11 -0700] rev 40024
add functions mk_imp, mk_all
huffman [Fri, 15 Oct 2010 08:52:53 -0700] rev 40023
move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman [Fri, 15 Oct 2010 08:07:20 -0700] rev 40022
simplify automation of induct proof
huffman [Fri, 15 Oct 2010 06:08:42 -0700] rev 40021
add function mk_adm
huffman [Fri, 15 Oct 2010 05:50:27 -0700] rev 40020
rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman [Thu, 14 Oct 2010 19:16:52 -0700] rev 40019
put constructor argument specs in constr_info type
huffman [Thu, 14 Oct 2010 14:42:05 -0700] rev 40018
avoid using Global_Theory.get_thm
huffman [Thu, 14 Oct 2010 13:46:27 -0700] rev 40017
include iso_info as part of constr_info type
huffman [Thu, 14 Oct 2010 13:28:31 -0700] rev 40016
remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman [Thu, 14 Oct 2010 10:16:46 -0700] rev 40015
add take_strict_thms field to take_info type
huffman [Thu, 14 Oct 2010 09:44:40 -0700] rev 40014
add record type synonym 'constr_info'
huffman [Thu, 14 Oct 2010 09:34:00 -0700] rev 40013
add function take_theorems
huffman [Thu, 14 Oct 2010 09:28:05 -0700] rev 40012
add type annotation to avoid warning
huffman [Wed, 13 Oct 2010 10:56:42 -0700] rev 40011
cleaned up Fun_Cpo.thy; deprecated a few theorem names
huffman [Wed, 13 Oct 2010 10:27:26 -0700] rev 40010
edit comments
huffman [Tue, 12 Oct 2010 09:32:21 -0700] rev 40009
remove unneeded lemmas Lift_exhaust, Lift_cases
huffman [Tue, 12 Oct 2010 09:08:27 -0700] rev 40008
move lemmas from Lift.thy to Cfun.thy
huffman [Tue, 12 Oct 2010 07:46:44 -0700] rev 40007
cleaned up Adm.thy
huffman [Tue, 12 Oct 2010 06:20:05 -0700] rev 40006
remove unneeded lemmas from Fun_Cpo.thy
huffman [Tue, 12 Oct 2010 05:48:32 -0700] rev 40005
remove unused lemmas
huffman [Tue, 12 Oct 2010 05:48:15 -0700] rev 40004
reformulate lemma cont2cont_lub and move to Cont.thy
huffman [Tue, 12 Oct 2010 05:25:21 -0700] rev 40003
remove unused lemmas cont_fst_snd_D1, cont_fst_snd_D2
huffman [Mon, 11 Oct 2010 21:35:31 -0700] rev 40002
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman [Mon, 11 Oct 2010 16:24:44 -0700] rev 40001
rename Ffun.thy to Fun_Cpo.thy
huffman [Mon, 11 Oct 2010 16:14:15 -0700] rev 40000
remove unused constant 'directed'
huffman [Mon, 11 Oct 2010 09:54:04 -0700] rev 39999
add HOLCF/Library/Defl_Bifinite.thy, which proves instance defl :: bifinite
paulson [Fri, 15 Oct 2010 17:21:37 +0100] rev 39998
merged
paulson [Fri, 15 Oct 2010 17:21:07 +0100] rev 39997
prevention of self-referential type environments
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 15 Oct 2010 21:50:26 +0900] rev 39996
FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 15 Oct 2010 21:47:45 +0900] rev 39995
FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 15 Oct 2010 21:46:45 +0900] rev 39994
FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
krauss [Thu, 14 Oct 2010 12:40:14 +0200] rev 39993
NEWS
krauss [Sun, 10 Oct 2010 22:50:25 +0200] rev 39992
removed output syntax "'a ~=> 'b" for "'a => 'b option"
krauss [Wed, 13 Oct 2010 09:56:00 +0200] rev 39991
reactivated
krauss [Tue, 12 Oct 2010 21:30:44 +0200] rev 39990
slightly more robust proof
huffman [Mon, 11 Oct 2010 08:32:09 -0700] rev 39989
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman [Mon, 11 Oct 2010 07:09:42 -0700] rev 39988
merged
huffman [Sat, 09 Oct 2010 07:24:49 -0700] rev 39987
move all bifinite class instances to Bifinite.thy
huffman [Fri, 08 Oct 2010 07:39:50 -0700] rev 39986
rename class 'sfp' to 'bifinite'
huffman [Thu, 07 Oct 2010 13:54:43 -0700] rev 39985
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman [Thu, 07 Oct 2010 13:33:06 -0700] rev 39984
add lemma typedef_ideal_completion
huffman [Thu, 07 Oct 2010 13:22:13 -0700] rev 39983
remove unused lemmas
huffman [Thu, 07 Oct 2010 13:19:45 -0700] rev 39982
remove Infinite_Set from ROOT.ML
huffman [Thu, 07 Oct 2010 13:18:48 -0700] rev 39981
remove some junk that made it in by accient
blanchet [Mon, 11 Oct 2010 18:03:47 +0700] rev 39980
"setup" in theory
blanchet [Mon, 11 Oct 2010 18:03:18 +0700] rev 39979
added "trace_meson" configuration option, replacing old-fashioned reference
blanchet [Mon, 11 Oct 2010 18:02:14 +0700] rev 39978
added "trace_metis" configuration option, replacing old-fashioned references
krauss [Sun, 10 Oct 2010 23:16:24 +0200] rev 39977
do not mention unqualified names, now that 'global' and 'local' are gone
nipkow [Sun, 10 Oct 2010 16:34:20 +0200] rev 39976
simplified proof
blanchet [Sun, 10 Oct 2010 18:42:13 +0700] rev 39975
avoid generating several formulas with the same name ("tfrees")
huffman [Wed, 06 Oct 2010 10:49:27 -0700] rev 39974
major reorganization/simplification of HOLCF type classes:
removed profinite/bifinite classes and approx function;
universal domain uses approx_chain locale instead of bifinite class;
ideal_completion locale does not use 'take' functions, requires countable basis instead;
replaced type 'udom alg_defl' with type 'sfp';
replaced class 'rep' with class 'sfp';
renamed REP('a) to SFP('a);
Brian Huffman <brianh@cs.pdx.edu> [Tue, 05 Oct 2010 17:53:00 -0700] rev 39973
add lemma finite_deflation_intro