src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML
Sun, 23 Mar 2014 15:46:21 +0100 wenzelm more sensible treatment of quasi-local variables (NB: Variable.add_fixes is only for term variables, while Variable.declare takes care of types within given terms);
Sat, 22 Mar 2014 18:19:57 +0100 wenzelm more antiquotations;
Fri, 21 Mar 2014 20:33:56 +0100 wenzelm more qualified names;
Fri, 21 Mar 2014 10:45:03 +0100 wenzelm tuned signature;
Wed, 12 Feb 2014 14:32:45 +0100 wenzelm merged, resolving some conflicts;
Wed, 12 Feb 2014 13:33:05 +0100 wenzelm tuned whitespace;
Wed, 12 Feb 2014 08:35:56 +0100 blanchet ported predicate compiler to 'ctr_sugar'
Sat, 14 Dec 2013 17:28:05 +0100 wenzelm proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
Wed, 25 Sep 2013 16:43:46 +0200 blanchet filled in gap in library offering
Thu, 30 May 2013 12:35:40 +0200 wenzelm standardized aliases;
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;
Thu, 28 Feb 2013 16:54:52 +0100 wenzelm just one HOLogic.Trueprop_conv, with regular exception CTERM;
Mon, 12 Nov 2012 23:24:40 +0100 haftmann dropped dead code
Mon, 09 Jul 2012 09:47:59 +0200 bulwahn adding a missing entry to predicate compiler's setup
Fri, 24 Feb 2012 18:46:01 +0100 haftmann dropped dead code
Fri, 24 Feb 2012 09:40:02 +0100 haftmann given up disfruitful branch
Thu, 23 Feb 2012 21:16:54 +0100 haftmann dropped dead code
Sat, 17 Dec 2011 12:10:37 +0100 wenzelm tuned signature;
Wed, 14 Dec 2011 21:54:32 +0100 wenzelm avoid fragile Sign.intern_const -- pass internal names directly;
Mon, 12 Dec 2011 23:05:21 +0100 wenzelm datatype dtyp with explicit sort information;
Wed, 30 Nov 2011 23:30:08 +0100 wenzelm discontinued obsolete datatype "alt_names";
Fri, 11 Nov 2011 12:10:49 +0100 bulwahn using more conventional names for monad plus operations
Fri, 11 Nov 2011 08:32:45 +0100 bulwahn adding CPS compilation to predicate compiler;
Wed, 17 Aug 2011 18:05:31 +0200 wenzelm modernized signature of Term.absfree/absdummy;
Thu, 09 Jun 2011 17:51:49 +0200 wenzelm simplified Name.variant -- discontinued builtin fold_map;
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;
Thu, 24 Mar 2011 15:29:31 +0100 bulwahn allowing special set comprehensions in values command; adding an example for special set comprehension in values
Fri, 17 Dec 2010 17:08:56 +0100 wenzelm renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
Wed, 01 Dec 2010 15:03:44 +0100 wenzelm more direct use of binder_types/body_type;
Mon, 25 Oct 2010 21:17:14 +0200 bulwahn renaming split_modeT' to split_modeT
Fri, 22 Oct 2010 18:38:59 +0200 bulwahn moving general functions from core_data to predicate_compile_aux
Thu, 21 Oct 2010 19:13:11 +0200 bulwahn adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
Thu, 21 Oct 2010 19:13:09 +0200 bulwahn splitting large core file into core_data, mode_inference and predicate_compile_proof
Thu, 21 Oct 2010 19:13:09 +0200 bulwahn added generator_dseq compilation for a sound depth-limited compilation with small value generators
Thu, 21 Oct 2010 19:13:07 +0200 bulwahn adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
Thu, 21 Oct 2010 19:13:06 +0200 bulwahn adding option smart_depth_limiting to predicate compiler
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
Thu, 23 Sep 2010 17:22:45 +0200 bulwahn removing unneccessary expansion procedure for elimination rules; removing obsolete elim preprocessing; tuned
Thu, 23 Sep 2010 17:22:44 +0200 bulwahn moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
Mon, 20 Sep 2010 09:26:15 +0200 bulwahn removing clone in code_prolog and predicate_compile_quickcheck
Wed, 15 Sep 2010 09:36:39 +0200 bulwahn adding option show_invalid_clauses for a more detailed message when modes are not inferred
Wed, 15 Sep 2010 09:36:38 +0200 bulwahn proposed modes for code_pred now supports modes for mutual predicates
Mon, 13 Sep 2010 16:44:19 +0200 bulwahn handling function types more carefully than in e98a06145530
Mon, 13 Sep 2010 16:44:18 +0200 bulwahn adding order on modes
Fri, 10 Sep 2010 17:53:25 +0200 bulwahn directly computing the values of interest instead of composing functions in an unintelligent way that causes exponential much garbage; using the latest theory
Tue, 07 Sep 2010 11:51:53 +0200 bulwahn raising an exception instead of guessing some reasonable behaviour for this function
Sat, 28 Aug 2010 16:14:32 +0200 haftmann formerly unnamed infix equality now named HOL.eq
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:19:24 +0200 haftmann tuned
Thu, 19 Aug 2010 11:02:14 +0200 haftmann use antiquotations for remaining unqualified constants in HOL
Thu, 01 Jul 2010 16:54:44 +0200 haftmann "prod" and "sum" replace "*" and "+" respectively
Thu, 10 Jun 2010 12:24:03 +0200 haftmann tuned quotes, antiquotations and whitespace
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;
Wed, 21 Apr 2010 12:10:52 +0200 bulwahn added switch detection to the predicate compiler
Wed, 21 Apr 2010 12:10:52 +0200 bulwahn only add relevant predicates to the list of extra modes
less more (0) -60 tip