src/HOL/Tools/Quotient/quotient_tacs.ML
Wed, 09 Sep 2015 20:57:21 +0200 wenzelm simplified simproc programming interfaces;
Mon, 27 Jul 2015 17:44:55 +0200 wenzelm tuned signature;
Sat, 18 Jul 2015 20:47:08 +0200 wenzelm prefer tactics with explicit context;
Sun, 05 Jul 2015 15:02:30 +0200 wenzelm simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
Fri, 06 Mar 2015 23:54:05 +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 22:05:01 +0100 wenzelm clarified signature;
Wed, 04 Mar 2015 20:50:20 +0100 wenzelm tuned;
Wed, 04 Mar 2015 19:53:18 +0100 wenzelm tuned signature -- prefer qualified names;
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);
Sun, 09 Nov 2014 14:08:00 +0100 wenzelm proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
Sat, 08 Nov 2014 21:31:51 +0100 wenzelm optional proof context for unify operations, for the sake of proper local options;
Sat, 16 Aug 2014 20:14:45 +0200 wenzelm updated to named_theorems;
Thu, 06 Mar 2014 15:40:33 +0100 blanchet renamed 'fun_rel' to 'rel_fun'
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;
Thu, 18 Apr 2013 17:07:01 +0200 wenzelm simplifier uses proper Proof.context instead of historic type simpset;
Tue, 03 Apr 2012 16:26:48 +0200 kuncar new package Lifting - initial commit
Tue, 14 Feb 2012 20:09:35 +0100 wenzelm simplified use of tacticals;
Wed, 07 Dec 2011 14:00:02 +0000 Christian Urban added a specific tactic and method that deal with partial equivalence relations
Thu, 24 Nov 2011 21:01:06 +0100 wenzelm modernized some old-style infix operations, which were left over from the time of ML proof scripts;
Sat, 05 Nov 2011 12:01:21 +0000 Christian Urban more use of global operations (see 98ec8b51af9c)
Thu, 27 Oct 2011 20:26:38 +0200 wenzelm simplified/standardized signatures;
Fri, 19 Aug 2011 10:23:16 +0900 Cezary Kaliszyk Quotient Package: Regularization: do not fail if no progress is made, leave the subgoal to the user. Injection: try assumptions before extensionality to avoid looping.
Wed, 20 Jul 2011 16:14:49 +0200 Cezary Kaliszyk Quotient Package: handle Bound variables in rep_abs_rsp_tac not only at top-level of the goal
Wed, 29 Jun 2011 20:39:41 +0200 wenzelm simplified/unified Simplifier.mk_solver;
Thu, 09 Jun 2011 23:12:02 +0200 wenzelm renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
Sat, 16 Apr 2011 16:15:37 +0200 wenzelm modernized structure Proof_Context;
Fri, 07 Jan 2011 21:26:49 +0100 wenzelm do not open ML structures;
Fri, 07 Jan 2011 15:35:00 +0100 wenzelm more precise parentheses and indentation;
Fri, 07 Jan 2011 14:58:15 +0100 wenzelm comments;
Fri, 17 Dec 2010 17:08:56 +0100 wenzelm renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
Wed, 01 Dec 2010 13:09:08 +0100 wenzelm just one Term.dest_funT;
Thu, 18 Nov 2010 17:01:15 +0100 haftmann map_fun combinator in theory Fun
Tue, 09 Nov 2010 14:02:14 +0100 haftmann slightly changed fun_map_def
Mon, 30 Aug 2010 09:35:30 +0200 haftmann merged
Sat, 28 Aug 2010 16:14:32 +0200 haftmann formerly unnamed infix equality now named HOL.eq
Mon, 30 Aug 2010 15:52:09 +0900 Cezary Kaliszyk Quotient Package: dont unfold mem_def, use rsp and prs instead
Sat, 28 Aug 2010 21:17:25 +0800 Christian Urban quotient package: added a list of pre-simplification rules for Ball, Bex and mem
Sat, 28 Aug 2010 20:24:40 +0800 Christian Urban quotient package: lemmas to be lifted and descended can be pre-simplified
Wed, 25 Aug 2010 18:46:22 +0200 wenzelm merged
Wed, 25 Aug 2010 20:04:49 +0800 Christian Urban tuned code
Wed, 25 Aug 2010 18:26:58 +0800 Christian Urban quotient package: deal correctly with frees in lifted theorems
Wed, 25 Aug 2010 18:36:22 +0200 wenzelm renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
Tue, 24 Aug 2010 22:38:45 +0800 Christian Urban use matching of types than just equality - this is needed in nominal to cope with type variables
Sun, 22 Aug 2010 14:01:25 +0800 Christian Urban use a smaller simpset in order to not solve refl-instances
Sun, 22 Aug 2010 12:58:03 +0800 Christian Urban allow for pre-simplification of lifted theorems
Sun, 22 Aug 2010 10:45:53 +0800 Christian Urban changed to a more convenient argument order
Wed, 11 Aug 2010 13:30:24 +0800 Christian Urban deleted duplicate lemma
Thu, 08 Jul 2010 16:19:24 +0200 haftmann tuned titles
Tue, 29 Jun 2010 01:38:29 +0100 Christian Urban separated the lifting and descending procedures in the quotient package
Mon, 28 Jun 2010 16:20:39 +0100 Christian Urban separation of translations in derive_qtrm / derive_rtrm (similarly for types)
Mon, 28 Jun 2010 09:48:36 +0200 Cezary Kaliszyk Quotient package reverse lifting
Mon, 28 Jun 2010 07:38:39 +0200 Cezary Kaliszyk Add reverse lifting flag to automated theorem derivation
Sat, 26 Jun 2010 08:23:40 +0100 Christian Urban streamlined the generation of quotient theorems out of raw theorems
Wed, 23 Jun 2010 08:44:44 +0200 Cezary Kaliszyk Quotient package now uses Partial Equivalence instead place of equivalence
Sat, 15 May 2010 21:50:05 +0200 wenzelm less pervasive names from structure Thm;
Sat, 15 May 2010 17:59:06 +0200 wenzelm incorporated further conversions and conversionals, after some minor tuning;
Wed, 12 May 2010 11:30:18 +0200 Cezary Kaliszyk Remove RANGE_WARN
less more (0) -60 tip