src/HOL/Nominal/Nominal.thy
Thu, 24 Sep 2015 13:33:42 +0200 wenzelm explicit indication of overloaded typedefs;
Fri, 26 Jun 2015 10:20:33 +0200 wenzelm tuned whitespace;
Thu, 25 Jun 2015 23:33:47 +0200 wenzelm tuned proofs;
Mon, 06 Apr 2015 23:14:05 +0200 wenzelm local setup of induction tools, with restricted access to auxiliary consts;
Thu, 18 Sep 2014 16:47:40 +0200 blanchet moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
Thu, 11 Sep 2014 21:11:03 +0200 blanchet fixed some spelling mistakes
Thu, 11 Sep 2014 19:32:36 +0200 blanchet updated news
Mon, 08 Sep 2014 23:04:18 +0200 blanchet ported old Nominal to use new datatypes
Thu, 20 Mar 2014 19:58:33 +0100 wenzelm more standard method_setup;
Tue, 19 Nov 2013 10:05:53 +0100 haftmann eliminiated neg_numeral in favour of - (numeral _)
Fri, 12 Oct 2012 18:58:20 +0200 wenzelm discontinued obsolete typedef (open) syntax;
Wed, 22 Aug 2012 22:55:41 +0200 wenzelm prefer ML_file over old uses;
Sun, 25 Mar 2012 20:15:39 +0200 huffman merged fork with new numeral representation (see NEWS)
Thu, 15 Mar 2012 22:08:53 +0100 wenzelm declare command keywords via theory header, including strict checking outside Pure;
Thu, 15 Mar 2012 19:02:34 +0100 wenzelm declare minor keywords via theory header;
Tue, 10 Jan 2012 23:49:53 +0100 berghofe Reverted several lemmas involving sets to the state before the
Sat, 24 Dec 2011 15:53:08 +0100 haftmann treatment of type constructor `set`
Wed, 30 Nov 2011 16:27:10 +0100 wenzelm prefer typedef without extra definition and alternative name;
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;
Thu, 08 Sep 2011 00:35:22 +0200 haftmann merged
Tue, 06 Sep 2011 22:04:14 +0200 haftmann merged
Mon, 05 Sep 2011 22:02:32 +0200 haftmann tuned
Sun, 04 Sep 2011 08:43:06 +0200 haftmann pseudo-definition for perms on sets; tuned
Sat, 03 Sep 2011 23:59:36 +0200 haftmann tuned specifications
Sat, 03 Sep 2011 17:32:34 +0200 haftmann tuned specifications and proofs
Sun, 28 Aug 2011 14:16:14 +0200 haftmann tuned
Mon, 21 Feb 2011 17:43:23 +0100 wenzelm modernized specifications;
Sat, 15 Jan 2011 12:38:56 +0100 berghofe Finally removed old primrec package, since Primrec.add_primrec_global
Fri, 14 Jan 2011 15:44:47 +0100 wenzelm eliminated global prems;
Wed, 29 Dec 2010 17:34:41 +0100 wenzelm explicit file specifications -- avoid secondary load path;
Wed, 29 Sep 2010 09:07:58 +0200 haftmann moved old_primrec source to nominal package, where it is still used
Mon, 13 Sep 2010 11:13:15 +0200 nipkow renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
Tue, 07 Sep 2010 10:05:19 +0200 nipkow expand_fun_eq -> ext_iff
Mon, 01 Mar 2010 13:42:31 +0100 haftmann merged
Mon, 01 Mar 2010 13:40:23 +0100 haftmann replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
Wed, 24 Feb 2010 21:59:21 +0100 wenzelm proper type syntax (cf. 7425aece4ee3);
Sat, 17 Oct 2009 14:43:18 +0200 wenzelm eliminated hard tabulators, guessing at each author's individual tab-width;
Mon, 21 Sep 2009 15:02:23 +0200 Christian Urban tuned some proofs
Fri, 03 Jul 2009 16:51:08 +0200 haftmann nominal.ML is nominal_datatype.ML
Fri, 19 Jun 2009 17:23:21 +0200 haftmann discontinued ancient tradition to suffix certain ML module names with "_package"
Sun, 26 Apr 2009 23:16:24 +0200 Christian Urban reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Sat, 25 Apr 2009 21:42:05 +0200 Christian Urban adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Wed, 04 Mar 2009 11:05:29 +0100 blanchet Merge.
Wed, 04 Mar 2009 10:45:52 +0100 blanchet Merge.
Thu, 26 Feb 2009 16:34:03 +0100 berghofe Added postprocessing rules for fresh_star.
Wed, 25 Feb 2009 19:34:00 +0100 berghofe Added lemmas for normalizing freshness results involving fresh_star.
Wed, 25 Feb 2009 11:02:25 +0100 berghofe Added equivariance lemmas for fresh_star.
Sat, 14 Feb 2009 08:45:16 +0100 nipkow more finiteness changes
Tue, 16 Dec 2008 20:32:41 +0000 Christian Urban changed the names of insert_eqvt and set_eqvt so that it is clear that they have preconditions
Thu, 04 Dec 2008 14:43:33 +0100 haftmann cleaned up binding module and related code
Tue, 21 Oct 2008 21:18:54 +0200 berghofe Added nominal_inductive2.ML
Fri, 26 Sep 2008 14:51:27 +0200 berghofe Added some more lemmas that are useful in proof of strong induction rule.
Mon, 22 Sep 2008 19:46:24 +0200 urbanc made the perm_simp tactic to understand options such as (no_asm)
Wed, 27 Aug 2008 04:47:30 +0200 urbanc added equivariance lemmas for ex1 and the
Sun, 27 Jul 2008 20:08:16 +0200 urbanc simplified a proof
Fri, 27 Jun 2008 00:37:30 +0200 urbanc added a lemma to at_swap_simps
Mon, 16 Jun 2008 17:54:39 +0200 wenzelm allE_Nil: only one copy, proven in regular theory source;
Thu, 08 May 2008 04:20:08 +0200 urbanc added at_set_avoiding lemmas
Wed, 07 May 2008 10:57:19 +0200 berghofe Adapted to encoding of sets as predicates
Fri, 02 May 2008 22:43:14 +0200 urbanc added more infrastructure for fresh_star
Fri, 02 May 2008 16:32:51 +0200 urbanc polished the proof for atm_prm_fresh and more lemmas for fresh_star
Thu, 03 Apr 2008 11:48:48 +0200 urbanc added generalised definitions for freshness of sets of atoms
Wed, 19 Mar 2008 22:50:42 +0100 wenzelm more antiquotations;
Mon, 18 Feb 2008 03:12:08 +0100 urbanc added eqvt-flag to subseteq-lemma
Thu, 24 Jan 2008 11:23:11 +0100 berghofe Added lemma at_fin_set_fresh.
Fri, 14 Sep 2007 13:32:07 +0200 urbanc reverted back to the old version of the equivariance lemma for ALL
Thu, 13 Sep 2007 23:58:38 +0200 urbanc some cleaning up to do with contexts
Thu, 13 Sep 2007 18:06:50 +0200 berghofe Added equivariance lemma for induct_implies.
Thu, 06 Sep 2007 16:28:42 +0200 urbanc trivial cleaning up
Wed, 11 Jul 2007 11:28:13 +0200 berghofe Adapted to new inductive definition package.
Thu, 14 Jun 2007 23:04:36 +0200 wenzelm tuned proofs: avoid implicit prems;
Thu, 31 May 2007 15:23:35 +0200 urbanc tuned the proof
Thu, 31 May 2007 14:47:20 +0200 urbanc introduced symmetric variants of the lemmas for alpha-equivalence
Sun, 20 May 2007 13:39:46 +0200 urbanc added lemma for permutations on strings
Mon, 07 May 2007 00:49:59 +0200 wenzelm simplified DataFun interfaces;
Thu, 03 May 2007 17:54:36 +0200 urbanc tuned some of the proofs and added the lemma fresh_bool
Wed, 02 May 2007 01:42:23 +0200 urbanc tuned some proofs and changed variable names in some definitions of Nominal.thy
Thu, 26 Apr 2007 14:24:08 +0200 wenzelm eliminated unnamed infixes, tuned syntax;
Tue, 24 Apr 2007 17:16:55 +0200 narboux fixes last commit
Tue, 24 Apr 2007 16:44:11 +0200 narboux add two lemmas dealing with freshness on permutations.
Tue, 24 Apr 2007 14:02:16 +0200 narboux oups : wrong commit
Tue, 24 Apr 2007 14:01:23 +0200 narboux adds op in front of an infix to fix SML compilation
Mon, 23 Apr 2007 18:38:42 +0200 urbanc simplified the proof of pt_set_eqvt (as suggested by Randy Pollack)
Sat, 21 Apr 2007 01:34:15 +0200 urbanc tuned the setup of fresh_fun
Fri, 20 Apr 2007 00:28:07 +0200 urbanc declared lemmas true_eqvt and false_eqvt to be equivariant (suggested by samth at ccs.neu.edu)
Thu, 19 Apr 2007 16:27:53 +0200 narboux add a tactic to generate fresh names
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, 16 Apr 2007 06:45:22 +0200 urbanc added a more usuable lemma for dealing with fresh_fun
Thu, 12 Apr 2007 15:46:12 +0200 urbanc tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
Sat, 07 Apr 2007 11:36:35 +0200 urbanc tuned slightly the previous commit
Sat, 07 Apr 2007 11:05:25 +0200 narboux perm_simp can now simplify using the rules (a,b) o a = b and (a,b) o b = a
Tue, 03 Apr 2007 17:05:52 +0200 narboux remove the lemma swap_fun which was not needed
Tue, 27 Mar 2007 19:39:40 +0200 urbanc added extended the lemma for equivariance of freshness
Sun, 25 Mar 2007 15:15:07 +0200 urbanc moving lemmas into appropriate sections
Fri, 23 Mar 2007 10:50:03 +0100 urbanc added the permutation operation on options to the list of equivariance lemmas
Thu, 22 Mar 2007 14:26:05 +0100 urbanc corrected the lemmas min_nat_eqvt and min_int_eqvt
Fri, 16 Mar 2007 17:09:18 +0100 urbanc added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
Tue, 06 Mar 2007 15:28:22 +0100 urbanc major update of the nominal package; there is now an infrastructure
Thu, 15 Feb 2007 18:18:21 +0100 narboux start adding the attribute eqvt to some lemmas of the nominal library
Tue, 13 Feb 2007 18:17:28 +0100 berghofe Added new file nominal_inductive.ML
Tue, 06 Feb 2007 00:41:54 +0100 urbanc moved the infrastructure from the nominal_tags file to nominal_thmdecls
Fri, 02 Feb 2007 17:16:16 +0100 urbanc added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
Mon, 27 Nov 2006 12:10:51 +0100 berghofe Implemented new "nominal_primrec" command for defining
Fri, 17 Nov 2006 17:32:30 +0100 urbanc added an intro lemma for freshness of products; set up
Wed, 15 Nov 2006 16:23:55 +0100 urbanc replaced exists_fresh lemma with a version formulated with obtains;
Mon, 13 Nov 2006 12:10:49 +0100 wenzelm added fresh_prodD, which is included fresh_prodD into mksimps setup;
Fri, 13 Oct 2006 15:01:34 +0200 urbanc added the missing freshness-lemmas for nat, int, char and string and
Sun, 01 Oct 2006 18:29:26 +0200 wenzelm moved theory Infinite_Set to Library;
Wed, 16 Aug 2006 16:44:41 +0200 urbanc added missing supp_nat lemma
Tue, 04 Jul 2006 15:22:54 +0200 berghofe - nominal_permeq.ML is now loaded before nominal_package.ML
Sun, 02 Jul 2006 17:27:10 +0200 urbanc added more infrastructure for the recursion combinator
Tue, 13 Jun 2006 15:07:47 +0200 paulson removal of the obsolete "infinite_nonempty"
Mon, 12 Jun 2006 20:32:33 +0200 urbanc added lemma fresh_unit to Nominal.thy
Mon, 12 Jun 2006 18:15:45 +0200 berghofe Removed comments around declaration of fresh_guess method.
Fri, 09 Jun 2006 17:32:38 +0200 berghofe unique_names no longer set to false (thanks to improved naming
Mon, 05 Jun 2006 19:09:40 +0200 urbanc added some further lemmas that deal with permutations and set-operators
Mon, 05 Jun 2006 18:38:41 +0200 urbanc added the lemma perm_diff
Thu, 01 Jun 2006 14:15:08 +0200 urbanc added the hack "reset NameSpace.unique_names" to Nominal.thy
Sat, 20 May 2006 23:36:56 +0200 wenzelm primrec (unchecked);
Mon, 15 May 2006 19:40:17 +0200 urbanc added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
less more (0) -120 tip