Tue, 01 Aug 2006 15:28:55 +0200 webertj comment (timing information for last example) added
Tue, 01 Aug 2006 14:58:43 +0200 webertj possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
Tue, 01 Aug 2006 13:51:16 +0200 obua removed skip
Tue, 01 Aug 2006 13:44:05 +0200 mengj Added in code to check too general axiom clauses.
Tue, 01 Aug 2006 11:32:43 +0200 wenzelm exported attrib;
Mon, 31 Jul 2006 21:06:40 +0200 webertj lin_arith_prover splits certain operators (e.g. min, max, abs)
Mon, 31 Jul 2006 20:56:49 +0200 webertj fixed a bug in function poly: decomposition of products
Mon, 31 Jul 2006 18:07:42 +0200 krauss Function package can now do automatic splits of overlapping datatype patterns
Mon, 31 Jul 2006 18:05:40 +0200 krauss Removed an "apply arith" where there are already "No Subgoals"
Mon, 31 Jul 2006 15:29:36 +0200 webertj code reformatted
Mon, 31 Jul 2006 14:08:42 +0200 berghofe Additional freshness constraints for FCB.
Sun, 30 Jul 2006 21:28:59 +0200 wenzelm proper Element.generalize_facts;
Sun, 30 Jul 2006 21:28:58 +0200 wenzelm add_consts: proper Sign.full_name;
Sun, 30 Jul 2006 21:28:57 +0200 wenzelm added generalize_facts;
Sun, 30 Jul 2006 21:28:56 +0200 wenzelm added maxidx_values;
Sun, 30 Jul 2006 21:28:55 +0200 wenzelm export: refrain from adjusting maxidx;
Sun, 30 Jul 2006 21:28:54 +0200 wenzelm adjust_maxidx: pass explicit lower bound;
Sun, 30 Jul 2006 21:28:52 +0200 wenzelm Thm.adjust_maxidx;
Sun, 30 Jul 2006 21:28:51 +0200 wenzelm removed unused add_in_order/add_once (cf. OrdList.insert);
Sun, 30 Jul 2006 21:28:50 +0200 wenzelm demod_rule: depend on context, proper Variable.import/export;
Sun, 30 Jul 2006 21:28:48 +0200 wenzelm tuned proofs;
Sun, 30 Jul 2006 05:53:10 +0200 webertj lin_arith_prover splits certain operators (e.g. min, max, abs)
Sun, 30 Jul 2006 02:06:01 +0200 webertj bugfix related to cancel_div_mod simproc
Sat, 29 Jul 2006 13:15:12 +0200 webertj lin_arith_prover splits certain operators (e.g. min, max, abs)
Sat, 29 Jul 2006 00:51:36 +0200 wenzelm rename legacy_pretty_thm to pretty_thm_legacy;
Sat, 29 Jul 2006 00:51:34 +0200 wenzelm tuned comment;
Sat, 29 Jul 2006 00:51:33 +0200 wenzelm added add_fixes_direct;
Sat, 29 Jul 2006 00:51:32 +0200 wenzelm prove: proper assumption context, more tactic arguments;
Sat, 29 Jul 2006 00:51:31 +0200 wenzelm added mk_conjunction_list;
Sat, 29 Jul 2006 00:51:29 +0200 wenzelm Goal.prove: more tactic arguments;
Fri, 28 Jul 2006 18:37:25 +0200 paulson Checking for unsound proofs. Tidying.
Fri, 28 Jul 2006 18:36:24 +0200 paulson "all theorems" mode forces definition-expansion off.
Fri, 28 Jul 2006 18:11:22 +0200 webertj title fixed
Thu, 27 Jul 2006 23:28:30 +0200 wenzelm replaced extern_skolem by slightly more simplistic revert_skolems;
Thu, 27 Jul 2006 23:28:28 +0200 wenzelm renamed ProofContext.fix_frees to Variable.fix_frees;
Thu, 27 Jul 2006 23:28:27 +0200 wenzelm replaced ProofContext.extern_skolem by slightly more simplistic ProofContext.revert_skolems;
Thu, 27 Jul 2006 23:28:26 +0200 wenzelm no_vars: based on Variable.import;
Thu, 27 Jul 2006 23:28:25 +0200 wenzelm added fix_frees (from Isar/proof_context.ML);
Thu, 27 Jul 2006 23:28:23 +0200 wenzelm declare_term_names: cover types as well;
Thu, 27 Jul 2006 15:33:21 +0200 wenzelm eliminated obsolete freeze_thaw;
Thu, 27 Jul 2006 14:21:57 +0200 webertj type annotation added to make SML/NJ happy
Thu, 27 Jul 2006 13:44:03 +0200 wenzelm "moved basic assumption operations from structure ProofContext to Assumption;"
Thu, 27 Jul 2006 13:43:13 +0200 wenzelm ProofContext.legacy_pretty_thm;
Thu, 27 Jul 2006 13:43:12 +0200 wenzelm added legacy_pretty_thm (with fall-back on ProtoPure.thy);
Thu, 27 Jul 2006 13:43:11 +0200 wenzelm Assumption.assume;
Thu, 27 Jul 2006 13:43:10 +0200 wenzelm removed obsolete is_fact (cf. Thm.no_prems);
Thu, 27 Jul 2006 13:43:09 +0200 wenzelm tuned interfaces;
Thu, 27 Jul 2006 13:43:08 +0200 wenzelm read_def_cterms (legacy version): Consts.certify;
Thu, 27 Jul 2006 13:43:07 +0200 wenzelm Assumption.assume;
Thu, 27 Jul 2006 13:43:06 +0200 wenzelm moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
Thu, 27 Jul 2006 13:43:05 +0200 wenzelm removed obsolete equal_abs_elim(_list);
Thu, 27 Jul 2006 13:43:04 +0200 wenzelm removed obsolete pretty_thm_no_quote;
Thu, 27 Jul 2006 13:43:03 +0200 wenzelm added Pure/assumption.ML;
Thu, 27 Jul 2006 13:43:01 +0200 wenzelm moved basic assumption operations from structure ProofContext to Assumption;
Thu, 27 Jul 2006 13:43:00 +0200 wenzelm tuned proofs;
Thu, 27 Jul 2006 13:42:59 +0200 wenzelm Local assumptions, parameterized by export rules.
Wed, 26 Jul 2006 19:37:44 +0200 wenzelm updated;
Wed, 26 Jul 2006 19:37:43 +0200 wenzelm import(T): result includes fixed types/terms;
Wed, 26 Jul 2006 19:37:42 +0200 wenzelm focus: result record includes (fixed) schematic variables;
Wed, 26 Jul 2006 19:37:41 +0200 wenzelm Variable.import(T): result includes fixed types/terms;
Wed, 26 Jul 2006 19:23:04 +0200 webertj linear arithmetic splits certain operators (e.g. min, max, abs)
Wed, 26 Jul 2006 13:31:07 +0200 haftmann added eval_term
Wed, 26 Jul 2006 11:32:55 +0200 wenzelm updated;
Wed, 26 Jul 2006 11:32:50 +0200 wenzelm fixed LaTeX problem;
Wed, 26 Jul 2006 10:09:25 +0200 haftmann added eval_term
Wed, 26 Jul 2006 09:50:23 +0200 nipkow Removed wrong sentence (Simon Funke)
Wed, 26 Jul 2006 00:44:49 +0200 wenzelm moved pprint functions to Isar/proof_display.ML;
Wed, 26 Jul 2006 00:44:48 +0200 wenzelm Tactical operations depending on local subgoal structure.
Wed, 26 Jul 2006 00:44:47 +0200 wenzelm moved pprint functions to Isar/proof_display.ML;
Wed, 26 Jul 2006 00:44:46 +0200 wenzelm export goal_tac (was internal refine_tac);
Wed, 26 Jul 2006 00:44:44 +0200 wenzelm added Pure/subgoal.ML;
Tue, 25 Jul 2006 23:17:42 +0200 wenzelm updated;
Tue, 25 Jul 2006 23:17:41 +0200 wenzelm raw symbols: disallow dot to avoid confusion in NameSpace.unpack;
Tue, 25 Jul 2006 21:18:12 +0200 wenzelm avoid Term.is_funtype;
Tue, 25 Jul 2006 21:18:11 +0200 wenzelm avoid structure Char;
Tue, 25 Jul 2006 21:18:09 +0200 wenzelm added variant_abs (from term.ML);
Tue, 25 Jul 2006 21:18:08 +0200 wenzelm added find_free (from term.ML);
Tue, 25 Jul 2006 21:18:07 +0200 wenzelm added is/to_ascii_lower/upper;
Tue, 25 Jul 2006 21:18:06 +0200 wenzelm is_funtype: do not export internal operation;
Tue, 25 Jul 2006 21:18:05 +0200 wenzelm tuned;
Tue, 25 Jul 2006 21:18:04 +0200 wenzelm use Term.add_vars instead of obsolete term_varnames;
Tue, 25 Jul 2006 21:18:02 +0200 wenzelm renamed add_term_varnames to Term.add_varnames (cf. Term.add_vars etc.);
Tue, 25 Jul 2006 21:18:01 +0200 wenzelm tuned ML code;
Tue, 25 Jul 2006 21:18:00 +0200 wenzelm renamed Term.variant_abs to Syntax.variant_abs;
Tue, 25 Jul 2006 21:17:58 +0200 wenzelm Drule.merge_rules;
Tue, 25 Jul 2006 16:51:26 +0200 haftmann renamed Name.give_names to Name.names and moved Name.alphanum to Symbol.alphanum
Tue, 25 Jul 2006 16:43:47 +0200 haftmann improvements for lazy code generation
Tue, 25 Jul 2006 16:43:33 +0200 haftmann fixed typo
Tue, 25 Jul 2006 16:43:32 +0200 haftmann added code generator serialization for Char
Tue, 25 Jul 2006 16:43:31 +0200 haftmann added notes on class_package.ML and codegen_package.ML
Sun, 23 Jul 2006 07:21:41 +0200 haftmann small adjustments
Sun, 23 Jul 2006 07:21:22 +0200 haftmann fixed bug for serialization for uminus on ints
Sun, 23 Jul 2006 07:20:52 +0200 haftmann small improvement in serialization for wfrec
Sun, 23 Jul 2006 07:20:26 +0200 haftmann added structure HOList
Sun, 23 Jul 2006 07:19:36 +0200 haftmann major simplifications for integers
Sun, 23 Jul 2006 07:19:26 +0200 haftmann tactic for prove_instance_arity now gets definition theorems as arguments
Fri, 21 Jul 2006 14:49:11 +0200 haftmann added term_of_string function
Fri, 21 Jul 2006 14:48:35 +0200 haftmann simplification for code generation for Integers
Fri, 21 Jul 2006 14:48:17 +0200 haftmann adaption to argument change in primrec_package
Fri, 21 Jul 2006 14:47:44 +0200 haftmann adaption to changes in class_package
Fri, 21 Jul 2006 14:47:22 +0200 haftmann hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
Fri, 21 Jul 2006 14:46:27 +0200 haftmann exported equation transformator
Fri, 21 Jul 2006 14:45:43 +0200 haftmann class package and codegen refinements
Fri, 21 Jul 2006 14:45:25 +0200 haftmann added give_names and alphanum
Fri, 21 Jul 2006 14:11:14 +0200 berghofe Some cases in "case ... of ..." expressions may now
Fri, 21 Jul 2006 11:34:01 +0200 berghofe - Added new "undefined" constant
Thu, 20 Jul 2006 14:35:37 +0200 wenzelm removed Variable.monomorphic;
Thu, 20 Jul 2006 14:34:57 +0200 webertj comments fixed, member function renamed
Wed, 19 Jul 2006 23:22:22 +0200 ballarin Change to algebra method.
Wed, 19 Jul 2006 19:25:58 +0200 ballarin Reimplemented algebra method; now controlled by attribute.
Wed, 19 Jul 2006 19:24:02 +0200 ballarin Strict dfs traversal of imported and registered identifiers.
Wed, 19 Jul 2006 14:16:36 +0200 haftmann added map_default, internal restructuring
Wed, 19 Jul 2006 12:12:08 +0200 wenzelm export is_tid;
Wed, 19 Jul 2006 12:12:07 +0200 wenzelm thm_of_proof: improved generation of variables;
Wed, 19 Jul 2006 12:12:06 +0200 wenzelm Sign.infer_types: Name.context;
Wed, 19 Jul 2006 12:12:05 +0200 wenzelm reorganize declarations (more efficient);
Wed, 19 Jul 2006 12:12:04 +0200 wenzelm Name.context for used'';
Wed, 19 Jul 2006 12:12:03 +0200 wenzelm added variant_frees;
Wed, 19 Jul 2006 12:12:02 +0200 wenzelm tuned;
Wed, 19 Jul 2006 12:12:01 +0200 wenzelm export make_context, is_declared;
Wed, 19 Jul 2006 12:12:00 +0200 wenzelm prove: Variable.declare_internal (more efficient);
Wed, 19 Jul 2006 12:11:59 +0200 wenzelm add_local: simplified interface, all frees are known'';
Wed, 19 Jul 2006 12:11:57 +0200 wenzelm Sign.infer_types: Name.context;
Wed, 19 Jul 2006 12:11:56 +0200 wenzelm renamed Variable.rename_wrt to Variable.variant_frees;
Wed, 19 Jul 2006 11:55:26 +0200 paulson Fixed the bugs introduced by the last commit! Output is now *identical* to that
Wed, 19 Jul 2006 02:35:22 +0200 webertj MiniSat proof trace format changed; MiniSat is now expected to produce a proof also for "trivial" problems
Tue, 18 Jul 2006 20:01:47 +0200 wenzelm thm_of_proof: tuned Name operations;
Tue, 18 Jul 2006 20:01:46 +0200 wenzelm print_statement: tuned Variable operations;
Tue, 18 Jul 2006 20:01:45 +0200 wenzelm added newly_fixed, focus;
Tue, 18 Jul 2006 20:01:44 +0200 wenzelm added declare_term_names;
Tue, 18 Jul 2006 20:01:42 +0200 wenzelm fold_proof_terms: canonical arguments;
Tue, 18 Jul 2006 20:01:41 +0200 wenzelm Term.declare_term_names;
Tue, 18 Jul 2006 17:10:22 +0200 berghofe Started implementing uniqueness proof for recursion
Tue, 18 Jul 2006 16:15:47 +0200 webertj typo (theorerms) fixed
Tue, 18 Jul 2006 14:53:27 +0200 webertj typo (theorerms) fixed
Tue, 18 Jul 2006 13:27:59 +0200 haftmann AList.join now with 'DUP' exception
Tue, 18 Jul 2006 08:48:11 +0200 haftmann added Table.map_default
Tue, 18 Jul 2006 02:22:38 +0200 wenzelm removed obsolete ML files;
Mon, 17 Jul 2006 18:42:38 +0200 wenzelm replaced butlast by Library.split_last;
Mon, 17 Jul 2006 18:42:37 +0200 wenzelm replaced butlast by Library.split_last;
Mon, 17 Jul 2006 15:16:17 +0200 webertj butlast removed (use fst o split_last instead)
Mon, 17 Jul 2006 01:28:17 +0200 webertj support for MiniSat proof traces added
Mon, 17 Jul 2006 00:37:06 +0200 webertj support for MiniSat proof traces added
Sun, 16 Jul 2006 14:26:22 +0200 paulson has_consts renamed to has_conn, now actually parses the first-order formula
Sat, 15 Jul 2006 18:17:47 +0200 webertj function butlast added
Sat, 15 Jul 2006 15:26:50 +0200 paulson Replaced a-lists by tables to improve efficiency
Sat, 15 Jul 2006 13:52:10 +0200 mengj Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
Sat, 15 Jul 2006 13:50:26 +0200 mengj Only include combinators if required by goals and user specified lemmas.
Fri, 14 Jul 2006 14:37:15 +0200 ballarin Term.term_lpo takes order on terms rather than strings as argument.
Fri, 14 Jul 2006 14:19:48 +0200 wenzelm keep/transaction: unified execution model (with debugging etc.);
Fri, 14 Jul 2006 13:51:30 +0200 webertj trivial whitespace changes
Fri, 14 Jul 2006 12:18:33 +0200 wenzelm simp method: depth_limit;
Thu, 13 Jul 2006 17:39:56 +0200 paulson "conjecture" must be lower case
Thu, 13 Jul 2006 13:42:05 +0200 wenzelm tuned insert_list;
Thu, 13 Jul 2006 13:42:02 +0200 wenzelm Name.context already declares empty names;
Thu, 13 Jul 2006 13:42:00 +0200 wenzelm strip_abs_eta: proper use of Name.context;
Thu, 13 Jul 2006 13:41:57 +0200 wenzelm do not export make_context;
Thu, 13 Jul 2006 13:41:53 +0200 wenzelm removed colon from sym category;
Thu, 13 Jul 2006 12:31:00 +0200 paulson fix to refl_clause_aux: added occurs check
Wed, 12 Jul 2006 21:19:27 +0200 wenzelm * Isar: ":" (colon) is no longer a symbolic identifier character;
Wed, 12 Jul 2006 21:19:24 +0200 wenzelm removed duplicate of Tactic.make_elim_preserve;
Wed, 12 Jul 2006 21:19:22 +0200 wenzelm removed obsolete adhoc_freeze_vars (may use Variable.import_terms instead);
Wed, 12 Jul 2006 21:19:19 +0200 wenzelm exported make_elim_preserve;
Wed, 12 Jul 2006 21:19:17 +0200 wenzelm prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
Wed, 12 Jul 2006 21:19:14 +0200 wenzelm simplified prove_conv;
Wed, 12 Jul 2006 19:59:14 +0200 wenzelm removed ':' from category of symbolic identifier chars;
Wed, 12 Jul 2006 19:59:13 +0200 wenzelm query/bang_colon: separate tokens;
Wed, 12 Jul 2006 17:40:52 +0200 haftmann purged website sources
Wed, 12 Jul 2006 17:00:33 +0200 haftmann added strip_abs_eta
Wed, 12 Jul 2006 17:00:32 +0200 haftmann added chop_prefix
Wed, 12 Jul 2006 17:00:31 +0200 haftmann class_of_param instead of class_of
Wed, 12 Jul 2006 17:00:30 +0200 haftmann adaptions in class_package
Wed, 12 Jul 2006 17:00:22 +0200 haftmann adaptions in codegen
Wed, 12 Jul 2006 00:34:54 +0200 wenzelm variants: special treatment of empty name;
Tue, 11 Jul 2006 23:49:32 +0200 wenzelm avoid reference to internal skolem;
Tue, 11 Jul 2006 23:00:39 +0200 wenzelm separate names filed (covers fixes/defaults);
Tue, 11 Jul 2006 23:00:39 +0200 wenzelm adapted Name.defaults_of;
Tue, 11 Jul 2006 23:00:37 +0200 wenzelm removed obsolete xless;
Tue, 11 Jul 2006 23:00:36 +0200 wenzelm clean: no special treatment of empty name;
Tue, 11 Jul 2006 23:00:35 +0200 wenzelm removed obsolete xless;
Tue, 11 Jul 2006 18:10:47 +0200 webertj replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
Tue, 11 Jul 2006 14:21:08 +0200 wenzelm uniform treatment of num/xnum;
Tue, 11 Jul 2006 14:21:07 +0200 wenzelm replaced read_radixint by read_intinf;
Tue, 11 Jul 2006 14:21:05 +0200 wenzelm removed str_to_int in favour of general Syntax.read_xnum;
Tue, 11 Jul 2006 14:21:04 +0200 wenzelm num/xnum: bin or hex;
Tue, 11 Jul 2006 12:24:23 +0200 wenzelm Name.internal;
Tue, 11 Jul 2006 12:22:58 +0200 wenzelm tuned;
Tue, 11 Jul 2006 12:17:30 +0200 wenzelm * Pure: structure Name;
Tue, 11 Jul 2006 12:17:17 +0200 wenzelm Names of basic logical entities (variables etc.).
Tue, 11 Jul 2006 12:17:16 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Tue, 11 Jul 2006 12:17:13 +0200 wenzelm Name.internal;
Tue, 11 Jul 2006 12:17:12 +0200 wenzelm adapted to more efficient Name/Variable implementation;
Tue, 11 Jul 2006 12:17:11 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Tue, 11 Jul 2006 12:17:09 +0200 wenzelm maintain Name.context for fixes/defaults;
Tue, 11 Jul 2006 12:17:08 +0200 wenzelm removed obsolete mem_ix;
Tue, 11 Jul 2006 12:17:07 +0200 wenzelm removed obsolete ins_ix, mem_ix, ins_term, mem_term;
Tue, 11 Jul 2006 12:17:06 +0200 wenzelm Name.dest_skolem;
Tue, 11 Jul 2006 12:17:05 +0200 wenzelm Name.bound;
Tue, 11 Jul 2006 12:17:04 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Tue, 11 Jul 2006 12:17:03 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Tue, 11 Jul 2006 12:17:02 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Tue, 11 Jul 2006 12:17:01 +0200 wenzelm Name.invent_list;
Tue, 11 Jul 2006 12:17:00 +0200 wenzelm added name.ML;
Tue, 11 Jul 2006 12:16:59 +0200 wenzelm Name.is_bound;
Tue, 11 Jul 2006 12:16:58 +0200 wenzelm removed obsolete mem_term;
Tue, 11 Jul 2006 12:16:57 +0200 wenzelm Name.internal;
Tue, 11 Jul 2006 12:16:54 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Tue, 11 Jul 2006 12:16:52 +0200 wenzelm let_simproc: activate Variable.import;
Tue, 11 Jul 2006 11:19:28 +0200 ballarin Witness theorems of interpretations now transfered to current theory.
Tue, 11 Jul 2006 11:17:09 +0200 ballarin New function transfer_witness lifting Thm.transfer to witnesses.
Tue, 11 Jul 2006 00:43:54 +0200 kleing hex and binary numerals (contributed by Rafal Kolanski)
Mon, 10 Jul 2006 21:02:29 +0200 webertj minor optimization wrt. certifying terms
Sat, 08 Jul 2006 14:12:13 +0200 wenzelm tuned;
Sat, 08 Jul 2006 14:01:40 +0200 wenzelm tuned;
Sat, 08 Jul 2006 14:01:31 +0200 wenzelm updated;
Sat, 08 Jul 2006 12:54:50 +0200 wenzelm tuned interface;
Sat, 08 Jul 2006 12:54:49 +0200 wenzelm tuned interface;
Sat, 08 Jul 2006 12:54:48 +0200 wenzelm removed dead code;
Sat, 08 Jul 2006 12:54:47 +0200 wenzelm Element.prove_witness: context;
Sat, 08 Jul 2006 12:54:46 +0200 wenzelm prove_witness: context;
Sat, 08 Jul 2006 12:54:45 +0200 wenzelm tuned exception handling;
Sat, 08 Jul 2006 12:54:44 +0200 wenzelm prove/prove_multi: context;
Sat, 08 Jul 2006 12:54:43 +0200 wenzelm simprocs: no theory argument -- use simpset context instead;
Sat, 08 Jul 2006 12:54:42 +0200 wenzelm distinct simproc/simpset: proper context;
Sat, 08 Jul 2006 12:54:41 +0200 wenzelm presburger_ss: proper context;
Sat, 08 Jul 2006 12:54:40 +0200 wenzelm presburger_ss: proper context;
Sat, 08 Jul 2006 12:54:39 +0200 wenzelm tuned;
Sat, 08 Jul 2006 12:54:38 +0200 wenzelm avoid Force_tac, which uses a different context;
Sat, 08 Jul 2006 12:54:37 +0200 wenzelm Goal.prove: context;
Sat, 08 Jul 2006 12:54:36 +0200 wenzelm tactic/method simpset: maintain proper context;
Sat, 08 Jul 2006 12:54:35 +0200 wenzelm Goal.prove_global;
Sat, 08 Jul 2006 12:54:33 +0200 wenzelm Goal.prove_global;
Sat, 08 Jul 2006 12:54:32 +0200 wenzelm simprocs: no theory argument -- use simpset context instead;
Sat, 08 Jul 2006 12:54:30 +0200 wenzelm simprocs: no theory argument -- use simpset context instead;
Sat, 08 Jul 2006 12:54:29 +0200 wenzelm updated;
Sat, 08 Jul 2006 12:54:28 +0200 wenzelm updated Goal.prove, Goal.prove_global;
Sat, 08 Jul 2006 12:54:27 +0200 wenzelm added some bits on variables;
Sat, 08 Jul 2006 12:54:26 +0200 wenzelm * Pure: structure Variable provides operations for proper treatment of fixed/schematic variables;
Fri, 07 Jul 2006 18:13:58 +0200 webertj "solver" reference added to make the SAT solver configurable
Fri, 07 Jul 2006 15:13:15 +0200 paulson Some tidying.
(0) -10000 -3000 -1000 -240 +240 +1000 +3000 +10000 +30000 tip