src/HOL/Auth/Message.thy
Fri, 18 Feb 2011 15:46:13 +0100 wenzelm modernized specifications;
Wed, 02 Feb 2011 14:11:26 +0000 paulson Introduction of metis calls and other cosmetic modifications.
Wed, 08 Sep 2010 13:25:32 +0100 paulson tidied using inductive_simps
Thu, 22 Jul 2010 18:08:39 +0200 wenzelm updated some headers;
Thu, 04 Mar 2010 17:08:41 +0000 paulson Simplified a couple of proofs and corrected a comment
Mon, 01 Mar 2010 13:40:23 +0100 haftmann replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
Thu, 11 Feb 2010 21:33:25 +0100 wenzelm modernized syntax/translations;
Mon, 08 Feb 2010 21:28:27 +0100 wenzelm modernized some syntax translations;
Thu, 24 Dec 2009 17:30:55 +0000 paulson tidied proofs
Sat, 17 Oct 2009 14:43:18 +0200 wenzelm eliminated hard tabulators, guessing at each author's individual tab-width;
Thu, 23 Jul 2009 18:44:09 +0200 wenzelm renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
Tue, 21 Jul 2009 11:09:50 +0200 haftmann Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
Fri, 20 Mar 2009 15:24:18 +0100 wenzelm eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
Mon, 16 Mar 2009 18:24:30 +0100 wenzelm simplified method setup;
Fri, 13 Mar 2009 19:58:26 +0100 wenzelm unified type Proof.method and pervasive METHOD combinators;
Mon, 27 Oct 2008 18:14:34 +0100 paulson metis proof
Mon, 16 Jun 2008 22:13:39 +0200 wenzelm pervasive RuleInsts;
Mon, 16 Jun 2008 17:54:35 +0200 wenzelm eliminated OldGoals.inst;
Wed, 11 Jun 2008 18:02:25 +0200 wenzelm OldGoals.inst;
Wed, 11 Jun 2008 15:40:44 +0200 wenzelm RuleInsts.res_inst_tac with proper context;
Tue, 10 Jun 2008 15:30:54 +0200 haftmann slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
Wed, 07 May 2008 10:59:02 +0200 berghofe Replaced blast by fast in proof of parts_singleton, since blast looped
Wed, 19 Mar 2008 22:50:42 +0100 wenzelm more antiquotations;
Wed, 01 Aug 2007 20:25:16 +0200 wenzelm tuned ML bindings (for multithreading);
Wed, 11 Jul 2007 11:14:51 +0200 berghofe Adapted to new inductive definition package.
Sun, 06 May 2007 21:49:32 +0200 haftmann dropped legacy ML binding
Fri, 09 Mar 2007 08:45:55 +0100 haftmann *** empty log message ***
Wed, 29 Nov 2006 15:44:51 +0100 wenzelm simplified method setup;
Wed, 20 Sep 2006 15:11:46 +0200 paulson tidied
Thu, 22 Dec 2005 14:22:11 +0100 paulson shorter proof
Fri, 30 Sep 2005 09:52:46 +0200 paulson theorems need names
Wed, 28 Sep 2005 11:15:33 +0200 paulson streamlined theory; conformance to recent publication
Wed, 13 Jul 2005 16:47:23 +0200 paulson tidied
Wed, 13 Jul 2005 15:06:20 +0200 paulson generlization of some "nat" theorems
Fri, 17 Jun 2005 16:12:49 +0200 haftmann migrated theory headers to new format
Sun, 11 Jul 2004 20:33:22 +0200 wenzelm local_cla/simpset_of;
Tue, 23 Sep 2003 15:41:33 +0200 paulson Removal of the Key_supply axiom (affects many possbility proofs) and minor
Thu, 04 Sep 2003 11:08:24 +0200 paulson new, separate specifications
Tue, 12 Aug 2003 13:35:03 +0200 paulson ZhouGollmann: new example (fair non-repudiation protocol)
Thu, 24 Jul 2003 16:36:29 +0200 paulson Tidying and replacement of some axioms by specifications
Mon, 05 May 2003 18:22:01 +0200 paulson improved presentation of HOL/Auth theories
Sat, 26 Apr 2003 12:38:42 +0200 paulson converting more HOL-Auth to new-style theories
Fri, 25 Apr 2003 11:18:14 +0200 paulson Changes required by the certified email protocol
Fri, 27 Apr 2001 17:16:21 +0200 paulson better treatment of methods: uses Method.ctxt_args to refer to current
Tue, 24 Apr 2001 12:19:01 +0200 paulson (rough) conversion of Auth/Recur to Isar format
Thu, 12 Apr 2001 12:45:05 +0200 paulson converted many HOL/Auth theories to Isar scripts
Mon, 09 Apr 2001 14:49:51 +0200 paulson new theorem Fake_parts_insert_in_Un
Thu, 29 Mar 2001 10:44:37 +0200 paulson misc tidying; changing the predicate isSymKey to the set symKeys
Mon, 05 Mar 2001 12:31:31 +0100 paulson a few basic X-symbols
Fri, 02 Mar 2001 13:18:31 +0100 paulson conversion of Message.thy to Isar format
Tue, 09 Jan 2001 15:29:17 +0100 nipkow `` -> ` and ``` -> ``
Thu, 24 Aug 2000 12:39:42 +0200 paulson xsymbols for {| and |}
Wed, 21 Jul 1999 15:22:11 +0200 paulson tweaked proofs to handle new freeness reasoning for data c onstructors
Thu, 10 Jun 1999 10:36:41 +0200 paulson new translation to allow images over Nonce
Fri, 16 Oct 1998 12:20:41 +0200 paulson parent is Main
Mon, 03 Aug 1998 10:37:34 +0200 paulson Better comments
Fri, 24 Jul 1998 13:03:20 +0200 berghofe Adapted to new datatype package.
Tue, 30 Jun 1998 20:51:15 +0200 berghofe Adapted to new inductive definition package.
Thu, 11 Sep 1997 12:24:28 +0200 paulson Split base cases from "msg" to "atomic" in order
Fri, 17 Jan 1997 12:49:31 +0100 paulson Now with Andy Gordon's treatment of freshness to replace newN/K
less more (0) -60 tip