src/Pure/thm.ML
Tue, 21 Jul 2009 10:24:57 +0200 wenzelm prefer simultaneous join -- for improved scheduling;
Sun, 19 Jul 2009 18:42:05 +0200 wenzelm support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
Fri, 17 Jul 2009 23:11:40 +0200 wenzelm tuned/modernized Envir.subst_XXX;
Fri, 17 Jul 2009 21:33:00 +0200 wenzelm tuned/modernized Envir operations;
Thu, 16 Jul 2009 22:58:45 +0200 wenzelm tuned incr_indexes;
Thu, 09 Jul 2009 22:48:12 +0200 wenzelm renamed structure TermSubst to Term_Subst;
Mon, 06 Jul 2009 22:42:27 +0200 wenzelm clarified strip_shyps: proper type witnesses for present sorts;
Mon, 06 Jul 2009 21:24:30 +0200 wenzelm structure Thm: less pervasive names;
Mon, 06 Jul 2009 20:36:38 +0200 wenzelm clarified Thm.of_class/of_sort/class_triv;
Mon, 06 Jul 2009 19:58:52 +0200 wenzelm renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
Thu, 02 Jul 2009 22:17:08 +0200 wenzelm strip_shyps: remove top sort, which is logically insignificant;
Thu, 02 Jul 2009 20:55:44 +0200 wenzelm added pro-forma proof constructor Inclass;
Wed, 25 Mar 2009 16:54:17 +0100 wenzelm fulfill_proof/thm_proof: pass whole proof_body, not just the projection to proof (which might be incomplete);
Tue, 24 Mar 2009 22:56:17 +0100 wenzelm status_of: need to include local promises as well!
Tue, 24 Mar 2009 21:24:53 +0100 wenzelm display derivation status of thms;
Tue, 17 Mar 2009 12:10:42 +0100 wenzelm eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
Mon, 16 Mar 2009 23:52:30 +0100 wenzelm substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
Thu, 12 Mar 2009 11:10:02 +0100 wenzelm renamed NameSpace.bind to NameSpace.define;
Sat, 07 Mar 2009 22:04:59 +0100 wenzelm moved Thm.def_name(_optional) to more_thm.ML;
Thu, 05 Mar 2009 19:48:02 +0100 wenzelm Thm.add_oracle interface: replaced old bstring by binding;
Tue, 27 Jan 2009 00:42:12 +0100 wenzelm thm_proof: replaced lazy by composed futures;
Sat, 10 Jan 2009 21:47:02 +0100 wenzelm future_result: avoid expensive norm_proof (consider usedir option -Q false in low-memory situations);
Sat, 10 Jan 2009 16:54:55 +0100 wenzelm added pending_groups -- accumulates task groups of local derivations only;
Thu, 01 Jan 2009 14:23:39 +0100 wenzelm avoid polymorphic equality;
Wed, 31 Dec 2008 18:53:17 +0100 wenzelm use regular Term.add_XXX etc.;
Wed, 31 Dec 2008 15:30:10 +0100 wenzelm moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
Sat, 06 Dec 2008 00:09:01 +0100 wenzelm renamed force_proof to join_proof;
Fri, 05 Dec 2008 20:38:40 +0100 wenzelm refined type deriv: replaced all_promises by max_promise (dependency limit) and open_promises (potentially unfinished/failed promises);
Fri, 05 Dec 2008 08:05:14 +0100 haftmann merged
Fri, 05 Dec 2008 08:04:53 +0100 haftmann dropped NameSpace.declare_base
Fri, 05 Dec 2008 00:23:37 +0100 wenzelm merged
Thu, 04 Dec 2008 23:02:56 +0100 wenzelm future proofs: pass actual futures to facilitate composite computations;
Thu, 04 Dec 2008 14:43:33 +0100 haftmann cleaned up binding module and related code
Sun, 23 Nov 2008 17:25:56 +0100 wenzelm future: norm_proof after make_result reduces memory requirements but increases runtime (factor 1.5 for both for HOL with proof terms);
Tue, 18 Nov 2008 22:25:36 +0100 wenzelm fulfill_proof/thm_proof: commuted lazyness;
Tue, 18 Nov 2008 18:25:52 +0100 wenzelm added force_proofs;
Mon, 17 Nov 2008 23:17:13 +0100 wenzelm tuned promise/fullfill;
Sun, 16 Nov 2008 22:12:43 +0100 wenzelm put_name/thm_proof: promises are filled with fulfilled proofs;
Sun, 16 Nov 2008 20:03:42 +0100 wenzelm clarified Thm.proof_body_of vs. Thm.proof_of;
Sat, 15 Nov 2008 21:31:25 +0100 wenzelm refined notion of derivation, consiting of promises and proof_body;
Thu, 23 Oct 2008 15:28:05 +0200 wenzelm renamed get_axiom_i to axiom, removed obsolete get_axiom;
Tue, 21 Oct 2008 16:53:10 +0200 wenzelm join results in isolation;
Thu, 16 Oct 2008 22:44:33 +0200 wenzelm added weaken_sorts;
Wed, 01 Oct 2008 12:18:18 +0200 wenzelm renamed promise to future, tuned related interfaces;
Wed, 01 Oct 2008 12:00:00 +0200 wenzelm more precise join_futures, improved termination;
Tue, 30 Sep 2008 22:02:44 +0200 wenzelm export explicit joint_futures, removed Theory.at_end hook;
Sun, 28 Sep 2008 12:42:35 +0200 wenzelm join earlier promises first;
Sun, 28 Sep 2008 12:23:44 +0200 wenzelm promise_result: enforce strictly sequential dependencies, via serial numbers;
Sat, 27 Sep 2008 18:18:05 +0200 wenzelm Future.release_results;
Sat, 27 Sep 2008 15:20:37 +0200 wenzelm promise: include check into future body, i.e. joined results are always valid;
Thu, 25 Sep 2008 20:34:19 +0200 wenzelm simplified promise;
Thu, 25 Sep 2008 14:35:02 +0200 wenzelm abtract types: plain datatype with opaque signature matching;
Thu, 25 Sep 2008 13:21:13 +0200 wenzelm explicit type OrdList.T;
Tue, 23 Sep 2008 15:48:52 +0200 wenzelm added dest_deriv, removed external type deriv;
Mon, 22 Sep 2008 15:26:14 +0200 wenzelm type thm: fully internal derivation, no longer exported;
Thu, 18 Sep 2008 19:39:44 +0200 wenzelm simplified oracle interface;
Thu, 18 Sep 2008 14:06:56 +0200 wenzelm added deriv.ML: Abstract derivations based on raw proof terms.
Wed, 27 Aug 2008 11:48:54 +0200 wenzelm type Properties.T;
Mon, 23 Jun 2008 23:45:49 +0200 wenzelm Logic.all/mk_equals/mk_implies;
Sun, 18 May 2008 15:04:09 +0200 wenzelm moved global pretty/string_of functions from Sign to Syntax;
Wed, 07 May 2008 10:59:47 +0200 berghofe eq_assumption now uses aeconv instead of aconv.
Tue, 15 Apr 2008 18:49:19 +0200 wenzelm Theory.eq_thy;
Sun, 13 Apr 2008 16:40:06 +0200 wenzelm simplified handling of sorts, removed unnecessary universal witness;
Sat, 12 Apr 2008 17:00:45 +0200 wenzelm rep_cterm/rep_thm: no longer dereference theory_ref;
Thu, 27 Mar 2008 14:41:09 +0100 wenzelm eliminated theory ProtoPure;
Mon, 21 Jan 2008 14:18:49 +0100 berghofe Removed Logic.auto_rename.
Thu, 04 Oct 2007 20:29:24 +0200 wenzelm replaced literal 'a by Name.aT;
Fri, 17 Aug 2007 23:10:45 +0200 wenzelm compress: proper check_thy;
Fri, 03 Aug 2007 16:28:22 +0200 wenzelm replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
Wed, 11 Jul 2007 11:59:21 +0200 berghofe Added function norm_proof for normalizing the proof term
Sun, 08 Jul 2007 19:52:05 +0200 wenzelm thm tag: Markup.property list;
Thu, 05 Jul 2007 20:01:38 +0200 wenzelm added type conv;
Mon, 25 Jun 2007 00:36:42 +0200 wenzelm added eta_long_conversion;
Fri, 08 Jun 2007 18:13:58 +0200 berghofe Adapted Proofterm.bicompose_proof to Larry's changes in
Thu, 31 May 2007 23:47:36 +0200 wenzelm simplified/unified list fold;
Thu, 10 May 2007 00:39:55 +0200 wenzelm added dest_fun/fun2/arg1;
Sat, 14 Apr 2007 17:36:17 +0200 wenzelm removed obsolete read_ctyp, read_def_cterm;
Wed, 04 Apr 2007 23:29:33 +0200 wenzelm rep_thm/cterm/ctyp: removed obsolete sign field;
Wed, 04 Apr 2007 00:11:16 +0200 wenzelm improved exception CTERM;
Tue, 03 Apr 2007 19:24:19 +0200 wenzelm avoid clash with Alice keywords;
Mon, 26 Feb 2007 23:18:30 +0100 wenzelm moved some non-kernel material to more_thm.ML;
Sun, 04 Feb 2007 22:02:17 +0100 wenzelm old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
Tue, 02 Jan 2007 22:43:04 +0100 wenzelm Term.lambda: abstract over arbitrary closed terms;
Tue, 12 Dec 2006 20:49:25 +0100 wenzelm tuned error messages;
Tue, 05 Dec 2006 00:30:38 +0100 wenzelm thm/prf: separate official name vs. additional tags;
Wed, 29 Nov 2006 04:11:09 +0100 wenzelm simplified Logic.count_prems;
Tue, 21 Nov 2006 18:07:33 +0100 wenzelm moved theorem kinds from PureThy to Thm;
Sun, 05 Nov 2006 21:44:35 +0100 wenzelm removed obsolete first_duplicate;
Tue, 31 Oct 2006 09:29:06 +0100 haftmann fixed type signature of Type.varify
Wed, 11 Oct 2006 10:49:36 +0200 haftmann abandoned findrep
Sat, 07 Oct 2006 01:31:12 +0200 wenzelm added def_name_optional;
Sun, 01 Oct 2006 18:29:34 +0200 wenzelm cterm_match: avoid recalculation of maxidx;
Thu, 21 Sep 2006 19:05:31 +0200 wenzelm added dest_binop;
Mon, 18 Sep 2006 19:39:11 +0200 wenzelm added dest_arg, i.e. a tuned version of #2 o dest_comb;
Fri, 15 Sep 2006 22:56:13 +0200 wenzelm renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
Fri, 15 Sep 2006 20:08:37 +0200 wenzelm instantiate: omit has_duplicates check, which is irrelevant for soundness;
Tue, 12 Sep 2006 12:12:57 +0200 wenzelm ctyp: maintain maxidx;
Tue, 08 Aug 2006 08:18:59 +0200 haftmann abandoned equal_list in favor for eq_list
Thu, 03 Aug 2006 17:30:38 +0200 wenzelm tuned;
Wed, 02 Aug 2006 22:26:41 +0200 wenzelm normalized Proof.context/method type aliases;
Sun, 30 Jul 2006 21:28:54 +0200 wenzelm adjust_maxidx: pass explicit lower bound;
Tue, 11 Jul 2006 12:16:54 +0200 wenzelm replaced Term.variant(list) by Name.variant(_list);
Sat, 08 Jul 2006 12:54:45 +0200 wenzelm tuned exception handling;
Tue, 04 Jul 2006 19:49:53 +0200 wenzelm varifyT: no longer pervasive;
Sat, 17 Jun 2006 19:37:51 +0200 wenzelm added generalize rule;
Tue, 13 Jun 2006 23:41:52 +0200 wenzelm added hyps_of;
Mon, 12 Jun 2006 21:19:00 +0200 wenzelm tuned Seq/Envir/Unify interfaces;
Mon, 01 May 2006 17:05:13 +0200 wenzelm class_triv: Sign.certify_class;
Sat, 29 Apr 2006 23:16:47 +0200 wenzelm added unconstrainT;
Wed, 26 Apr 2006 22:38:16 +0200 wenzelm curried Seq.cons;
Thu, 13 Apr 2006 12:01:07 +0200 wenzelm added maxidx_of;
Tue, 21 Mar 2006 12:18:15 +0100 wenzelm avoid polymorphic equality;
Fri, 10 Mar 2006 04:03:48 +0100 mengj Shortened the exception messages from assume.
Sat, 11 Feb 2006 17:17:47 +0100 wenzelm tuned;
Tue, 07 Feb 2006 19:56:51 +0100 wenzelm removed obsolete sign_of_cterm;
Mon, 06 Feb 2006 20:59:11 +0100 wenzelm union_tpairs: Library.merge;
Sat, 21 Jan 2006 23:02:25 +0100 wenzelm simplified type attribute;
Fri, 23 Dec 2005 15:16:52 +0100 wenzelm turned bicompose_no_flatten into compose_no_flatten, without elimination;
Thu, 22 Dec 2005 00:29:20 +0100 wenzelm added bicompose_no_flatten, which refrains from
Fri, 16 Dec 2005 09:00:11 +0100 haftmann re-arranged tuples (theory * 'a) to ('a * theory) in Pure
Wed, 16 Nov 2005 17:45:30 +0100 wenzelm tuned Pattern.match/unify;
Thu, 10 Nov 2005 20:57:11 +0100 wenzelm renamed Thm.cgoal_of to Thm.cprem_of;
Wed, 09 Nov 2005 16:26:41 +0100 wenzelm Thm.varifyT': natural argument order;
Fri, 28 Oct 2005 22:28:04 +0200 wenzelm added cgoal_of;
Sat, 15 Oct 2005 00:09:20 +0200 wenzelm tuned comment;
Thu, 29 Sep 2005 00:59:01 +0200 wenzelm abstract_rule: tuned exception msgs;
Thu, 15 Sep 2005 17:16:56 +0200 wenzelm TableFun/Symtab: curried lookup and update;
Tue, 13 Sep 2005 22:19:29 +0200 wenzelm added simple_fact;
Thu, 01 Sep 2005 18:48:50 +0200 wenzelm curried_lookup/update;
Wed, 31 Aug 2005 15:46:40 +0200 wenzelm refer to theory instead of low-level tsig;
Mon, 29 Aug 2005 16:18:04 +0200 wenzelm use AList operations;
Mon, 01 Aug 2005 19:20:45 +0200 wenzelm Compress.term;
Thu, 28 Jul 2005 15:20:02 +0200 wenzelm added weaken, adjust_maxidx_thm;
Tue, 19 Jul 2005 17:21:56 +0200 wenzelm tuned instantiate (avoid subst_atomic, subst_atomic_types);
Thu, 14 Jul 2005 19:28:31 +0200 wenzelm invoke_oracle: do not keep theory value, but theory_ref;
Wed, 06 Jul 2005 20:00:34 +0200 wenzelm added full_prop_of: includes tpairs;
Wed, 06 Jul 2005 10:41:45 +0200 wenzelm tuned maxidx;
Mon, 04 Jul 2005 17:07:13 +0200 wenzelm dest_ctyp: raise exception for non-constructor;
Fri, 01 Jul 2005 14:42:00 +0200 wenzelm ctyp: added 'sorts' field;
Wed, 29 Jun 2005 15:13:31 +0200 wenzelm more efficient treatment of shyps and hyps (use ordered lists);
Fri, 17 Jun 2005 18:33:08 +0200 wenzelm accomodate identification of type Sign.sg and theory;
Thu, 09 Jun 2005 12:04:53 +0200 wenzelm axioms: NameSpace.table;
Sun, 05 Jun 2005 23:07:25 +0200 wenzelm Type.freeze;
Tue, 31 May 2005 11:53:26 +0200 wenzelm added eq_thms;
Sun, 22 May 2005 16:51:12 +0200 wenzelm tuned terms_of_tpairs;
Thu, 21 Apr 2005 19:12:03 +0200 berghofe - Eliminated nodup_vars check.
Thu, 07 Apr 2005 09:28:03 +0200 wenzelm added get_axiom_i, invoke_oracle_i;
Fri, 04 Mar 2005 15:07:34 +0100 skalberg Removed practically all references to Library.foldr.
Thu, 03 Mar 2005 12:43:01 +0100 skalberg Move towards standard functions.
Sun, 13 Feb 2005 17:15:14 +0100 skalberg Deleted Library.option type.
Mon, 24 Jan 2005 12:41:06 +0100 paulson thin_tac now works on P==>Q
Tue, 26 Oct 2004 16:34:19 +0200 berghofe Changed function cabs to also allow abstraction over Vars.
Thu, 29 Jul 2004 17:45:21 +0200 berghofe - optimized nodup_vars check in capply
Sat, 29 May 2004 15:03:59 +0200 wenzelm improved output; refer to Pretty.pp;
Fri, 21 May 2004 21:27:10 +0200 wenzelm adapted names of some sort ops;
Mon, 21 Oct 2002 17:04:47 +0200 berghofe Changed handling of flex-flex constraints: now stored in separate
Thu, 10 Oct 2002 19:02:23 +0200 nipkow added failure trace information to pattern unification
Mon, 07 Oct 2002 19:01:51 +0200 nipkow take/drop -> splitAt
Tue, 27 Aug 2002 11:05:31 +0200 wenzelm added proof_of;
Thu, 28 Feb 2002 19:24:00 +0100 wenzelm moved match_bvs, match_bvars, renAbs to term.ML;
Thu, 21 Feb 2002 20:11:05 +0100 wenzelm fixed get_name_tags in order to work with hyps;
Thu, 17 Jan 2002 21:04:16 +0100 wenzelm added prop_of: thm -> term (at last!);
Fri, 14 Dec 2001 11:54:13 +0100 wenzelm varifyT' returns newly introduces variables;
Thu, 04 Oct 2001 23:27:01 +0200 wenzelm major_prem_of: Logic.strip_assums_concl;
Fri, 28 Sep 2001 11:04:44 +0200 berghofe Exchanged % and %%.
Thu, 13 Sep 2001 16:26:16 +0200 berghofe Fixed proof term bug in permute_prems.
Fri, 31 Aug 2001 16:13:00 +0200 berghofe Replaced old derivations by proof terms.
Wed, 03 Jan 2001 21:18:31 +0100 wenzelm Thm: dest_comb, dest_abs, capply, cabs no longer global;
Fri, 17 Nov 2000 18:50:01 +0100 wenzelm Envir.beta_norm;
Tue, 07 Nov 2000 17:52:12 +0100 berghofe - Moved rewriting functions to meta_simplifier.ML
Fri, 27 Oct 2000 16:25:21 +0200 wenzelm back to 1.167, due to Emacs/CVS casualty!!;
Fri, 27 Oct 2000 15:23:39 +0200 wenzelm *** empty log message ***
Thu, 07 Sep 2000 15:31:09 +0200 nipkow Added meaningful output to cong-error msg.
Tue, 29 Aug 2000 15:13:10 +0200 nipkow *** empty log message ***
Wed, 16 Aug 2000 18:10:15 +0200 nipkow Fixed completeness bug in simplifier: congruence rules could preclude
Wed, 02 Aug 2000 19:39:48 +0200 wenzelm derivations: maintain oracle flag;
Sun, 30 Jul 2000 12:50:07 +0200 wenzelm added sign_of_cterm;
Sat, 03 Jun 2000 23:58:37 +0200 wenzelm fixed Thm.eq_thm: use Sign.joinable;
Wed, 10 May 2000 21:03:12 +0200 wenzelm dest_mss: sort procs wrt. names;
Mon, 08 May 2000 11:45:57 +0200 wenzelm tuned msg;
Thu, 30 Mar 2000 14:19:33 +0200 wenzelm read_def_cterms: use Sign.read_def_terms;
Fri, 10 Mar 2000 14:58:25 +0100 berghofe Envir now uses Vartab instead of association lists.
Sun, 27 Feb 2000 15:13:44 +0100 wenzelm added major_prem_of;
Thu, 24 Feb 2000 16:04:25 +0100 wenzelm tuned;
Thu, 24 Feb 2000 15:58:10 +0100 wenzelm capply, cabs: Sign.nodup_vars;
Mon, 17 Jan 2000 14:10:32 +0100 paulson Thm.instantiate no longer normalizes, but Drule.instantiate does
Thu, 16 Dec 1999 17:01:16 +0100 paulson SOUNDNESS BUG FIX for rotate_rule. The original code did not expect
Fri, 22 Oct 1999 20:25:00 +0200 wenzelm debug_simp;
Wed, 29 Sep 1999 14:02:33 +0200 wenzelm removed implies_intr_shyps;
Thu, 09 Sep 1999 19:01:37 +0200 wenzelm added no_prems;
Thu, 09 Sep 1999 12:25:01 +0200 wenzelm removed obsolete comment;
Mon, 23 Aug 1999 16:13:42 +0200 nipkow Now rewrite rules with flexible heads are allowed.
Mon, 23 Aug 1999 09:36:05 +0200 nipkow Corrected two busg in the simplifier.
Wed, 18 Aug 1999 20:39:41 +0200 wenzelm (*no fix_shyps*);
Wed, 18 Aug 1999 10:54:44 +0200 paulson new primitive rule permute_prems to underlie defer_tac and rotate_prems
Fri, 23 Jul 1999 16:52:45 +0200 wenzelm tuned add_term_varnames;
Thu, 08 Jul 1999 18:30:00 +0200 wenzelm improved error msgs of instantiate;
Tue, 06 Jul 1999 21:03:03 +0200 wenzelm added clear_mss;
Sat, 05 Jun 1999 20:30:29 +0200 wenzelm varifyT': observe additional 'fixed' tfrees;
Thu, 29 Apr 1999 18:33:31 +0200 nipkow Eta contraction is now performed all the time during rewriting.
Wed, 17 Mar 1999 16:33:00 +0100 wenzelm qualify Theory.sign_of etc.;
Wed, 17 Mar 1999 13:32:20 +0100 wenzelm added def_name;
Tue, 12 Jan 1999 13:39:21 +0100 wenzelm signature BASIC_THM;
Thu, 08 Oct 1998 11:59:17 +0200 nipkow Further improvement of the simplifier.
Wed, 07 Oct 1998 18:17:37 +0200 nipkow Tuned simplifier not to re-normalized already normalized terms.
Fri, 18 Sep 1998 14:32:49 +0200 paulson improved error messages
Wed, 19 Aug 1998 17:04:21 +0200 wenzelm assume: adjust_maxidx;
Wed, 19 Aug 1998 10:37:56 +0200 paulson The warning "Rewrite rule from different theory" is ALWAYS printed, even if
Fri, 05 Jun 1998 14:29:54 +0200 wenzelm Object.T;
Wed, 29 Apr 1998 11:20:53 +0200 wenzelm tuned get_ax (uses ancestry);
Wed, 22 Apr 1998 14:04:35 +0200 nipkow Tried to speed up the rewriter by eta-contracting all patterns beforehand and
Sat, 04 Apr 1998 11:40:18 +0200 wenzelm tuned trace msgs;
Thu, 12 Mar 1998 12:49:24 +0100 nipkow Made mutual simplification of prems a special case.
Tue, 10 Mar 1998 16:47:26 +0100 nipkow Asm_full_simp_tac now reorients asm c = t to t = c.
Tue, 10 Mar 1998 13:24:11 +0100 nipkow New simplifier flag for mutual simplification.
Fri, 06 Mar 1998 16:05:04 +0100 nipkow Removed superfluous `op'
Wed, 04 Mar 1998 13:16:05 +0100 nipkow Reorganized simplifier. May now reorient rules.
Sat, 28 Feb 1998 15:40:03 +0100 nipkow Tried to reorganize rewriter a little. More to be done.
Fri, 30 Jan 1998 11:33:01 +0100 wenzelm improved tracing of rewrite rule application;
Fri, 19 Dec 1997 09:58:42 +0100 wenzelm Term.termless;
Fri, 12 Dec 1997 17:14:58 +0100 wenzelm tuned msg;
Thu, 27 Nov 1997 19:36:51 +0100 wenzelm removed read_cterms;
Wed, 26 Nov 1997 16:38:04 +0100 wenzelm added crep_cterm;
Mon, 24 Nov 1997 16:43:43 +0100 nipkow Added read_def_cterms for simultaneous reading/typing of terms under
Fri, 21 Nov 1997 15:27:43 +0100 wenzelm changed Sequence interface (now Seq, in seq.ML);
Thu, 20 Nov 1997 15:06:57 +0100 wenzelm added transfer_sg;
Thu, 20 Nov 1997 12:51:55 +0100 wenzelm tuned infer_types interface;
Thu, 06 Nov 1997 16:41:08 +0100 wenzelm deriv: eliminated references to theory;
Tue, 04 Nov 1997 16:17:04 +0100 wenzelm type object = exn (enhance readability);
Tue, 04 Nov 1997 12:58:10 +0100 nipkow logic: loops -> rewrite_rule_ok
Thu, 30 Oct 1997 16:57:09 +0100 nipkow Modified trace output routines of simplifier.
Thu, 30 Oct 1997 09:59:38 +0100 wenzelm tuned simp trace;
Tue, 28 Oct 1997 17:30:47 +0100 wenzelm added name_of_thm;
Fri, 24 Oct 1997 17:15:59 +0200 wenzelm eq_thm (from drule.ML);
Thu, 23 Oct 1997 12:09:31 +0200 wenzelm tuned;
Wed, 22 Oct 1997 11:36:43 +0200 wenzelm tuned;
Tue, 21 Oct 1997 18:09:13 +0200 wenzelm sg_ref: automatic adjustment of thms of draft theories;
Thu, 16 Oct 1997 14:00:20 +0200 wenzelm added transfer: theory -> thm -> thm;
Thu, 16 Oct 1997 13:45:16 +0200 nipkow The simplifier has been improved a little: equations s=t which used to be
Thu, 09 Oct 1997 14:52:36 +0200 wenzelm fixed get_axiom, invoke_oracle;
less more (0) -240 tip