Fri, 23 Dec 2005 20:02:30 +0100 wenzelm tuned proofs;
Fri, 23 Dec 2005 18:36:27 +0100 wenzelm removed obsolete atomize_old;
Fri, 23 Dec 2005 18:36:26 +0100 wenzelm removed obsolete induct_atomize_old;
Fri, 23 Dec 2005 17:37:54 +0100 paulson the "skolem" attribute and better initialization of the clause database
Fri, 23 Dec 2005 17:36:00 +0100 paulson blacklist of prolific theorems (must be replaced by an attribute later
Fri, 23 Dec 2005 17:34:46 +0100 paulson tidied
Fri, 23 Dec 2005 15:21:05 +0100 wenzelm tuned;
Fri, 23 Dec 2005 15:18:13 +0100 wenzelm * Provers/induct: support simultaneous goals with mutual rules;
Fri, 23 Dec 2005 15:16:58 +0100 wenzelm induct etc.: admit multiple rules;
Fri, 23 Dec 2005 15:16:56 +0100 wenzelm backpatching of Substring.full;
Fri, 23 Dec 2005 15:16:55 +0100 wenzelm goal/qed: proper treatment of two levels of conjunctions;
Fri, 23 Dec 2005 15:16:53 +0100 wenzelm Logic.mk_conjunction_list;
Fri, 23 Dec 2005 15:16:52 +0100 wenzelm turned bicompose_no_flatten into compose_no_flatten, without elimination;
Fri, 23 Dec 2005 15:16:49 +0100 wenzelm CONJUNCTS: full nesting (again), PRECISE_CONJUNCTS: outer level of nesting;
Fri, 23 Dec 2005 15:16:48 +0100 wenzelm added mk_conjunction_list2;
Fri, 23 Dec 2005 15:16:46 +0100 wenzelm conj_elim_precise: proper treatment of nested conjunctions;
Fri, 23 Dec 2005 15:16:46 +0100 wenzelm Thm.compose_no_flatten;
Fri, 23 Dec 2005 15:16:44 +0100 wenzelm proper treatment of nested conjunctions, i.e. simultaneous goals and mutual rules;
Fri, 23 Dec 2005 14:33:28 +0100 haftmann is_prefix
Thu, 22 Dec 2005 19:08:15 +0100 haftmann slight improvements
Thu, 22 Dec 2005 17:57:09 +0100 nipkow more lemmas
Thu, 22 Dec 2005 14:22:11 +0100 paulson shorter proof
Thu, 22 Dec 2005 13:31:58 +0100 berghofe Tuned syntax for perm.
Thu, 22 Dec 2005 13:00:53 +0100 nipkow new lemmas
Thu, 22 Dec 2005 12:27:10 +0100 paulson Fixed a use of an outdated Substring function
Thu, 22 Dec 2005 00:41:26 +0100 wenzelm tuned;
Thu, 22 Dec 2005 00:29:22 +0100 wenzelm exh_casedist2: norm_hhf_eq;
Thu, 22 Dec 2005 00:29:20 +0100 wenzelm added bicompose_no_flatten, which refrains from
Thu, 22 Dec 2005 00:29:19 +0100 wenzelm bicompose_proof: no_flatten;
Thu, 22 Dec 2005 00:29:18 +0100 wenzelm conclude/SELECT: Thm.bicompose_no_flatten avoids unsolicited modification of result;
Thu, 22 Dec 2005 00:29:17 +0100 wenzelm Transform mutual rule into projection.
Thu, 22 Dec 2005 00:29:15 +0100 wenzelm added Provers/project_rule.ML
Thu, 22 Dec 2005 00:29:14 +0100 wenzelm structure ProjectRule;
Thu, 22 Dec 2005 00:29:12 +0100 wenzelm * induct: improved treatment of mutual goals and mutual rules;
Thu, 22 Dec 2005 00:29:11 +0100 wenzelm updated auxiliary facts for induct method;
Thu, 22 Dec 2005 00:29:10 +0100 wenzelm prop_tr': proper handling of aprop marked as bound;
Thu, 22 Dec 2005 00:29:09 +0100 wenzelm consume: expand defs in prems of concls;
Thu, 22 Dec 2005 00:29:08 +0100 wenzelm cases: main is_proper flag;
Thu, 22 Dec 2005 00:29:07 +0100 wenzelm auto cases: marked improper;
Thu, 22 Dec 2005 00:29:06 +0100 wenzelm conjunction_tac: single goal;
Thu, 22 Dec 2005 00:29:04 +0100 wenzelm CONJUNCTS2;
Thu, 22 Dec 2005 00:29:03 +0100 wenzelm rule_context: numbered cases;
Thu, 22 Dec 2005 00:29:01 +0100 wenzelm conjunction_tac: single goal;
Thu, 22 Dec 2005 00:29:00 +0100 wenzelm renamed imp_cong' to imp_cong_rule;
Thu, 22 Dec 2005 00:28:58 +0100 wenzelm mk_conjunction: proper treatment of bounds;
Thu, 22 Dec 2005 00:28:54 +0100 wenzelm fixed has_internal;
Thu, 22 Dec 2005 00:28:53 +0100 wenzelm Tactic.precise_conjunction_tac;
Thu, 22 Dec 2005 00:28:52 +0100 wenzelm added locale meta_conjunction_syntax and various conjunction rules;
Thu, 22 Dec 2005 00:28:51 +0100 wenzelm simplified setup: removed dest_concls, local_impI, conjI;
Thu, 22 Dec 2005 00:28:49 +0100 wenzelm induct_rulify;
Thu, 22 Dec 2005 00:28:47 +0100 wenzelm actually produce projected rules;
Thu, 22 Dec 2005 00:28:46 +0100 wenzelm *.inducts holds all projected rules;
Thu, 22 Dec 2005 00:28:44 +0100 wenzelm tuned;
Thu, 22 Dec 2005 00:28:43 +0100 wenzelm tuned induct proofs;
Thu, 22 Dec 2005 00:28:41 +0100 wenzelm mutual induct with *.inducts;
Thu, 22 Dec 2005 00:28:39 +0100 wenzelm wf_induct_rule: consumes 1;
Thu, 22 Dec 2005 00:28:36 +0100 wenzelm structure ProjectRule;
Thu, 22 Dec 2005 00:28:34 +0100 wenzelm updated auxiliary facts for induct method;
Wed, 21 Dec 2005 17:00:57 +0100 haftmann slight improvements
Wed, 21 Dec 2005 15:19:16 +0100 haftmann slight improvements in name handling
Wed, 21 Dec 2005 15:18:57 +0100 haftmann slight refinements
Wed, 21 Dec 2005 15:18:36 +0100 haftmann added eq_ord
Wed, 21 Dec 2005 15:18:17 +0100 haftmann slight clean ups
Wed, 21 Dec 2005 13:25:20 +0100 haftmann discontinued unflat in favour of burrow and burrow_split
Wed, 21 Dec 2005 12:06:08 +0100 paulson new hash table module in HOL/Too/s
Wed, 21 Dec 2005 12:05:47 +0100 paulson modified suffix for [iff] attribute
Wed, 21 Dec 2005 12:02:57 +0100 paulson removed or modified some instances of [iff]
Tue, 20 Dec 2005 22:06:00 +0100 wenzelm tuned;
Tue, 20 Dec 2005 09:02:41 +0100 haftmann added .cvsignore
Tue, 20 Dec 2005 09:01:48 +0100 haftmann removed improper .cvsignore
Tue, 20 Dec 2005 08:58:36 +0100 haftmann removed superfluos is_prefix functions
Tue, 20 Dec 2005 08:38:43 +0100 haftmann resolved shadowing of Library.find_first
Tue, 20 Dec 2005 08:38:10 +0100 haftmann removed infix prefix, introduces burrow
Tue, 20 Dec 2005 04:29:25 +0100 mengj Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
Tue, 20 Dec 2005 04:27:46 +0100 mengj Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
Mon, 19 Dec 2005 16:07:19 +0100 urbanc made the changes according to Florian's re-arranging of
Mon, 19 Dec 2005 15:29:51 +0100 urbanc added proofs to show that every atom-kind combination
Mon, 19 Dec 2005 12:58:15 +0100 urbanc added thms to perm_compose (so far only composition
Mon, 19 Dec 2005 12:09:56 +0100 urbanc tuned one comment
Mon, 19 Dec 2005 12:08:16 +0100 urbanc fixed a bug that occured when more than one atom-type
Mon, 19 Dec 2005 09:19:08 +0100 nipkow fixed proof
Sun, 18 Dec 2005 20:10:15 +0100 urbanc more cleaning up - this time of the cp-instance
Sun, 18 Dec 2005 14:36:42 +0100 urbanc improved the finite-support proof
Sun, 18 Dec 2005 13:38:06 +0100 urbanc improved the code for showing that a type is
Sat, 17 Dec 2005 01:58:41 +0100 wenzelm simp del: empty_Collect_eq;
Sat, 17 Dec 2005 01:00:40 +0100 wenzelm sort_distinct;
Sat, 17 Dec 2005 01:00:38 +0100 wenzelm added sort_distinct;
Fri, 16 Dec 2005 18:22:58 +0100 urbanc added container-lemma fresh_eqvt
Fri, 16 Dec 2005 18:20:59 +0100 urbanc I think the earlier version was completely broken
Fri, 16 Dec 2005 18:20:03 +0100 urbanc tuned more proofs
Fri, 16 Dec 2005 16:59:32 +0100 nipkow new lemmas
Fri, 16 Dec 2005 16:00:58 +0100 haftmann re-arranged tuples (theory * 'a) to ('a * theory) in Pure
Fri, 16 Dec 2005 15:52:05 +0100 haftmann re-arranged tuples (theory * 'a) to ('a * theory) in Pure
Fri, 16 Dec 2005 12:15:54 +0100 paulson hashing to eliminate the output of duplicate clauses
Fri, 16 Dec 2005 11:51:24 +0100 paulson hash tables from SML/NJ
Fri, 16 Dec 2005 09:00:11 +0100 haftmann re-arranged tuples (theory * 'a) to ('a * theory) in Pure
Thu, 15 Dec 2005 21:51:31 +0100 urbanc there was a small error in the last commit - fixed now
Thu, 15 Dec 2005 21:49:14 +0100 urbanc tuned more proof and added in-file documentation
Thu, 15 Dec 2005 19:42:03 +0100 wenzelm improved proofs;
Thu, 15 Dec 2005 19:42:02 +0100 wenzelm acc_induct: proper tags;
Thu, 15 Dec 2005 19:42:00 +0100 wenzelm removed obsolete/unused setup_induction;
Thu, 15 Dec 2005 18:13:40 +0100 urbanc tuned the proof of transitivity/narrowing
Thu, 15 Dec 2005 15:26:37 +0100 paulson No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
Thu, 15 Dec 2005 11:53:38 +0100 urbanc made further tunings
Thu, 15 Dec 2005 05:47:26 +0100 mengj Added functions to test equivalence between clauses.
Wed, 14 Dec 2005 22:05:22 +0100 webertj ex/Sudoku.thy
Wed, 14 Dec 2005 18:01:50 +0100 paulson deleted redundant (looping!) simprule
Wed, 14 Dec 2005 16:14:41 +0100 paulson modified example for new clauses
Wed, 14 Dec 2005 16:14:26 +0100 paulson removal of some redundancies (e.g. one-point-rules) in clause production
Wed, 14 Dec 2005 16:13:09 +0100 paulson removed unused function repeat_RS
Wed, 14 Dec 2005 06:19:33 +0100 mengj Changed literals' ordering and the functions for sorting literals.
Wed, 14 Dec 2005 01:40:43 +0100 mengj 1. changed fol_type, it's not a string type anymore.
Wed, 14 Dec 2005 01:39:41 +0100 mengj changed ATP input files' names and location.
Tue, 13 Dec 2005 19:32:36 +0100 wenzelm Potentially infinite lists as greatest fixed-point.
Tue, 13 Dec 2005 19:32:06 +0100 wenzelm Provers/induct: coinduct;
Tue, 13 Dec 2005 19:32:05 +0100 wenzelm tuned proofs;
Tue, 13 Dec 2005 19:32:04 +0100 wenzelm added HOL/Library/Coinductive_List.thy;
Tue, 13 Dec 2005 18:11:21 +0100 urbanc added a fresh_left lemma that contains all instantiation
Tue, 13 Dec 2005 16:30:50 +0100 urbanc initial commit (not to be seen by the public)
Tue, 13 Dec 2005 16:25:10 +0100 chaieb simpset for computation in raw_arith_tac added just as comment, nothing changed!
Tue, 13 Dec 2005 16:24:12 +0100 chaieb deals with Suc in mod expressions
Tue, 13 Dec 2005 15:46:41 +0100 wenzelm Poplog/pml provides a proper print function already!
Tue, 13 Dec 2005 15:34:21 +0100 paulson meson no longer does these examples
Tue, 13 Dec 2005 15:34:02 +0100 paulson now generates the name "append"
Tue, 13 Dec 2005 15:27:43 +0100 paulson removal of functional reflexivity axioms
Tue, 13 Dec 2005 10:39:32 +0100 berghofe list_of_indset now also generates code for set type.
Mon, 12 Dec 2005 17:24:06 +0100 haftmann added generic name mangler
Mon, 12 Dec 2005 15:37:35 +0100 haftmann improvement in eq handling
Mon, 12 Dec 2005 15:37:05 +0100 haftmann improvements in class and eq handling
Mon, 12 Dec 2005 15:36:46 +0100 haftmann added dummy 'print' to non-polyml systems
Sun, 11 Dec 2005 11:57:01 +0100 urbanc ISAR-fied some proofs
Sun, 11 Dec 2005 01:21:26 +0100 urbanc completed the sn proof and changed the manual
Sat, 10 Dec 2005 00:11:35 +0100 urbanc changed the types in accordance with Florian's changes
Fri, 09 Dec 2005 15:25:52 +0100 haftmann substantial improvements for class code generation
Fri, 09 Dec 2005 15:25:29 +0100 haftmann improved extraction interface
Fri, 09 Dec 2005 12:38:49 +0100 urbanc tuned
Fri, 09 Dec 2005 09:06:45 +0100 haftmann oriented result pairs in PureThy
Thu, 08 Dec 2005 20:16:17 +0100 wenzelm tuned;
Thu, 08 Dec 2005 20:16:10 +0100 wenzelm removed Syntax.deskolem;
Thu, 08 Dec 2005 20:16:04 +0100 wenzelm swap: no longer pervasive;
Thu, 08 Dec 2005 20:15:57 +0100 wenzelm replaced swap by contrapos_np;
Thu, 08 Dec 2005 20:15:50 +0100 wenzelm tuned proofs;
Thu, 08 Dec 2005 20:15:41 +0100 wenzelm Cla.swap;
Thu, 08 Dec 2005 20:15:34 +0100 wenzelm removed hint for Classical.swap, which is not really user-level anyway;
Thu, 08 Dec 2005 12:50:04 +0100 wenzelm tuned sources and proofs
Thu, 08 Dec 2005 12:50:03 +0100 wenzelm Classical.swap;
Thu, 08 Dec 2005 12:50:02 +0100 wenzelm * HOL: Theorem 'swap' is no longer bound at the ML top-level.
Thu, 08 Dec 2005 10:17:21 +0100 berghofe Adapted to new type of PureThy.add_defs_i.
Wed, 07 Dec 2005 16:47:04 +0100 wenzelm replaced swap by Classical.swap;
Wed, 07 Dec 2005 15:23:22 +0100 wenzelm avoid unportable tail;
Wed, 07 Dec 2005 12:33:38 +0100 urbanc tuned
Tue, 06 Dec 2005 17:26:26 +0100 haftmann removed thms 'swap' and 'nth_map' from ML toplevel
Tue, 06 Dec 2005 17:11:40 +0100 haftmann improved serialization of classes to haskell
Tue, 06 Dec 2005 16:07:25 +0100 haftmann improved class handling
Tue, 06 Dec 2005 16:07:10 +0100 haftmann added 'dig' combinator
Tue, 06 Dec 2005 09:04:09 +0100 haftmann re-oriented some result tuples in PureThy
Tue, 06 Dec 2005 06:22:14 +0100 mengj Added more functions for new type embedding of HOL clauses.
Tue, 06 Dec 2005 06:21:07 +0100 mengj Added new type embedding methods for translating HOL clauses.
Mon, 05 Dec 2005 18:19:49 +0100 urbanc added code to say that discrete types (nat, bool, char)
Mon, 05 Dec 2005 17:02:20 +0100 urbanc tuned
Mon, 05 Dec 2005 15:55:19 +0100 urbanc transitivity should be now in a reasonable state. But
Mon, 05 Dec 2005 10:33:30 +0100 urbanc tuned
Mon, 05 Dec 2005 10:32:37 +0100 urbanc ISAR-fied two proofs
Mon, 05 Dec 2005 00:39:18 +0100 berghofe Adapted to new type of store_thmss(_atts).
Mon, 05 Dec 2005 00:38:07 +0100 berghofe Added store_thmss_atts to signature again.
Sun, 04 Dec 2005 12:40:39 +0100 urbanc tuned
Sun, 04 Dec 2005 12:25:57 +0100 urbanc tuned
Sun, 04 Dec 2005 12:14:40 +0100 urbanc tuned
Sun, 04 Dec 2005 12:14:27 +0100 urbanc added an Isar-proof for the abs_ALL lemma
Sun, 04 Dec 2005 12:14:03 +0100 urbanc tuning
Fri, 02 Dec 2005 22:57:36 +0100 wenzelm defines: beta/eta contract lhs;
Fri, 02 Dec 2005 22:54:52 +0100 wenzelm asts_to_terms: builtin free_fixed translation makes local binders work properly;
Fri, 02 Dec 2005 22:54:51 +0100 wenzelm mixfix_args: 1 for binders;
Fri, 02 Dec 2005 22:54:50 +0100 wenzelm removed fixed_tr: now handled in syntax module;
Fri, 02 Dec 2005 22:54:48 +0100 wenzelm added mixfixT;
Fri, 02 Dec 2005 22:54:47 +0100 wenzelm defs: beta/eta contract lhs;
Fri, 02 Dec 2005 22:54:45 +0100 wenzelm abs_def: beta/eta contract lhs;
Fri, 02 Dec 2005 16:43:42 +0100 krauss Added recdef congruence rules for bounded quantifiers and commonly used
Fri, 02 Dec 2005 16:05:31 +0100 haftmann various improvements
Fri, 02 Dec 2005 16:05:12 +0100 haftmann adjusted to improved code generator interface
Fri, 02 Dec 2005 16:04:48 +0100 haftmann added perhaps option combinator
Fri, 02 Dec 2005 16:04:29 +0100 haftmann adopted keyword for code generator
Fri, 02 Dec 2005 13:10:12 +0100 berghofe Factored out proof for normalization of applications (norm_list).
Fri, 02 Dec 2005 08:06:59 +0100 haftmann introduced new map2, fold
Thu, 01 Dec 2005 22:43:15 +0100 kleing typo
Thu, 01 Dec 2005 22:04:27 +0100 wenzelm simprocs: static evaluation of simpset;
Thu, 01 Dec 2005 22:03:06 +0100 wenzelm tuned;
Thu, 01 Dec 2005 22:03:05 +0100 wenzelm tuned msg;
Thu, 01 Dec 2005 22:03:04 +0100 wenzelm ProofContext.lthms_containing: suppress obvious duplicates;
Thu, 01 Dec 2005 22:03:01 +0100 wenzelm unfold_tac: static evaluation of simpset;
Thu, 01 Dec 2005 18:45:24 +0100 wenzelm superceded by timestart|stop.bash;
Thu, 01 Dec 2005 18:44:47 +0100 wenzelm cpu time = user + system;
Thu, 01 Dec 2005 18:41:46 +0100 wenzelm replaced lib/scripts/showtime by more advanced lib/scripts/timestart|stop.bash;
Thu, 01 Dec 2005 18:41:44 +0100 berghofe assoc_consts and assoc_types now check number of arguments in template.
Thu, 01 Dec 2005 18:39:08 +0100 berghofe Added new component "sorts" to record datatype_info.
Thu, 01 Dec 2005 18:37:39 +0100 wenzelm timestop - report timing based on environment (cf. timestart.bash);
Thu, 01 Dec 2005 18:37:22 +0100 wenzelm timestart - setup bash environment for timing;
Thu, 01 Dec 2005 17:07:50 +0100 berghofe Improved norm_proof to handle proofs containing term (type) variables
Thu, 01 Dec 2005 15:45:54 +0100 paulson restoring the old status of subset_refl
Thu, 01 Dec 2005 08:28:02 +0100 haftmann oriented pairs theory * 'a to 'a * theory
Thu, 01 Dec 2005 06:28:41 +0100 urbanc initial cleanup to use the new induction method
Thu, 01 Dec 2005 05:20:13 +0100 urbanc cleaned up further the proofs (diamond still needs work);
Thu, 01 Dec 2005 04:46:17 +0100 urbanc changed "fresh:" to "avoiding:" and cleaned up the weakening example
Wed, 30 Nov 2005 22:52:50 +0100 wenzelm match_bind(_i): return terms;
Wed, 30 Nov 2005 22:52:49 +0100 wenzelm method 'fact': SIMPLE_METHOD, i.e. insert facts;
Wed, 30 Nov 2005 22:52:46 +0100 wenzelm simulaneous 'def';
Wed, 30 Nov 2005 21:51:23 +0100 urbanc added facilities to prove the pt and fs instances
Wed, 30 Nov 2005 19:08:51 +0100 urbanc started to change the transitivity/narrowing case:
Wed, 30 Nov 2005 18:37:12 +0100 urbanc changed everything until the interesting transitivity_narrowing
Wed, 30 Nov 2005 18:13:31 +0100 haftmann minor improvements
Wed, 30 Nov 2005 17:56:08 +0100 urbanc modified almost everything for the new nominal_induct
Wed, 30 Nov 2005 16:59:19 +0100 berghofe Changed order of predicate arguments and quantifiers in strong induction rule.
Wed, 30 Nov 2005 15:30:08 +0100 urbanc fixed the lemma where the new names generated by nominal_induct
Wed, 30 Nov 2005 15:27:30 +0100 urbanc added one clause for substitution in the lambda-case and
Wed, 30 Nov 2005 15:24:32 +0100 wenzelm added rename_params_rule: recover orginal fresh names in subgoals/cases;
Wed, 30 Nov 2005 15:03:15 +0100 urbanc changed induction principle and everything to conform with the
Wed, 30 Nov 2005 14:27:50 +0100 wenzelm fresh: frees instead of terms, rename corresponding params in rule;
Wed, 30 Nov 2005 14:27:09 +0100 urbanc adapted to the new nominal_induction
Wed, 30 Nov 2005 12:28:47 +0100 urbanc changed \<sim> of permutation equality to \<triangleq>
Wed, 30 Nov 2005 12:23:35 +0100 wenzelm fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
Wed, 30 Nov 2005 01:01:15 +0100 huffman reimplement Case expression pattern matching to support lazy patterns
Wed, 30 Nov 2005 00:59:04 +0100 huffman add definitions as defs, not axioms
Wed, 30 Nov 2005 00:56:01 +0100 huffman changed section names
Wed, 30 Nov 2005 00:55:24 +0100 huffman add xsymbol syntax for u type constructor
Wed, 30 Nov 2005 00:53:30 +0100 huffman add constant unit_when
Wed, 30 Nov 2005 00:46:08 +0100 wenzelm proper treatment of tuple/tuple_fun -- nest to the left!
Tue, 29 Nov 2005 23:00:20 +0100 wenzelm moved nth_list to Pure/library.ML;
Tue, 29 Nov 2005 23:00:03 +0100 wenzelm added nth_list;
Tue, 29 Nov 2005 22:52:19 +0100 wenzelm added mk_split;
Tue, 29 Nov 2005 19:26:38 +0100 urbanc changed the order of the induction variable and the context
Tue, 29 Nov 2005 18:09:12 +0100 wenzelm reworked version with proper support for defs, fixes, fresh specification;
Tue, 29 Nov 2005 16:05:10 +0100 haftmann added haskell serializer
Tue, 29 Nov 2005 16:04:57 +0100 haftmann exported customized syntax interface
Tue, 29 Nov 2005 12:26:22 +0100 berghofe Corrected atom class constraints in strong induction rule.
Tue, 29 Nov 2005 01:37:01 +0100 urbanc made some of the theorem look-ups static (by using
Mon, 28 Nov 2005 13:43:56 +0100 haftmann added (curried) fold2
Mon, 28 Nov 2005 10:58:40 +0100 wenzelm added proof of measure_induct_rule;
Mon, 28 Nov 2005 07:17:07 +0100 mengj Added flags for setting and detecting whether a problem needs combinators.
Mon, 28 Nov 2005 07:16:16 +0100 mengj Only output types if !keep_types is true.
Mon, 28 Nov 2005 07:15:38 +0100 mengj Added two functions for CNF translation, used by other files.
(0) -10000 -3000 -1000 -240 +240 +1000 +3000 +10000 +30000 tip