src/HOL/Nominal/nominal_thmdecls.ML
Fri, 21 Mar 2014 20:33:56 +0100 wenzelm more qualified names;
Sat, 11 Jan 2014 14:12:33 +0100 wenzelm tuned signature;
Thu, 18 Apr 2013 17:07:01 +0200 wenzelm simplifier uses proper Proof.context instead of historic type simpset;
Mon, 02 May 2011 16:33:21 +0200 wenzelm added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
Sat, 16 Apr 2011 16:15:37 +0200 wenzelm modernized structure Proof_Context;
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;
Sat, 28 Aug 2010 16:14:32 +0200 haftmann formerly unnamed infix equality now named HOL.eq
Fri, 27 Aug 2010 17:59:40 +0200 wenzelm more antiquotations;
Fri, 27 Aug 2010 17:23:57 +0200 wenzelm eliminated Unsynchronized.ref in favour of configuration option;
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
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
Sun, 08 Nov 2009 16:30:41 +0100 wenzelm adapted Generic_Data, Proof_Data;
Sat, 17 Oct 2009 14:43:18 +0200 wenzelm eliminated hard tabulators, guessing at each author's individual tab-width;
Tue, 29 Sep 2009 16:24:36 +0200 wenzelm explicit indication of Unsynchronized.ref;
Fri, 28 Aug 2009 18:18:30 +0200 wenzelm modernized messages -- eliminated old Display.print_cterm;
Tue, 21 Jul 2009 01:03:18 +0200 wenzelm proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
Mon, 27 Apr 2009 00:29:54 +0200 Christian Urban tuned
Sun, 26 Apr 2009 00:42:49 +0200 Christian Urban deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Sun, 15 Mar 2009 15:59:44 +0100 wenzelm simplified attribute setup;
Sun, 08 Mar 2009 17:26:14 +0100 wenzelm moved basic algebra of long names from structure NameSpace to Long_Name;
Thu, 05 Mar 2009 12:08:00 +0100 wenzelm renamed NameSpace.base to NameSpace.base_name;
Wed, 25 Feb 2009 11:07:10 +0100 berghofe Replaced Logic.unvarify by Variable.import_terms to make declaration of
Wed, 21 Jan 2009 18:27:43 +0100 haftmann binding replaces bstring
Sat, 17 May 2008 13:54:30 +0200 wenzelm structure Display: less pervasive operations;
Tue, 15 Apr 2008 16:12:01 +0200 wenzelm proper dynamic facts for eqvts, freshs, bijs;
Tue, 25 Mar 2008 21:59:48 +0100 wenzelm update_context: always store as "Nominal.eqvts";
Thu, 20 Mar 2008 00:20:44 +0100 wenzelm simplified get_thm(s): back to plain name argument;
Wed, 19 Mar 2008 22:28:08 +0100 wenzelm auxiliary dynamic_thm(s) for fact lookup;
Thu, 13 Sep 2007 23:58:38 +0200 urbanc some cleaning up to do with contexts
Tue, 14 Aug 2007 23:22:51 +0200 wenzelm avoid low-level tsig;
Tue, 14 Aug 2007 15:09:33 +0200 narboux fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
Sun, 29 Jul 2007 14:29:54 +0200 wenzelm renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
Mon, 07 May 2007 00:49:59 +0200 wenzelm simplified DataFun interfaces;
Mon, 16 Apr 2007 07:32:23 +0200 urbanc improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
Mon, 02 Apr 2007 23:24:52 +0200 narboux rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
Wed, 28 Mar 2007 18:25:23 +0200 urbanc the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
Thu, 22 Mar 2007 11:17:32 +0100 urbanc clarified an error-message
Mon, 12 Mar 2007 16:25:39 +0100 narboux removes some unused code that used to try to derive a simpler version of the eqvt lemmas
Tue, 06 Mar 2007 15:28:22 +0100 urbanc major update of the nominal package; there is now an infrastructure
Thu, 08 Feb 2007 17:22:22 +0100 urbanc added get-functions
Tue, 06 Feb 2007 15:12:18 +0100 urbanc fixed two stupid bugs of SML to do with the value restriction and missing type
Tue, 06 Feb 2007 00:41:54 +0100 urbanc moved the infrastructure from the nominal_tags file to nominal_thmdecls
less more (0) tip