src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML
Fri, 25 Sep 2015 20:37:59 +0200 wenzelm moved remaining display.ML to more_thm.ML;
Sat, 25 Jul 2015 23:41:53 +0200 wenzelm updated to infer_instantiate;
Sat, 18 Jul 2015 20:47:08 +0200 wenzelm prefer tactics with explicit context;
Fri, 06 Mar 2015 23:53:36 +0100 wenzelm clarified context;
Fri, 06 Mar 2015 15:58:56 +0100 wenzelm Thm.cterm_of and Thm.ctyp_of operate on local context;
Wed, 04 Mar 2015 19:53:18 +0100 wenzelm tuned signature -- prefer qualified names;
Wed, 11 Feb 2015 11:42:39 +0100 wenzelm proper context;
Tue, 10 Feb 2015 14:48:26 +0100 wenzelm proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
Mon, 10 Nov 2014 21:49:48 +0100 wenzelm proper context for assume_tac (atac remains as fall-back without context);
Fri, 21 Mar 2014 10:45:03 +0100 wenzelm tuned signature;
Wed, 12 Feb 2014 13:33:05 +0100 wenzelm tuned whitespace;
Mon, 10 Feb 2014 14:33:47 +0100 wenzelm comments;
Wed, 01 Jan 2014 14:29:22 +0100 wenzelm clarified simplifier context;
Sat, 14 Dec 2013 17:28:05 +0100 wenzelm proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
Thu, 30 May 2013 12:35:40 +0200 wenzelm standardized aliases;
Fri, 24 May 2013 17:00:46 +0200 wenzelm tuned signature;
Thu, 18 Apr 2013 17:07:01 +0200 wenzelm simplifier uses proper Proof.context instead of historic type simpset;
Wed, 27 Mar 2013 14:50:30 +0100 wenzelm clarified Skip_Proof.cheat_tac: more standard tactic;
Mon, 12 Nov 2012 23:24:40 +0100 haftmann dropped dead code
Tue, 13 Mar 2012 20:04:24 +0100 wenzelm more explicit indication of def names;
Thu, 09 Jun 2011 16:34:49 +0200 wenzelm discontinued Name.variant to emphasize that this is old-style / indirect;
Sun, 15 May 2011 18:59:27 +0200 wenzelm eliminated obsolete "assert" function, including divergent (unused!?) clone in Predicate_Compile_Proof;
Sat, 16 Apr 2011 16:15:37 +0200 wenzelm modernized structure Proof_Context;
Fri, 17 Dec 2010 17:08:56 +0100 wenzelm renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
Thu, 30 Sep 2010 15:37:11 +0200 bulwahn applying case beta reduction to case term before matching in predicate compile function flattening; moving case beta reduction function to Predicate_Compile_Aux
Wed, 29 Sep 2010 10:33:15 +0200 bulwahn adding splitting of conjuncts in assumptions as forward rule on theorems; replacing term transformation for splitting conjuncts by theorem transformation; removing obsolete functions; tuned
Mon, 27 Sep 2010 12:22:57 +0200 bulwahn handling nested cases more elegant by requiring less new constants
Mon, 20 Sep 2010 16:05:25 +0200 wenzelm renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
Thu, 16 Sep 2010 13:49:17 +0200 bulwahn improving replacing higher order arguments to work with tuples
Tue, 31 Aug 2010 08:00:54 +0200 bulwahn avoiding warning for a duplicate rewrite rule in preprocessing of the predicate compiler
Fri, 27 Aug 2010 10:56:46 +0200 haftmann formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
Thu, 19 Aug 2010 16:08:59 +0200 haftmann tuned quotes
Thu, 19 Aug 2010 11:02:14 +0200 haftmann use antiquotations for remaining unqualified constants in HOL
Wed, 21 Jul 2010 18:11:51 +0200 bulwahn putting proof in the right context; adding if rewriting; tuned
Mon, 03 May 2010 14:25:56 +0200 wenzelm renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
Mon, 29 Mar 2010 17:30:48 +0200 bulwahn removing fishing for split thm in the predicate compiler
Mon, 22 Mar 2010 19:29:11 +0100 wenzelm eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
Mon, 22 Mar 2010 08:30:13 +0100 bulwahn improving handling of case expressions in predicate rewriting
Mon, 22 Mar 2010 00:48:56 +0100 wenzelm replaced PureThy.add_axioms by more basic Drule.add_axiom, which is old-style nonetheless;
Sat, 20 Mar 2010 17:33:11 +0100 wenzelm renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
Mon, 01 Mar 2010 09:47:44 +0100 bulwahn made smlnj happy
Tue, 23 Feb 2010 13:36:15 +0100 bulwahn adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
Wed, 20 Jan 2010 11:56:45 +0100 bulwahn refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
Thu, 05 Nov 2009 16:10:49 +0100 wenzelm eliminated funny record patterns and made SML/NJ happy;
Tue, 03 Nov 2009 10:24:05 +0100 bulwahn recursively replacing abstractions by new definitions in the predicate compiler
Fri, 30 Oct 2009 09:55:15 +0100 bulwahn renamed rpred to random
Wed, 28 Oct 2009 00:07:51 +0100 wenzelm proper headers;
Tue, 27 Oct 2009 09:02:22 +0100 bulwahn including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
less more (0) tip