| Tue, 13 Dec 2016 11:51:42 +0100 | 
wenzelm | 
more symbols;
 | 
file |
diff |
annotate
 | 
| Wed, 20 Jul 2016 22:36:10 +0200 | 
wenzelm | 
provide Pure.simp/simp_all, which only know about meta-equality;
 | 
file |
diff |
annotate
 | 
| Thu, 02 Jun 2016 16:23:10 +0200 | 
wenzelm | 
avoid warnings on duplicate rules in the given list;
 | 
file |
diff |
annotate
 | 
| Fri, 08 Apr 2016 20:15:20 +0200 | 
wenzelm | 
eliminated unused simproc identifier;
 | 
file |
diff |
annotate
 | 
| Wed, 16 Dec 2015 16:31:36 +0100 | 
wenzelm | 
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 | 
file |
diff |
annotate
 | 
| Sun, 13 Dec 2015 21:56:15 +0100 | 
wenzelm | 
more general types Proof.method / context_tactic;
 | 
file |
diff |
annotate
 | 
| Fri, 25 Sep 2015 20:37:59 +0200 | 
wenzelm | 
moved remaining display.ML to more_thm.ML;
 | 
file |
diff |
annotate
 | 
| Wed, 09 Sep 2015 23:01:27 +0200 | 
wenzelm | 
clarified declaration flags, like 'declaration' command;
 | 
file |
diff |
annotate
 | 
| Wed, 09 Sep 2015 20:57:21 +0200 | 
wenzelm | 
simplified simproc programming interfaces;
 | 
file |
diff |
annotate
 | 
| Wed, 02 Sep 2015 23:31:41 +0200 | 
wenzelm | 
eliminated pointless cterms;
 | 
file |
diff |
annotate
 | 
| Fri, 03 Apr 2015 19:56:51 +0200 | 
wenzelm | 
more uniform "verbose" option to print name space;
 | 
file |
diff |
annotate
 | 
| Fri, 06 Mar 2015 15:58:56 +0100 | 
wenzelm | 
Thm.cterm_of and Thm.ctyp_of operate on local context;
 | 
file |
diff |
annotate
 | 
| Sun, 01 Mar 2015 23:35:41 +0100 | 
wenzelm | 
added Proof_Context.cterm_of/ctyp_of convenience;
 | 
file |
diff |
annotate
 | 
| Tue, 10 Feb 2015 14:48:26 +0100 | 
wenzelm | 
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 | 
file |
diff |
annotate
 | 
| Mon, 10 Nov 2014 21:49:48 +0100 | 
wenzelm | 
proper context for assume_tac (atac remains as fall-back without context);
 | 
file |
diff |
annotate
 | 
| Sun, 09 Nov 2014 17:04:14 +0100 | 
wenzelm | 
proper context for match_tac etc.;
 | 
file |
diff |
annotate
 | 
| Wed, 27 Aug 2014 14:54:32 +0200 | 
wenzelm | 
more explicit Method.modifier with reported position;
 | 
file |
diff |
annotate
 | 
| Tue, 19 Aug 2014 17:00:44 +0200 | 
wenzelm | 
added PARALLEL_ALLGOALS convenience;
 | 
file |
diff |
annotate
 | 
| Thu, 10 Apr 2014 12:22:29 +0200 | 
wenzelm | 
added simproc markup, which also indicates legacy simprocs outside the name space;
 | 
file |
diff |
annotate
 | 
| Tue, 18 Mar 2014 15:29:58 +0100 | 
wenzelm | 
more antiquotations;
 | 
file |
diff |
annotate
 | 
| Wed, 12 Mar 2014 22:57:50 +0100 | 
wenzelm | 
tuned signature -- clarified module name;
 | 
file |
diff |
annotate
 | 
| Sun, 12 Jan 2014 18:34:00 +0100 | 
wenzelm | 
proper context for clear_simpset: preserve dounds, depth;
 | 
file |
diff |
annotate
 | 
| Thu, 12 Dec 2013 22:38:25 +0100 | 
wenzelm | 
clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
 | 
file |
diff |
annotate
 | 
| Thu, 12 Dec 2013 21:14:33 +0100 | 
wenzelm | 
generic trace operations for main steps of Simplifier;
 | 
file |
diff |
annotate
 | 
| Thu, 12 Dec 2013 17:34:50 +0100 | 
wenzelm | 
tuned signature;
 | 
file |
diff |
annotate
 | 
| Mon, 11 Nov 2013 20:50:12 +0100 | 
wenzelm | 
tuned signature -- removed obsolete Addsimprocs, Delsimprocs;
 | 
file |
diff |
annotate
 | 
| Fri, 23 Aug 2013 20:35:50 +0200 | 
wenzelm | 
added Theory.setup convenience;
 | 
file |
diff |
annotate
 | 
| Wed, 26 Jun 2013 21:48:23 +0200 | 
wenzelm | 
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 | 
file |
diff |
annotate
 | 
| Thu, 18 Apr 2013 17:07:01 +0200 | 
wenzelm | 
simplifier uses proper Proof.context instead of historic type simpset;
 | 
file |
diff |
annotate
 | 
| Wed, 10 Apr 2013 17:27:38 +0200 | 
wenzelm | 
obsolete -- tools should refer to proper Proof.context;
 | 
file |
diff |
annotate
 | 
| Sat, 30 Mar 2013 17:13:21 +0100 | 
wenzelm | 
more formal cong_name;
 | 
file |
diff |
annotate
 | 
| Sat, 30 Mar 2013 13:40:19 +0100 | 
wenzelm | 
more item markup;
 | 
file |
diff |
annotate
 | 
| Fri, 29 Mar 2013 22:14:27 +0100 | 
wenzelm | 
Pretty.item markup for improved readability of lists of items;
 | 
file |
diff |
annotate
 | 
| Sat, 17 Nov 2012 17:55:52 +0100 | 
wenzelm | 
tuned signature;
 | 
file |
diff |
annotate
 | 
| Sat, 11 Aug 2012 22:17:46 +0200 | 
wenzelm | 
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 | 
file |
diff |
annotate
 | 
| Sat, 14 Apr 2012 17:15:57 +0200 | 
wenzelm | 
outermost SELECT_GOAL potentially improves performance;
 | 
file |
diff |
annotate
 | 
| Sat, 31 Mar 2012 19:26:23 +0200 | 
wenzelm | 
tuned signature;
 | 
file |
diff |
annotate
 | 
| Sun, 18 Mar 2012 13:04:22 +0100 | 
wenzelm | 
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 | 
file |
diff |
annotate
 | 
| Sat, 17 Mar 2012 23:55:03 +0100 | 
wenzelm | 
proper naming of simprocs according to actual target context;
 | 
file |
diff |
annotate
 | 
| Sat, 03 Mar 2012 21:52:15 +0100 | 
wenzelm | 
tuned;
 | 
file |
diff |
annotate
 | 
| Tue, 14 Feb 2012 19:29:54 +0100 | 
wenzelm | 
tuned signature;
 | 
file |
diff |
annotate
 | 
| 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;
 | 
file |
diff |
annotate
 | 
| Wed, 23 Nov 2011 22:59:39 +0100 | 
wenzelm | 
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 | 
file |
diff |
annotate
 | 
| Sun, 06 Nov 2011 21:51:46 +0100 | 
wenzelm | 
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 | 
file |
diff |
annotate
 | 
| Thu, 03 Nov 2011 22:15:47 +0100 | 
wenzelm | 
tuned -- Variable.declare_term is already part of Variable.auto_fixes;
 | 
file |
diff |
annotate
 | 
| Fri, 28 Oct 2011 22:17:30 +0200 | 
wenzelm | 
uniform Local_Theory.declaration with explicit params;
 | 
file |
diff |
annotate
 | 
| Fri, 28 Oct 2011 17:15:52 +0200 | 
wenzelm | 
tuned signature -- refined terminology;
 | 
file |
diff |
annotate
 | 
| Mon, 08 Aug 2011 17:23:15 +0200 | 
wenzelm | 
misc tuning -- eliminated old-fashioned rep_thm;
 | 
file |
diff |
annotate
 | 
| Wed, 29 Jun 2011 21:34:16 +0200 | 
wenzelm | 
tuned signature;
 | 
file |
diff |
annotate
 | 
| Wed, 29 Jun 2011 20:39:41 +0200 | 
wenzelm | 
simplified/unified Simplifier.mk_solver;
 | 
file |
diff |
annotate
 | 
| Mon, 27 Jun 2011 16:53:31 +0200 | 
wenzelm | 
ML antiquotations are managed as theory data, with proper name space and entity markup;
 | 
file |
diff |
annotate
 | 
| Fri, 13 May 2011 23:58:40 +0200 | 
wenzelm | 
clarified map_simpset versus Simplifier.map_simpset_global;
 | 
file |
diff |
annotate
 | 
| Fri, 13 May 2011 22:55:00 +0200 | 
wenzelm | 
proper Proof.context for classical tactics;
 | 
file |
diff |
annotate
 | 
| Sat, 23 Apr 2011 17:02:12 +0200 | 
wenzelm | 
added Name_Space.check/get convenience;
 | 
file |
diff |
annotate
 | 
| Sat, 23 Apr 2011 16:30:00 +0200 | 
wenzelm | 
clarified check_simproc (with report) vs. the_simproc;
 | 
file |
diff |
annotate
 | 
| Sat, 23 Apr 2011 13:53:09 +0200 | 
wenzelm | 
proper binding/report of defined simprocs;
 | 
file |
diff |
annotate
 | 
| Sun, 17 Apr 2011 19:54:04 +0200 | 
wenzelm | 
report Name_Space.declare/define, relatively to context;
 | 
file |
diff |
annotate
 | 
| Sat, 16 Apr 2011 23:41:25 +0200 | 
wenzelm | 
PARALLEL_GOALS for method "simp_all";
 | 
file |
diff |
annotate
 | 
| Sat, 16 Apr 2011 15:47:52 +0200 | 
wenzelm | 
modernized structure Proof_Context;
 | 
file |
diff |
annotate
 | 
| Wed, 22 Dec 2010 20:08:40 +0100 | 
haftmann | 
tuned comment
 | 
file |
diff |
annotate
 |