new and updated protocol proofs by Giamp Bella
authorpaulson
Wed, 01 Feb 2006 15:22:02 +0100
changeset 18886 9f27383426db
parent 18885 ee8b5c36ba2b
child 18887 6ad81e3fa478
new and updated protocol proofs by Giamp Bella
src/HOL/Auth/KerberosIV.thy
src/HOL/Auth/KerberosIV_Gets.thy
src/HOL/Auth/KerberosV.thy
src/HOL/Auth/Kerberos_BAN.thy
src/HOL/Auth/Kerberos_BAN_Gets.thy
src/HOL/Auth/NS_Shared.thy
src/HOL/Auth/OtwayReesBella.thy
src/HOL/Auth/ROOT.ML
src/HOL/Auth/Smartcard/EventSC.thy
src/HOL/Auth/Smartcard/ShoupRubin.thy
src/HOL/Auth/Smartcard/ShoupRubinBella.thy
src/HOL/Auth/Smartcard/Smartcard.thy
src/HOL/IsaMakefile
--- a/src/HOL/Auth/KerberosIV.thy	Wed Feb 01 12:23:14 2006 +0100
+++ b/src/HOL/Auth/KerberosIV.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -8,6 +8,8 @@
 
 theory KerberosIV imports Public begin
 
+text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+
 syntax
   Kas :: agent
   Tgs :: agent  --{*the two servers are translations...*}
@@ -22,26 +24,26 @@
   Tgs_not_bad [iff]: "Tgs \<notin> bad"
    --{*Tgs is secure --- we already know that Kas is secure*}
 
-(*The current time is just the length of the trace!*)
+(*The current time is the length of the trace*)
 syntax
     CT :: "event list=>nat"
 
-    ExpirAuth :: "[nat, event list] => bool"
+    expiredAK :: "[nat, event list] => bool"
 
-    ExpirServ :: "[nat, event list] => bool"
+    expiredSK :: "[nat, event list] => bool"
 
-    ExpirAutc :: "[nat, event list] => bool"
+    expiredA :: "[nat, event list] => bool"
 
-    RecentResp :: "[nat, nat] => bool"
+    valid :: "[nat, nat] => bool" ("valid _ wrt _")
 
 
 constdefs
- (* AuthKeys are those contained in an AuthTicket *)
-    AuthKeys :: "event list => key set"
-    "AuthKeys evs == {AuthKey. \<exists>A Peer Tk. Says Kas A
-                        (Crypt (shrK A) {|Key AuthKey, Agent Peer, Tk,
-                   (Crypt (shrK Peer) {|Agent A, Agent Peer, Key AuthKey, Tk|})
-                  |}) \<in> set evs}"
+ (* authKeys are those contained in an authTicket *)
+    authKeys :: "event list => key set"
+    "authKeys evs == {authK. \<exists>A Peer Ta. Says Kas A
+                        (Crypt (shrK A) \<lbrace>Key authK, Agent Peer, Number Ta,
+               (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key authK, Number Ta\<rbrace>)
+                  \<rbrace>) \<in> set evs}"
 
  (* A is the true creator of X if she has sent X and X never appeared on
     the trace before this event. Recall that traces grow from head. *)
@@ -51,75 +53,84 @@
       \<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
       X \<notin> parts (spies (takeWhile (% z. z  \<noteq> Says A B Y) (rev evs)))"
 
+ (* Yields the subtrace of a given trace from its beginning to a given event *)
+  before :: "[event, event list] => event list" ("before _ on _")
+   "before ev on evs ==  takeWhile (% z. z ~= ev) (rev evs)"
+
+ (* States than an event really appears only once on a trace *)
+  Unique :: "[event, event list] => bool" ("Unique _ on _")
+   "Unique ev on evs == 
+      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
+
 
 consts
     (*Duration of the authentication key*)
-    AuthLife   :: nat
+    authKlife   :: nat
 
     (*Duration of the service key*)
-    ServLife   :: nat
+    servKlife   :: nat
 
     (*Duration of an authenticator*)
-    AutcLife   :: nat
+    authlife   :: nat
 
     (*Upper bound on the time of reaction of a server*)
-    RespLife   :: nat
+    replylife   :: nat
 
-specification (AuthLife)
-  AuthLife_LB [iff]: "2 \<le> AuthLife"
+specification (authKlife)
+  authKlife_LB [iff]: "2 \<le> authKlife"
     by blast
 
-specification (ServLife)
-  ServLife_LB [iff]: "2 \<le> ServLife"
+specification (servKlife)
+  servKlife_LB [iff]: "2 + authKlife \<le> servKlife"
     by blast
 
-specification (AutcLife)
-  AutcLife_LB [iff]: "Suc 0 \<le> AutcLife"
+specification (authlife)
+  authlife_LB [iff]: "Suc 0 \<le> authlife"
     by blast
 
-specification (RespLife)
-  RespLife_LB [iff]: "Suc 0 \<le> RespLife"
+specification (replylife)
+  replylife_LB [iff]: "Suc 0 \<le> replylife"
     by blast
 
 translations
    "CT" == "length "
 
-   "ExpirAuth T evs" == "AuthLife + T < CT evs"
+   "expiredAK Ta evs" == "authKlife + Ta < CT evs"
 
-   "ExpirServ T evs" == "ServLife + T < CT evs"
+   "expiredSK Ts evs" == "servKlife + Ts < CT evs"
 
-   "ExpirAutc T evs" == "AutcLife + T < CT evs"
+   "expiredA T evs" == "authlife + T < CT evs"
 
-   "RecentResp T1 T2" == "T1 <= RespLife + T2"
+   "valid T1 wrt T2" == "T1 <= replylife + T2"
 
 (*---------------------------------------------------------------------*)
 
 
-(* Predicate formalising the association between AuthKeys and ServKeys *)
+(* Predicate formalising the association between authKeys and servKeys *)
 constdefs
-  KeyCryptKey :: "[key, key, event list] => bool"
-  "KeyCryptKey AuthKey ServKey evs ==
-     \<exists>A B tt.
-       Says Tgs A (Crypt AuthKey
-                     {|Key ServKey, Agent B, tt,
-                       Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |})
+  AKcryptSK :: "[key, key, event list] => bool"
+  "AKcryptSK authK servK evs ==
+     \<exists>A B Ts.
+       Says Tgs A (Crypt authK
+                     \<lbrace>Key servK, Agent B, Number Ts,
+                       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>)
          \<in> set evs"
 
 consts
 
-kerberos   :: "event list set"
-inductive "kerberos"
+kerbIV   :: "event list set"
+inductive "kerbIV"
   intros
 
-   Nil:  "[] \<in> kerberos"
+   Nil:  "[] \<in> kerbIV"
 
-   Fake: "[| evsf \<in> kerberos;  X \<in> synth (analz (spies evsf)) |]
-          ==> Says Spy B X  # evsf \<in> kerberos"
+   Fake: "\<lbrakk> evsf \<in> kerbIV;  X \<in> synth (analz (spies evsf)) \<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> kerbIV"
 
 (* FROM the initiator *)
-   K1:   "[| evs1 \<in> kerberos |]
-          ==> Says A Kas {|Agent A, Agent Tgs, Number (CT evs1)|} # evs1
-          \<in> kerberos"
+   K1:   "\<lbrakk> evs1 \<in> kerbIV \<rbrakk>
+          \<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
+          \<in> kerbIV"
 
 (* Adding the timestamp serves to A in K3 to check that
    she doesn't get a reply too late. This kind of timeouts are ordinary.
@@ -128,61 +139,62 @@
 (*---------------------------------------------------------------------*)
 
 (*FROM Kas *)
-   K2:  "[| evs2 \<in> kerberos; Key AuthKey \<notin> used evs2; AuthKey \<in> symKeys;
-            Says A' Kas {|Agent A, Agent Tgs, Number Ta|} \<in> set evs2 |]
-          ==> Says Kas A
-                (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number (CT evs2),
-                      (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey,
-                          Number (CT evs2)|})|}) # evs2 \<in> kerberos"
+   K2:  "\<lbrakk> evs2 \<in> kerbIV; Key authK \<notin> used evs2; authK \<in> symKeys;
+            Says A' Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Kas A
+                (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2),
+                      (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
+                          Number (CT evs2)\<rbrace>)\<rbrace>) # evs2 \<in> kerbIV"
 (*
-  The internal encryption builds the AuthTicket.
+  The internal encryption builds the authTicket.
   The timestamp doesn't change inside the two encryptions: the external copy
   will be used by the initiator in K3; the one inside the
-  AuthTicket by Tgs in K4.
+  authTicket by Tgs in K4.
 *)
 
 (*---------------------------------------------------------------------*)
 
 (* FROM the initiator *)
-   K3:  "[| evs3 \<in> kerberos;
-            Says A Kas {|Agent A, Agent Tgs, Number Ta|} \<in> set evs3;
-            Says Kas' A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-              AuthTicket|}) \<in> set evs3;
-            RecentResp Tk Ta
-         |]
-          ==> Says A Tgs {|AuthTicket,
-                           (Crypt AuthKey {|Agent A, Number (CT evs3)|}),
-                           Agent B|} # evs3 \<in> kerberos"
-(*The two events amongst the premises allow A to accept only those AuthKeys
+   K3:  "\<lbrakk> evs3 \<in> kerbIV;
+            Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
+            Says Kas' A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+              authTicket\<rbrace>) \<in> set evs3;
+            valid Ta wrt T1
+         \<rbrakk>
+          \<Longrightarrow> Says A Tgs \<lbrace>authTicket,
+                           (Crypt authK \<lbrace>Agent A, Number (CT evs3)\<rbrace>),
+                           Agent B\<rbrace> # evs3 \<in> kerbIV"
+(*The two events amongst the premises allow A to accept only those authKeys
   that are not issued late. *)
 
 (*---------------------------------------------------------------------*)
 
 (* FROM Tgs *)
 (* Note that the last temporal check is not mentioned in the original MIT
-   specification. Adding it strengthens the guarantees assessed by the
-   protocol. Theorems that exploit it have the suffix `_refined'
+   specification. Adding it makes many goals "available" to the peers. 
+   Theorems that exploit it have the suffix `_u', which stands for updated 
+   protocol.
 *)
-   K4:  "[| evs4 \<in> kerberos; Key ServKey \<notin> used evs4; ServKey \<in> symKeys;
-            B \<noteq> Tgs;  AuthKey \<in> symKeys;
-            Says A' Tgs {|
-             (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey,
-				 Number Tk|}),
-             (Crypt AuthKey {|Agent A, Number Ta1|}), Agent B|}
+   K4:  "\<lbrakk> evs4 \<in> kerbIV; Key servK \<notin> used evs4; servK \<in> symKeys;
+            B \<noteq> Tgs;  authK \<in> symKeys;
+            Says A' Tgs \<lbrace>
+             (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
+				 Number Ta\<rbrace>),
+             (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
 	        \<in> set evs4;
-            ~ ExpirAuth Tk evs4;
-            ~ ExpirAutc Ta1 evs4;
-            ServLife + (CT evs4) <= AuthLife + Tk
-         |]
-          ==> Says Tgs A
-                (Crypt AuthKey {|Key ServKey, Agent B, Number (CT evs4),
-			       Crypt (shrK B) {|Agent A, Agent B, Key ServKey,
-		 			        Number (CT evs4)|} |})
-	        # evs4 \<in> kerberos"
+            \<not> expiredAK Ta evs4;
+            \<not> expiredA T2 evs4;
+            servKlife + (CT evs4) <= authKlife + Ta
+         \<rbrakk>
+          \<Longrightarrow> Says Tgs A
+                (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
+			       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+		 			        Number (CT evs4)\<rbrace> \<rbrace>)
+	        # evs4 \<in> kerbIV"
 (* Tgs creates a new session key per each request for a service, without
    checking if there is still a fresh one for that service.
-   The cipher under Tgs' key is the AuthTicket, the cipher under B's key
-   is the ServTicket, which is built now.
+   The cipher under Tgs' key is the authTicket, the cipher under B's key
+   is the servTicket, which is built now.
    NOTE that the last temporal check is not present in the MIT specification.
 
 *)
@@ -190,56 +202,56 @@
 (*---------------------------------------------------------------------*)
 
 (* FROM the initiator *)
-   K5:  "[| evs5 \<in> kerberos; AuthKey \<in> symKeys; ServKey \<in> symKeys;
+   K5:  "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
             Says A Tgs
-                {|AuthTicket, Crypt AuthKey {|Agent A, Number Ta1|},
-		  Agent B|}
+                \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
+		  Agent B\<rbrace>
               \<in> set evs5;
             Says Tgs' A
-             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
+             (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
                 \<in> set evs5;
-            RecentResp Tt Ta1 |]
-          ==> Says A B {|ServTicket,
-			 Crypt ServKey {|Agent A, Number (CT evs5)|} |}
-               # evs5 \<in> kerberos"
+            valid Ts wrt T2 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>servTicket,
+			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+               # evs5 \<in> kerbIV"
 (* Checks similar to those in K3. *)
 
 (*---------------------------------------------------------------------*)
 
 (* FROM the responder*)
-    K6:  "[| evs6 \<in> kerberos;
-            Says A' B {|
-              (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}),
-              (Crypt ServKey {|Agent A, Number Ta2|})|}
+    K6:  "\<lbrakk> evs6 \<in> kerbIV;
+            Says A' B \<lbrace>
+              (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
+              (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
             \<in> set evs6;
-            ~ ExpirServ Tt evs6;
-            ~ ExpirAutc Ta2 evs6
-         |]
-          ==> Says B A (Crypt ServKey (Number Ta2))
-               # evs6 \<in> kerberos"
+            \<not> expiredSK Ts evs6;
+            \<not> expiredA T3 evs6
+         \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt servK (Number T3))
+               # evs6 \<in> kerbIV"
 (* Checks similar to those in K4. *)
 
 (*---------------------------------------------------------------------*)
 
-(* Leaking an AuthKey... *)
-   Oops1: "[| evsO1 \<in> kerberos;  A \<noteq> Spy;
+(* Leaking an authK... *)
+   Oops1: "\<lbrakk> evsO1 \<in> kerbIV;  A \<noteq> Spy;
               Says Kas A
-                (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-                                  AuthTicket|})  \<in> set evsO1;
-              ExpirAuth Tk evsO1 |]
-          ==> Says A Spy {|Agent A, Agent Tgs, Number Tk, Key AuthKey|}
-               # evsO1 \<in> kerberos"
+                (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                                  authTicket\<rbrace>)  \<in> set evsO1;
+              expiredAK Ta evsO1 \<rbrakk>
+          \<Longrightarrow> Says A Spy \<lbrace>Agent A, Agent Tgs, Number Ta, Key authK\<rbrace>
+               # evsO1 \<in> kerbIV"
 
 (*---------------------------------------------------------------------*)
 
-(*Leaking a ServKey... *)
-   Oops2: "[| evsO2 \<in> kerberos;  A \<noteq> Spy;
+(*Leaking a servK... *)
+   Oops2: "\<lbrakk> evsO2 \<in> kerbIV;  A \<noteq> Spy;
               Says Tgs A
-                (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
+                (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
                    \<in> set evsO2;
-              ExpirServ Tt evsO2 |]
-          ==> Says A Spy {|Agent A, Agent B, Number Tt, Key ServKey|}
-               # evsO2 \<in> kerberos"
+              expiredSK Ts evsO2 \<rbrakk>
+          \<Longrightarrow> Says A Spy \<lbrace>Agent A, Agent B, Number Ts, Key servK\<rbrace>
+               # evsO2 \<in> kerbIV"
 
 (*---------------------------------------------------------------------*)
 
@@ -249,7 +261,7 @@
 declare Fake_parts_insert_in_Un [dest]
 
 
-subsection{*Lemmas about Lists*}
+subsection{*Lemmas about lists, for reasoning about  Issues*}
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
@@ -284,104 +296,106 @@
 lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
 
 
-subsection{*Lemmas about @{term AuthKeys}*}
+subsection{*Lemmas about @{term authKeys}*}
 
-lemma AuthKeys_empty: "AuthKeys [] = {}"
-apply (unfold AuthKeys_def)
+lemma authKeys_empty: "authKeys [] = {}"
+apply (unfold authKeys_def)
 apply (simp (no_asm))
 done
 
-lemma AuthKeys_not_insert:
- "(\<forall>A Tk akey Peer.
-   ev \<noteq> Says Kas A (Crypt (shrK A) {|akey, Agent Peer, Tk,
-              (Crypt (shrK Peer) {|Agent A, Agent Peer, akey, Tk|})|}))
-       ==> AuthKeys (ev # evs) = AuthKeys evs"
-by (unfold AuthKeys_def, auto)
+lemma authKeys_not_insert:
+ "(\<forall>A Ta akey Peer.
+   ev \<noteq> Says Kas A (Crypt (shrK A) \<lbrace>akey, Agent Peer, Ta,
+              (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, akey, Ta\<rbrace>)\<rbrace>))
+       \<Longrightarrow> authKeys (ev # evs) = authKeys evs"
+by (unfold authKeys_def, auto)
 
-lemma AuthKeys_insert:
-  "AuthKeys
-     (Says Kas A (Crypt (shrK A) {|Key K, Agent Peer, Number Tk,
-      (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K, Number Tk|})|}) # evs)
-       = insert K (AuthKeys evs)"
-by (unfold AuthKeys_def, auto)
+lemma authKeys_insert:
+  "authKeys
+     (Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Peer, Number Ta,
+      (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K, Number Ta\<rbrace>)\<rbrace>) # evs)
+       = insert K (authKeys evs)"
+by (unfold authKeys_def, auto)
 
-lemma AuthKeys_simp:
-   "K \<in> AuthKeys
-    (Says Kas A (Crypt (shrK A) {|Key K', Agent Peer, Number Tk,
-     (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K', Number Tk|})|}) # evs)
-        ==> K = K' | K \<in> AuthKeys evs"
-by (unfold AuthKeys_def, auto)
+lemma authKeys_simp:
+   "K \<in> authKeys
+    (Says Kas A (Crypt (shrK A) \<lbrace>Key K', Agent Peer, Number Ta,
+     (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K', Number Ta\<rbrace>)\<rbrace>) # evs)
+        \<Longrightarrow> K = K' | K \<in> authKeys evs"
+by (unfold authKeys_def, auto)
 
-lemma AuthKeysI:
-   "Says Kas A (Crypt (shrK A) {|Key K, Agent Tgs, Number Tk,
-     (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key K, Number Tk|})|}) \<in> set evs
-        ==> K \<in> AuthKeys evs"
-by (unfold AuthKeys_def, auto)
+lemma authKeysI:
+   "Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Tgs, Number Ta,
+     (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key K, Number Ta\<rbrace>)\<rbrace>) \<in> set evs
+        \<Longrightarrow> K \<in> authKeys evs"
+by (unfold authKeys_def, auto)
 
-lemma AuthKeys_used: "K \<in> AuthKeys evs ==> Key K \<in> used evs"
-by (simp add: AuthKeys_def, blast)
+lemma authKeys_used: "K \<in> authKeys evs \<Longrightarrow> Key K \<in> used evs"
+by (simp add: authKeys_def, blast)
 
 
 subsection{*Forwarding Lemmas*}
 
 text{*--For reasoning about the encrypted portion of message K3--*}
 lemma K3_msg_in_parts_spies:
-     "Says Kas' A (Crypt KeyA {|AuthKey, Peer, Tk, AuthTicket|})
-               \<in> set evs ==> AuthTicket \<in> parts (spies evs)"
-by blast
+     "Says Kas' A (Crypt KeyA \<lbrace>authK, Peer, Ta, authTicket\<rbrace>)
+               \<in> set evs \<Longrightarrow> authTicket \<in> parts (spies evs)"
+apply blast
+done
 
 lemma Oops_range_spies1:
-     "[| Says Kas A (Crypt KeyA {|Key AuthKey, Peer, Tk, AuthTicket|})
+     "\<lbrakk> Says Kas A (Crypt KeyA \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>)
            \<in> set evs ;
-         evs \<in> kerberos |] ==> AuthKey \<notin> range shrK & AuthKey \<in> symKeys"
+         evs \<in> kerbIV \<rbrakk> \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys"
 apply (erule rev_mp)
-apply (erule kerberos.induct, auto)
+apply (erule kerbIV.induct, auto)
 done
 
 text{*--For reasoning about the encrypted portion of message K5--*}
 lemma K5_msg_in_parts_spies:
-     "Says Tgs' A (Crypt AuthKey {|ServKey, Agent B, Tt, ServTicket|})
-               \<in> set evs ==> ServTicket \<in> parts (spies evs)"
-by blast
+     "Says Tgs' A (Crypt authK \<lbrace>servK, Agent B, Ts, servTicket\<rbrace>)
+               \<in> set evs \<Longrightarrow> servTicket \<in> parts (spies evs)"
+apply blast
+done
 
 lemma Oops_range_spies2:
-     "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
+     "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>)
            \<in> set evs ;
-         evs \<in> kerberos |] ==> ServKey \<notin> range shrK & ServKey \<in> symKeys"
+         evs \<in> kerbIV \<rbrakk> \<Longrightarrow> servK \<notin> range shrK & servK \<in> symKeys"
 apply (erule rev_mp)
-apply (erule kerberos.induct, auto)
+apply (erule kerbIV.induct, auto)
 done
 
-lemma Says_ticket_in_parts_spies:
-     "Says S A (Crypt K {|SesKey, B, TimeStamp, Ticket|}) \<in> set evs
-      ==> Ticket \<in> parts (spies evs)"
-by blast
-
+lemma Says_ticket_parts:
+     "Says S A (Crypt K \<lbrace>SesKey, B, TimeStamp, Ticket\<rbrace>) \<in> set evs
+      \<Longrightarrow> Ticket \<in> parts (spies evs)"
+apply blast
+done
 
 (*Spy never sees another agent's shared key! (unless it's lost at start)*)
 lemma Spy_see_shrK [simp]:
-     "evs \<in> kerberos ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
-apply (erule kerberos.induct)
+     "evs \<in> kerbIV \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 apply (blast+)
 done
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> kerberos ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
+     "evs \<in> kerbIV \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[| Key (shrK A) \<in> parts (spies evs);  evs \<in> kerberos |] ==> A:bad"
+     "\<lbrakk> Key (shrK A) \<in> parts (spies evs);  evs \<in> kerbIV \<rbrakk> \<Longrightarrow> A:bad"
 by (blast dest: Spy_see_shrK)
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
 
 text{*Nobody can have used non-existent keys!*}
 lemma new_keys_not_used [simp]:
-    "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerberos|]
-     ==> K \<notin> keysFor (parts (spies evs))"
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbIV\<rbrakk>
+     \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake*}
@@ -393,121 +407,270 @@
 (*Earlier, all protocol proofs declared this theorem.
   But few of them actually need it! (Another is Yahalom) *)
 lemma new_keys_not_analzd:
- "[|evs \<in> kerberos; K \<in> symKeys; Key K \<notin> used evs|]
-  ==> K \<notin> keysFor (analz (spies evs))"
+ "\<lbrakk>evs \<in> kerbIV; K \<in> symKeys; Key K \<notin> used evs\<rbrakk>
+  \<Longrightarrow> K \<notin> keysFor (analz (spies evs))"
 by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD])
 
 
+
+subsection{*Lemmas for reasoning about predicate "before"*}
+
+lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_evs_rev: "used evs = used (rev evs)"
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply (simp add: used_Says_rev)
+apply (simp add: used_Gets_rev)
+apply (simp add: used_Notes_rev)
+done
+
+lemma used_takeWhile_used [rule_format]: 
+      "x : used (takeWhile P X) --> x : used X"
+apply (induct_tac "X")
+apply simp
+apply (induct_tac "a")
+apply (simp_all add: used_Nil)
+apply (blast dest!: initState_into_used)+
+done
+
+lemma set_evs_rev: "set evs = set (rev evs)"
+apply auto
+done
+
+lemma takeWhile_void [rule_format]:
+      "x \<notin> set evs \<longrightarrow> takeWhile (\<lambda>z. z \<noteq> x) evs = evs"
+apply auto
+done
+
+
 subsection{*Regularity Lemmas*}
 text{*These concern the form of items passed in messages*}
 
-text{*Describes the form of AuthKey, AuthTicket, and K sent by Kas*}
+text{*Describes the form of all components sent by Kas*}
 lemma Says_Kas_message_form:
-     "[| Says Kas A (Crypt K {|Key AuthKey, Agent Peer, Tk, AuthTicket|})
+     "\<lbrakk> Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
            \<in> set evs;
-         evs \<in> kerberos |]
-      ==> AuthKey \<notin> range shrK & AuthKey \<in> AuthKeys evs & AuthKey \<in> symKeys & 
-  AuthTicket = (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}) &
-             K = shrK A  & Peer = Tgs"
+         evs \<in> kerbIV \<rbrakk> \<Longrightarrow>  
+  K = shrK A  & Peer = Tgs &
+  authK \<notin> range shrK & authK \<in> authKeys evs & authK \<in> symKeys & 
+  authTicket = (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>) &
+  Key authK \<notin> used(before 
+           Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
+                   on evs) &
+  Ta = CT (before 
+           Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
+           on evs)"
+apply (unfold before_def)
 apply (erule rev_mp)
-apply (erule kerberos.induct)
-apply (simp_all (no_asm) add: AuthKeys_def AuthKeys_insert)
-apply (blast+)
+apply (erule kerbIV.induct)
+apply (simp_all (no_asm) add: authKeys_def authKeys_insert, blast, blast)
+txt{*K2*}
+apply (simp (no_asm) add: takeWhile_tail)
+apply (rule conjI)
+apply clarify
+apply (rule conjI)
+apply clarify
+apply (rule conjI)
+apply blast
+apply (rule conjI)
+apply clarify
+apply (rule conjI)
+txt{*subcase: used before*}
+apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
+                   used_takeWhile_used)
+txt{*subcase: CT before*}
+apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
+apply blast
+txt{*rest*}
+apply blast+
 done
 
+
+
 (*This lemma is essential for proving Says_Tgs_message_form:
 
-  the session key AuthKey
+  the session key authK
   supplied by Kas in the authentication ticket
   cannot be a long-term key!
 
-  Generalised to any session keys (both AuthKey and ServKey).
+  Generalised to any session keys (both authK and servK).
 *)
 lemma SesKey_is_session_key:
-     "[| Crypt (shrK Tgs_B) {|Agent A, Agent Tgs_B, Key SesKey, Number T|}
+     "\<lbrakk> Crypt (shrK Tgs_B) \<lbrace>Agent A, Agent Tgs_B, Key SesKey, Number T\<rbrace>
             \<in> parts (spies evs); Tgs_B \<notin> bad;
-         evs \<in> kerberos |]
-      ==> SesKey \<notin> range shrK"
+         evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> SesKey \<notin> range shrK"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
 done
 
-lemma A_trusts_AuthTicket:
-     "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}
+lemma authTicket_authentic:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
            \<in> parts (spies evs);
-         evs \<in> kerberos |]
-      ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk,
-                 Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}|})
+         evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                 Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
             \<in> set evs"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake, K4*}
 apply (blast+)
 done
 
-lemma AuthTicket_crypt_AuthKey:
-     "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}
+lemma authTicket_crypt_authK:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
            \<in> parts (spies evs);
-         evs \<in> kerberos |]
-      ==> AuthKey \<in> AuthKeys evs"
-apply (frule A_trusts_AuthTicket, assumption)
-apply (simp (no_asm) add: AuthKeys_def)
+         evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> authK \<in> authKeys evs"
+apply (frule authTicket_authentic, assumption)
+apply (simp (no_asm) add: authKeys_def)
 apply blast
 done
 
-text{*Describes the form of ServKey, ServTicket and AuthKey sent by Tgs*}
+text{*Describes the form of servK, servTicket and authK sent by Tgs*}
 lemma Says_Tgs_message_form:
-     "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
+     "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
            \<in> set evs;
-         evs \<in> kerberos |]
-   ==> B \<noteq> Tgs & 
-       ServKey \<notin> range shrK & ServKey \<notin> AuthKeys evs & ServKey \<in> symKeys &
-       ServTicket = (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}) &
-       AuthKey \<notin> range shrK & AuthKey \<in> AuthKeys evs & AuthKey \<in> symKeys"
+         evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> B \<noteq> Tgs & 
+      authK \<notin> range shrK & authK \<in> authKeys evs & authK \<in> symKeys &
+      servK \<notin> range shrK & servK \<notin> authKeys evs & servK \<in> symKeys &
+      servTicket = (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>) &
+      Key servK \<notin> used (before
+        Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+                        on evs) &
+      Ts = CT(before 
+        Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+              on evs) "
+apply (unfold before_def)
 apply (erule rev_mp)
-apply (erule kerberos.induct)
-apply (simp_all add: AuthKeys_insert AuthKeys_not_insert AuthKeys_empty AuthKeys_simp, blast, auto)
-txt{*Three subcases of Message 4*}
-apply (blast dest!: AuthKeys_used Says_Kas_message_form)
+apply (erule kerbIV.induct)
+apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast)
+txt{*We need this simplification only for Message 4*}
+apply (simp (no_asm) add: takeWhile_tail)
+apply auto
+txt{*Five subcases of Message 4*}
 apply (blast dest!: SesKey_is_session_key)
-apply (blast dest: AuthTicket_crypt_AuthKey)
+apply (blast dest: authTicket_crypt_authK)
+apply (blast dest!: authKeys_used Says_Kas_message_form)
+txt{*subcase: used before*}
+apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
+                   used_takeWhile_used)
+txt{*subcase: CT before*}
+apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
+done
+
+lemma authTicket_form:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         A \<notin> bad;
+         evs \<in> kerbIV \<rbrakk>
+    \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys & 
+        authTicket = Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>"
+apply (erule rev_mp)
+apply (erule kerbIV.induct)
+apply (frule_tac [7] K5_msg_in_parts_spies)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
+apply (blast+)
 done
 
-text{*Authenticity of AuthKey for A:
-     If a certain encrypted message appears then it originated with Kas*}
-lemma A_trusts_AuthKey:
-     "[| Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|}
+text{* This form holds also over an authTicket, but is not needed below.*}
+lemma servTicket_form:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>
+              \<in> parts (spies evs);
+            Key authK \<notin> analz (spies evs);
+            evs \<in> kerbIV \<rbrakk>
+         \<Longrightarrow> servK \<notin> range shrK & servK \<in> symKeys & 
+    (\<exists>A. servTicket = Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [7] K5_msg_in_parts_spies)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
+done
+
+text{* Essentially the same as @{text authTicket_form} *}
+lemma Says_kas_message_form:
+     "\<lbrakk> Says Kas' A (Crypt (shrK A)
+              \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>) \<in> set evs;
+         evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys & 
+          authTicket =
+                  Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>
+          | authTicket \<in> analz (spies evs)"
+by (blast dest: analz_shrK_Decrypt authTicket_form
+                Says_imp_spies [THEN analz.Inj])
+
+lemma Says_tgs_message_form:
+ "\<lbrakk> Says Tgs' A (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>)
+       \<in> set evs;  authK \<in> symKeys;
+     evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> servK \<notin> range shrK &
+      (\<exists>A. servTicket =
+	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+       | servTicket \<in> analz (spies evs)"
+apply (frule Says_imp_spies [THEN analz.Inj], auto)
+ apply (force dest!: servTicket_form)
+apply (frule analz_into_parts)
+apply (frule servTicket_form, auto)
+done
+
+
+subsection{*Authenticity theorems: confirm origin of sensitive messages*}
+
+lemma authK_authentic:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>
            \<in> parts (spies evs);
-         A \<notin> bad;  evs \<in> kerberos |]
-      ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|})
+         A \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>)
             \<in> set evs"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake*}
 apply blast
 txt{*K4*}
-apply (blast dest!: A_trusts_AuthTicket [THEN Says_Kas_message_form])
+apply (blast dest!: authTicket_authentic [THEN Says_Kas_message_form])
 done
 
-
 text{*If a certain encrypted message appears then it originated with Tgs*}
-lemma A_trusts_K4:
-     "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}
+lemma servK_authentic:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
-         Key AuthKey \<notin> analz (spies evs);
-         AuthKey \<notin> range shrK;
-         evs \<in> kerberos |]
- ==> \<exists>A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
+         Key authK \<notin> analz (spies evs);
+         authK \<notin> range shrK;
+         evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists>A. Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
        \<in> set evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake*}
@@ -518,269 +681,485 @@
 apply auto
 done
 
-lemma AuthTicket_form:
-     "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}
+lemma servK_authentic_bis:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
-         A \<notin> bad;
-         evs \<in> kerberos |]
-    ==> AuthKey \<notin> range shrK & AuthKey \<in> symKeys & 
-        AuthTicket = Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}"
+         Key authK \<notin> analz (spies evs);
+         B \<noteq> Tgs;
+         evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists>A. Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-apply (blast+)
+txt{*Fake*}
+apply blast
+txt{*K4*}
+apply blast
 done
 
-text{* This form holds also over an AuthTicket, but is not needed below.*}
-lemma ServTicket_form:
-     "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}
-              \<in> parts (spies evs);
-            Key AuthKey \<notin> analz (spies evs);
-            evs \<in> kerberos |]
-         ==> ServKey \<notin> range shrK & ServKey \<in> symKeys & 
-    (\<exists>A. ServTicket = Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})"
+text{*Authenticity of servK for B*}
+lemma servTicket_authentic_Tgs:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs); B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists>authK.
+       Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+       \<in> set evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (erule kerbIV.induct)
+apply (frule_tac [7] K5_msg_in_parts_spies)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
+apply blast+
+done
+
+text{*Anticipated here from next subsection*}
+lemma K4_imp_K2:
+"\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+      \<in> set evs;  evs \<in> kerbIV\<rbrakk>
+   \<Longrightarrow> \<exists>Ta. Says Kas A
+        (Crypt (shrK A)
+         \<lbrace>Key authK, Agent Tgs, Number Ta,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+        \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, auto)
+apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
+done
+
+text{*Anticipated here from next subsection*}
+lemma u_K4_imp_K2:
+"\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+      \<in> set evs; evs \<in> kerbIV\<rbrakk>
+   \<Longrightarrow> \<exists>Ta. (Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+             \<in> set evs
+          & servKlife + Ts <= authKlife + Ta)"
+apply (erule rev_mp)
+apply (erule kerbIV.induct)
+apply (frule_tac [7] K5_msg_in_parts_spies)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, auto)
+apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
 done
 
-text{* Essentially the same as @{text AuthTicket_form} *}
-lemma Says_kas_message_form:
-     "[| Says Kas' A (Crypt (shrK A)
-              {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \<in> set evs;
-         evs \<in> kerberos |]
-      ==> AuthKey \<notin> range shrK & AuthKey \<in> symKeys & 
-          AuthTicket =
-                  Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}
-          | AuthTicket \<in> analz (spies evs)"
-by (blast dest: analz_shrK_Decrypt AuthTicket_form
-                Says_imp_spies [THEN analz.Inj])
+lemma servTicket_authentic_Kas:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> \<exists>authK Ta.
+       Says Kas A
+         (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+            Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+        \<in> set evs"
+apply (blast dest!: servTicket_authentic_Tgs K4_imp_K2)
+done
+
+lemma u_servTicket_authentic_Kas:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> \<exists>authK Ta. Says Kas A (Crypt(shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+             \<in> set evs
+           & servKlife + Ts <= authKlife + Ta"
+apply (blast dest!: servTicket_authentic_Tgs u_K4_imp_K2)
+done
+
+lemma servTicket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists>Ta authK.
+     Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                   Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+       \<in> set evs
+     & Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+       \<in> set evs"
+apply (blast dest: servTicket_authentic_Tgs K4_imp_K2)
+done
+
+lemma u_servTicket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists>Ta authK.
+     (Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                   Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+       \<in> set evs
+     & Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+       \<in> set evs
+     & servKlife + Ts <= authKlife + Ta)"
+apply (blast dest: servTicket_authentic_Tgs u_K4_imp_K2)
+done
+
+lemma u_NotexpiredSK_NotexpiredAK:
+     "\<lbrakk> \<not> expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \<rbrakk>
+      \<Longrightarrow> \<not> expiredAK Ta evs"
+apply (blast dest: leI le_trans dest: leD)
+done
 
 
-lemma Says_tgs_message_form:
- "[| Says Tgs' A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
-       \<in> set evs;  AuthKey \<in> symKeys;
-     evs \<in> kerberos |]
-  ==> ServKey \<notin> range shrK &
-      (\<exists>A. ServTicket =
-	      Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})
-       | ServTicket \<in> analz (spies evs)"
-apply (frule Says_imp_spies [THEN analz.Inj], auto)
- apply (force dest!: ServTicket_form)
-apply (frule analz_into_parts)
-apply (frule ServTicket_form, auto)
+subsection{* Reliability: friendly agents send something if something else happened*}
+
+lemma K3_imp_K2:
+     "\<lbrakk> Says A Tgs
+             \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>
+           \<in> set evs;
+         A \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<exists>Ta. Says Kas A (Crypt (shrK A)
+                      \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
+                   \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV.induct)
+apply (frule_tac [7] K5_msg_in_parts_spies)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast, blast)
+apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN authK_authentic])
 done
 
-subsection{*Unicity Theorems*}
-
-text{* The session key, if secure, uniquely identifies the Ticket
-   whether AuthTicket or ServTicket. As a matter of fact, one can read
-   also Tgs in the place of B.                                     *}
-
-lemma unique_CryptKey:
-     "[| Crypt (shrK B)  {|Agent A,  Agent B,  Key SesKey, T|}
+text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
+lemma Key_unique_SesKey:
+     "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T, Ticket\<rbrace>
            \<in> parts (spies evs);
-         Crypt (shrK B') {|Agent A', Agent B', Key SesKey, T'|}
+         Crypt K' \<lbrace>Key SesKey,  Agent B', T', Ticket'\<rbrace>
            \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
-         evs \<in> kerberos |]
-      ==> A=A' & B=B' & T=T'"
+         evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> K=K' & B=B' & T=T' & Ticket=Ticket'"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake, K2, K4*}
 apply (blast+)
 done
 
-text{*An AuthKey is encrypted by one and only one Shared key.
-  A ServKey is encrypted by one and only one AuthKey.
-*}
-lemma Key_unique_SesKey:
-     "[| Crypt K  {|Key SesKey,  Agent B, T, Ticket|}
+lemma Tgs_authenticates_A:
+  "\<lbrakk>  Crypt authK \<lbrace>Agent A, Number T2\<rbrace> \<in> parts (spies evs); 
+      Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
            \<in> parts (spies evs);
-         Crypt K' {|Key SesKey,  Agent B', T', Ticket'|}
-           \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
-         evs \<in> kerberos |]
-      ==> K=K' & B=B' & T=T' & Ticket=Ticket'"
+      Key authK \<notin> analz (spies evs); A \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists> B. Says A Tgs \<lbrace>
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>,
+          Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B \<rbrace> \<in> set evs"  
+apply (drule authTicket_authentic, assumption, rotate_tac 4)
+apply (erule rev_mp, erule rev_mp, erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*Fake*}
+apply blast
+txt{*K2*}
+apply (force dest!: Crypt_imp_keysFor)
+txt{*K3*}
+apply (blast dest: Key_unique_SesKey)
+txt{*K5*}
+txt{*If authKa were compromised, so would be authK*}
+apply (case_tac "Key authKa \<in> analz (spies evs5)")
+apply (force dest!: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+txt{*Besides, since authKa originated with Kas anyway...*}
+apply (clarify, drule K3_imp_K2, assumption, assumption)
+apply (clarify, drule Says_Kas_message_form, assumption)
+txt{*...it cannot be a shared key*. Therefore @{text servK_authentic} applies. 
+     Contradition: Tgs used authK as a servkey, 
+     while Kas used it as an authkey*}
+apply (blast dest: servK_authentic Says_Tgs_message_form)
+done
+
+lemma Says_K5:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                                     servTicket\<rbrace>) \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad; B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> Says A B \<lbrace>servTicket, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+apply blast
+txt{*K3*}
+apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
+txt{*K4*}
+apply (force dest!: Crypt_imp_keysFor)
+txt{*K5*}
+apply (blast dest: Key_unique_SesKey)
+done
+
+text{*Anticipated here from next subsection*}
+lemma unique_CryptKey:
+     "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SesKey, T'\<rbrace>
+           \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
+         evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> A=A' & B=B' & T=T'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake, K2, K4*}
 apply (blast+)
 done
 
+lemma Says_K6:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                                     servTicket\<rbrace>) \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad; B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp))
+apply blast
+apply (force dest!: Crypt_imp_keysFor, clarify)
+apply (frule Says_Tgs_message_form, assumption, clarify) (*PROOF FAILED if omitted*)
+apply (blast dest: unique_CryptKey)
+done
+
+text{*Needs a unicity theorem, hence moved here*}
+lemma servK_authentic_ter:
+ "\<lbrakk> Says Kas A
+    (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>) \<in> set evs;
+     Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+       \<in> parts (spies evs);
+     Key authK \<notin> analz (spies evs);
+     evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
+apply (frule Says_Kas_message_form, assumption)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [7] K5_msg_in_parts_spies)
+apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
+txt{*K2 and K4 remain*}
+prefer 2 apply (blast dest!: unique_CryptKey)
+apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
+done
+
+
+subsection{*Unicity Theorems*}
+
+text{* The session key, if secure, uniquely identifies the Ticket
+   whether authTicket or servTicket. As a matter of fact, one can read
+   also Tgs in the place of B.                                     *}
+
 
 (*
   At reception of any message mentioning A, Kas associates shrK A with
-  a new AuthKey. Realistic, as the user gets a new AuthKey at each login.
-  Similarly, at reception of any message mentioning an AuthKey
+  a new authK. Realistic, as the user gets a new authK at each login.
+  Similarly, at reception of any message mentioning an authK
   (a legitimate user could make several requests to Tgs - by K3), Tgs
-  associates it with a new ServKey.
+  associates it with a new servK.
 
   Therefore, a goal like
 
-   "evs \<in> kerberos
-     ==> Key Kc \<notin> analz (spies evs) -->
+   "evs \<in> kerbIV
+     \<Longrightarrow> Key Kc \<notin> analz (spies evs) \<longrightarrow>
            (\<exists>K' B' T' Ticket'. \<forall>K B T Ticket.
-            Crypt Kc {|Key K, Agent B, T, Ticket|}
-             \<in> parts (spies evs) --> K=K' & B=B' & T=T' & Ticket=Ticket')"
+            Crypt Kc \<lbrace>Key K, Agent B, T, Ticket\<rbrace>
+             \<in> parts (spies evs) \<longrightarrow> K=K' & B=B' & T=T' & Ticket=Ticket')"
 
   would fail on the K2 and K4 cases.
 *)
 
-lemma unique_AuthKeys:
-     "[| Says Kas A
-              (Crypt Ka {|Key AuthKey, Agent Tgs, Tk, X|}) \<in> set evs;
+lemma unique_authKeys:
+     "\<lbrakk> Says Kas A
+              (Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta, X\<rbrace>) \<in> set evs;
          Says Kas A'
-              (Crypt Ka' {|Key AuthKey, Agent Tgs, Tk', X'|}) \<in> set evs;
-         evs \<in> kerberos |] ==> A=A' & Ka=Ka' & Tk=Tk' & X=X'"
+              (Crypt Ka' \<lbrace>Key authK, Agent Tgs, Ta', X'\<rbrace>) \<in> set evs;
+         evs \<in> kerbIV \<rbrakk> \<Longrightarrow> A=A' & Ka=Ka' & Ta=Ta' & X=X'"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*K2*}
 apply blast
 done
 
-text{* ServKey uniquely identifies the message from Tgs *}
-lemma unique_ServKeys:
-     "[| Says Tgs A
-              (Crypt K {|Key ServKey, Agent B, Tt, X|}) \<in> set evs;
+text{* servK uniquely identifies the message from Tgs *}
+lemma unique_servKeys:
+     "\<lbrakk> Says Tgs A
+              (Crypt K \<lbrace>Key servK, Agent B, Ts, X\<rbrace>) \<in> set evs;
          Says Tgs A'
-              (Crypt K' {|Key ServKey, Agent B', Tt', X'|}) \<in> set evs;
-         evs \<in> kerberos |] ==> A=A' & B=B' & K=K' & Tt=Tt' & X=X'"
+              (Crypt K' \<lbrace>Key servK, Agent B', Ts', X'\<rbrace>) \<in> set evs;
+         evs \<in> kerbIV \<rbrakk> \<Longrightarrow> A=A' & B=B' & K=K' & Ts=Ts' & X=X'"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*K4*}
 apply blast
 done
 
+text{* Revised unicity theorems *}
 
-subsection{*Lemmas About the Predicate @{term KeyCryptKey}*}
+lemma Kas_Unique:
+     "\<lbrakk> Says Kas A
+              (Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>) \<in> set evs;
+        evs \<in> kerbIV \<rbrakk> \<Longrightarrow> 
+   Unique (Says Kas A (Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>)) 
+   on evs"
+apply (erule rev_mp, erule kerbIV.induct, simp_all add: Unique_def)
+apply blast
+done
 
-lemma not_KeyCryptKey_Nil [iff]: "~ KeyCryptKey AuthKey ServKey []"
-by (simp add: KeyCryptKey_def)
+lemma Tgs_Unique:
+     "\<lbrakk> Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>) \<in> set evs;
+        evs \<in> kerbIV \<rbrakk> \<Longrightarrow> 
+  Unique (Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>)) 
+  on evs"
+apply (erule rev_mp, erule kerbIV.induct, simp_all add: Unique_def)
+apply blast
+done
 
-lemma KeyCryptKeyI:
- "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |}) \<in> set evs;
-     evs \<in> kerberos |] ==> KeyCryptKey AuthKey ServKey evs"
-apply (unfold KeyCryptKey_def)
+
+subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
+
+lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
+by (simp add: AKcryptSK_def)
+
+lemma AKcryptSKI:
+ "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>) \<in> set evs;
+     evs \<in> kerbIV \<rbrakk> \<Longrightarrow> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
 apply (blast dest: Says_Tgs_message_form)
 done
 
-lemma KeyCryptKey_Says [simp]:
-   "KeyCryptKey AuthKey ServKey (Says S A X # evs) =
+lemma AKcryptSK_Says [simp]:
+   "AKcryptSK authK servK (Says S A X # evs) =
      (Tgs = S &
-      (\<exists>B tt. X = Crypt AuthKey
-                {|Key ServKey, Agent B, tt,
-                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |})
-     | KeyCryptKey AuthKey ServKey evs)"
-apply (unfold KeyCryptKey_def)
+      (\<exists>B Ts. X = Crypt authK
+                \<lbrace>Key servK, Agent B, Number Ts,
+                  Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>)
+     | AKcryptSK authK servK evs)"
+apply (unfold AKcryptSK_def)
 apply (simp (no_asm))
 apply blast
 done
 
-(*A fresh AuthKey cannot be associated with any other
+(*A fresh authK cannot be associated with any other
   (with respect to a given trace). *)
-lemma Auth_fresh_not_KeyCryptKey:
-     "[| Key AuthKey \<notin> used evs; evs \<in> kerberos |]
-      ==> ~ KeyCryptKey AuthKey ServKey evs"
-apply (unfold KeyCryptKey_def)
+lemma Auth_fresh_not_AKcryptSK:
+     "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
 done
 
-(*A fresh ServKey cannot be associated with any other
+(*A fresh servK cannot be associated with any other
   (with respect to a given trace). *)
-lemma Serv_fresh_not_KeyCryptKey:
- "Key ServKey \<notin> used evs ==> ~ KeyCryptKey AuthKey ServKey evs"
-apply (unfold KeyCryptKey_def, blast)
+lemma Serv_fresh_not_AKcryptSK:
+ "Key servK \<notin> used evs \<Longrightarrow> \<not> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def, blast)
 done
 
-lemma AuthKey_not_KeyCryptKey:
-     "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, tk|}
-           \<in> parts (spies evs);  evs \<in> kerberos |]
-      ==> ~ KeyCryptKey K AuthKey evs"
+lemma authK_not_AKcryptSK:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, tk\<rbrace>
+           \<in> parts (spies evs);  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK K authK evs"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 txt{*Fake*}
 apply blast
 txt{*K2: by freshness*}
-apply (simp add: KeyCryptKey_def)
+apply (simp add: AKcryptSK_def)
 txt{*K4*}
 apply (blast+)
 done
 
 text{*A secure serverkey cannot have been used to encrypt others*}
-lemma ServKey_not_KeyCryptKey:
- "[| Crypt (shrK B) {|Agent A, Agent B, Key SK, tt|} \<in> parts (spies evs);
+lemma servK_not_AKcryptSK:
+ "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, Number Ts\<rbrace> \<in> parts (spies evs);
      Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
-     B \<noteq> Tgs;  evs \<in> kerberos |]
-  ==> ~ KeyCryptKey SK K evs"
+     B \<noteq> Tgs;  evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> \<not> AKcryptSK SK K evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
 txt{*K4 splits into distinct subcases*}
 apply auto
-txt{*ServKey can't have been enclosed in two certificates*}
+txt{*servK can't have been enclosed in two certificates*}
  prefer 2 apply (blast dest: unique_CryptKey)
-txt{*ServKey is fresh and so could not have been used, by
+txt{*servK is fresh and so could not have been used, by
    @{text new_keys_not_used}*}
-apply (force dest!: Crypt_imp_invKey_keysFor simp add: KeyCryptKey_def)
+apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
 done
 
-text{*Long term keys are not issued as ServKeys*}
-lemma shrK_not_KeyCryptKey:
-     "evs \<in> kerberos ==> ~ KeyCryptKey K (shrK A) evs"
-apply (unfold KeyCryptKey_def)
-apply (erule kerberos.induct)
+text{*Long term keys are not issued as servKeys*}
+lemma shrK_not_AKcryptSK:
+     "evs \<in> kerbIV \<Longrightarrow> \<not> AKcryptSK K (shrK A) evs"
+apply (unfold AKcryptSK_def)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, auto)
 done
 
-text{*The Tgs message associates ServKey with AuthKey and therefore not with any
-  other key AuthKey.*}
-lemma Says_Tgs_KeyCryptKey:
-     "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |})
+text{*The Tgs message associates servK with authK and therefore not with any
+  other key authK.*}
+lemma Says_Tgs_AKcryptSK:
+     "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>)
            \<in> set evs;
-         AuthKey' \<noteq> AuthKey;  evs \<in> kerberos |]
-      ==> ~ KeyCryptKey AuthKey' ServKey evs"
-apply (unfold KeyCryptKey_def)
-apply (blast dest: unique_ServKeys)
+         authK' \<noteq> authK;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK' servK evs"
+apply (unfold AKcryptSK_def)
+apply (blast dest: unique_servKeys)
 done
 
-lemma KeyCryptKey_not_KeyCryptKey:
-     "[| KeyCryptKey AuthKey ServKey evs;  evs \<in> kerberos |]
-      ==> ~ KeyCryptKey ServKey K evs"
+text{*Equivalently*}
+lemma not_different_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;
+        authK' \<noteq> authK;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK' servK evs  \<and> servK \<in> symKeys"
+apply (simp add: AKcryptSK_def)
+apply (blast dest: unique_servKeys Says_Tgs_message_form)
+done
+
+lemma AKcryptSK_not_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK K evs"
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, safe)
 txt{*K4 splits into subcases*}
 apply simp_all
-prefer 4 apply (blast dest!: AuthKey_not_KeyCryptKey)
-txt{*ServKey is fresh and so could not have been used, by
+prefer 4 apply (blast dest!: authK_not_AKcryptSK)
+txt{*servK is fresh and so could not have been used, by
    @{text new_keys_not_used}*}
  prefer 2 
- apply (force dest!: Crypt_imp_invKey_keysFor simp add: KeyCryptKey_def)
+ apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
 txt{*Others by freshness*}
 apply (blast+)
 done
@@ -791,31 +1170,31 @@
 text{*We take some pains to express the property
   as a logical equivalence so that the simplifier can apply it.*}
 lemma Key_analz_image_Key_lemma:
-     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K:KK | Key K \<in> analz H)
-      ==>
-      P --> (Key K \<in> analz (Key`KK Un H)) = (K:KK | Key K \<in> analz H)"
+     "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
+      \<Longrightarrow>
+      P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) = (K:KK | Key K \<in> analz H)"
 by (blast intro: analz_mono [THEN subsetD])
 
 
-lemma KeyCryptKey_analz_insert:
-     "[| KeyCryptKey K K' evs; K \<in> symKeys; evs \<in> kerberos |]
-      ==> Key K' \<in> analz (insert (Key K) (spies evs))"
-apply (simp add: KeyCryptKey_def, clarify)
+lemma AKcryptSK_analz_insert:
+     "\<lbrakk> AKcryptSK K K' evs; K \<in> symKeys; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key K' \<in> analz (insert (Key K) (spies evs))"
+apply (simp add: AKcryptSK_def, clarify)
 apply (drule Says_imp_spies [THEN analz.Inj, THEN analz_insertI], auto)
 done
 
-lemma AuthKeys_are_not_KeyCryptKey:
-     "[| K \<in> AuthKeys evs Un range shrK;  evs \<in> kerberos |]
-      ==> \<forall>SK. ~ KeyCryptKey SK K evs"
-apply (simp add: KeyCryptKey_def)
-apply (blast dest: Says_Tgs_message_form)
+lemma authKeys_are_not_AKcryptSK:
+     "\<lbrakk> K \<in> authKeys evs Un range shrK;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<forall>SK. \<not> AKcryptSK SK K evs \<and> K \<in> symKeys"
+apply (simp add: authKeys_def AKcryptSK_def)
+apply (blast dest: Says_Kas_message_form Says_Tgs_message_form)
 done
 
-lemma not_AuthKeys_not_KeyCryptKey:
-     "[| K \<notin> AuthKeys evs;
-         K \<notin> range shrK; evs \<in> kerberos |]
-      ==> \<forall>SK. ~ KeyCryptKey K SK evs"
-apply (simp add: KeyCryptKey_def)
+lemma not_authKeys_not_AKcryptSK:
+     "\<lbrakk> K \<notin> authKeys evs;
+         K \<notin> range shrK; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> \<forall>SK. \<not> AKcryptSK K SK evs"
+apply (simp add: AKcryptSK_def)
 apply (blast dest: Says_Tgs_message_form)
 done
 
@@ -823,28 +1202,27 @@
 subsection{*Secrecy Theorems*}
 
 text{*For the Oops2 case of the next theorem*}
-lemma Oops2_not_KeyCryptKey:
-     "[| evs \<in> kerberos;
-         Says Tgs A (Crypt AuthKey
-                     {|Key ServKey, Agent B, Number Tt, ServTicket|})
-           \<in> set evs |]
-      ==> ~ KeyCryptKey ServKey SK evs"
-apply (blast dest: KeyCryptKeyI KeyCryptKey_not_KeyCryptKey)
+lemma Oops2_not_AKcryptSK:
+     "\<lbrakk> evs \<in> kerbIV;
+         Says Tgs A (Crypt authK
+                     \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK SK evs"
+apply (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
 done
-
-
+   
 text{* Big simplification law for keys SK that are not crypted by keys in KK
  It helps prove three, otherwise hard, facts about keys. These facts are
  exploited as simplification laws for analz, and also "limit the damage"
  in case of loss of a key to the spy. See ESORICS98.
  [simplified by LCP] *}
 lemma Key_analz_image_Key [rule_format (no_asm)]:
-     "evs \<in> kerberos ==>
-      (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) -->
-       (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs)   -->
+     "evs \<in> kerbIV \<Longrightarrow>
+      (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
+       (\<forall>K \<in> KK. \<not> AKcryptSK K SK evs)   \<longrightarrow>
        (Key SK \<in> analz (Key`KK Un (spies evs))) =
        (SK \<in> KK | Key SK \<in> analz (spies evs)))"
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (frule_tac [10] Oops_range_spies2)
 apply (frule_tac [9] Oops_range_spies1)
 apply (frule_tac [7] Says_tgs_message_form)
@@ -852,46 +1230,42 @@
 apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
 txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
  the induction hypothesis*}
-apply (case_tac [11] "KeyCryptKey AuthKey SK evsO1")
-apply (case_tac [8] "KeyCryptKey ServKey SK evs5")
+apply (case_tac [11] "AKcryptSK authK SK evsO1")
+apply (case_tac [8] "AKcryptSK servK SK evs5")
 apply (simp_all del: image_insert
-          add: analz_image_freshK_simps KeyCryptKey_Says)
+        add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
+             Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
+       Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
 txt{*Fake*} 
 apply spy_analz
-apply (simp_all del: image_insert
-          add: shrK_not_KeyCryptKey
-               Oops2_not_KeyCryptKey Auth_fresh_not_KeyCryptKey
-               Serv_fresh_not_KeyCryptKey Says_Tgs_KeyCryptKey Spy_analz_shrK)
-  --{*Splitting the @{text simp_all} into two parts makes it faster.*}
 txt{*K2*}
 apply blast 
 txt{*K3*}
 apply blast 
 txt{*K4*}
-apply (blast dest!: AuthKey_not_KeyCryptKey)
+apply (blast dest!: authK_not_AKcryptSK)
 txt{*K5*}
-apply (case_tac "Key ServKey \<in> analz (spies evs5) ")
-txt{*If ServKey is compromised then the result follows directly...*}
+apply (case_tac "Key servK \<in> analz (spies evs5) ")
+txt{*If servK is compromised then the result follows directly...*}
 apply (simp (no_asm_simp) add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
-txt{*...therefore ServKey is uncompromised.*}
-txt{*The KeyCryptKey ServKey SK evs5 case leads to a contradiction.*}
-apply (blast elim!: ServKey_not_KeyCryptKey [THEN [2] rev_notE] del: allE ballE)
+txt{*...therefore servK is uncompromised.*}
+txt{*The AKcryptSK servK SK evs5 case leads to a contradiction.*}
+apply (blast elim!: servK_not_AKcryptSK [THEN [2] rev_notE] del: allE ballE)
 txt{*Another K5 case*}
 apply blast 
 txt{*Oops1*}
 apply simp 
-apply (blast dest!: KeyCryptKey_analz_insert)
+apply (blast dest!: AKcryptSK_analz_insert)
 done
 
 text{* First simplification law for analz: no session keys encrypt
 authentication keys or shared keys. *}
 lemma analz_insert_freshK1:
-     "[| evs \<in> kerberos;  K \<in> (AuthKeys evs) Un range shrK;
-         K \<in> symKeys;
-         SesKey \<notin> range shrK |]
-      ==> (Key K \<in> analz (insert (Key SesKey) (spies evs))) =
+     "\<lbrakk> evs \<in> kerbIV;  K \<in> authKeys evs Un range shrK;
+        SesKey \<notin> range shrK \<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key SesKey) (spies evs))) =
           (K = SesKey | Key K \<in> analz (spies evs))"
-apply (frule AuthKeys_are_not_KeyCryptKey, assumption)
+apply (frule authKeys_are_not_AKcryptSK, assumption)
 apply (simp del: image_insert
             add: analz_image_freshK_simps add: Key_analz_image_Key)
 done
@@ -899,53 +1273,59 @@
 
 text{* Second simplification law for analz: no service keys encrypt any other keys.*}
 lemma analz_insert_freshK2:
-     "[| evs \<in> kerberos;  ServKey \<notin> (AuthKeys evs); ServKey \<notin> range shrK;
-         K \<in> symKeys|]
-      ==> (Key K \<in> analz (insert (Key ServKey) (spies evs))) =
-          (K = ServKey | Key K \<in> analz (spies evs))"
-apply (frule not_AuthKeys_not_KeyCryptKey, assumption, assumption)
+     "\<lbrakk> evs \<in> kerbIV;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
+        K \<in> symKeys \<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key servK) (spies evs))) =
+          (K = servK | Key K \<in> analz (spies evs))"
+apply (frule not_authKeys_not_AKcryptSK, assumption, assumption)
 apply (simp del: image_insert
             add: analz_image_freshK_simps add: Key_analz_image_Key)
 done
 
 
-text{* Third simplification law for analz: only one authentication key encrypts a
-certain service key.*}
+text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
+
 lemma analz_insert_freshK3:
- "[| Says Tgs A
-            (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
-        \<in> set evs;  ServKey \<in> symKeys;
-     AuthKey \<noteq> AuthKey'; AuthKey' \<notin> range shrK; evs \<in> kerberos |]
-        ==> (Key ServKey \<in> analz (insert (Key AuthKey') (spies evs))) =
-                (ServKey = AuthKey' | Key ServKey \<in> analz (spies evs))"
-apply (drule_tac AuthKey' = AuthKey' in Says_Tgs_KeyCryptKey, blast, assumption)
+ "\<lbrakk> AKcryptSK authK servK evs;
+    authK' \<noteq> authK; authK' \<notin> range shrK; evs \<in> kerbIV \<rbrakk>
+        \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
+                (servK = authK' | Key servK \<in> analz (spies evs))"
+apply (drule_tac authK' = authK' in not_different_AKcryptSK, blast, assumption)
 apply (simp del: image_insert
             add: analz_image_freshK_simps add: Key_analz_image_Key)
 done
 
+lemma analz_insert_freshK3_bis:
+ "\<lbrakk> Says Tgs A
+            (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+        \<in> set evs; 
+     authK \<noteq> authK'; authK' \<notin> range shrK; evs \<in> kerbIV \<rbrakk>
+        \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
+                (servK = authK' | Key servK \<in> analz (spies evs))"
+apply (frule AKcryptSKI, assumption)
+apply (simp add: analz_insert_freshK3)
+done
 
 text{*a weakness of the protocol*}
-lemma AuthKey_compromises_ServKey:
-     "[| Says Tgs A
-              (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
-           \<in> set evs;  AuthKey \<in> symKeys;
-         Key AuthKey \<in> analz (spies evs); evs \<in> kerberos |]
-      ==> Key ServKey \<in> analz (spies evs)"
+lemma authK_compromises_servK:
+     "\<lbrakk> Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs;  authK \<in> symKeys;
+         Key authK \<in> analz (spies evs); evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<in> analz (spies evs)"
 by (force dest: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
 
-
-subsection{* Guarantees for Kas *}
-lemma ServKey_notin_AuthKeysD:
-     "[| Crypt AuthKey {|Key ServKey, Agent B, Tt,
-                      Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|}
+lemma servK_notin_authKeysD:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts,
+                      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>\<rbrace>
            \<in> parts (spies evs);
-         Key ServKey \<notin> analz (spies evs);
-         B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> ServKey \<notin> AuthKeys evs"
+         Key servK \<notin> analz (spies evs);
+         B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> servK \<notin> authKeys evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (simp add: AuthKeys_def)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (simp add: authKeys_def)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
 apply (blast+)
@@ -955,15 +1335,15 @@
 text{*If Spy sees the Authentication Key sent in msg K2, then
     the Key has expired.*}
 lemma Confidentiality_Kas_lemma [rule_format]:
-     "[| AuthKey \<in> symKeys; A \<notin> bad;  evs \<in> kerberos |]
-      ==> Says Kas A
+     "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Says Kas A
                (Crypt (shrK A)
-                  {|Key AuthKey, Agent Tgs, Number Tk,
-          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-            \<in> set evs -->
-          Key AuthKey \<in> analz (spies evs) -->
-          ExpirAuth Tk evs"
-apply (erule kerberos.induct)
+                  \<lbrace>Key authK, Agent Tgs, Number Ta,
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+            \<in> set evs \<longrightarrow>
+          Key authK \<in> analz (spies evs) \<longrightarrow>
+          expiredAK Ta evs"
+apply (erule kerbIV.induct)
 apply (frule_tac [10] Oops_range_spies2)
 apply (frule_tac [9] Oops_range_spies1)
 apply (frule_tac [7] Says_tgs_message_form)
@@ -977,584 +1357,569 @@
 txt{*K4*}
 apply blast
 txt{*Level 8: K5*}
-apply (blast dest: ServKey_notin_AuthKeysD Says_Kas_message_form intro: less_SucI)
+apply (blast dest: servK_notin_authKeysD Says_Kas_message_form intro: less_SucI)
 txt{*Oops1*}
-apply (blast dest!: unique_AuthKeys intro: less_SucI)
+apply (blast dest!: unique_authKeys intro: less_SucI)
 txt{*Oops2*}
 apply (blast dest: Says_Tgs_message_form Says_Kas_message_form)
 done
 
 lemma Confidentiality_Kas:
-     "[| Says Kas A
-              (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})
+     "\<lbrakk> Says Kas A
+              (Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
            \<in> set evs;
-         ~ ExpirAuth Tk evs;
-         A \<notin> bad;  evs \<in> kerberos |]
-      ==> Key AuthKey \<notin> analz (spies evs)"
+         \<not> expiredAK Ta evs;
+         A \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key authK \<notin> analz (spies evs)"
 by (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
 
-
-subsection{* Guarantees for Tgs *}
-
 text{*If Spy sees the Service Key sent in msg K4, then
     the Key has expired.*}
 
 lemma Confidentiality_lemma [rule_format]:
-     "[| Says Tgs A
-	    (Crypt AuthKey
-	       {|Key ServKey, Agent B, Number Tt,
-		 Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})
+     "\<lbrakk> Says Tgs A
+	    (Crypt authK
+	       \<lbrace>Key servK, Agent B, Number Ts,
+		 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
 	   \<in> set evs;
-	Key AuthKey \<notin> analz (spies evs);
-        ServKey \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Key ServKey \<in> analz (spies evs) -->
-	  ExpirServ Tt evs"
+	Key authK \<notin> analz (spies evs);
+        servK \<in> symKeys;
+	A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
+	  expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct)
+apply (erule kerbIV.induct)
 apply (rule_tac [9] impI)+;
   --{*The Oops1 case is unusual: must simplify
     @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
    @{text analz_mono_contra} weaken it to
    @{term "Authkey \<notin> analz (spies evs)"},
-  for we then conclude @{term "AuthKey \<noteq> AuthKeya"}.*}
+  for we then conclude @{term "authK \<noteq> authKa"}.*}
 apply analz_mono_contra
 apply (frule_tac [10] Oops_range_spies2)
 apply (frule_tac [9] Oops_range_spies1)
 apply (frule_tac [7] Says_tgs_message_form)
 apply (frule_tac [5] Says_kas_message_form)
 apply (safe del: impI conjI impCE)
-apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3 pushes)
+apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
 txt{*Fake*}
 apply spy_analz
 txt{*K2*}
 apply (blast intro: parts_insertI less_SucI)
 txt{*K4*}
-apply (blast dest: A_trusts_AuthTicket Confidentiality_Kas)
+apply (blast dest: authTicket_authentic Confidentiality_Kas)
 txt{*Oops2*}
   prefer 3
   apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
-txt{*Oops1*}
+txt{*Oops1*} 
  prefer 2
- apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
-txt{*K5.  Not clear how this step could be integrated with the main
-       simplification step.*}
+apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
+txt{*K5. Not obvious how this step could be integrated with the main
+       simplification step. Done in KerberosV.thy *}
 apply clarify
 apply (erule_tac V = "Says Aa Tgs ?X \<in> set ?evs" in thin_rl)
-apply (frule Says_imp_spies [THEN parts.Inj, THEN ServKey_notin_AuthKeysD])
+apply (frule Says_imp_spies [THEN parts.Inj, THEN servK_notin_authKeysD])
 apply (assumption, blast, assumption)
 apply (simp add: analz_insert_freshK2)
 apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
 done
 
 
-text{* In the real world Tgs can't check wheter AuthKey is secure! *}
-lemma Confidentiality_Tgs1:
-     "[| Says Tgs A
-              (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
+text{* In the real world Tgs can't check wheter authK is secure! *}
+lemma Confidentiality_Tgs:
+     "\<lbrakk> Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
            \<in> set evs;
-         Key AuthKey \<notin> analz (spies evs);
-         ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Key ServKey \<notin> analz (spies evs)"
+         Key authK \<notin> analz (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 apply (blast dest: Says_Tgs_message_form Confidentiality_lemma)
 done
 
 text{* In the real world Tgs CAN check what Kas sends! *}
-lemma Confidentiality_Tgs2:
-     "[| Says Kas A
-               (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})
+lemma Confidentiality_Tgs_bis:
+     "\<lbrakk> Says Kas A
+               (Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
            \<in> set evs;
          Says Tgs A
-              (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
+              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
            \<in> set evs;
-         ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Key ServKey \<notin> analz (spies evs)"
-apply (blast dest!: Confidentiality_Kas Confidentiality_Tgs1)
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
 done
 
 text{*Most general form*}
-lemmas Confidentiality_Tgs3 = A_trusts_AuthTicket [THEN Confidentiality_Tgs2]
+lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
+
+lemmas Confidentiality_Auth_A = authK_authentic [THEN Confidentiality_Kas]
+
+text{*Needs a confidentiality guarantee, hence moved here.
+      Authenticity of servK for A*}
+lemma servK_authentic_bis_r:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; A \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow>Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
+apply (blast dest: authK_authentic Confidentiality_Auth_A servK_authentic_ter)
+done
+
+lemma Confidentiality_Serv_A:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (drule authK_authentic, assumption, assumption)
+apply (blast dest: Confidentiality_Kas Says_Kas_message_form servK_authentic_ter Confidentiality_Tgs_bis)
+done
+
+lemma Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (frule authK_authentic)
+apply (frule_tac [3] Confidentiality_Kas)
+apply (frule_tac [6] servTicket_authentic, auto)
+apply (blast dest!: Confidentiality_Tgs_bis dest: Says_Kas_message_form servK_authentic unique_servKeys unique_authKeys)
+done
+(*
+The proof above is fast.  It can be done in one command in 17 secs:
+apply (blast dest: authK_authentic servK_authentic
+                               Says_Kas_message_form servTicket_authentic
+                               unique_servKeys unique_authKeys
+                               Confidentiality_Kas
+                               Confidentiality_Tgs_bis)
+It is very brittle: we can't use this command partway
+through the script above.
+*)
+
+lemma u_Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad;  B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest: u_servTicket_authentic u_NotexpiredSK_NotexpiredAK Confidentiality_Tgs_bis)
+done
 
 
-subsection{* Guarantees for Alice *}
+
+subsection{*Parties authentication: each party verifies "the identity of
+       another party who generated some data" (quoted from Neuman and Ts'o).*}
 
-lemmas Confidentiality_Auth_A = A_trusts_AuthKey [THEN Confidentiality_Kas]
+text{*These guarantees don't assess whether two parties agree on
+         the same session key: sending a message containing a key
+         doesn't a priori state knowledge of the key.*}
+
 
-lemma A_trusts_K4_bis:
- "[| Says Kas A
-       (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \<in> set evs;
-     Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}
+text{*@{text Tgs_authenticates_A} can be found above*}
+
+lemma A_authenticates_Tgs:
+ "\<lbrakk> Says Kas A
+    (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>) \<in> set evs;
+     Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
        \<in> parts (spies evs);
-     Key AuthKey \<notin> analz (spies evs);
-     evs \<in> kerberos |]
- ==> Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
+     Key authK \<notin> analz (spies evs);
+     evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
        \<in> set evs"
 apply (frule Says_Kas_message_form, assumption)
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
+apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
 txt{*K2 and K4 remain*}
 prefer 2 apply (blast dest!: unique_CryptKey)
-apply (blast dest!: A_trusts_K4 Says_Tgs_message_form AuthKeys_used)
-done
-
-
-lemma Confidentiality_Serv_A:
-     "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
-           \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-           \<in> parts (spies evs);
-         ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Key ServKey \<notin> analz (spies evs)"
-apply (drule A_trusts_AuthKey, assumption, assumption)
-apply (blast dest: Confidentiality_Kas Says_Kas_message_form A_trusts_K4_bis Confidentiality_Tgs2)
+apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
 done
 
 
-subsection{* Guarantees for Bob *}
-text{* Theorems for the refined model have suffix "refined"                *}
-
-lemma K4_imp_K2:
-"[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
-      \<in> set evs;  evs \<in> kerberos|]
-   ==> \<exists>Tk. Says Kas A
-        (Crypt (shrK A)
-         {|Key AuthKey, Agent Tgs, Number Tk,
-           Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-        \<in> set evs"
-apply (erule rev_mp)
-apply (erule kerberos.induct)
-apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, auto)
-apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN A_trusts_AuthTicket])
+lemma B_authenticates_A:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+        Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+        Key servK \<notin> analz (spies evs);
+        A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>,
+               Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs"
+apply (blast dest: servTicket_authentic_Tgs intro: Says_K5)
 done
 
-lemma K4_imp_K2_refined:
-"[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
-      \<in> set evs; evs \<in> kerberos|]
-   ==> \<exists>Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-           Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-             \<in> set evs
-          & ServLife + Tt <= AuthLife + Tk)"
-apply (erule rev_mp)
-apply (erule kerberos.induct)
-apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, auto)
-apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN A_trusts_AuthTicket])
-done
-
-text{*Authenticity of ServKey for B*}
-lemma B_trusts_ServKey:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}
-           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
-         evs \<in> kerberos |]
- ==> \<exists>AuthKey.
-       Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt,
-                   Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|})
-       \<in> set evs"
-apply (erule rev_mp)
-apply (erule kerberos.induct)
-apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-apply (blast+)
+text{*The second assumption tells B what kind of key servK is.*}
+lemma B_authenticates_A_r:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>,
+                  Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
+apply (blast intro: Says_K5 dest: Confidentiality_B servTicket_authentic_Tgs)
 done
 
-lemma B_trusts_ServTicket_Kas:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
-         evs \<in> kerberos |]
-  ==> \<exists>AuthKey Tk.
-       Says Kas A
-         (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-            Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-        \<in> set evs"
-by (blast dest!: B_trusts_ServKey K4_imp_K2)
-
-lemma B_trusts_ServTicket_Kas_refined:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
-         evs \<in> kerberos |]
-  ==> \<exists>AuthKey Tk. Says Kas A (Crypt(shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-           Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-             \<in> set evs
-           & ServLife + Tt <= AuthLife + Tk"
-by (blast dest!: B_trusts_ServKey K4_imp_K2_refined)
-
-lemma B_trusts_ServTicket:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
-         evs \<in> kerberos |]
- ==> \<exists>Tk AuthKey.
-     Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-                   Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-       \<in> set evs
-     & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
-                   Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})
-       \<in> set evs"
-by (blast dest: B_trusts_ServKey K4_imp_K2)
+text{* @{text u_B_authenticates_A} would be the same as @{text B_authenticates_A} because the servK confidentiality assumption is yet unrelaxed*}
 
-lemma B_trusts_ServTicket_refined:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
-         evs \<in> kerberos |]
- ==> \<exists>Tk AuthKey.
-     (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
-                   Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
-       \<in> set evs
-     & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
-                   Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})
-       \<in> set evs
-     & ServLife + Tt <= AuthLife + Tk)"
-by (blast dest: B_trusts_ServKey K4_imp_K2_refined)
-
-
-lemma NotExpirServ_NotExpirAuth_refined:
-     "[| ~ ExpirServ Tt evs; ServLife + Tt <= AuthLife + Tk |]
-      ==> ~ ExpirAuth Tk evs"
-by (blast dest: leI le_trans dest: leD)
-
-
-lemma Confidentiality_B:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-           \<in> parts (spies evs);
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
+lemma u_B_authenticates_A_r:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
            \<in> parts (spies evs);
-         ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Key ServKey \<notin> analz (spies evs)"
-apply (frule A_trusts_AuthKey)
-apply (frule_tac [3] Confidentiality_Kas)
-apply (frule_tac [6] B_trusts_ServTicket, auto)
-apply (blast dest!: Confidentiality_Tgs2 dest: Says_Kas_message_form A_trusts_K4 unique_ServKeys unique_AuthKeys)
-done
-(*
-The proof above is fast.  It can be done in one command in 17 secs:
-apply (blast dest: A_trusts_AuthKey A_trusts_K4
-                               Says_Kas_message_form B_trusts_ServTicket
-                               unique_ServKeys unique_AuthKeys
-                               Confidentiality_Kas
-                               Confidentiality_Tgs2)
-It is very brittle: we can't use this command partway
-through the script above.
-*)
-
-
-
-text{*Most general form -- only for refined model! *}
-lemma Confidentiality_B_refined:
-     "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);
-         ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Key ServKey \<notin> analz (spies evs)"
-apply (blast dest: B_trusts_ServTicket_refined NotExpirServ_NotExpirAuth_refined Confidentiality_Tgs2)
+         \<not> expiredSK Ts evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>,
+                  Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
+apply (blast intro: Says_K5 dest: u_Confidentiality_B servTicket_authentic_Tgs)
 done
 
-
-subsection{* Authenticity theorems *}
-
-text{*1. Session Keys authenticity: they originated with servers.*}
-
-text{*Authenticity of ServKey for A*}
-lemma A_trusts_ServKey:
-     "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
+lemma A_authenticates_B:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
            \<in> parts (spies evs);
-         ~ ExpirAuth Tk evs; A \<notin> bad; evs \<in> kerberos |]
- ==>Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
-       \<in> set evs"
-by (blast dest: A_trusts_AuthKey Confidentiality_Auth_A A_trusts_K4_bis)
-
-text{*Note: requires a temporal check*}
-
-
-(***2. Parties authenticity: each party verifies "the identity of
-       another party who generated some data" (quoted from Neuman & Ts'o).***)
-
-       (*These guarantees don't assess whether two parties agree on
-         the same session key: sending a message containing a key
-         doesn't a priori state knowledge of the key.***)
-
-text{*B checks authenticity of A by theorems @{text A_Authenticity} and
-  @{text A_authenticity_refined} *}
-lemma Says_Auth:
-     "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
-         Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
-                                     ServTicket|}) \<in> set evs;
-         Key ServKey \<notin> analz (spies evs);
-         A \<notin> bad; B \<notin> bad; evs \<in> kerberos |]
- ==> Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \<in> set evs"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
-apply (frule_tac [5] Says_ticket_in_parts_spies)
-apply (frule_tac [7] Says_ticket_in_parts_spies)
-apply (simp_all (no_asm_simp) add: all_conj_distrib)
-apply blast
-txt{*K3*}
-apply (blast dest: A_trusts_AuthKey Says_Kas_message_form Says_Tgs_message_form)
-txt{*K4*}
-apply (force dest!: Crypt_imp_keysFor)
-txt{*K5*}
-apply (blast dest: Key_unique_SesKey)
+         Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (frule authK_authentic)
+apply assumption+
+apply (frule servK_authentic)
+prefer 2 apply (blast dest: authK_authentic Says_Kas_message_form)
+apply assumption+
+apply (blast dest: K4_imp_K2 Key_unique_SesKey intro!: Says_K6)
+(*Single command proof: slower!
+apply (blast dest: authK_authentic servK_authentic Says_Kas_message_form Key_unique_SesKey K4_imp_K2 intro!: Says_K6)
+*)
 done
 
-text{*The second assumption tells B what kind of key ServKey is.*}
-lemma A_Authenticity:
-     "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
-         Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-           \<in> parts (spies evs);
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
-           \<in> parts (spies evs);
-         ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},
-                  Crypt ServKey {|Agent A, Number Ta|} |} \<in> set evs"
-by (blast intro: Says_Auth dest: Confidentiality_B Key_unique_SesKey B_trusts_ServKey)
-
-text{*Stronger form in the refined model*}
-lemma A_Authenticity_refined:
-     "[| Crypt ServKey {|Agent A, Number Ta2|} \<in> parts (spies evs);
-         Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
+lemma A_authenticates_B_r:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
-         ~ ExpirServ Tt evs;
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},
-                  Crypt ServKey {|Agent A, Number Ta2|} |} \<in> set evs"
-by (blast dest: Confidentiality_B_refined B_trusts_ServKey Key_unique_SesKey intro: Says_Auth)
-
-
-text{*A checks authenticity of B by theorem @{text B_authenticity}*}
-
-lemma Says_K6:
-     "[| Crypt ServKey (Number Ta) \<in> parts (spies evs);
-         Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
-                                     ServTicket|}) \<in> set evs;
-         Key ServKey \<notin> analz (spies evs);
-         A \<notin> bad; B \<notin> bad; evs \<in> kerberos |]
-      ==> Says B A (Crypt ServKey (Number Ta)) \<in> set evs"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
-apply (frule_tac [5] Says_ticket_in_parts_spies)
-apply (frule_tac [7] Says_ticket_in_parts_spies)
-apply (simp_all (no_asm_simp))
-apply blast
-apply (force dest!: Crypt_imp_keysFor, clarify)
-apply (frule Says_Tgs_message_form, assumption, clarify) (*PROOF FAILED if omitted*)
-apply (blast dest: unique_CryptKey)
-done
-
-lemma K4_trustworthy:
-     "[| Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|}
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
            \<in> parts (spies evs);
-         Key AuthKey \<notin> analz (spies evs);  AuthKey \<notin> range shrK;
-         evs \<in> kerberos |]
-  ==> \<exists>A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|})
-              \<in> set evs"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
-apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-apply (blast+)
-done
-
-lemma B_Authenticity:
-     "[| Crypt ServKey (Number Ta) \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-           \<in> parts (spies evs);
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
-           \<in> parts (spies evs);
-         ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> Says B A (Crypt ServKey (Number Ta)) \<in> set evs"
-apply (frule A_trusts_AuthKey)
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (frule authK_authentic)
 apply (frule_tac [3] Says_Kas_message_form)
 apply (frule_tac [4] Confidentiality_Kas)
-apply (frule_tac [7] K4_trustworthy)
+apply (frule_tac [7] servK_authentic)
 prefer 8 apply blast
 apply (erule_tac [9] exE)
 apply (frule_tac [9] K4_imp_K2)
-txt{*Yes the proof's a mess, but I don't know how to improve it.*}
 apply assumption+
-apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs1
+apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs
 )
 done
 
 
-(***3. Parties' knowledge of session keys. A knows a session key if she
-       used it to build a cipher.***)
+subsection{* Key distribution guarantees
+       An agent knows a session key if he used it to issue a cipher.
+       These guarantees also convey a stronger form of 
+       authentication - non-injective agreement on the session key*}
+
+
+lemma Kas_Issues_A:
+   "\<lbrakk> Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>) \<in> set evs;
+      evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> Kas Issues A with (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>) 
+          on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule kerbIV.induct)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*K2*}
+apply (simp add: takeWhile_tail)
+apply (blast dest: authK_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD])
+done
+
+lemma A_authenticates_and_keydist_to_Kas:
+  "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace> \<in> parts (spies evs);
+     A \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> Kas Issues A with (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>) 
+          on evs"
+apply (blast dest: authK_authentic Kas_Issues_A)
+done
 
-lemma B_Knows_B_Knows_ServKey_lemma:
-     "[| Says B A (Crypt ServKey (Number Ta)) \<in> set evs;
-         Key ServKey \<notin> analz (spies evs);
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> B Issues A with (Crypt ServKey (Number Ta)) on evs"
+lemma honest_never_says_newer_timestamp_in_auth:
+     "\<lbrakk> (CT evs) \<le> T; A \<notin> bad; Number T \<in> parts {X}; evs \<in> kerbIV \<rbrakk> 
+     \<Longrightarrow> \<forall> B Y.  Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV.induct)
+apply (simp_all)
+apply force+
+done
+
+lemma honest_never_says_current_timestamp_in_auth:
+     "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; evs \<in> kerbIV \<rbrakk> 
+     \<Longrightarrow> \<forall> A B Y. A \<notin> bad \<longrightarrow> Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
+apply (frule eq_imp_le)
+apply (blast dest: honest_never_says_newer_timestamp_in_auth)
+done
+
+lemma A_trusts_secure_authenticator:
+    "\<lbrakk> Crypt K \<lbrace>Agent A, Number T\<rbrace> \<in> parts (spies evs);
+       Key K \<notin> analz (spies evs); evs \<in> kerbIV \<rbrakk>
+\<Longrightarrow> \<exists> B X. Says A Tgs \<lbrace>X, Crypt K \<lbrace>Agent A, Number T\<rbrace>, Agent B\<rbrace> \<in> set evs \<or> 
+           Says A B \<lbrace>X, Crypt K \<lbrace>Agent A, Number T\<rbrace>\<rbrace> \<in> set evs";
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all add: all_conj_distrib)
+apply blast+
+done
+
+lemma A_Issues_Tgs:
+  "\<lbrakk> Says A Tgs \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>
+       \<in> set evs; 
+     Key authK \<notin> analz (spies evs);  
+     A \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> A Issues Tgs with (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>) on evs"
 apply (simp (no_asm) add: Issues_def)
 apply (rule exI)
 apply (rule conjI, assumption)
 apply (simp (no_asm))
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
-apply (frule_tac [5] Says_ticket_in_parts_spies)
-apply (frule_tac [7] Says_ticket_in_parts_spies)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*fake*}
 apply blast
-txt{*K6 requires numerous lemmas*}
+txt{*K3*}
+(*
+apply clarify
+apply (drule Says_imp_knows_Spy [THEN parts.Inj, THEN authK_authentic, THEN Says_Kas_message_form], assumption, assumption, assumption)
+*)
 apply (simp add: takeWhile_tail)
-apply (blast dest: B_trusts_ServTicket parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: Says_K6)
+apply auto
+apply (force dest!: authK_authentic Says_Kas_message_form)
+apply (drule parts_spies_takeWhile_mono [THEN subsetD, THEN parts_spies_evs_revD2 [THEN subsetD]])
+apply (drule A_trusts_secure_authenticator, assumption, assumption)
+apply (simp add: honest_never_says_current_timestamp_in_auth)
 done
-(*Key ServKey \<notin> analz (spies evs) could be relaxed by Confidentiality_B
-  but this is irrelevant because B knows what he knows!                  *)
-
-lemma B_Knows_B_Knows_ServKey:
-     "[| Says B A (Crypt ServKey (Number Ta)) \<in> set evs;
-         Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-            \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-            \<in> parts (spies evs);
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
-           \<in> parts (spies evs);
-         ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> B Issues A with (Crypt ServKey (Number Ta)) on evs"
-by (blast dest!: Confidentiality_B B_Knows_B_Knows_ServKey_lemma)
 
-lemma B_Knows_B_Knows_ServKey_refined:
-     "[| Says B A (Crypt ServKey (Number Ta)) \<in> set evs;
-         Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-            \<in> parts (spies evs);
-         ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> B Issues A with (Crypt ServKey (Number Ta)) on evs"
-by (blast dest!: Confidentiality_B_refined B_Knows_B_Knows_ServKey_lemma)
-
-lemma A_Knows_B_Knows_ServKey:
-     "[| Crypt ServKey (Number Ta) \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-           \<in> parts (spies evs);
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
+lemma Tgs_authenticates_and_keydist_to_A:
+  "\<lbrakk>  Crypt authK \<lbrace>Agent A, Number T2\<rbrace> \<in> parts (spies evs); 
+      Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
            \<in> parts (spies evs);
-         ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;
-         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |]
-      ==> B Issues A with (Crypt ServKey (Number Ta)) on evs"
-by (blast dest!: B_Authenticity Confidentiality_Serv_A B_Knows_B_Knows_ServKey_lemma)
-
-lemma K3_imp_K2:
-     "[| Says A Tgs
-             {|AuthTicket, Crypt AuthKey {|Agent A, Number Ta|}, Agent B|}
-           \<in> set evs;
-         A \<notin> bad;  evs \<in> kerberos |]
-      ==> \<exists>Tk. Says Kas A (Crypt (shrK A)
-                      {|Key AuthKey, Agent Tgs, Tk, AuthTicket|})
-                   \<in> set evs"
-apply (erule rev_mp)
-apply (erule kerberos.induct)
-apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast, blast)
-apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN A_trusts_AuthKey])
+     Key authK \<notin> analz (spies evs);  
+     A \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> A Issues Tgs with (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>) on evs"
+apply (blast dest: A_Issues_Tgs Tgs_authenticates_A)
 done
 
-lemma K4_trustworthy':
-     "[| Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
-           \<in> parts (spies evs);
-         Says Kas A (Crypt (shrK A)
-                     {|Key AuthKey, Agent Tgs, Tk, AuthTicket|})
-         \<in> set evs;
-         Key AuthKey \<notin> analz (spies evs);
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> Says Tgs A (Crypt AuthKey
-                     {|Key ServKey, Agent B, Number Tt, ServTicket|})
-         \<in> set evs"
-apply (erule rev_mp)
+lemma Tgs_Issues_A:
+    "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket \<rbrace>)
+         \<in> set evs; 
+       Key authK \<notin> analz (spies evs);  evs \<in> kerbIV \<rbrakk>
+  \<Longrightarrow> Tgs Issues A with 
+          (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket \<rbrace>) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
-apply (frule_tac [7] K5_msg_in_parts_spies)
-apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
-apply (force dest!: Crypt_imp_keysFor)
-apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN A_trusts_AuthTicket] unique_AuthKeys)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*K4*}
+apply (simp add: takeWhile_tail)
+(*Last two thms installed only to derive authK \<notin> range shrK*)
+apply (blast dest: servK_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] authTicket_authentic Says_Kas_message_form)
 done
 
-lemma A_Knows_A_Knows_ServKey_lemma:
-     "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|}
-           \<in> set evs;
-         Key ServKey \<notin> analz (spies evs);
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
+lemma A_authenticates_and_keydist_to_Tgs:
+"\<lbrakk>Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace> \<in> parts (spies evs);
+  Key authK \<notin> analz (spies evs); B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+ \<Longrightarrow> \<exists>A. Tgs Issues A with 
+          (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket \<rbrace>) on evs"
+apply (blast dest: Tgs_Issues_A servK_authentic_bis)
+done
+
+
+
+lemma B_Issues_A:
+     "\<lbrakk> Says B A (Crypt servK (Number T3)) \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
 apply (simp (no_asm) add: Issues_def)
 apply (rule exI)
 apply (rule conjI, assumption)
 apply (simp (no_asm))
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos.induct, analz_mono_contra)
-apply (frule_tac [5] Says_ticket_in_parts_spies)
-apply (frule_tac [7] Says_ticket_in_parts_spies)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+apply blast
+txt{*K6 requires numerous lemmas*}
+apply (simp add: takeWhile_tail)
+apply (blast dest: servTicket_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: Says_K6)
+done
+
+lemma B_Issues_A_r:
+     "\<lbrakk> Says B A (Crypt servK (Number T3)) \<in> set evs;
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+            \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+            \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
+apply (blast dest!: Confidentiality_B B_Issues_A)
+done
+
+lemma u_B_Issues_A_r:
+     "\<lbrakk> Says B A (Crypt servK (Number T3)) \<in> set evs;
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+            \<in> parts (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
+apply (blast dest!: u_Confidentiality_B B_Issues_A)
+done
+
+lemma A_authenticates_and_keydist_to_B:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
+apply (blast dest!: A_authenticates_B B_Issues_A)
+done
+
+lemma A_authenticates_and_keydist_to_B_r:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
+apply (blast dest!: A_authenticates_B_r Confidentiality_Serv_A B_Issues_A)
+done
+
+
+lemma A_Issues_B:
+     "\<lbrakk> Says A B \<lbrace>servTicket, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace>
+           \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp))
 apply clarify
 txt{*K5*}
 apply auto
 apply (simp add: takeWhile_tail)
 txt{*Level 15: case study necessary because the assumption doesn't state
-  the form of ServTicket. The guarantee becomes stronger.*}
+  the form of servTicket. The guarantee becomes stronger.*}
 apply (blast dest: Says_imp_spies [THEN analz.Inj, THEN analz_Decrypt']
-                   K3_imp_K2 K4_trustworthy'
+                   K3_imp_K2 servK_authentic_ter
                    parts_spies_takeWhile_mono [THEN subsetD]
                    parts_spies_evs_revD2 [THEN subsetD]
-             intro: Says_Auth)
+             intro: Says_K5)
 apply (simp add: takeWhile_tail)
 done
 
-lemma A_Knows_A_Knows_ServKey:
-     "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|}
+lemma A_Issues_B_r:
+     "\<lbrakk> Says A B \<lbrace>servTicket, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace>
            \<in> set evs;
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+apply (blast dest!: Confidentiality_Serv_A A_Issues_B)
+done
+
+lemma B_authenticates_and_keydist_to_A:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
            \<in> parts (spies evs);
-         ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
-by (blast dest!: Confidentiality_Serv_A A_Knows_A_Knows_ServKey_lemma)
+         Key servK \<notin> analz (spies evs);
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+apply (blast dest: B_authenticates_A A_Issues_B)
+done
 
-lemma B_Knows_A_Knows_ServKey:
-     "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
-         Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
+lemma B_authenticates_and_keydist_to_A_r:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
            \<in> parts (spies evs);
-         Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
            \<in> parts (spies evs);
-         Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+apply (blast dest: B_authenticates_A Confidentiality_B A_Issues_B)
+done
+
+text{* @{text u_B_authenticates_and_keydist_to_A} would be the same as @{text B_authenticates_and_keydist_to_A} because the
+ servK confidentiality assumption is yet unrelaxed*}
+
+lemma u_B_authenticates_and_keydist_to_A_r:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
            \<in> parts (spies evs);
-         ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
-by (blast dest: A_Authenticity Confidentiality_B A_Knows_A_Knows_ServKey_lemma)
+         \<not> expiredSK Ts evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbIV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+apply (blast dest: u_B_authenticates_A_r u_Confidentiality_B A_Issues_B)
+done
 
 
-lemma B_Knows_A_Knows_ServKey_refined:
-     "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
-         Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
-           \<in> parts (spies evs);
-         ~ ExpirServ Tt evs;
-         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos |]
-   ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
-by (blast dest: A_Authenticity_refined Confidentiality_B_refined A_Knows_A_Knows_ServKey_lemma)
+
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/KerberosIV_Gets.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,1548 @@
+(*  Title:      HOL/Auth/KerberosIV
+    ID:         $Id$
+    Author:     Giampaolo Bella, Cambridge University Computer Laboratory
+    Copyright   1998  University of Cambridge
+*)
+
+header{*The Kerberos Protocol, Version IV*}
+
+theory KerberosIV_Gets imports Public begin
+
+text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+
+syntax
+  Kas :: agent
+  Tgs :: agent  --{*the two servers are translations...*}
+
+
+translations
+  "Kas"       == "Server "
+  "Tgs"       == "Friend 0"
+
+
+axioms
+  Tgs_not_bad [iff]: "Tgs \<notin> bad"
+   --{*Tgs is secure --- we already know that Kas is secure*}
+
+(*The current time is just the length of the trace!*)
+syntax
+    CT :: "event list=>nat"
+
+    expiredAK :: "[nat, event list] => bool"
+
+    expiredSK :: "[nat, event list] => bool"
+
+    expiredA :: "[nat, event list] => bool"
+
+    valid :: "[nat, nat] => bool" ("valid _ wrt _")
+
+
+constdefs
+ (* authKeys are those contained in an authTicket *)
+    authKeys :: "event list => key set"
+    "authKeys evs == {authK. \<exists>A Peer Ta. Says Kas A
+                        (Crypt (shrK A) \<lbrace>Key authK, Agent Peer, Number Ta,
+               (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key authK, Number Ta\<rbrace>)
+                  \<rbrace>) \<in> set evs}"
+
+ (* States than an event really appears only once on a trace *)
+  Unique :: "[event, event list] => bool" ("Unique _ on _")
+   "Unique ev on evs == 
+      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
+
+
+consts
+    (*Duration of the authentication key*)
+    authKlife   :: nat
+
+    (*Duration of the service key*)
+    servKlife   :: nat
+
+    (*Duration of an authenticator*)
+    authlife   :: nat
+
+    (*Upper bound on the time of reaction of a server*)
+    replylife   :: nat
+
+specification (authKlife)
+  authKlife_LB [iff]: "2 \<le> authKlife"
+    by blast
+
+specification (servKlife)
+  servKlife_LB [iff]: "2 + authKlife \<le> servKlife"
+    by blast
+
+specification (authlife)
+  authlife_LB [iff]: "Suc 0 \<le> authlife"
+    by blast
+
+specification (replylife)
+  replylife_LB [iff]: "Suc 0 \<le> replylife"
+    by blast
+
+translations
+   "CT" == "length "
+
+   "expiredAK Ta evs" == "authKlife + Ta < CT evs"
+
+   "expiredSK Ts evs" == "servKlife + Ts < CT evs"
+
+   "expiredA T evs" == "authlife + T < CT evs"
+
+   "valid T1 wrt T2" == "T1 <= replylife + T2"
+
+(*---------------------------------------------------------------------*)
+
+
+(* Predicate formalising the association between authKeys and servKeys *)
+constdefs
+  AKcryptSK :: "[key, key, event list] => bool"
+  "AKcryptSK authK servK evs ==
+     \<exists>A B Ts.
+       Says Tgs A (Crypt authK
+                     \<lbrace>Key servK, Agent B, Number Ts,
+                       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>)
+         \<in> set evs"
+
+consts
+
+kerbIV_gets   :: "event list set"
+inductive "kerbIV_gets"
+  intros
+
+   Nil:  "[] \<in> kerbIV_gets"
+
+   Fake: "\<lbrakk> evsf \<in> kerbIV_gets;  X \<in> synth (analz (spies evsf)) \<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> kerbIV_gets"
+
+   Reception: "\<lbrakk> evsr \<in> kerbIV_gets;  Says A B X \<in> set evsr \<rbrakk>
+                \<Longrightarrow> Gets B X # evsr \<in> kerbIV_gets"
+
+(* FROM the initiator *)
+   K1:   "\<lbrakk> evs1 \<in> kerbIV_gets \<rbrakk>
+          \<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
+          \<in> kerbIV_gets"
+
+(* Adding the timestamp serves to A in K3 to check that
+   she doesn't get a reply too late. This kind of timeouts are ordinary.
+   If a server's reply is late, then it is likely to be fake. *)
+
+(*---------------------------------------------------------------------*)
+
+(*FROM Kas *)
+   K2:  "\<lbrakk> evs2 \<in> kerbIV_gets; Key authK \<notin> used evs2; authK \<in> symKeys;
+            Gets Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Kas A
+                (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2),
+                      (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
+                          Number (CT evs2)\<rbrace>)\<rbrace>) # evs2 \<in> kerbIV_gets"
+(*
+  The internal encryption builds the authTicket.
+  The timestamp doesn't change inside the two encryptions: the external copy
+  will be used by the initiator in K3; the one inside the
+  authTicket by Tgs in K4.
+*)
+
+(*---------------------------------------------------------------------*)
+
+(* FROM the initiator *)
+   K3:  "\<lbrakk> evs3 \<in> kerbIV_gets;
+            Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
+            Gets A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+              authTicket\<rbrace>) \<in> set evs3;
+            valid Ta wrt T1
+         \<rbrakk>
+          \<Longrightarrow> Says A Tgs \<lbrace>authTicket,
+                           (Crypt authK \<lbrace>Agent A, Number (CT evs3)\<rbrace>),
+                           Agent B\<rbrace> # evs3 \<in> kerbIV_gets"
+(*The two events amongst the premises allow A to accept only those authKeys
+  that are not issued late. *)
+
+(*---------------------------------------------------------------------*)
+
+(* FROM Tgs *)
+(* Note that the last temporal check is not mentioned in the original MIT
+   specification. Adding it makes many goals "available" to the peers. 
+   Theorems that exploit it have the suffix `_u', which stands for updated 
+   protocol.
+*)
+   K4:  "\<lbrakk> evs4 \<in> kerbIV_gets; Key servK \<notin> used evs4; servK \<in> symKeys;
+            B \<noteq> Tgs;  authK \<in> symKeys;
+            Gets Tgs \<lbrace>
+             (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
+				 Number Ta\<rbrace>),
+             (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
+	        \<in> set evs4;
+            \<not> expiredAK Ta evs4;
+            \<not> expiredA T2 evs4;
+            servKlife + (CT evs4) <= authKlife + Ta
+         \<rbrakk>
+          \<Longrightarrow> Says Tgs A
+                (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
+			       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+		 			        Number (CT evs4)\<rbrace> \<rbrace>)
+	        # evs4 \<in> kerbIV_gets"
+(* Tgs creates a new session key per each request for a service, without
+   checking if there is still a fresh one for that service.
+   The cipher under Tgs' key is the authTicket, the cipher under B's key
+   is the servTicket, which is built now.
+   NOTE that the last temporal check is not present in the MIT specification.
+
+*)
+
+(*---------------------------------------------------------------------*)
+
+(* FROM the initiator *)
+   K5:  "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
+            Says A Tgs
+                \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
+		  Agent B\<rbrace>
+              \<in> set evs5;
+            Gets A
+             (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+                \<in> set evs5;
+            valid Ts wrt T2 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>servTicket,
+			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+               # evs5 \<in> kerbIV_gets"
+(* Checks similar to those in K3. *)
+
+(*---------------------------------------------------------------------*)
+
+(* FROM the responder*)
+    K6:  "\<lbrakk> evs6 \<in> kerbIV_gets;
+            Gets B \<lbrace>
+              (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
+              (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
+            \<in> set evs6;
+            \<not> expiredSK Ts evs6;
+            \<not> expiredA T3 evs6
+         \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt servK (Number T3))
+               # evs6 \<in> kerbIV_gets"
+(* Checks similar to those in K4. *)
+
+(*---------------------------------------------------------------------*)
+
+(* Leaking an authK... *)
+   Oops1: "\<lbrakk> evsO1 \<in> kerbIV_gets;  A \<noteq> Spy;
+              Says Kas A
+                (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                                  authTicket\<rbrace>)  \<in> set evsO1;
+              expiredAK Ta evsO1 \<rbrakk>
+          \<Longrightarrow> Says A Spy \<lbrace>Agent A, Agent Tgs, Number Ta, Key authK\<rbrace>
+               # evsO1 \<in> kerbIV_gets"
+
+(*---------------------------------------------------------------------*)
+
+(*Leaking a servK... *)
+   Oops2: "\<lbrakk> evsO2 \<in> kerbIV_gets;  A \<noteq> Spy;
+              Says Tgs A
+                (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+                   \<in> set evsO2;
+              expiredSK Ts evsO2 \<rbrakk>
+          \<Longrightarrow> Says A Spy \<lbrace>Agent A, Agent B, Number Ts, Key servK\<rbrace>
+               # evsO2 \<in> kerbIV_gets"
+
+(*---------------------------------------------------------------------*)
+
+declare Says_imp_knows_Spy [THEN parts.Inj, dest]
+declare parts.Body [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un [dest]
+
+subsection{*Lemmas about reception event*}
+
+lemma Gets_imp_Says :
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply auto
+done
+
+lemma Gets_imp_knows_Spy: 
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> kerbIV_gets \<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+done
+
+(*Needed for force to work for example in new_keys_not_used*)
+declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
+
+lemma Gets_imp_knows:
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> kerbIV_gets \<rbrakk>  \<Longrightarrow> X \<in> knows B evs"
+apply (case_tac "B = Spy")
+apply (blast dest!: Gets_imp_knows_Spy)
+apply (blast dest!: Gets_imp_knows_agents)
+done
+
+subsection{*Lemmas about @{term authKeys}*}
+
+lemma authKeys_empty: "authKeys [] = {}"
+apply (unfold authKeys_def)
+apply (simp (no_asm))
+done
+
+lemma authKeys_not_insert:
+ "(\<forall>A Ta akey Peer.
+   ev \<noteq> Says Kas A (Crypt (shrK A) \<lbrace>akey, Agent Peer, Ta,
+              (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, akey, Ta\<rbrace>)\<rbrace>))
+       \<Longrightarrow> authKeys (ev # evs) = authKeys evs"
+by (unfold authKeys_def, auto)
+
+lemma authKeys_insert:
+  "authKeys
+     (Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Peer, Number Ta,
+      (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K, Number Ta\<rbrace>)\<rbrace>) # evs)
+       = insert K (authKeys evs)"
+by (unfold authKeys_def, auto)
+
+lemma authKeys_simp:
+   "K \<in> authKeys
+    (Says Kas A (Crypt (shrK A) \<lbrace>Key K', Agent Peer, Number Ta,
+     (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K', Number Ta\<rbrace>)\<rbrace>) # evs)
+        \<Longrightarrow> K = K' | K \<in> authKeys evs"
+by (unfold authKeys_def, auto)
+
+lemma authKeysI:
+   "Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Tgs, Number Ta,
+     (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key K, Number Ta\<rbrace>)\<rbrace>) \<in> set evs
+        \<Longrightarrow> K \<in> authKeys evs"
+by (unfold authKeys_def, auto)
+
+lemma authKeys_used: "K \<in> authKeys evs \<Longrightarrow> Key K \<in> used evs"
+by (simp add: authKeys_def, blast)
+
+
+subsection{*Forwarding Lemmas*}
+
+lemma Says_ticket_parts:
+     "Says S A (Crypt K \<lbrace>SesKey, B, TimeStamp, Ticket\<rbrace>) \<in> set evs
+      \<Longrightarrow> Ticket \<in> parts (spies evs)"
+apply blast
+done
+
+lemma Gets_ticket_parts:
+     "\<lbrakk>Gets A (Crypt K \<lbrace>SesKey, Peer, Ta, Ticket\<rbrace>) \<in> set evs; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Ticket \<in> parts (spies evs)"
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj])
+done
+
+lemma Oops_range_spies1:
+     "\<lbrakk> Says Kas A (Crypt KeyA \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>)
+           \<in> set evs ;
+         evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, auto)
+done
+
+lemma Oops_range_spies2:
+     "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>)
+           \<in> set evs ;
+         evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> servK \<notin> range shrK & servK \<in> symKeys"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, auto)
+done
+
+
+(*Spy never sees another agent's shared key! (unless it's lost at start)*)
+lemma Spy_see_shrK [simp]:
+     "evs \<in> kerbIV_gets \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+apply (blast+)
+done
+
+lemma Spy_analz_shrK [simp]:
+     "evs \<in> kerbIV_gets \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
+by auto
+
+lemma Spy_see_shrK_D [dest!]:
+     "\<lbrakk> Key (shrK A) \<in> parts (spies evs);  evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> A:bad"
+by (blast dest: Spy_see_shrK)
+lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
+
+text{*Nobody can have used non-existent keys!*}
+lemma new_keys_not_used [simp]:
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbIV_gets\<rbrakk>
+     \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake*}
+apply (force dest!: keysFor_parts_insert)
+txt{*Others*}
+apply (force dest!: analz_shrK_Decrypt)+
+done
+
+(*Earlier, all protocol proofs declared this theorem.
+  But few of them actually need it! (Another is Yahalom) *)
+lemma new_keys_not_analzd:
+ "\<lbrakk>evs \<in> kerbIV_gets; K \<in> symKeys; Key K \<notin> used evs\<rbrakk>
+  \<Longrightarrow> K \<notin> keysFor (analz (spies evs))"
+by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD])
+
+
+subsection{*Regularity Lemmas*}
+text{*These concern the form of items passed in messages*}
+
+text{*Describes the form of all components sent by Kas*}
+
+lemma Says_Kas_message_form:
+     "\<lbrakk> Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
+           \<in> set evs;
+         evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow>  
+  K = shrK A  & Peer = Tgs &
+  authK \<notin> range shrK & authK \<in> authKeys evs & authK \<in> symKeys & 
+  authTicket = (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>)"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (simp_all (no_asm) add: authKeys_def authKeys_insert)
+apply blast+
+done
+
+
+lemma SesKey_is_session_key:
+     "\<lbrakk> Crypt (shrK Tgs_B) \<lbrace>Agent A, Agent Tgs_B, Key SesKey, Number T\<rbrace>
+            \<in> parts (spies evs); Tgs_B \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> SesKey \<notin> range shrK"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
+done
+
+lemma authTicket_authentic:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                 Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+            \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake, K4*}
+apply (blast+)
+done
+
+lemma authTicket_crypt_authK:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> authK \<in> authKeys evs"
+apply (frule authTicket_authentic, assumption)
+apply (simp (no_asm) add: authKeys_def)
+apply blast
+done
+
+lemma Says_Tgs_message_form:
+     "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs;
+         evs \<in> kerbIV_gets \<rbrakk>
+  \<Longrightarrow> B \<noteq> Tgs & 
+      authK \<notin> range shrK & authK \<in> authKeys evs & authK \<in> symKeys &
+      servK \<notin> range shrK & servK \<notin> authKeys evs & servK \<in> symKeys &
+      servTicket = (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>)"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast, auto)
+txt{*Three subcases of Message 4*}
+apply (blast dest!: SesKey_is_session_key)
+apply (blast dest: authTicket_crypt_authK)
+apply (blast dest!: authKeys_used Says_Kas_message_form)
+done
+
+
+lemma authTicket_form:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         A \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+    \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys & 
+        authTicket = Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+apply blast+
+done
+
+text{* This form holds also over an authTicket, but is not needed below.*}
+lemma servTicket_form:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>
+              \<in> parts (spies evs);
+            Key authK \<notin> analz (spies evs);
+            evs \<in> kerbIV_gets \<rbrakk>
+         \<Longrightarrow> servK \<notin> range shrK & servK \<in> symKeys & 
+    (\<exists>A. servTicket = Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
+done
+
+text{* Essentially the same as @{text authTicket_form} *}
+lemma Says_kas_message_form:
+     "\<lbrakk> Gets A (Crypt (shrK A)
+              \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>) \<in> set evs;
+         evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys & 
+          authTicket =
+                  Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>
+          | authTicket \<in> analz (spies evs)"
+by (blast dest: analz_shrK_Decrypt authTicket_form
+                Gets_imp_knows_Spy [THEN analz.Inj])
+
+lemma Says_tgs_message_form:
+ "\<lbrakk> Gets A (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>)
+       \<in> set evs;  authK \<in> symKeys;
+     evs \<in> kerbIV_gets \<rbrakk>
+  \<Longrightarrow> servK \<notin> range shrK &
+      (\<exists>A. servTicket =
+	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+       | servTicket \<in> analz (spies evs)"
+apply (frule Gets_imp_knows_Spy [THEN analz.Inj], auto)
+ apply (force dest!: servTicket_form)
+apply (frule analz_into_parts)
+apply (frule servTicket_form, auto)
+done
+
+
+subsection{*Authenticity theorems: confirm origin of sensitive messages*}
+
+lemma authK_authentic:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         A \<notin> bad;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>)
+            \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake*}
+apply blast
+txt{*K4*}
+apply (blast dest!: authTicket_authentic [THEN Says_Kas_message_form])
+done
+
+text{*If a certain encrypted message appears then it originated with Tgs*}
+lemma servK_authentic:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs);
+         authK \<notin> range shrK;
+         evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists>A. Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake*}
+apply blast
+txt{*K2*}
+apply blast
+txt{*K4*}
+apply auto
+done
+
+lemma servK_authentic_bis:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs);
+         B \<noteq> Tgs;
+         evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists>A. Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake*}
+apply blast
+txt{*K4*}
+apply blast
+done
+
+text{*Authenticity of servK for B*}
+lemma servTicket_authentic_Tgs:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs); B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists>authK.
+       Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+apply blast+
+done
+
+text{*Anticipated here from next subsection*}
+lemma K4_imp_K2:
+"\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+      \<in> set evs;  evs \<in> kerbIV_gets\<rbrakk>
+   \<Longrightarrow> \<exists>Ta. Says Kas A
+        (Crypt (shrK A)
+         \<lbrace>Key authK, Agent Tgs, Number Ta,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+        \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, auto)
+apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
+done
+
+text{*Anticipated here from next subsection*}
+lemma u_K4_imp_K2:
+"\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+      \<in> set evs; evs \<in> kerbIV_gets\<rbrakk>
+   \<Longrightarrow> \<exists>Ta. (Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+             \<in> set evs
+          & servKlife + Ts <= authKlife + Ta)"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, auto)
+apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
+done
+
+lemma servTicket_authentic_Kas:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+  \<Longrightarrow> \<exists>authK Ta.
+       Says Kas A
+         (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+            Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+        \<in> set evs"
+apply (blast dest!: servTicket_authentic_Tgs K4_imp_K2)
+done
+
+lemma u_servTicket_authentic_Kas:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+  \<Longrightarrow> \<exists>authK Ta. Says Kas A (Crypt(shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+             \<in> set evs
+           & servKlife + Ts <= authKlife + Ta"
+apply (blast dest!: servTicket_authentic_Tgs u_K4_imp_K2)
+done
+
+lemma servTicket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists>Ta authK.
+     Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                   Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+       \<in> set evs
+     & Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+       \<in> set evs"
+apply (blast dest: servTicket_authentic_Tgs K4_imp_K2)
+done
+
+lemma u_servTicket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists>Ta authK.
+     (Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
+                   Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+       \<in> set evs
+     & Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+       \<in> set evs
+     & servKlife + Ts <= authKlife + Ta)"
+apply (blast dest: servTicket_authentic_Tgs u_K4_imp_K2)
+done
+
+lemma u_NotexpiredSK_NotexpiredAK:
+     "\<lbrakk> \<not> expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \<rbrakk>
+      \<Longrightarrow> \<not> expiredAK Ta evs"
+apply (blast dest: leI le_trans dest: leD)
+done
+
+
+subsection{* Reliability: friendly agents send something if something else happened*}
+
+lemma K3_imp_K2:
+     "\<lbrakk> Says A Tgs
+             \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>
+           \<in> set evs;
+         A \<notin> bad;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<exists>Ta. Says Kas A (Crypt (shrK A)
+                      \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
+                   \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, blast, blast)
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN authK_authentic])
+done
+
+text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
+lemma Key_unique_SesKey:
+     "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T, Ticket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt K' \<lbrace>Key SesKey,  Agent B', T', Ticket'\<rbrace>
+           \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
+         evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> K=K' & B=B' & T=T' & Ticket=Ticket'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake, K2, K4*}
+apply (blast+)
+done
+
+lemma Tgs_authenticates_A:
+  "\<lbrakk>  Crypt authK \<lbrace>Agent A, Number T2\<rbrace> \<in> parts (spies evs); 
+      Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+      Key authK \<notin> analz (spies evs); A \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists> B. Says A Tgs \<lbrace>
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>,
+          Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B \<rbrace> \<in> set evs"  
+apply (drule authTicket_authentic, assumption, rotate_tac 4)
+apply (erule rev_mp, erule rev_mp, erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [6] Gets_ticket_parts)
+apply (frule_tac [9] Gets_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*Fake*}
+apply blast
+txt{*K2*}
+apply (force dest!: Crypt_imp_keysFor)
+txt{*K3*}
+apply (blast dest: Key_unique_SesKey)
+txt{*K5*}
+txt{*If authKa were compromised, so would be authK*}
+apply (case_tac "Key authKa \<in> analz (spies evs5)")
+apply (force dest!: Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+txt{*Besides, since authKa originated with Kas anyway...*}
+apply (clarify, drule K3_imp_K2, assumption, assumption)
+apply (clarify, drule Says_Kas_message_form, assumption)
+txt{*...it cannot be a shared key*. Therefore @{term servK_authentic} applies. 
+     Contradition: Tgs used authK as a servkey, 
+     while Kas used it as an authkey*}
+apply (blast dest: servK_authentic Says_Tgs_message_form)
+done
+
+lemma Says_K5:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                                     servTicket\<rbrace>) \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad; B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> Says A B \<lbrace>servTicket, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [6] Gets_ticket_parts)
+apply (frule_tac [9] Gets_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+apply blast
+txt{*K3*}
+apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
+txt{*K4*}
+apply (force dest!: Crypt_imp_keysFor)
+txt{*K5*}
+apply (blast dest: Key_unique_SesKey)
+done
+
+text{*Anticipated here from next subsection*}
+lemma unique_CryptKey:
+     "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SesKey, T'\<rbrace>
+           \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
+         evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> A=A' & B=B' & T=T'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake, K2, K4*}
+apply (blast+)
+done
+
+lemma Says_K6:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts,
+                                     servTicket\<rbrace>) \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad; B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts)
+apply (simp_all (no_asm_simp))
+apply blast
+apply (force dest!: Crypt_imp_keysFor, clarify)
+apply (frule Says_Tgs_message_form, assumption, clarify) (*PROOF FAILED if omitted*)
+apply (blast dest: unique_CryptKey)
+done
+
+text{*Needs a unicity theorem, hence moved here*}
+lemma servK_authentic_ter:
+ "\<lbrakk> Says Kas A
+    (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>) \<in> set evs;
+     Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+       \<in> parts (spies evs);
+     Key authK \<notin> analz (spies evs);
+     evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
+apply (frule Says_Kas_message_form, assumption)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
+txt{*K2 and K4 remain*}
+prefer 2 apply (blast dest!: unique_CryptKey)
+apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
+done
+
+
+subsection{*Unicity Theorems*}
+
+text{* The session key, if secure, uniquely identifies the Ticket
+   whether authTicket or servTicket. As a matter of fact, one can read
+   also Tgs in the place of B.                                     *}
+
+
+lemma unique_authKeys:
+     "\<lbrakk> Says Kas A
+              (Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta, X\<rbrace>) \<in> set evs;
+         Says Kas A'
+              (Crypt Ka' \<lbrace>Key authK, Agent Tgs, Ta', X'\<rbrace>) \<in> set evs;
+         evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> A=A' & Ka=Ka' & Ta=Ta' & X=X'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*K2*}
+apply blast
+done
+
+text{* servK uniquely identifies the message from Tgs *}
+lemma unique_servKeys:
+     "\<lbrakk> Says Tgs A
+              (Crypt K \<lbrace>Key servK, Agent B, Ts, X\<rbrace>) \<in> set evs;
+         Says Tgs A'
+              (Crypt K' \<lbrace>Key servK, Agent B', Ts', X'\<rbrace>) \<in> set evs;
+         evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> A=A' & B=B' & K=K' & Ts=Ts' & X=X'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*K4*}
+apply blast
+done
+
+text{* Revised unicity theorems *}
+
+lemma Kas_Unique:
+     "\<lbrakk> Says Kas A
+              (Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>) \<in> set evs;
+        evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> 
+   Unique (Says Kas A (Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>)) 
+   on evs"
+apply (erule rev_mp, erule kerbIV_gets.induct, simp_all add: Unique_def)
+apply blast
+done
+
+lemma Tgs_Unique:
+     "\<lbrakk> Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>) \<in> set evs;
+        evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> 
+  Unique (Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>)) 
+  on evs"
+apply (erule rev_mp, erule kerbIV_gets.induct, simp_all add: Unique_def)
+apply blast
+done
+
+
+subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
+
+lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
+by (simp add: AKcryptSK_def)
+
+lemma AKcryptSKI:
+ "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>) \<in> set evs;
+     evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
+apply (blast dest: Says_Tgs_message_form)
+done
+
+lemma AKcryptSK_Says [simp]:
+   "AKcryptSK authK servK (Says S A X # evs) =
+     (Tgs = S &
+      (\<exists>B Ts. X = Crypt authK
+                \<lbrace>Key servK, Agent B, Number Ts,
+                  Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>)
+     | AKcryptSK authK servK evs)"
+apply (unfold AKcryptSK_def)
+apply (simp (no_asm))
+apply blast
+done
+
+(*A fresh authK cannot be associated with any other
+  (with respect to a given trace). *)
+lemma Auth_fresh_not_AKcryptSK:
+     "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
+done
+
+(*A fresh servK cannot be associated with any other
+  (with respect to a given trace). *)
+lemma Serv_fresh_not_AKcryptSK:
+ "Key servK \<notin> used evs \<Longrightarrow> \<not> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def, blast)
+done
+
+lemma authK_not_AKcryptSK:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, tk\<rbrace>
+           \<in> parts (spies evs);  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK K authK evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+txt{*Fake*}
+apply blast
+txt{*Reception*}
+apply (simp add: AKcryptSK_def)
+txt{*K2: by freshness*}
+apply (simp add: AKcryptSK_def)
+txt{*K4*}
+apply (blast+)
+done
+
+text{*A secure serverkey cannot have been used to encrypt others*}
+lemma servK_not_AKcryptSK:
+ "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, Number Ts\<rbrace> \<in> parts (spies evs);
+     Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
+     B \<noteq> Tgs;  evs \<in> kerbIV_gets \<rbrakk>
+  \<Longrightarrow> \<not> AKcryptSK SK K evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
+txt{*Reception*}
+apply (simp add: AKcryptSK_def)
+txt{*K4 splits into distinct subcases*}
+apply auto
+txt{*servK can't have been enclosed in two certificates*}
+ prefer 2 apply (blast dest: unique_CryptKey)
+txt{*servK is fresh and so could not have been used, by
+   @{text new_keys_not_used}*}
+apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
+done
+
+text{*Long term keys are not issued as servKeys*}
+lemma shrK_not_AKcryptSK:
+     "evs \<in> kerbIV_gets \<Longrightarrow> \<not> AKcryptSK K (shrK A) evs"
+apply (unfold AKcryptSK_def)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, auto)
+done
+
+text{*The Tgs message associates servK with authK and therefore not with any
+  other key authK.*}
+lemma Says_Tgs_AKcryptSK:
+     "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>)
+           \<in> set evs;
+         authK' \<noteq> authK;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK' servK evs"
+apply (unfold AKcryptSK_def)
+apply (blast dest: unique_servKeys)
+done
+
+text{*Equivalently*}
+lemma not_different_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;
+        authK' \<noteq> authK;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK' servK evs  \<and> servK \<in> symKeys"
+apply (simp add: AKcryptSK_def)
+apply (blast dest: unique_servKeys Says_Tgs_message_form)
+done
+
+lemma AKcryptSK_not_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK K evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts)
+txt{*Reception*}
+prefer 3 apply (simp add: AKcryptSK_def)
+apply (simp_all, safe)
+txt{*K4 splits into subcases*}
+prefer 4 apply (blast dest!: authK_not_AKcryptSK)
+txt{*servK is fresh and so could not have been used, by
+   @{text new_keys_not_used}*}
+ prefer 2 
+ apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
+txt{*Others by freshness*}
+apply (blast+)
+done
+
+text{*The only session keys that can be found with the help of session keys are
+  those sent by Tgs in step K4.  *}
+
+text{*We take some pains to express the property
+  as a logical equivalence so that the simplifier can apply it.*}
+lemma Key_analz_image_Key_lemma:
+     "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
+      \<Longrightarrow>
+      P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) = (K:KK | Key K \<in> analz H)"
+by (blast intro: analz_mono [THEN subsetD])
+
+
+lemma AKcryptSK_analz_insert:
+     "\<lbrakk> AKcryptSK K K' evs; K \<in> symKeys; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key K' \<in> analz (insert (Key K) (spies evs))"
+apply (simp add: AKcryptSK_def, clarify)
+apply (drule Says_imp_spies [THEN analz.Inj, THEN analz_insertI], auto)
+done
+
+lemma authKeys_are_not_AKcryptSK:
+     "\<lbrakk> K \<in> authKeys evs Un range shrK;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<forall>SK. \<not> AKcryptSK SK K evs \<and> K \<in> symKeys"
+apply (simp add: authKeys_def AKcryptSK_def)
+apply (blast dest: Says_Kas_message_form Says_Tgs_message_form)
+done
+
+lemma not_authKeys_not_AKcryptSK:
+     "\<lbrakk> K \<notin> authKeys evs;
+         K \<notin> range shrK; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> \<forall>SK. \<not> AKcryptSK K SK evs"
+apply (simp add: AKcryptSK_def)
+apply (blast dest: Says_Tgs_message_form)
+done
+
+
+subsection{*Secrecy Theorems*}
+
+text{*For the Oops2 case of the next theorem*}
+lemma Oops2_not_AKcryptSK:
+     "\<lbrakk> evs \<in> kerbIV_gets;
+         Says Tgs A (Crypt authK
+                     \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK SK evs"
+apply (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
+done
+   
+text{* Big simplification law for keys SK that are not crypted by keys in KK
+ It helps prove three, otherwise hard, facts about keys. These facts are
+ exploited as simplification laws for analz, and also "limit the damage"
+ in case of loss of a key to the spy. See ESORICS98. *}
+lemma Key_analz_image_Key [rule_format (no_asm)]:
+     "evs \<in> kerbIV_gets \<Longrightarrow>
+      (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
+       (\<forall>K \<in> KK. \<not> AKcryptSK K SK evs)   \<longrightarrow>
+       (Key SK \<in> analz (Key`KK Un (spies evs))) =
+       (SK \<in> KK | Key SK \<in> analz (spies evs)))"
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [11] Oops_range_spies2)
+apply (frule_tac [10] Oops_range_spies1)
+apply (frule_tac [8] Says_tgs_message_form)
+apply (frule_tac [6] Says_kas_message_form)
+apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
+txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
+ the induction hypothesis*}
+apply (case_tac [12] "AKcryptSK authK SK evsO1")
+apply (case_tac [9] "AKcryptSK servK SK evs5")
+apply (simp_all del: image_insert
+        add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
+             Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
+       Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
+  --{*18 seconds on a 1.8GHz machine??*}
+txt{*Fake*} 
+apply spy_analz
+txt{*Reception*}
+apply (simp add: AKcryptSK_def)
+txt{*K2*}
+apply blast 
+txt{*K3*}
+apply blast 
+txt{*K4*}
+apply (blast dest!: authK_not_AKcryptSK)
+txt{*K5*}
+apply (case_tac "Key servK \<in> analz (spies evs5) ")
+txt{*If servK is compromised then the result follows directly...*}
+apply (simp (no_asm_simp) add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
+txt{*...therefore servK is uncompromised.*}
+txt{*The AKcryptSK servK SK evs5 case leads to a contradiction.*}
+apply (blast elim!: servK_not_AKcryptSK [THEN [2] rev_notE] del: allE ballE)
+txt{*Another K5 case*}
+apply blast 
+txt{*Oops1*}
+apply simp 
+apply (blast dest!: AKcryptSK_analz_insert)
+done
+
+text{* First simplification law for analz: no session keys encrypt
+authentication keys or shared keys. *}
+lemma analz_insert_freshK1:
+     "\<lbrakk> evs \<in> kerbIV_gets;  K \<in> authKeys evs Un range shrK;
+        SesKey \<notin> range shrK \<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key SesKey) (spies evs))) =
+          (K = SesKey | Key K \<in> analz (spies evs))"
+apply (frule authKeys_are_not_AKcryptSK, assumption)
+apply (simp del: image_insert
+            add: analz_image_freshK_simps add: Key_analz_image_Key)
+done
+
+
+text{* Second simplification law for analz: no service keys encrypt any other keys.*}
+lemma analz_insert_freshK2:
+     "\<lbrakk> evs \<in> kerbIV_gets;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
+        K \<in> symKeys \<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key servK) (spies evs))) =
+          (K = servK | Key K \<in> analz (spies evs))"
+apply (frule not_authKeys_not_AKcryptSK, assumption, assumption)
+apply (simp del: image_insert
+            add: analz_image_freshK_simps add: Key_analz_image_Key)
+done
+
+
+text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
+
+lemma analz_insert_freshK3:
+ "\<lbrakk> AKcryptSK authK servK evs;
+    authK' \<noteq> authK; authK' \<notin> range shrK; evs \<in> kerbIV_gets \<rbrakk>
+        \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
+                (servK = authK' | Key servK \<in> analz (spies evs))"
+apply (drule_tac authK' = authK' in not_different_AKcryptSK, blast, assumption)
+apply (simp del: image_insert
+            add: analz_image_freshK_simps add: Key_analz_image_Key)
+done
+
+lemma analz_insert_freshK3_bis:
+ "\<lbrakk> Says Tgs A
+            (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+        \<in> set evs; 
+     authK \<noteq> authK'; authK' \<notin> range shrK; evs \<in> kerbIV_gets \<rbrakk>
+        \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
+                (servK = authK' | Key servK \<in> analz (spies evs))"
+apply (frule AKcryptSKI, assumption)
+apply (simp add: analz_insert_freshK3)
+done
+
+text{*a weakness of the protocol*}
+lemma authK_compromises_servK:
+     "\<lbrakk> Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs;  authK \<in> symKeys;
+         Key authK \<in> analz (spies evs); evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<in> analz (spies evs)"
+by (force dest: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+
+lemma servK_notin_authKeysD:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts,
+                      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>\<rbrace>
+           \<in> parts (spies evs);
+         Key servK \<notin> analz (spies evs);
+         B \<noteq> Tgs; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> servK \<notin> authKeys evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (simp add: authKeys_def)
+apply (erule kerbIV_gets.induct, analz_mono_contra)
+apply (frule_tac [8] Gets_ticket_parts)
+apply (frule_tac [6] Gets_ticket_parts, simp_all)
+apply (blast+)
+done
+
+
+text{*If Spy sees the Authentication Key sent in msg K2, then
+    the Key has expired.*}
+lemma Confidentiality_Kas_lemma [rule_format]:
+     "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Says Kas A
+               (Crypt (shrK A)
+                  \<lbrace>Key authK, Agent Tgs, Number Ta,
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+            \<in> set evs \<longrightarrow>
+          Key authK \<in> analz (spies evs) \<longrightarrow>
+          expiredAK Ta evs"
+apply (erule kerbIV_gets.induct)
+apply (frule_tac [11] Oops_range_spies2)
+apply (frule_tac [10] Oops_range_spies1)
+apply (frule_tac [8] Says_tgs_message_form)
+apply (frule_tac [6] Says_kas_message_form)
+apply (safe del: impI conjI impCE)
+apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*K2*}
+apply blast
+txt{*K4*}
+apply blast
+txt{*Level 8: K5*}
+apply (blast dest: servK_notin_authKeysD Says_Kas_message_form intro: less_SucI)
+txt{*Oops1*}
+apply (blast dest!: unique_authKeys intro: less_SucI)
+txt{*Oops2*}
+apply (blast dest: Says_Tgs_message_form Says_Kas_message_form)
+done
+
+lemma Confidentiality_Kas:
+     "\<lbrakk> Says Kas A
+              (Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
+           \<in> set evs;
+         \<not> expiredAK Ta evs;
+         A \<notin> bad;  evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key authK \<notin> analz (spies evs)"
+by (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
+
+text{*If Spy sees the Service Key sent in msg K4, then
+    the Key has expired.*}
+
+lemma Confidentiality_lemma [rule_format]:
+     "\<lbrakk> Says Tgs A
+	    (Crypt authK
+	       \<lbrace>Key servK, Agent B, Number Ts,
+		 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+	   \<in> set evs;
+	Key authK \<notin> analz (spies evs);
+        servK \<in> symKeys;
+	A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
+	  expiredSK Ts evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply (rule_tac [10] impI)+;
+  --{*The Oops1 case is unusual: must simplify
+    @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
+   @{text analz_mono_contra} weaken it to
+   @{term "Authkey \<notin> analz (spies evs)"},
+  for we then conclude @{term "authK \<noteq> authKa"}.*}
+apply analz_mono_contra
+apply (frule_tac [11] Oops_range_spies2)
+apply (frule_tac [10] Oops_range_spies1)
+apply (frule_tac [8] Says_tgs_message_form)
+apply (frule_tac [6] Says_kas_message_form)
+apply (safe del: impI conjI impCE)
+apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*K2*}
+apply (blast intro: parts_insertI less_SucI)
+txt{*K4*}
+apply (blast dest: authTicket_authentic Confidentiality_Kas)
+txt{*Oops2*}
+  prefer 3
+  apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
+txt{*Oops1*}
+ prefer 2
+apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
+txt{*K5. Not clear how this step could be integrated with the main
+       simplification step. Done in KerberosV.thy*}
+apply clarify
+apply (erule_tac V = "Says Aa Tgs ?X \<in> set ?evs" in thin_rl)
+apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN servK_notin_authKeysD])
+apply (assumption, assumption, blast, assumption)
+apply (simp add: analz_insert_freshK2)
+apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
+done
+
+
+text{* In the real world Tgs can't check wheter authK is secure! *}
+lemma Confidentiality_Tgs:
+     "\<lbrakk> Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs;
+         Key authK \<notin> analz (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest: Says_Tgs_message_form Confidentiality_lemma)
+done
+
+text{* In the real world Tgs CAN check what Kas sends! *}
+lemma Confidentiality_Tgs_bis:
+     "\<lbrakk> Says Kas A
+               (Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
+           \<in> set evs;
+         Says Tgs A
+              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+           \<in> set evs;
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
+done
+
+text{*Most general form*}
+lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
+
+lemmas Confidentiality_Auth_A = authK_authentic [THEN Confidentiality_Kas]
+
+text{*Needs a confidentiality guarantee, hence moved here.
+      Authenticity of servK for A*}
+lemma servK_authentic_bis_r:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; A \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow>Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs"
+apply (blast dest: authK_authentic Confidentiality_Auth_A servK_authentic_ter)
+done
+
+lemma Confidentiality_Serv_A:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (drule authK_authentic, assumption, assumption)
+apply (blast dest: Confidentiality_Kas Says_Kas_message_form servK_authentic_ter Confidentiality_Tgs_bis)
+done
+
+lemma Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (frule authK_authentic)
+apply (frule_tac [3] Confidentiality_Kas)
+apply (frule_tac [6] servTicket_authentic, auto)
+apply (blast dest!: Confidentiality_Tgs_bis dest: Says_Kas_message_form servK_authentic unique_servKeys unique_authKeys)
+done
+
+lemma u_Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad;  B \<noteq> Tgs; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest: u_servTicket_authentic u_NotexpiredSK_NotexpiredAK Confidentiality_Tgs_bis)
+done
+
+
+
+subsection{*2. Parties' strong authentication: 
+       non-injective agreement on the session key. The same guarantees also
+       express key distribution, hence their names*}
+
+text{*Authentication here still is weak agreement - of B with A*}
+lemma A_authenticates_B:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (frule authK_authentic)
+apply assumption+
+apply (frule servK_authentic)
+prefer 2 apply (blast dest: authK_authentic Says_Kas_message_form)
+apply assumption+
+apply (blast dest: K4_imp_K2 Key_unique_SesKey intro!: Says_K6)
+(*Single command proof: slower!
+apply (blast dest: authK_authentic servK_authentic Says_Kas_message_form Key_unique_SesKey K4_imp_K2 intro!: Says_K6)
+*)
+done
+
+(*These two have never been proved, because never were they needed before!*)
+lemma shrK_in_initState_Server[iff]:  "Key (shrK A) \<in> initState Kas"
+by (induct_tac "A", auto)
+lemma shrK_in_knows_Server [iff]: "Key (shrK A) \<in> knows Kas evs"
+by (simp add: initState_subset_knows [THEN subsetD])
+(*Because of our simple model of Tgs, the equivalent for it required an axiom*)
+
+
+lemma A_authenticates_and_keydist_to_Kas:
+   "\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>) \<in> set evs;
+      A \<notin> bad;  evs \<in> kerbIV_gets \<rbrakk>
+  \<Longrightarrow> Says Kas A (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>) \<in> set evs
+  \<and> Key authK \<in> analz(knows Kas evs)"
+apply (force dest!: authK_authentic Says_imp_knows [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+done
+
+
+lemma K3_imp_Gets:
+  "\<lbrakk> Says A Tgs \<lbrace>Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>,
+                 Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace> 
+      \<in> set evs;  A \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow>  Gets A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, 
+                 Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>)
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply auto
+apply (blast dest: authTicket_form)
+done
+
+lemma Tgs_authenticates_and_keydist_to_A:
+  "\<lbrakk>  Gets Tgs \<lbrace>
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>,
+          Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B \<rbrace> \<in> set evs;
+      Key authK \<notin> analz (spies evs); A \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists> B. Says A Tgs \<lbrace>
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>,
+          Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B \<rbrace> \<in> set evs
+ \<and>  Key authK \<in> analz (knows A evs)"  
+apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst], assumption)
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN parts.Fst], assumption)
+apply (drule Tgs_authenticates_A, assumption+, simp)
+apply (force dest!: K3_imp_Gets Gets_imp_knows [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+done
+
+lemma K4_imp_Gets:
+  "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists> Ta X. 
+     Gets Tgs \<lbrace>Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>, X\<rbrace>
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply auto
+done
+
+lemma A_authenticates_and_keydist_to_Tgs:
+ "\<lbrakk>  Gets A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
+       \<in> set evs;
+     Gets A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs;
+     Key authK \<notin> analz (spies evs); A \<notin> bad;
+     evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
+       \<in> set evs
+  \<and> Key authK \<in> analz (knows Tgs evs)
+  \<and> Key servK \<in> analz (knows Tgs evs)"
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj], assumption)
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj], assumption)
+apply (frule authK_authentic, assumption+)
+apply (drule servK_authentic_ter, assumption+)
+apply (frule K4_imp_Gets, assumption, erule exE, erule exE)
+apply (drule Gets_imp_knows [THEN analz.Inj, THEN analz.Fst, THEN analz.Decrypt, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst], assumption, force)
+apply (frule Says_imp_knows [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+apply (force dest!: authK_authentic Says_Kas_message_form)
+apply simp
+done
+
+lemma K5_imp_Gets:
+  "\<lbrakk> Says A B \<lbrace>servTicket, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs;
+    A \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists> authK Ts authTicket T2.
+    Gets A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>) \<in> set evs
+ \<and> Says A Tgs \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>  \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply auto
+done 
+
+lemma K3_imp_Gets:
+  "\<lbrakk> Says A Tgs \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>
+       \<in> set evs;
+    A \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> \<exists> Ta. Gets A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>) \<in> set evs";
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply auto
+done 
+
+lemma B_authenticates_and_keydist_to_A:
+     "\<lbrakk> Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>,
+                Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs;
+        Key servK \<notin> analz (spies evs);
+        A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>,
+               Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs
+  \<and> Key servK \<in> analz (knows A evs)"
+apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst, THEN servTicket_authentic_Tgs], assumption+)  
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd], assumption)
+apply (erule exE, drule Says_K5, assumption+)
+apply (frule K5_imp_Gets, assumption+)
+apply clarify
+apply (drule K3_imp_Gets, assumption+)
+apply (erule exE)
+apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN authK_authentic, THEN Says_Kas_message_form], assumption+, clarify)
+apply (force dest!: Gets_imp_knows [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
+done
+
+
+lemma K6_imp_Gets:
+  "\<lbrakk> Says B A (Crypt servK (Number T3)) \<in> set evs;
+     B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+\<Longrightarrow> \<exists> Ts X. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>,X\<rbrace>
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbIV_gets.induct)
+apply auto
+done
+
+
+lemma A_authenticates_and_keydist_to_B:
+  "\<lbrakk> Gets A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>,
+             Crypt servK (Number T3)\<rbrace> \<in> set evs;
+     Gets A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
+           \<in> set evs;
+     Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
+     A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs 
+   \<and> Key servK \<in> analz (knows B evs)"
+apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst], assumption)
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd], assumption)
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj], assumption)
+apply (drule A_authenticates_B, assumption+)
+apply (force dest!: K6_imp_Gets Gets_imp_knows [THEN analz.Inj, THEN analz.Fst, THEN analz.Decrypt, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
+done
+
+ 
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/KerberosV.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,1654 @@
+(*  ID:         $Id$
+    Author:     Giampaolo Bella, Catania University
+*)
+
+header{*The Kerberos Protocol, Version V*}
+
+theory KerberosV imports Public begin
+
+text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+
+syntax
+  Kas :: agent
+  Tgs :: agent  --{*the two servers are translations...*}
+
+
+translations
+  "Kas"       == "Server "
+  "Tgs"       == "Friend 0"
+
+
+axioms
+  Tgs_not_bad [iff]: "Tgs \<notin> bad"
+   --{*Tgs is secure --- we already know that Kas is secure*}
+
+(*The current time is just the length of the trace!*)
+syntax
+    CT :: "event list=>nat"
+
+    expiredAK :: "[nat, event list] => bool"
+
+    expiredSK :: "[nat, event list] => bool"
+
+    expiredA :: "[nat, event list] => bool"
+
+    valid :: "[nat, nat] => bool"  ("valid _ wrt _")
+
+
+constdefs
+ (* authKeys are those contained in an authTicket *)
+    authKeys :: "event list => key set"
+    "authKeys evs == {authK. \<exists>A Peer Ta. 
+        Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Peer, Ta\<rbrace>,
+                     Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key authK, Ta\<rbrace>
+                   \<rbrace> \<in> set evs}"
+
+ (* A is the true creator of X if she has sent X and X never appeared on
+    the trace before this event. Recall that traces grow from head. *)
+  Issues :: "[agent, agent, msg, event list] => bool"
+             ("_ Issues _ with _ on _")
+   "A Issues B with X on evs ==
+      \<exists>Y. Says A B Y \<in> set evs \<and> X \<in> parts {Y} \<and>
+      X \<notin> parts (spies (takeWhile (% z. z  \<noteq> Says A B Y) (rev evs)))"
+
+
+consts
+    (*Duration of the authentication key*)
+    authKlife   :: nat
+
+    (*Duration of the service key*)
+    servKlife   :: nat
+
+    (*Duration of an authenticator*)
+    authlife   :: nat
+
+    (*Upper bound on the time of reaction of a server*)
+    replylife   :: nat
+
+specification (authKlife)
+  authKlife_LB [iff]: "2 \<le> authKlife"
+    by blast
+
+specification (servKlife)
+  servKlife_LB [iff]: "2 + authKlife \<le> servKlife"
+    by blast
+
+specification (authlife)
+  authlife_LB [iff]: "Suc 0 \<le> authlife"
+    by blast
+
+specification (replylife)
+  replylife_LB [iff]: "Suc 0 \<le> replylife"
+    by blast
+
+translations
+   "CT" == "length "
+
+   "expiredAK T evs" == "authKlife + T < CT evs"
+
+   "expiredSK T evs" == "servKlife + T < CT evs"
+
+   "expiredA T evs" == "authlife + T < CT evs"
+
+   "valid T1 wrt T2" == "T1 <= replylife + T2"
+
+(*---------------------------------------------------------------------*)
+
+
+(* Predicate formalising the association between authKeys and servKeys *)
+constdefs
+  AKcryptSK :: "[key, key, event list] => bool"
+  "AKcryptSK authK servK evs ==
+     \<exists>A B tt.
+       Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>,
+                    Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, tt\<rbrace> \<rbrace>
+         \<in> set evs"
+
+consts
+
+kerbV   :: "event list set"
+inductive "kerbV"
+  intros
+
+   Nil:  "[] \<in> kerbV"
+
+   Fake: "\<lbrakk> evsf \<in> kerbV;  X \<in> synth (analz (spies evsf)) \<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> kerbV"
+
+
+(*Authentication phase*)
+   KV1:   "\<lbrakk> evs1 \<in> kerbV \<rbrakk>
+          \<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
+          \<in> kerbV"
+   (*Unlike version IV, authTicket is not re-encrypted*)
+   KV2:  "\<lbrakk> evs2 \<in> kerbV; Key authK \<notin> used evs2; authK \<in> symKeys;
+            Says A' Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Kas A \<lbrace>
+          Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2)\<rbrace>,
+        Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number (CT evs2)\<rbrace> 
+                         \<rbrace> # evs2 \<in> kerbV"
+
+
+(* Authorisation phase *)
+   KV3:  "\<lbrakk> evs3 \<in> kerbV; A \<noteq> Kas; A \<noteq> Tgs;
+            Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
+            Says Kas' A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+                          authTicket\<rbrace> \<in> set evs3;
+            valid Ta wrt T1
+         \<rbrakk>
+          \<Longrightarrow> Says A Tgs \<lbrace>authTicket,
+                           (Crypt authK \<lbrace>Agent A, Number (CT evs3)\<rbrace>),
+                           Agent B\<rbrace> # evs3 \<in> kerbV"
+   (*Unlike version IV, servTicket is not re-encrypted*)
+   KV4:  "\<lbrakk> evs4 \<in> kerbV; Key servK \<notin> used evs4; servK \<in> symKeys;
+            B \<noteq> Tgs;  authK \<in> symKeys;
+            Says A' Tgs \<lbrace>
+             (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
+				 Number Ta\<rbrace>),
+             (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
+	        \<in> set evs4;
+            \<not> expiredAK Ta evs4;
+            \<not> expiredA T2 evs4;
+            servKlife + (CT evs4) <= authKlife + Ta
+         \<rbrakk>
+          \<Longrightarrow> Says Tgs A \<lbrace>
+             Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4)\<rbrace>,
+   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number (CT evs4)\<rbrace> 
+                          \<rbrace> # evs4 \<in> kerbV"
+
+
+(*Service phase*)
+   KV5:  "\<lbrakk> evs5 \<in> kerbV; authK \<in> symKeys; servK \<in> symKeys;
+            A \<noteq> Kas; A \<noteq> Tgs;
+            Says A Tgs
+                \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
+		  Agent B\<rbrace>
+              \<in> set evs5;
+            Says Tgs' A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+                          servTicket\<rbrace>
+                \<in> set evs5;
+            valid Ts wrt T2 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>servTicket,
+			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+               # evs5 \<in> kerbV"
+
+    KV6:  "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
+            Says A' B \<lbrace>
+              (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
+              (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
+            \<in> set evs6;
+            \<not> expiredSK Ts evs6;
+            \<not> expiredA T3 evs6
+         \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt servK (Number Ta2))
+               # evs6 \<in> kerbV"
+
+
+
+(* Leaking an authK... *)
+   Oops1:"\<lbrakk> evsO1 \<in> kerbV;  A \<noteq> Spy;
+             Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+                          authTicket\<rbrace>  \<in> set evsO1;
+              expiredAK Ta evsO1 \<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent Tgs, Number Ta, Key authK\<rbrace>
+               # evsO1 \<in> kerbV"
+
+(*Leaking a servK... *)
+   Oops2: "\<lbrakk> evsO2 \<in> kerbV;  A \<noteq> Spy;
+              Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+                           servTicket\<rbrace>  \<in> set evsO2;
+              expiredSK Ts evsO2 \<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Number Ts, Key servK\<rbrace>
+               # evsO2 \<in> kerbV"
+
+
+
+declare Says_imp_knows_Spy [THEN parts.Inj, dest]
+declare parts.Body [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un [dest]
+
+
+
+subsection{*Lemmas about lists, for reasoning about  Issues*}
+
+lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
+          (if A:bad then insert X (spies evs) else spies evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_evs_rev: "spies evs = spies (rev evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a")
+apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
+done
+
+lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono]
+
+lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+done
+
+lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
+
+
+subsection{*Lemmas about @{term authKeys}*}
+
+lemma authKeys_empty: "authKeys [] = {}"
+apply (unfold authKeys_def)
+apply (simp (no_asm))
+done
+
+lemma authKeys_not_insert:
+ "(\<forall>A Ta akey Peer.
+   ev \<noteq> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>akey, Agent Peer, Ta\<rbrace>,
+                     Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, akey, Ta\<rbrace> \<rbrace>)
+       \<Longrightarrow> authKeys (ev # evs) = authKeys evs"
+apply (unfold authKeys_def, auto)
+done
+
+lemma authKeys_insert:
+  "authKeys
+     (Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent Peer, Number Ta\<rbrace>,
+         Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K, Number Ta\<rbrace> \<rbrace> # evs)
+       = insert K (authKeys evs)"
+apply (unfold authKeys_def, auto)
+done
+
+lemma authKeys_simp:
+   "K \<in> authKeys
+    (Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key K', Agent Peer, Number Ta\<rbrace>,
+        Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K', Number Ta\<rbrace> \<rbrace> # evs)
+        \<Longrightarrow> K = K' | K \<in> authKeys evs"
+apply (unfold authKeys_def, auto)
+done
+
+lemma authKeysI:
+   "Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent Tgs, Number Ta\<rbrace>,
+         Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key K, Number Ta\<rbrace> \<rbrace> \<in> set evs
+        \<Longrightarrow> K \<in> authKeys evs"
+apply (unfold authKeys_def, auto)
+done
+
+lemma authKeys_used: "K \<in> authKeys evs \<Longrightarrow> Key K \<in> used evs"
+apply (simp add: authKeys_def, blast)
+done
+
+
+subsection{*Forwarding Lemmas*}
+
+lemma Says_ticket_parts:
+     "Says S A \<lbrace>Crypt K \<lbrace>SesKey, B, TimeStamp\<rbrace>, Ticket\<rbrace>
+               \<in> set evs \<Longrightarrow> Ticket \<in> parts (spies evs)"
+apply blast
+done
+
+lemma Says_ticket_analz:
+     "Says S A \<lbrace>Crypt K \<lbrace>SesKey, B, TimeStamp\<rbrace>, Ticket\<rbrace>
+               \<in> set evs \<Longrightarrow> Ticket \<in> analz (spies evs)"
+apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd])
+done
+
+lemma Oops_range_spies1:
+     "\<lbrakk> Says Kas A \<lbrace>Crypt KeyA \<lbrace>Key authK, Peer, Ta\<rbrace>, authTicket\<rbrace>
+           \<in> set evs ;
+         evs \<in> kerbV \<rbrakk> \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys"
+apply (erule rev_mp)
+apply (erule kerbV.induct, auto)
+done
+
+lemma Oops_range_spies2:
+     "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, servTicket\<rbrace>
+           \<in> set evs ;
+         evs \<in> kerbV \<rbrakk> \<Longrightarrow> servK \<notin> range shrK \<and> servK \<in> symKeys"
+apply (erule rev_mp)
+apply (erule kerbV.induct, auto)
+done
+
+
+(*Spy never sees another agent's shared key! (unless it's lost at start)*)
+lemma Spy_see_shrK [simp]:
+     "evs \<in> kerbV \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+apply (blast+)
+done
+
+lemma Spy_analz_shrK [simp]:
+     "evs \<in> kerbV \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
+by auto
+
+lemma Spy_see_shrK_D [dest!]:
+     "\<lbrakk> Key (shrK A) \<in> parts (spies evs);  evs \<in> kerbV \<rbrakk> \<Longrightarrow> A:bad"
+by (blast dest: Spy_see_shrK)
+lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
+
+text{*Nobody can have used non-existent keys!*}
+lemma new_keys_not_used [simp]:
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbV\<rbrakk>
+     \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+txt{*Fake*}
+apply (force dest!: keysFor_parts_insert)
+txt{*Others*}
+apply (force dest!: analz_shrK_Decrypt)+
+done
+
+(*Earlier, all protocol proofs declared this theorem.
+  But few of them actually need it! (Another is Yahalom) *)
+lemma new_keys_not_analzd:
+ "\<lbrakk>evs \<in> kerbV; K \<in> symKeys; Key K \<notin> used evs\<rbrakk>
+  \<Longrightarrow> K \<notin> keysFor (analz (spies evs))"
+by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD])
+
+
+
+subsection{*Regularity Lemmas*}
+text{*These concern the form of items passed in messages*}
+
+text{*Describes the form of all components sent by Kas*}
+lemma Says_Kas_message_form:
+     "\<lbrakk> Says Kas A \<lbrace>Crypt K \<lbrace>Key authK, Agent Peer, Ta\<rbrace>, authTicket\<rbrace>
+           \<in> set evs;
+         evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> authK \<notin> range shrK \<and> authK \<in> authKeys evs \<and> authK \<in> symKeys \<and> 
+  authTicket = (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>) \<and>
+             K = shrK A  \<and> Peer = Tgs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all (no_asm) add: authKeys_def authKeys_insert)
+apply blast+
+done
+
+
+
+(*This lemma is essential for proving Says_Tgs_message_form:
+
+  the session key authK
+  supplied by Kas in the authentication ticket
+  cannot be a long-term key!
+
+  Generalised to any session keys (both authK and servK).
+*)
+lemma SesKey_is_session_key:
+     "\<lbrakk> Crypt (shrK Tgs_B) \<lbrace>Agent A, Agent Tgs_B, Key SesKey, Number T\<rbrace>
+            \<in> parts (spies evs); Tgs_B \<notin> bad;
+         evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> SesKey \<notin> range shrK"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
+done
+
+lemma authTicket_authentic:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>
+           \<in> parts (spies evs);
+         evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>,
+                 Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>\<rbrace>
+            \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+txt{*Fake, K4*}
+apply (blast+)
+done
+
+lemma authTicket_crypt_authK:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> authK \<in> authKeys evs"
+apply (frule authTicket_authentic, assumption)
+apply (simp (no_asm) add: authKeys_def)
+apply blast
+done
+
+text{*Describes the form of servK, servTicket and authK sent by Tgs*}
+lemma Says_Tgs_message_form:
+     "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, servTicket\<rbrace>
+           \<in> set evs;
+         evs \<in> kerbV \<rbrakk>
+   \<Longrightarrow> B \<noteq> Tgs \<and> 
+       servK \<notin> range shrK \<and> servK \<notin> authKeys evs \<and> servK \<in> symKeys \<and>
+       servTicket = (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>) \<and>
+       authK \<notin> range shrK \<and> authK \<in> authKeys evs \<and> authK \<in> symKeys"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast, auto)
+txt{*Three subcases of Message 4*}
+apply (blast dest!: authKeys_used Says_Kas_message_form)
+apply (blast dest!: SesKey_is_session_key)
+apply (blast dest: authTicket_crypt_authK)
+done
+
+
+
+(*
+lemma authTicket_form:
+lemma servTicket_form:
+lemma Says_kas_message_form:
+lemma Says_tgs_message_form:
+
+cannot be proved for version V, but a new proof strategy can be used in their
+place. The new strategy merely says that both the authTicket and the servTicket
+are in parts and in analz as soon as they appear, using lemmas Says_ticket_parts and Says_ticket_analz. 
+The new strategy always lets the simplifier solve cases K3 and K5, saving on
+long dedicated analyses, which seemed unavoidable. For this reason, lemma 
+servK_notin_authKeysD is no longer needed.
+*)
+
+subsection{*Authenticity theorems: confirm origin of sensitive messages*}
+
+lemma authK_authentic:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>
+           \<in> parts (spies evs);
+         A \<notin> bad;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<exists> AT. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>, AT\<rbrace>
+            \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+apply blast+
+done
+
+text{*If a certain encrypted message appears then it originated with Tgs*}
+lemma servK_authentic:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs);
+         authK \<notin> range shrK;
+         evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> \<exists>A ST. Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, ST\<rbrace>
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+apply blast+
+done
+
+lemma servK_authentic_bis:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs);
+         B \<noteq> Tgs;
+         evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> \<exists>A ST. Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, ST\<rbrace>
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast+)
+done
+
+text{*Authenticity of servK for B*}
+lemma servTicket_authentic_Tgs:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> \<exists>authK.
+       Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>,
+                    Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>\<rbrace>
+       \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast+)
+done
+
+text{*Anticipated here from next subsection*}
+lemma K4_imp_K2:
+"\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+      \<in> set evs;  evs \<in> kerbV\<rbrakk>
+   \<Longrightarrow> \<exists>Ta. Says Kas A
+        \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
+        \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, auto)
+apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
+done
+
+text{*Anticipated here from next subsection*}
+lemma u_K4_imp_K2:
+"\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>  \<in> set evs; evs \<in> kerbV\<rbrakk>
+   \<Longrightarrow> \<exists>Ta. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+             Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
+             \<in> set evs
+          \<and> servKlife + Ts <= authKlife + Ta"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, auto)
+apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
+done
+
+lemma servTicket_authentic_Kas:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<exists>authK Ta.
+       Says Kas A
+         \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
+        \<in> set evs"
+apply (blast dest!: servTicket_authentic_Tgs K4_imp_K2)
+done
+
+lemma u_servTicket_authentic_Kas:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<exists>authK Ta.
+       Says Kas A
+         \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
+        \<in> set evs \<and> 
+      servKlife + Ts <= authKlife + Ta"
+apply (blast dest!: servTicket_authentic_Tgs u_K4_imp_K2)
+done
+
+lemma servTicket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> \<exists>Ta authK.
+     Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+                  Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>  \<in> set evs
+     \<and> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+                  Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
+       \<in> set evs"
+apply (blast dest: servTicket_authentic_Tgs K4_imp_K2)
+done
+
+lemma u_servTicket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
+         evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> \<exists>Ta authK.
+     Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+                  Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace> \<in> set evs
+     \<and> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
+       \<in> set evs
+     \<and> servKlife + Ts <= authKlife + Ta"
+apply (blast dest: servTicket_authentic_Tgs u_K4_imp_K2)
+done
+
+lemma u_NotexpiredSK_NotexpiredAK:
+     "\<lbrakk> \<not> expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \<rbrakk>
+      \<Longrightarrow> \<not> expiredAK Ta evs"
+apply (blast dest: leI le_trans dest: leD)
+done
+
+
+subsection{* Reliability: friendly agents send somthing if something else happened*}
+
+lemma K3_imp_K2:
+     "\<lbrakk> Says A Tgs
+             \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>
+           \<in> set evs;
+         A \<notin> bad;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<exists>Ta AT. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, 
+                               AT\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast, blast)
+apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authK_authentic])
+done
+
+text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
+lemma Key_unique_SesKey:
+     "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T\<rbrace>
+           \<in> parts (spies evs);
+         Crypt K' \<lbrace>Key SesKey,  Agent B', T'\<rbrace>
+           \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
+         evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> K=K' \<and> B=B' \<and> T=T'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+txt{*Fake, K2, K4*}
+apply (blast+)
+done
+
+text{*This inevitably has an existential form in version V*}
+lemma Says_K5:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+                                     servTicket\<rbrace> \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+apply blast
+txt{*K3*}
+apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
+txt{*K4*}
+apply (force dest!: Crypt_imp_keysFor)
+txt{*K5*}
+apply (blast dest: Key_unique_SesKey)
+done
+
+text{*Anticipated here from next subsection*}
+lemma unique_CryptKey:
+     "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SesKey, T'\<rbrace>
+           \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
+         evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> A=A' & B=B' & T=T'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+txt{*Fake, K2, K4*}
+apply (blast+)
+done
+
+lemma Says_K6:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+                      servTicket\<rbrace> \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (frule Says_Tgs_message_form, assumption, clarify)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts)
+apply (simp_all (no_asm_simp))
+
+txt{*fake*}
+apply blast
+txt{*K4*}
+apply (force dest!: Crypt_imp_keysFor, clarify)
+txt{*K6*}
+apply (drule  Says_imp_spies [THEN parts.Inj, THEN parts.Fst])
+apply (drule  Says_imp_spies [THEN parts.Inj, THEN parts.Snd])
+apply (blast dest!: unique_CryptKey)
+done
+
+text{*Needs a unicity theorem, hence moved here*}
+lemma servK_authentic_ter:
+ "\<lbrakk> Says Kas A
+       \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, authTicket\<rbrace> \<in> set evs;
+     Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
+       \<in> parts (spies evs);
+     Key authK \<notin> analz (spies evs);
+     evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, 
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace> \<rbrace>
+       \<in> set evs"
+apply (frule Says_Kas_message_form, assumption)
+apply clarify
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
+txt{*K2 and K4 remain*}
+apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
+apply (blast dest!: unique_CryptKey)
+done
+
+
+subsection{*Unicity Theorems*}
+
+text{* The session key, if secure, uniquely identifies the Ticket
+   whether authTicket or servTicket. As a matter of fact, one can read
+   also Tgs in the place of B.                                     *}
+
+
+lemma unique_authKeys:
+     "\<lbrakk> Says Kas A
+              \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, X\<rbrace> \<in> set evs;
+         Says Kas A'
+              \<lbrace>Crypt Ka' \<lbrace>Key authK, Agent Tgs, Ta'\<rbrace>, X'\<rbrace> \<in> set evs;
+         evs \<in> kerbV \<rbrakk> \<Longrightarrow> A=A' \<and> Ka=Ka' \<and> Ta=Ta' \<and> X=X'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+apply blast+
+done
+
+text{* servK uniquely identifies the message from Tgs *}
+lemma unique_servKeys:
+     "\<lbrakk> Says Tgs A
+              \<lbrace>Crypt K \<lbrace>Key servK, Agent B, Ts\<rbrace>, X\<rbrace> \<in> set evs;
+         Says Tgs A'
+              \<lbrace>Crypt K' \<lbrace>Key servK, Agent B', Ts'\<rbrace>, X'\<rbrace> \<in> set evs;
+         evs \<in> kerbV \<rbrakk> \<Longrightarrow> A=A' \<and> B=B' \<and> K=K' \<and> Ts=Ts' \<and> X=X'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+apply blast+
+done
+
+subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
+
+lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
+apply (simp add: AKcryptSK_def)
+done
+
+lemma AKcryptSKI:
+ "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>, X \<rbrace> \<in> set evs;
+     evs \<in> kerbV \<rbrakk> \<Longrightarrow> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
+apply (blast dest: Says_Tgs_message_form)
+done
+
+lemma AKcryptSK_Says [simp]:
+   "AKcryptSK authK servK (Says S A X # evs) =
+     (S = Tgs \<and>
+      (\<exists>B tt. X = \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>,
+                    Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, tt\<rbrace> \<rbrace>)
+     | AKcryptSK authK servK evs)"
+apply (unfold AKcryptSK_def)
+apply (simp (no_asm))
+apply blast
+done
+
+lemma AKcryptSK_Notes [simp]:
+   "AKcryptSK authK servK (Notes A X # evs) =
+      AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
+apply (simp (no_asm))
+done
+
+(*A fresh authK cannot be associated with any other
+  (with respect to a given trace). *)
+lemma Auth_fresh_not_AKcryptSK:
+     "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def)
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
+done
+
+(*A fresh servK cannot be associated with any other
+  (with respect to a given trace). *)
+lemma Serv_fresh_not_AKcryptSK:
+ "Key servK \<notin> used evs \<Longrightarrow> \<not> AKcryptSK authK servK evs"
+apply (unfold AKcryptSK_def, blast)
+done
+
+lemma authK_not_AKcryptSK:
+     "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, tk\<rbrace>
+           \<in> parts (spies evs);  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK K authK evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all)
+txt{*Fake*}
+apply blast
+txt{*K2: by freshness*}
+apply (simp add: AKcryptSK_def)
+apply blast
+txt{*K4*}
+apply blast
+done
+
+text{*A secure serverkey cannot have been used to encrypt others*}
+lemma servK_not_AKcryptSK:
+ "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, tt\<rbrace> \<in> parts (spies evs);
+     Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
+     B \<noteq> Tgs;  evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<not> AKcryptSK SK K evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
+txt{*K4 splits into distinct subcases*}
+apply auto
+txt{*servK can't have been enclosed in two certificates*}
+ prefer 2 apply (blast dest: unique_CryptKey)
+txt{*servK is fresh and so could not have been used, by
+   @{text new_keys_not_used}*}
+apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
+done
+
+text{*Long term keys are not issued as servKeys*}
+lemma shrK_not_AKcryptSK:
+     "evs \<in> kerbV \<Longrightarrow> \<not> AKcryptSK K (shrK A) evs"
+apply (unfold AKcryptSK_def)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, auto)
+done
+
+text{*The Tgs message associates servK with authK and therefore not with any
+  other key authK.*}
+lemma Says_Tgs_AKcryptSK:
+     "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>, X \<rbrace>
+           \<in> set evs;
+         authK' \<noteq> authK;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK' servK evs"
+apply (unfold AKcryptSK_def)
+apply (blast dest: unique_servKeys)
+done
+
+lemma AKcryptSK_not_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK K evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts)
+apply (simp_all, safe)
+txt{*K4 splits into subcases*}
+(*apply simp_all*)
+prefer 4 apply (blast dest!: authK_not_AKcryptSK)
+txt{*servK is fresh and so could not have been used, by
+   @{text new_keys_not_used}*}
+ prefer 2 
+ apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
+txt{*Others by freshness*}
+apply (blast+)
+done
+
+lemma not_different_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;
+        authK' \<noteq> authK;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK authK' servK evs  \<and> servK \<in> symKeys"
+apply (simp add: AKcryptSK_def)
+apply (blast dest: unique_servKeys Says_Tgs_message_form)
+done
+
+lemma AKcryptSK_not_AKcryptSK:
+     "\<lbrakk> AKcryptSK authK servK evs;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK K evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts, simp_all, safe)
+txt{*K4 splits into subcases*}
+apply simp_all
+prefer 4 apply (blast dest!: authK_not_AKcryptSK)
+txt{*servK is fresh and so could not have been used, by
+   @{text new_keys_not_used}*}
+ prefer 2 
+ apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
+txt{*Others by freshness*}
+apply (blast+)
+done
+
+text{*The only session keys that can be found with the help of session keys are
+  those sent by Tgs in step K4.  *}
+
+text{*We take some pains to express the property
+  as a logical equivalence so that the simplifier can apply it.*}
+lemma Key_analz_image_Key_lemma:
+     "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
+      \<Longrightarrow>
+      P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) = (K:KK | Key K \<in> analz H)"
+by (blast intro: analz_mono [THEN subsetD])
+
+
+lemma AKcryptSK_analz_insert:
+     "\<lbrakk> AKcryptSK K K' evs; K \<in> symKeys; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key K' \<in> analz (insert (Key K) (spies evs))"
+apply (simp add: AKcryptSK_def, clarify)
+apply (drule Says_imp_spies [THEN analz.Inj, THEN analz_insertI], auto)
+done
+
+lemma authKeys_are_not_AKcryptSK:
+     "\<lbrakk> K \<in> authKeys evs Un range shrK;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<forall>SK. \<not> AKcryptSK SK K evs \<and> K \<in> symKeys"
+apply (simp add: authKeys_def AKcryptSK_def)
+apply (blast dest: Says_Kas_message_form Says_Tgs_message_form)
+done
+
+lemma not_authKeys_not_AKcryptSK:
+     "\<lbrakk> K \<notin> authKeys evs;
+         K \<notin> range shrK; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> \<forall>SK. \<not> AKcryptSK K SK evs"
+apply (simp add: AKcryptSK_def)
+apply (blast dest: Says_Tgs_message_form)
+done
+
+
+subsection{*Secrecy Theorems*}
+
+text{*For the Oops2 case of the next theorem*}
+lemma Oops2_not_AKcryptSK:
+     "\<lbrakk> evs \<in> kerbV;
+         Says Tgs A \<lbrace>Crypt authK
+                     \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+           \<in> set evs \<rbrakk>
+      \<Longrightarrow> \<not> AKcryptSK servK SK evs"
+apply (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
+done
+   
+text{* Big simplification law for keys SK that are not crypted by keys in KK
+ It helps prove three, otherwise hard, facts about keys. These facts are
+ exploited as simplification laws for analz, and also "limit the damage"
+ in case of loss of a key to the spy. See ESORICS98.*}
+lemma Key_analz_image_Key [rule_format (no_asm)]:
+     "evs \<in> kerbV \<Longrightarrow>
+      (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
+       (\<forall>K \<in> KK. \<not> AKcryptSK K SK evs)   \<longrightarrow>
+       (Key SK \<in> analz (Key`KK Un (spies evs))) =
+       (SK \<in> KK | Key SK \<in> analz (spies evs)))"
+apply (erule kerbV.induct)
+apply (frule_tac [10] Oops_range_spies2)
+apply (frule_tac [9] Oops_range_spies1)
+(*Used to apply Says_tgs_message form, which is no longer available. 
+  Instead\<dots>*)
+apply (drule_tac [7] Says_ticket_analz)
+(*Used to apply Says_kas_message form, which is no longer available. 
+  Instead\<dots>*)
+apply (drule_tac [5] Says_ticket_analz)
+apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
+txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
+ the induction hypothesis*}
+apply (case_tac [9] "AKcryptSK authK SK evsO1")
+apply (case_tac [7] "AKcryptSK servK SK evs5")
+apply (simp_all del: image_insert
+          add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
+               Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
+               Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
+txt{*Fake*} 
+apply spy_analz
+txt{*K2*}
+apply blast 
+txt{*Cases K3 and K5 solved by the simplifier thanks to the ticket being in 
+analz - this strategy is new wrt version IV*} 
+txt{*K4*}
+apply (blast dest!: authK_not_AKcryptSK)
+txt{*Oops1*}
+apply clarify
+apply simp
+apply (blast dest!: AKcryptSK_analz_insert)
+done
+
+text{* First simplification law for analz: no session keys encrypt
+authentication keys or shared keys. *}
+lemma analz_insert_freshK1:
+     "\<lbrakk> evs \<in> kerbV;  K \<in> authKeys evs Un range shrK;
+        SesKey \<notin> range shrK \<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key SesKey) (spies evs))) =
+          (K = SesKey | Key K \<in> analz (spies evs))"
+apply (frule authKeys_are_not_AKcryptSK, assumption)
+apply (simp del: image_insert
+            add: analz_image_freshK_simps add: Key_analz_image_Key)
+done
+
+
+text{* Second simplification law for analz: no service keys encrypt any other keys.*}
+lemma analz_insert_freshK2:
+     "\<lbrakk> evs \<in> kerbV;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
+        K \<in> symKeys \<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key servK) (spies evs))) =
+          (K = servK | Key K \<in> analz (spies evs))"
+apply (frule not_authKeys_not_AKcryptSK, assumption, assumption)
+apply (simp del: image_insert
+            add: analz_image_freshK_simps add: Key_analz_image_Key)
+done
+
+
+text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
+
+lemma analz_insert_freshK3:
+ "\<lbrakk> AKcryptSK authK servK evs;
+    authK' \<noteq> authK; authK' \<notin> range shrK; evs \<in> kerbV \<rbrakk>
+        \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
+                (servK = authK' | Key servK \<in> analz (spies evs))"
+apply (drule_tac authK' = authK' in not_different_AKcryptSK, blast, assumption)
+apply (simp del: image_insert
+            add: analz_image_freshK_simps add: Key_analz_image_Key)
+done
+
+lemma analz_insert_freshK3_bis:
+ "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+        \<in> set evs; 
+     authK \<noteq> authK'; authK' \<notin> range shrK; evs \<in> kerbV \<rbrakk>
+        \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
+                (servK = authK' | Key servK \<in> analz (spies evs))"
+apply (frule AKcryptSKI, assumption)
+apply (simp add: analz_insert_freshK3)
+done
+
+text{*a weakness of the protocol*}
+lemma authK_compromises_servK:
+     "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+        \<in> set evs;  authK \<in> symKeys;
+         Key authK \<in> analz (spies evs); evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<in> analz (spies evs)"
+apply (force dest: Says_imp_spies [THEN analz.Inj, THEN analz.Fst, THEN analz.Decrypt, THEN analz.Fst])
+done
+
+text{*lemma @{text servK_notin_authKeysD} not needed in version V*}
+
+text{*If Spy sees the Authentication Key sent in msg K2, then
+    the Key has expired.*}
+lemma Confidentiality_Kas_lemma [rule_format]:
+     "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Says Kas A
+               \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
+          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>
+            \<in> set evs \<longrightarrow>
+          Key authK \<in> analz (spies evs) \<longrightarrow>
+          expiredAK Ta evs"
+apply (erule kerbV.induct)
+apply (frule_tac [10] Oops_range_spies2)
+apply (frule_tac [9] Oops_range_spies1)
+apply (frule_tac [7] Says_ticket_analz)
+apply (frule_tac [5] Says_ticket_analz)
+apply (safe del: impI conjI impCE)
+apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*K2*}
+apply blast
+txt{*K4*}
+apply blast
+txt{*Oops1*}
+apply (blast dest!: unique_authKeys intro: less_SucI)
+txt{*Oops2*}
+apply (blast dest: Says_Tgs_message_form Says_Kas_message_form)
+done
+
+lemma Confidentiality_Kas:
+     "\<lbrakk> Says Kas A
+              \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>, authTicket\<rbrace>
+           \<in> set evs;
+        \<not> expiredAK Ta evs;
+        A \<notin> bad;  evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key authK \<notin> analz (spies evs)"
+apply (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
+done
+
+text{*If Spy sees the Service Key sent in msg K4, then
+    the Key has expired.*}
+
+lemma Confidentiality_lemma [rule_format]:
+     "\<lbrakk> Says Tgs A
+	    \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
+	   \<in> set evs;
+	Key authK \<notin> analz (spies evs);
+        servK \<in> symKeys;
+	A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
+	  expiredSK Ts evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (rule_tac [9] impI)+;
+  --{*The Oops1 case is unusual: must simplify
+    @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
+   @{text analz_mono_contra} weaken it to
+   @{term "Authkey \<notin> analz (spies evs)"},
+  for we then conclude @{term "authK \<noteq> authKa"}.*}
+apply analz_mono_contra
+apply (frule_tac [10] Oops_range_spies2)
+apply (frule_tac [9] Oops_range_spies1)
+apply (frule_tac [7] Says_ticket_analz)
+apply (frule_tac [5] Says_ticket_analz)
+apply (safe del: impI conjI impCE)
+apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*K2*}
+apply (blast intro: parts_insertI less_SucI)
+txt{*K4*}
+apply (blast dest: authTicket_authentic Confidentiality_Kas)
+txt{*Oops1*}
+ apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
+txt{*Oops2*}
+  apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
+done
+
+
+text{* In the real world Tgs can't check wheter authK is secure! *}
+lemma Confidentiality_Tgs:
+     "\<lbrakk> Says Tgs A
+              \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+           \<in> set evs;
+         Key authK \<notin> analz (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest: Says_Tgs_message_form Confidentiality_lemma)
+done
+
+text{* In the real world Tgs CAN check what Kas sends! *}
+lemma Confidentiality_Tgs_bis:
+     "\<lbrakk> Says Kas A
+               \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>, authTicket\<rbrace>
+           \<in> set evs;
+         Says Tgs A
+              \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+           \<in> set evs;
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
+done
+
+text{*Most general form*}
+lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
+
+lemmas Confidentiality_Auth_A = authK_authentic [THEN exE, THEN Confidentiality_Kas]
+
+text{*Needs a confidentiality guarantee, hence moved here.
+      Authenticity of servK for A*}
+lemma servK_authentic_bis_r:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; A \<notin> bad; evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, 
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>
+       \<in> set evs"
+apply (frule authK_authentic, assumption, assumption)
+apply (erule exE)
+apply (drule Confidentiality_Auth_A, assumption, assumption)
+apply (blast, assumption, assumption, assumption)
+apply (blast dest:  servK_authentic_ter)
+done
+
+lemma Confidentiality_Serv_A:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (drule authK_authentic, assumption, assumption)
+apply (blast dest: Confidentiality_Kas Says_Kas_message_form servK_authentic_ter Confidentiality_Tgs_bis)
+done
+
+lemma Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (frule authK_authentic)
+apply (erule_tac [3] exE)
+apply (frule_tac [3] Confidentiality_Kas)
+apply (frule_tac [6] servTicket_authentic, auto)
+apply (blast dest!: Confidentiality_Tgs_bis dest: Says_Kas_message_form servK_authentic unique_servKeys unique_authKeys)
+done
+
+lemma u_Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad;  B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Key servK \<notin> analz (spies evs)"
+apply (blast dest: u_servTicket_authentic u_NotexpiredSK_NotexpiredAK Confidentiality_Tgs_bis)
+done
+
+
+
+subsection{*Parties authentication: each party verifies "the identity of
+       another party who generated some data" (quoted from Neuman and Ts'o).*}
+
+text{*These guarantees don't assess whether two parties agree on
+      the same session key: sending a message containing a key
+      doesn't a priori state knowledge of the key.*}
+
+
+text{*These didn't have existential form in version IV*}
+lemma B_authenticates_A:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+        Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+        Key servK \<notin> analz (spies evs);
+        A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
+apply (blast dest: servTicket_authentic_Tgs intro: Says_K5)
+done
+
+text{*The second assumption tells B what kind of key servK is.*}
+lemma B_authenticates_A_r:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
+apply (blast intro: Says_K5 dest: Confidentiality_B servTicket_authentic_Tgs)
+done
+
+text{* @{text u_B_authenticates_A} would be the same as @{text B_authenticates_A} because the
+ servK confidentiality assumption is yet unrelaxed*}
+
+lemma u_B_authenticates_A_r:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredSK Ts evs;
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
+apply (blast intro: Says_K5 dest: u_Confidentiality_B servTicket_authentic_Tgs)
+done
+
+lemma A_authenticates_B:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (frule authK_authentic)
+apply assumption+
+apply (frule servK_authentic)
+prefer 2 apply (blast dest: authK_authentic Says_Kas_message_form)
+apply assumption+
+apply clarify
+apply (blast dest: K4_imp_K2 Key_unique_SesKey intro!: Says_K6)
+(*Single command proof: much slower!
+apply (blast dest: authK_authentic servK_authentic Says_Kas_message_form Key_unique_SesKey K4_imp_K2 intro!: Says_K6)
+*)
+done
+
+lemma A_authenticates_B_r:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
+         A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
+apply (frule authK_authentic)
+apply (erule_tac [3] exE)
+apply (frule_tac [3] Says_Kas_message_form)
+apply (frule_tac [4] Confidentiality_Kas)
+apply (frule_tac [7] servK_authentic)
+prefer 8 apply blast
+apply (erule_tac [9] exE)
+apply (erule_tac [9] exE)
+apply (frule_tac [9] K4_imp_K2)
+apply assumption+
+apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs
+)
+done
+
+
+
+
+subsection{*Parties' knowledge of session keys. 
+       An agent knows a session key if he used it to issue a cipher. These
+       guarantees can be interpreted both in terms of key distribution
+       and of non-injective agreement on the session key.*}
+
+lemma Kas_Issues_A:
+   "\<lbrakk> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>, authTicket\<rbrace> \<in> set evs;
+      evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> Kas Issues A with (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>)
+          on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*K2*}
+apply (simp add: takeWhile_tail)
+apply (blast dest: authK_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD])
+done
+
+lemma A_authenticates_and_keydist_to_Kas:
+  "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace> \<in> parts (spies evs);
+     A \<notin> bad; evs \<in> kerbV \<rbrakk>
+ \<Longrightarrow> Kas Issues A with (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>) 
+          on evs"
+by (blast dest!: authK_authentic Kas_Issues_A)
+
+lemma Tgs_Issues_A:
+    "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
+         \<in> set evs; 
+       Key authK \<notin> analz (spies evs);  evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> Tgs Issues A with 
+          (Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [5] Says_ticket_parts)
+apply (frule_tac [7] Says_ticket_parts)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+apply (simp add: takeWhile_tail)
+(*Last two thms installed only to derive authK \<notin> range shrK*)
+apply (blast dest: servK_authentic parts_spies_takeWhile_mono [THEN subsetD]
+      parts_spies_evs_revD2 [THEN subsetD] authTicket_authentic 
+      Says_Kas_message_form)
+done
+
+lemma A_authenticates_and_keydist_to_Tgs:
+     "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+       Key authK \<notin> analz (spies evs); B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> \<exists>A. Tgs Issues A with 
+          (Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>) on evs"
+by (blast dest: Tgs_Issues_A servK_authentic_bis)
+
+lemma B_Issues_A:
+     "\<lbrakk> Says B A (Crypt servK (Number T3)) \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+apply blast
+txt{*K6 requires numerous lemmas*}
+apply (simp add: takeWhile_tail)
+apply (blast intro: Says_K6 dest: servTicket_authentic 
+        parts_spies_takeWhile_mono [THEN subsetD] 
+        parts_spies_evs_revD2 [THEN subsetD])
+done
+
+lemma A_authenticates_and_keydist_to_B:
+     "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
+         Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
+           \<in> parts (spies evs);
+         Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
+by (blast dest!: A_authenticates_B B_Issues_A)
+
+
+(*Must use \<le> rather than =, otherwise it cannot be proved inductively!*)
+(*This is too strong for version V but would hold for version IV if only B 
+  in K6 said a fresh timestamp.
+lemma honest_never_says_newer_timestamp:
+     "\<lbrakk> (CT evs) \<le> T ; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> \<forall> A B. A \<noteq> Spy \<longrightarrow> Says A B X \<notin> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all)
+apply force
+apply force
+txt{*clarifying case K3*}
+apply (rule impI)
+apply (rule impI)
+apply (frule Suc_leD)
+apply (clarify)
+txt{*cannot solve K3 or K5 because the spy might send CT evs as authTicket
+or servTicket, which the honest agent would forward*}
+prefer 2 apply force
+prefer 4 apply force
+prefer 4 apply force
+txt{*cannot solve K6 unless B updates the timestamp - rather than bouncing T3*}
+oops
+*)
+
+
+text{*But can prove a less general fact conerning only authenticators!*}
+lemma honest_never_says_newer_timestamp_in_auth:
+     "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; A \<notin> bad; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all)
+apply force+
+done
+
+lemma honest_never_says_current_timestamp_in_auth:
+     "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; A \<notin> bad; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
+apply (frule eq_imp_le)
+apply (blast dest: honest_never_says_newer_timestamp_in_auth)
+done
+
+
+
+lemma A_Issues_B:
+     "\<lbrakk> Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs;
+         Key servK \<notin> analz (spies evs);
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule kerbV.induct, analz_mono_contra)
+apply (frule_tac [7] Says_ticket_parts)
+apply (frule_tac [5] Says_ticket_parts)
+apply (simp_all (no_asm_simp))
+txt{*K5*}
+apply auto
+apply (simp add: takeWhile_tail)
+txt{*Level 15: case study necessary because the assumption doesn't state
+  the form of servTicket. The guarantee becomes stronger.*}
+prefer 2 apply (simp add: takeWhile_tail)
+(**This single command of version IV...
+apply (blast dest: Says_imp_spies [THEN analz.Inj, THEN analz_Decrypt']
+                   K3_imp_K2 K4_trustworthy'
+                   parts_spies_takeWhile_mono [THEN subsetD]
+                   parts_spies_evs_revD2 [THEN subsetD]
+             intro: Says_Auth)
+...expands as follows - including extra exE because of new form of lemmas*)
+apply (frule K3_imp_K2, assumption, assumption, erule exE, erule exE)
+apply (case_tac "Key authK \<in> analz (spies evs5)")
+apply (drule Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Fst, THEN analz_Decrypt', THEN analz.Fst], assumption, assumption, simp)
+apply (frule K3_imp_K2, assumption, assumption, erule exE, erule exE)
+apply (drule Says_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst])
+apply (frule servK_authentic_ter, blast, assumption+)
+apply (drule parts_spies_takeWhile_mono [THEN subsetD])
+apply (drule parts_spies_evs_revD2 [THEN subsetD])
+txt{* @{term Says_K5} closes the proof in version IV because it is clear which 
+servTicket an authenticator appears with in msg 5. In version V an authenticator can appear with any item that the spy could replace the servTicket with*}
+apply (frule Says_K5, blast, assumption, assumption, assumption, assumption, erule exE)
+txt{*We need to state that an honest agent wouldn't send the wrong timestamp
+within an authenticator, wathever it is paired with*}
+apply (simp add: honest_never_says_current_timestamp_in_auth)
+done
+
+lemma B_authenticates_and_keydist_to_A:
+     "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
+           \<in> parts (spies evs);
+         Key servK \<notin> analz (spies evs);
+         B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
+by (blast dest: B_authenticates_A A_Issues_B)
+
+
+
+subsection{*
+Novel guarantees, never studied before. Because honest agents always say
+the right timestamp in authenticators, we can prove unicity guarantees based 
+exactly on timestamps. Classical unicity guarantees are based on nonces.
+Of course assuming the agent to be different from the Spy, rather than not in 
+bad, would suffice below. Similar guarantees must also hold of
+Kerberos IV.*}
+
+text{*Notice that an honest agent can send the same timestamp on two
+different traces of the same length, but not on the same trace!*}
+
+lemma unique_timestamp_authenticator1:
+     "\<lbrakk> Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs;
+         Says A Kas' \<lbrace>Agent A, Agent Tgs', Number T1\<rbrace> \<in> set evs;
+         A \<notin>bad; evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> Kas=Kas' \<and> Tgs=Tgs'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all, blast)
+apply auto
+apply (simp_all add: honest_never_says_current_timestamp_in_auth)
+done
+
+lemma unique_timestamp_authenticator2:
+     "\<lbrakk> Says A Tgs \<lbrace>AT, Crypt AK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace> \<in> set evs;
+     Says A Tgs' \<lbrace>AT', Crypt AK' \<lbrace>Agent A, Number T2\<rbrace>, Agent B'\<rbrace> \<in> set evs;
+         A \<notin>bad; evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> Tgs=Tgs' \<and> AT=AT' \<and> AK=AK' \<and> B=B'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all, blast)
+apply auto
+apply (simp_all add: honest_never_says_current_timestamp_in_auth)
+done
+
+lemma unique_timestamp_authenticator3:
+     "\<lbrakk> Says A B \<lbrace>ST, Crypt SK \<lbrace>Agent A, Number T\<rbrace>\<rbrace> \<in> set evs;
+         Says A B' \<lbrace>ST', Crypt SK' \<lbrace>Agent A, Number T\<rbrace>\<rbrace> \<in> set evs;
+         A \<notin>bad; evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> B=B' \<and> ST=ST' \<and> SK=SK'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (simp_all, blast)
+apply (auto simp add: honest_never_says_current_timestamp_in_auth)
+done
+
+text{*The second part of the message is treated as an authenticator by the last
+simplification step, even if it is not an authenticator!*}
+lemma unique_timestamp_authticket:
+     "\<lbrakk> Says Kas A \<lbrace>X, Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key AK, T\<rbrace>\<rbrace> \<in> set evs;
+       Says Kas A' \<lbrace>X', Crypt (shrK Tgs') \<lbrace>Agent A', Agent Tgs', Key AK', T\<rbrace>\<rbrace> \<in> set evs;
+         evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> A=A' \<and> X=X' \<and> Tgs=Tgs' \<and> AK=AK'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (auto simp add: honest_never_says_current_timestamp_in_auth)
+done
+
+text{*The second part of the message is treated as an authenticator by the last
+simplification step, even if it is not an authenticator!*}
+lemma unique_timestamp_servticket:
+     "\<lbrakk> Says Tgs A \<lbrace>X, Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, T\<rbrace>\<rbrace> \<in> set evs;
+       Says Tgs A' \<lbrace>X', Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SK', T\<rbrace>\<rbrace> \<in> set evs;
+         evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> A=A' \<and> X=X' \<and> B=B' \<and> SK=SK'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (auto simp add: honest_never_says_current_timestamp_in_auth)
+done
+
+(*Uses assumption K6's assumption that B \<noteq> Kas, otherwise B should say
+fresh timestamp*)
+lemma Kas_never_says_newer_timestamp:
+     "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> \<forall> A. Says Kas A X \<notin> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct, auto)
+done
+
+lemma Kas_never_says_current_timestamp:
+     "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> \<forall> A. Says Kas A X \<notin> set evs"
+apply (frule eq_imp_le)
+apply (blast dest: Kas_never_says_newer_timestamp)
+done
+
+lemma unique_timestamp_msg2:
+     "\<lbrakk> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key AK, Agent Tgs, T\<rbrace>, AT\<rbrace> \<in> set evs;
+     Says Kas A' \<lbrace>Crypt (shrK A') \<lbrace>Key AK', Agent Tgs', T\<rbrace>, AT'\<rbrace> \<in> set evs;
+         evs \<in> kerbV \<rbrakk>
+  \<Longrightarrow> A=A' \<and> AK=AK' \<and> Tgs=Tgs' \<and> AT=AT'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (auto simp add: Kas_never_says_current_timestamp)
+done
+
+(*Uses assumption K6's assumption that B \<noteq> Tgs, otherwise B should say
+fresh timestamp*)
+lemma Tgs_never_says_newer_timestamp:
+     "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> \<forall> A. Says Tgs A X \<notin> set evs"
+apply (erule rev_mp)
+apply (erule kerbV.induct, auto)
+done
+
+lemma Tgs_never_says_current_timestamp:
+     "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
+     \<Longrightarrow> \<forall> A. Says Tgs A X \<notin> set evs"
+apply (frule eq_imp_le)
+apply (blast dest: Tgs_never_says_newer_timestamp)
+done
+
+
+lemma unique_timestamp_msg4:
+     "\<lbrakk> Says Tgs A \<lbrace>Crypt (shrK A) \<lbrace>Key SK, Agent B, T\<rbrace>, ST\<rbrace> \<in> set evs;
+       Says Tgs A' \<lbrace>Crypt (shrK A') \<lbrace>Key SK', Agent B', T\<rbrace>, ST'\<rbrace> \<in> set evs;
+         evs \<in> kerbV \<rbrakk> 
+  \<Longrightarrow> A=A' \<and> SK=SK' \<and> B=B' \<and> ST=ST'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule kerbV.induct)
+apply (auto simp add: Tgs_never_says_current_timestamp)
+done
+ 
+end
--- a/src/HOL/Auth/Kerberos_BAN.thy	Wed Feb 01 12:23:14 2006 +0100
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -2,8 +2,6 @@
     ID:         $Id$
     Author:     Giampaolo Bella, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
-
-Tidied and converted to Isar by lcp.
 *)
 
 header{*The Kerberos Protocol, BAN Version*}
@@ -14,91 +12,112 @@
   Burrows, Abadi and Needham (1989).  A Logic of Authentication.
   Proc. Royal Soc. 426
 
-  Confidentiality (secrecy) and authentication properties rely on
-  temporal checks: strong guarantees in a little abstracted - but
-  very realistic - model.
+  Confidentiality (secrecy) and authentication properties are also
+  given in a termporal version: strong guarantees in a little abstracted 
+  - but very realistic - model.
 *}
 
-(* Temporal modelization: session keys can be leaked
-                          ONLY when they have expired *)
+(* Temporal model of accidents: session keys can be leaked
+                                ONLY when they have expired *)
 
 syntax
     CT :: "event list=>nat"
-    Expired :: "[nat, event list] => bool"
-    RecentAuth :: "[nat, event list] => bool"
+    expiredK :: "[nat, event list] => bool"
+    expiredA :: "[nat, event list] => bool"
 
 consts
 
     (*Duration of the session key*)
-    SesKeyLife   :: nat
+    sesKlife   :: nat
 
     (*Duration of the authenticator*)
-    AutLife :: nat
+    authlife :: nat
 
 text{*The ticket should remain fresh for two journeys on the network at least*}
-specification (SesKeyLife)
-  SesKeyLife_LB [iff]: "2 \<le> SesKeyLife"
+specification (sesKlife)
+  sesKlife_LB [iff]: "2 \<le> sesKlife"
     by blast
 
 text{*The authenticator only for one journey*}
-specification (AutLife)
-  AutLife_LB [iff]:    "Suc 0 \<le> AutLife"
+specification (authlife)
+  authlife_LB [iff]:    "0 < authlife"
     by blast
 
 
 translations
    "CT" == "length "
 
-   "Expired T evs" == "SesKeyLife + T < CT evs"
+   "expiredK T evs" == "sesKlife + T < CT evs"
+
+   "expiredA T evs" == "authlife + T < CT evs"
+
+
+constdefs
 
-   "RecentAuth T evs" == "CT evs \<le> AutLife + T"
+ (* A is the true creator of X if she has sent X and X never appeared on
+    the trace before this event. Recall that traces grow from head. *)
+  Issues :: "[agent, agent, msg, event list] => bool"
+             ("_ Issues _ with _ on _")
+   "A Issues B with X on evs ==
+      \<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
+      X \<notin> parts (spies (takeWhile (% z. z  \<noteq> Says A B Y) (rev evs)))"
 
-consts  kerberos_ban   :: "event list set"
-inductive "kerberos_ban"
+ (* Yields the subtrace of a given trace from its beginning to a given event *)
+  before :: "[event, event list] => event list" ("before _ on _")
+   "before ev on evs ==  takeWhile (% z. z ~= ev) (rev evs)"
+
+ (* States than an event really appears only once on a trace *)
+  Unique :: "[event, event list] => bool" ("Unique _ on _")
+   "Unique ev on evs == 
+      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
+
+
+consts  bankerberos   :: "event list set"
+inductive "bankerberos"
  intros
 
-   Nil:  "[] \<in> kerberos_ban"
-
-   Fake: "[| evsf \<in> kerberos_ban;  X \<in> synth (analz (spies evsf)) |]
-	  ==> Says Spy B X # evsf \<in> kerberos_ban"
+   Nil:  "[] \<in> bankerberos"
 
-
-   Kb1:  "[| evs1 \<in> kerberos_ban |]
-	  ==> Says A Server {|Agent A, Agent B|} # evs1
-		\<in>  kerberos_ban"
+   Fake: "\<lbrakk> evsf \<in> bankerberos;  X \<in> synth (analz (spies evsf)) \<rbrakk>
+	  \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
 
 
-   Kb2:  "[| evs2 \<in> kerberos_ban;  Key KAB \<notin> used evs2;  KAB \<in> symKeys;
-	     Says A' Server {|Agent A, Agent B|} \<in> set evs2 |]
-	  ==> Says Server A
+   BK1:  "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
+	  \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+		\<in>  bankerberos"
+
+
+   BK2:  "\<lbrakk> evs2 \<in> bankerberos;  Key K \<notin> used evs2; K \<in> symKeys;
+	     Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+	  \<Longrightarrow> Says Server A
 		(Crypt (shrK A)
-		   {|Number (CT evs2), Agent B, Key KAB,
-		    (Crypt (shrK B) {|Number (CT evs2), Agent A, Key KAB|})|})
-		# evs2 \<in> kerberos_ban"
+		   \<lbrace>Number (CT evs2), Agent B, Key K,
+		    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+		# evs2 \<in> bankerberos"
 
 
-   Kb3:  "[| evs3 \<in> kerberos_ban;
-	     Says S A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
+   BK3:  "\<lbrakk> evs3 \<in> bankerberos;
+	     Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
 	       \<in> set evs3;
-	     Says A Server {|Agent A, Agent B|} \<in> set evs3;
-	     ~ Expired Ts evs3 |]
-	  ==> Says A B {|X, Crypt K {|Agent A, Number (CT evs3)|} |}
-	       # evs3 \<in> kerberos_ban"
+	     Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+	     \<not> expiredK Tk evs3 \<rbrakk>
+	  \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+	       # evs3 \<in> bankerberos"
 
 
-   Kb4:  "[| evs4 \<in> kerberos_ban;
-	     Says A' B {|(Crypt (shrK B) {|Number Ts, Agent A, Key K|}),
-			 (Crypt K {|Agent A, Number Ta|}) |}: set evs4;
-	     ~ Expired Ts evs4;  RecentAuth Ta evs4 |]
-	  ==> Says B A (Crypt K (Number Ta)) # evs4
-		\<in> kerberos_ban"
+   BK4:  "\<lbrakk> evs4 \<in> bankerberos;
+	     Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+			 (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+	     \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
+	  \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+		\<in> bankerberos"
 
 	(*Old session keys may become compromised*)
-   Oops: "[| evso \<in> kerberos_ban;
-	     Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
+   Oops: "\<lbrakk> evso \<in> bankerberos;
+         Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
 	       \<in> set evso;
-	     Expired Ts evso |]
-	  ==> Notes Spy {|Number Ts, Key K|} # evso \<in> kerberos_ban"
+	     expiredK Tk evso \<rbrakk>
+	  \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -107,65 +126,151 @@
 declare Fake_parts_insert_in_Un [dest]
 
 text{*A "possibility property": there are traces that reach the end.*}
-lemma "[|Key K \<notin> used []; K \<in> symKeys|]
-       ==> \<exists>Timestamp. \<exists>evs \<in> kerberos_ban.
+lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk>
+       \<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerberos.
              Says B A (Crypt K (Number Timestamp))
                   \<in> set evs"
-apply (cut_tac SesKeyLife_LB)
+apply (cut_tac sesKlife_LB)
 apply (intro exI bexI)
 apply (rule_tac [2]
-           kerberos_ban.Nil [THEN kerberos_ban.Kb1, THEN kerberos_ban.Kb2,
-                             THEN kerberos_ban.Kb3, THEN kerberos_ban.Kb4])
+           bankerberos.Nil [THEN bankerberos.BK1, THEN bankerberos.BK2,
+                             THEN bankerberos.BK3, THEN bankerberos.BK4])
 apply (possibility, simp_all (no_asm_simp) add: used_Cons)
 done
 
+subsection{*Lemmas for reasoning about predicate "Issues"*}
 
-(**** Inductive proofs about kerberos_ban ****)
+lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
+          (if A:bad then insert X (spies evs) else spies evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_evs_rev: "spies evs = spies (rev evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a")
+apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
+done
+
+lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono]
+
+lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+done
+
+lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
+
+
+text{*Lemmas for reasoning about predicate "before"*}
+lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
 
-text{*Forwarding Lemma for reasoning about the encrypted portion of message Kb3*}
-lemma Kb3_msg_in_parts_spies:
-     "Says S A (Crypt KA {|Timestamp, B, K, X|}) \<in> set evs
-      ==> X \<in> parts (spies evs)"
-by blast
+lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_evs_rev: "used evs = used (rev evs)"
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply (simp add: used_Says_rev)
+apply (simp add: used_Gets_rev)
+apply (simp add: used_Notes_rev)
+done
+
+lemma used_takeWhile_used [rule_format]: 
+      "x : used (takeWhile P X) --> x : used X"
+apply (induct_tac "X")
+apply simp
+apply (induct_tac "a")
+apply (simp_all add: used_Nil)
+apply (blast dest!: initState_into_used)+
+done
+
+lemma set_evs_rev: "set evs = set (rev evs)"
+apply auto
+done
+
+lemma takeWhile_void [rule_format]:
+      "x \<notin> set evs \<longrightarrow> takeWhile (\<lambda>z. z \<noteq> x) evs = evs"
+apply auto
+done
+
+(**** Inductive proofs about bankerberos ****)
+
+text{*Forwarding Lemma for reasoning about the encrypted portion of message BK3*}
+lemma BK3_msg_in_parts_spies:
+     "Says S A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs
+      \<Longrightarrow> X \<in> parts (spies evs)"
+apply blast
+done
 
 lemma Oops_parts_spies:
-     "Says Server A (Crypt (shrK A) {|Timestamp, B, K, X|}) \<in> set evs
-      ==> K \<in> parts (spies evs)"
-by blast
-
+     "Says Server A (Crypt (shrK A) \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs
+      \<Longrightarrow> K \<in> parts (spies evs)"
+apply blast
+done
 
 text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
 lemma Spy_see_shrK [simp]:
-     "evs \<in> kerberos_ban ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
-apply (erule kerberos_ban.induct)
+     "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
+apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all, blast+)
+apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast+)
 done
 
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> kerberos_ban ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
-by auto
+     "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
+apply auto
+done
 
 lemma Spy_see_shrK_D [dest!]:
-     "[| Key (shrK A) \<in> parts (spies evs);
-                evs \<in> kerberos_ban |] ==> A:bad"
-by (blast dest: Spy_see_shrK)
+     "\<lbrakk> Key (shrK A) \<in> parts (spies evs);
+                evs \<in> bankerberos \<rbrakk> \<Longrightarrow> A:bad"
+apply (blast dest: Spy_see_shrK)
+done
 
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D,  dest!]
 
 
 text{*Nobody can have used non-existent keys!*}
 lemma new_keys_not_used [simp]:
-    "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerberos_ban|]
-     ==> K \<notin> keysFor (parts (spies evs))"
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerberos\<rbrakk>
+     \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
 apply (erule rev_mp)
-apply (erule kerberos_ban.induct)
+apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)
+apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all)
 txt{*Fake*}
 apply (force dest!: keysFor_parts_insert)
-txt{*Kb2, Kb3, Kb4*}
+txt{*BK2, BK3, BK4*}
 apply (force dest!: analz_shrK_Decrypt)+
 done
 
@@ -173,47 +278,60 @@
 
 text{*Describes the form of K, X and K' when the Server sends this message.*}
 lemma Says_Server_message_form:
-     "[| Says Server A (Crypt K' {|Number Ts, Agent B, Key K, X|})
-         \<in> set evs; evs \<in> kerberos_ban |]
-      ==> K \<notin> range shrK &
-          X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|}) &
-          K' = shrK A"
+     "\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+         \<in> set evs; evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> K' = shrK A & K \<notin> range shrK &
+          Ticket = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>) &
+          Key K \<notin> used(before
+                  Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+                  on evs) &
+          Tk = CT(before 
+                  Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+                  on evs)"
+apply (unfold before_def)
 apply (erule rev_mp)
-apply (erule kerberos_ban.induct, auto)
+apply (erule bankerberos.induct, simp_all)
+txt{*We need this simplification only for Message 2*}
+apply (simp (no_asm) add: takeWhile_tail)
+apply auto
+txt{*Two subcases of Message 2. Subcase: used before*}
+apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
+                   used_takeWhile_used)
+txt{*subcase: CT before*}
+apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
 done
 
 
 text{*If the encrypted message appears then it originated with the Server
   PROVIDED that A is NOT compromised!
-
-  This shows implicitly the FRESHNESS OF THE SESSION KEY to A
+  This allows A to verify freshness of the session key.
 *}
-lemma A_trusts_K_by_Kb2:
-     "[| Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}
+lemma Kab_authentic:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
            \<in> parts (spies evs);
-         A \<notin> bad;  evs \<in> kerberos_ban |]
-       ==> Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
+         A \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+       \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
              \<in> set evs"
 apply (erule rev_mp)
-apply (erule kerberos_ban.induct)
+apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all, blast)
+apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast)
 done
 
 
 text{*If the TICKET appears then it originated with the Server*}
 text{*FRESHNESS OF THE SESSION KEY to B*}
-lemma B_trusts_K_by_Kb3:
-     "[| Crypt (shrK B) {|Number Ts, Agent A, Key K|} \<in> parts (spies evs);
-         B \<notin> bad;  evs \<in> kerberos_ban |]
-       ==> Says Server A
-            (Crypt (shrK A) {|Number Ts, Agent B, Key K,
-                          Crypt (shrK B) {|Number Ts, Agent A, Key K|}|})
+lemma ticket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs);
+         B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+       \<Longrightarrow> Says Server A
+            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
+                          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
            \<in> set evs"
 apply (erule rev_mp)
-apply (erule kerberos_ban.induct)
+apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all, blast)
+apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast)
 done
 
 
@@ -221,15 +339,15 @@
   OR     reduces it to the Fake case.
   Use @{text Says_Server_message_form} if applicable.*}
 lemma Says_S_message_form:
-     "[| Says S A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
+     "\<lbrakk> Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
             \<in> set evs;
-         evs \<in> kerberos_ban |]
- ==> (K \<notin> range shrK & X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|}))
+         evs \<in> bankerberos \<rbrakk>
+ \<Longrightarrow> (K \<notin> range shrK & X = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>))
           | X \<in> analz (spies evs)"
 apply (case_tac "A \<in> bad")
 apply (force dest!: Says_imp_spies [THEN analz.Inj])
 apply (frule Says_imp_spies [THEN parts.Inj])
-apply (blast dest!: A_trusts_K_by_Kb2 Says_Server_message_form)
+apply (blast dest!: Kab_authentic Says_Server_message_form)
 done
 
 
@@ -237,7 +355,7 @@
 (****
  The following is to prove theorems of the form
 
-  Key K \<in> analz (insert (Key KAB) (spies evs)) ==>
+  Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow>
   Key K \<in> analz (spies evs)
 
  A more general formula must be proved inductively.
@@ -246,169 +364,366 @@
 
 text{* Session keys are not used to encrypt other session keys *}
 lemma analz_image_freshK [rule_format (no_asm)]:
-     "evs \<in> kerberos_ban ==>
-   \<forall>K KK. KK \<subseteq> - (range shrK) -->
+     "evs \<in> bankerberos \<Longrightarrow>
+   \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
           (Key K \<in> analz (Key`KK Un (spies evs))) =
           (K \<in> KK | Key K \<in> analz (spies evs))"
-apply (erule kerberos_ban.induct)
+apply (erule bankerberos.induct)
 apply (drule_tac [7] Says_Server_message_form)
 apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz, auto) 
 done
 
 
 lemma analz_insert_freshK:
-     "[| evs \<in> kerberos_ban;  KAB \<notin> range shrK |] ==>
+     "\<lbrakk> evs \<in> bankerberos;  KAB \<notin> range shrK \<rbrakk> \<Longrightarrow>
       (Key K \<in> analz (insert (Key KAB) (spies evs))) =
       (K = KAB | Key K \<in> analz (spies evs))"
-by (simp only: analz_image_freshK analz_image_freshK_simps)
-
+apply (simp only: analz_image_freshK analz_image_freshK_simps)
+done
 
 text{* The session key K uniquely identifies the message *}
 lemma unique_session_keys:
-     "[| Says Server A
-           (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) \<in> set evs;
+     "\<lbrakk> Says Server A
+           (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;
          Says Server A'
-          (Crypt (shrK A') {|Number Ts', Agent B', Key K, X'|}) \<in> set evs;
-         evs \<in> kerberos_ban |] ==> A=A' & Ts=Ts' & B=B' & X = X'"
+          (Crypt (shrK A') \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs;
+         evs \<in> bankerberos \<rbrakk> \<Longrightarrow> A=A' & Tk=Tk' & B=B' & X = X'"
 apply (erule rev_mp)
 apply (erule rev_mp)
-apply (erule kerberos_ban.induct)
+apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)
-txt{*Kb2: it can't be a new key*}
+apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all)
+txt{*BK2: it can't be a new key*}
+apply blast
+done
+
+lemma Server_Unique:
+     "\<lbrakk> Says Server A
+            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
+        evs \<in> bankerberos \<rbrakk> \<Longrightarrow> 
+   Unique Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+   on evs"
+apply (erule rev_mp, erule bankerberos.induct, simp_all add: Unique_def)
 apply blast
 done
 
 
-text{* Lemma: the session key sent in msg Kb2 would be EXPIRED
+subsection{*Non-temporal guarantees, explicitly relying on non-occurrence of
+oops events - refined below by temporal guarantees*}
+
+text{*Non temporal treatment of confidentiality*}
+
+text{* Lemma: the session key sent in msg BK2 would be lost by oops
     if the spy could see it! *}
+lemma lemma_conf [rule_format (no_asm)]:
+     "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+  \<Longrightarrow> Says Server A
+          (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
+                            Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
+         \<in> set evs \<longrightarrow>
+      Key K \<in> analz (spies evs) \<longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<in> set evs"
+apply (erule bankerberos.induct)
+apply (frule_tac [7] Says_Server_message_form)
+apply (frule_tac [5] Says_S_message_form [THEN disjE])
+apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*BK2*}
+apply (blast intro: parts_insertI)
+txt{*BK3*}
+apply (case_tac "Aa \<in> bad")
+ prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
+apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)
+txt{*Oops*}
+apply (blast dest: unique_session_keys)
+done
+
+
+text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.*}
+lemma Confidentiality_S:
+     "\<lbrakk> Says Server A
+          (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
+apply (frule Says_Server_message_form, assumption)
+apply (blast intro: lemma_conf)
+done
+
+text{*Confidentiality for Alice*}
+lemma Confidentiality_A:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
+apply (blast dest!: Kab_authentic Confidentiality_S)
+done
+
+text{*Confidentiality for Bob*}
+lemma Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
+          \<in> parts (spies evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
+apply (blast dest!: ticket_authentic Confidentiality_S)
+done
+
+text{*Non temporal treatment of authentication*}
 
-lemma lemma2 [rule_format (no_asm)]:
-     "[| A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban |]
-  ==> Says Server A
-          (Crypt (shrK A) {|Number Ts, Agent B, Key K,
-                            Crypt (shrK B) {|Number Ts, Agent A, Key K|}|})
-         \<in> set evs -->
-      Key K \<in> analz (spies evs) --> Expired Ts evs"
-apply (erule kerberos_ban.induct)
+text{*Lemmas @{text lemma_A} and @{text lemma_B} in fact are common to both temporal and non-temporal treatments*}
+lemma lemma_A [rule_format]:
+     "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow>
+         Key K \<notin> analz (spies evs) \<longrightarrow>
+         Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
+         \<in> set evs \<longrightarrow>
+          Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs) \<longrightarrow>
+         Says A B \<lbrace>X, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace>
+             \<in> set evs"
+apply (erule bankerberos.induct)
+apply (frule_tac [7] Oops_parts_spies)
+apply (frule_tac [5] Says_S_message_form)
+apply (frule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*Fake*}
+apply blast
+txt{*BK2*}
+apply (force dest: Crypt_imp_invKey_keysFor)
+txt{*BK3*}
+apply (blast dest: Kab_authentic unique_session_keys)
+done
+lemma lemma_B [rule_format]:
+     "\<lbrakk> B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (spies evs) \<longrightarrow>
+          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
+          \<in> set evs \<longrightarrow>
+          Crypt K (Number Ta) \<in> parts (spies evs) \<longrightarrow>
+          Says B A (Crypt K (Number Ta)) \<in> set evs"
+apply (erule bankerberos.induct)
+apply (frule_tac [7] Oops_parts_spies)
+apply (frule_tac [5] Says_S_message_form)
+apply (drule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*Fake*}
+apply blast
+txt{*BK2*} 
+apply (force dest: Crypt_imp_invKey_keysFor)
+txt{*BK4*}
+apply (blast dest: ticket_authentic unique_session_keys
+                   Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad)
+done
+
+
+text{*The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+
+
+text{*Authentication of A to B*}
+lemma B_authenticates_A_r:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+apply (blast dest!: ticket_authentic
+          intro!: lemma_A
+          elim!: Confidentiality_S [THEN [2] rev_notE])
+done
+
+
+text{*Authentication of B to A*}
+lemma A_authenticates_B_r:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
+        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
+apply (blast dest!: Kab_authentic
+          intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])
+done
+
+lemma B_authenticates_A:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
+        Key K \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+apply (blast dest!: ticket_authentic intro!: lemma_A)
+done
+
+lemma A_authenticates_B:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
+        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+        Key K \<notin> analz (spies evs);
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
+apply (blast dest!: Kab_authentic intro!: lemma_B)
+done
+
+subsection{*Temporal guarantees, relying on a temporal check that insures that
+no oops event occurred. These are available in the sense of goal availability*}
+
+
+text{*Temporal treatment of confidentiality*}
+
+text{* Lemma: the session key sent in msg BK2 would be EXPIRED
+    if the spy could see it! *}
+lemma lemma_conf_temporal [rule_format (no_asm)]:
+     "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+  \<Longrightarrow> Says Server A
+          (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
+                            Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
+         \<in> set evs \<longrightarrow>
+      Key K \<in> analz (spies evs) \<longrightarrow> expiredK Tk evs"
+apply (erule bankerberos.induct)
 apply (frule_tac [7] Says_Server_message_form)
 apply (frule_tac [5] Says_S_message_form [THEN disjE])
 apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
 txt{*Fake*}
 apply spy_analz
-txt{*Kb2*}
+txt{*BK2*}
 apply (blast intro: parts_insertI less_SucI)
-txt{*Kb3*}
+txt{*BK3*}
 apply (case_tac "Aa \<in> bad")
- prefer 2 apply (blast dest: A_trusts_K_by_Kb2 unique_session_keys)
+ prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
 apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)
-txt{*Oops: PROOF FAILED if addIs below*}
+txt{*Oops: PROOF FAILS if unsafe intro below*}
 apply (blast dest: unique_session_keys intro!: less_SucI)
 done
 
 
-text{*Confidentiality for the Server: Spy does not see the keys sent in msg Kb2
+text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
 as long as they have not expired.*}
-lemma Confidentiality_S:
-     "[| Says Server A
-          (Crypt K' {|Number T, Agent B, Key K, X|}) \<in> set evs;
-         ~ Expired T evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban
-      |] ==> Key K \<notin> analz (spies evs)"
+lemma Confidentiality_S_temporal:
+     "\<lbrakk> Says Server A
+          (Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs;
+         \<not> expiredK T evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
 apply (frule Says_Server_message_form, assumption)
-apply (blast intro: lemma2)
+apply (blast intro: lemma_conf_temporal)
 done
 
-(**** THE COUNTERPART OF CONFIDENTIALITY
-      [|...; Expired Ts evs; ...|] ==> Key K \<in> analz (spies evs)
-      WOULD HOLD ONLY IF AN OOPS OCCURRED! ---> Nothing to prove!   ****)
-
-
 text{*Confidentiality for Alice*}
-lemma Confidentiality_A:
-     "[| Crypt (shrK A) {|Number T, Agent B, Key K, X|} \<in> parts (spies evs);
-         ~ Expired T evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban
-      |] ==> Key K \<notin> analz (spies evs)"
-by (blast dest!: A_trusts_K_by_Kb2 Confidentiality_S)
-
+lemma Confidentiality_A_temporal:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+         \<not> expiredK T evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
+apply (blast dest!: Kab_authentic Confidentiality_S_temporal)
+done
 
 text{*Confidentiality for Bob*}
-lemma Confidentiality_B:
-     "[| Crypt (shrK B) {|Number Tk, Agent A, Key K|}
+lemma Confidentiality_B_temporal:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
           \<in> parts (spies evs);
-        ~ Expired Tk evs;
-        A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban
-      |] ==> Key K \<notin> analz (spies evs)"
-by (blast dest!: B_trusts_K_by_Kb3 Confidentiality_S)
+        \<not> expiredK Tk evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
+apply (blast dest!: ticket_authentic Confidentiality_S_temporal)
+done
+
+text{*Temporal treatment of authentication*}
+
+text{*Authentication of A to B*}
+lemma B_authenticates_A_temporal:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
+         \<in> parts (spies evs);
+         \<not> expiredK Tk evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+apply (blast dest!: ticket_authentic
+          intro!: lemma_A
+          elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
+done
+
+text{*Authentication of B to A*}
+lemma A_authenticates_B_temporal:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
+         Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
+         \<in> parts (spies evs);
+         \<not> expiredK Tk evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
+apply (blast dest!: Kab_authentic
+          intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
+done
+
+subsection{*Treatment of the key distribution goal using trace inspection. All
+guarantees are in non-temporal form, hence non available, though their temporal
+form is trivial to derive. These guarantees also convey a stronger form of 
+authentication - non-injective agreement on the session key*}
 
 
-lemma lemma_B [rule_format]:
-     "[| B \<notin> bad;  evs \<in> kerberos_ban |]
-      ==> Key K \<notin> analz (spies evs) -->
-          Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
-          \<in> set evs -->
-          Crypt K (Number Ta) \<in> parts (spies evs) -->
-          Says B A (Crypt K (Number Ta)) \<in> set evs"
-apply (erule kerberos_ban.induct)
-apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Says_S_message_form)
-apply (drule_tac [6] Kb3_msg_in_parts_spies, analz_mono_contra)
-apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+lemma B_Issues_A:
+     "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
+         Key K \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt K (Number Ta)) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule bankerberos.induct, analz_mono_contra)
+apply (simp_all (no_asm_simp))
+txt{*fake*}
 apply blast
-txt{*Kb2*} 
-apply (force dest: Crypt_imp_invKey_keysFor)
-txt{*Kb4*}
-apply (blast dest: B_trusts_K_by_Kb3 unique_session_keys
-                   Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad)
+txt{*K4 obviously is the non-trivial case*}
+apply (simp add: takeWhile_tail)
+apply (blast dest: ticket_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: A_authenticates_B_temporal)
+done
+
+lemma A_authenticates_and_keydist_to_B:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
+        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+         Key K \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; evs \<in> bankerberos \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt K (Number Ta)) on evs"
+apply (blast dest!: A_authenticates_B B_Issues_A)
 done
 
 
-text{*Authentication of B to A*}
-lemma Authentication_B:
-     "[| Crypt K (Number Ta) \<in> parts (spies evs);
-         Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}
-         \<in> parts (spies evs);
-         ~ Expired Ts evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban |]
-      ==> Says B A (Crypt K (Number Ta)) \<in> set evs"
-by (blast dest!: A_trusts_K_by_Kb2
-          intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])
-
-lemma lemma_A [rule_format]:
-     "[| A \<notin> bad; B \<notin> bad; evs \<in> kerberos_ban |]
-      ==>
-         Key K \<notin> analz (spies evs) -->
-         Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
-         \<in> set evs -->
-          Crypt K {|Agent A, Number Ta|} \<in> parts (spies evs) -->
-         Says A B {|X, Crypt K {|Agent A, Number Ta|}|}
-             \<in> set evs"
-apply (erule kerberos_ban.induct)
-apply (frule_tac [7] Oops_parts_spies)
-apply (frule_tac [5] Says_S_message_form)
-apply (frule_tac [6] Kb3_msg_in_parts_spies, analz_mono_contra)
-apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+lemma A_Issues_B:
+     "\<lbrakk> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace>
+           \<in> set evs;
+         Key K \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule bankerberos.induct, analz_mono_contra)
+apply (simp_all (no_asm_simp))
+txt{*fake*}
 apply blast
-txt{*Kb2*}
-apply (force dest: Crypt_imp_invKey_keysFor)
-txt{*Kb3*}
-apply (blast dest: A_trusts_K_by_Kb2 unique_session_keys)
+txt{*K3 is the non trivial case*}
+apply (simp add: takeWhile_tail)
+apply auto (*Technically unnecessary, merely clarifies the subgoal as it is presemted in the book*)
+apply (blast dest: Kab_authentic Says_Server_message_form parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] 
+             intro!: B_authenticates_A)
 done
 
-text{*Authentication of A to B*}
-lemma Authentication_A:
-     "[| Crypt K {|Agent A, Number Ta|} \<in> parts (spies evs);
-         Crypt (shrK B) {|Number Ts, Agent A, Key K|}
-         \<in> parts (spies evs);
-         ~ Expired Ts evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban |]
-      ==> Says A B {|Crypt (shrK B) {|Number Ts, Agent A, Key K|},
-                     Crypt K {|Agent A, Number Ta|}|} \<in> set evs"
-by (blast dest!: B_trusts_K_by_Kb3
-          intro!: lemma_A
-          elim!: Confidentiality_S [THEN [2] rev_notE])
+
+lemma B_authenticates_and_keydist_to_A:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
+        Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
+        Key K \<notin> analz (spies evs);
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) on evs"
+apply (blast dest: B_authenticates_A A_Issues_B)
+done
+
+
+
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,720 @@
+(*  ID:         $Id$
+    Author:     Giampaolo Bella, Catania University
+*)
+
+header{*The Kerberos Protocol, BAN Version, with Gets event*}
+
+theory Kerberos_BAN_Gets imports Public begin
+
+text{*From page 251 of
+  Burrows, Abadi and Needham (1989).  A Logic of Authentication.
+  Proc. Royal Soc. 426
+
+  Confidentiality (secrecy) and authentication properties rely on
+  temporal checks: strong guarantees in a little abstracted - but
+  very realistic - model.
+*}
+
+(* Temporal modelization: session keys can be leaked
+                          ONLY when they have expired *)
+
+syntax
+    CT :: "event list=>nat"
+    expiredK :: "[nat, event list] => bool"
+    expiredA :: "[nat, event list] => bool"
+
+consts
+
+    (*Duration of the session key*)
+    sesKlife   :: nat
+
+    (*Duration of the authenticator*)
+    authlife :: nat
+
+text{*The ticket should remain fresh for two journeys on the network at least*}
+text{*The Gets event causes longer traces for the protocol to reach its end*}
+specification (sesKlife)
+  sesKlife_LB [iff]: "4 \<le> sesKlife"
+    by blast
+
+text{*The authenticator only for one journey*}
+text{*The Gets event causes longer traces for the protocol to reach its end*}
+specification (authlife)
+  authlife_LB [iff]:    "2 \<le> authlife"
+    by blast
+
+
+translations
+   "CT" == "length "
+
+   "expiredK T evs" == "sesKlife + T < CT evs"
+
+   "expiredA T evs" == "authlife + T < CT evs"
+
+
+constdefs
+ (* Yields the subtrace of a given trace from its beginning to a given event *)
+  before :: "[event, event list] => event list" ("before _ on _")
+   "before ev on evs ==  takeWhile (% z. z ~= ev) (rev evs)"
+
+ (* States than an event really appears only once on a trace *)
+  Unique :: "[event, event list] => bool" ("Unique _ on _")
+   "Unique ev on evs == 
+      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
+
+
+consts  bankerb_gets   :: "event list set"
+inductive "bankerb_gets"
+ intros
+
+   Nil:  "[] \<in> bankerb_gets"
+
+   Fake: "\<lbrakk> evsf \<in> bankerb_gets;  X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
+	  \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
+
+   Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
+                \<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"
+
+   BK1:  "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
+	  \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+		\<in>  bankerb_gets"
+
+
+   BK2:  "\<lbrakk> evs2 \<in> bankerb_gets;  Key K \<notin> used evs2; K \<in> symKeys;
+	     Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+	  \<Longrightarrow> Says Server A
+		(Crypt (shrK A)
+		   \<lbrace>Number (CT evs2), Agent B, Key K,
+		    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+		# evs2 \<in> bankerb_gets"
+
+
+   BK3:  "\<lbrakk> evs3 \<in> bankerb_gets;
+	     Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+	       \<in> set evs3;
+	     Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+	     \<not> expiredK Tk evs3 \<rbrakk>
+	  \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+	       # evs3 \<in> bankerb_gets"
+
+
+   BK4:  "\<lbrakk> evs4 \<in> bankerb_gets;
+	     Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+			 (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+	     \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
+	  \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+		\<in> bankerb_gets"
+
+	(*Old session keys may become compromised*)
+   Oops: "\<lbrakk> evso \<in> bankerb_gets;
+         Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+	       \<in> set evso;
+	     expiredK Tk evso \<rbrakk>
+	  \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
+
+
+declare Says_imp_knows_Spy [THEN parts.Inj, dest]
+declare parts.Body [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un [dest]
+declare knows_Spy_partsEs [elim]
+
+
+text{*A "possibility property": there are traces that reach the end.*}
+lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk>
+       \<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerb_gets.
+             Says B A (Crypt K (Number Timestamp))
+                  \<in> set evs"
+apply (cut_tac sesKlife_LB)
+apply (cut_tac authlife_LB)
+apply (intro exI bexI)
+apply (rule_tac [2]
+           bankerb_gets.Nil [THEN bankerb_gets.BK1, THEN bankerb_gets.Reception,
+                            THEN bankerb_gets.BK2, THEN bankerb_gets.Reception,
+                            THEN bankerb_gets.BK3, THEN bankerb_gets.Reception,
+                            THEN bankerb_gets.BK4])
+apply (possibility, simp_all (no_asm_simp) add: used_Cons)
+done
+
+
+text{*Lemmas about reception invariant: if a message is received it certainly
+was sent*}
+lemma Gets_imp_Says :
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply auto
+done
+
+lemma Gets_imp_knows_Spy: 
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+done
+
+lemma Gets_imp_knows_Spy_parts[dest]:
+    "\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk>  \<Longrightarrow> X \<in> parts (knows Spy evs)"
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj])
+done
+
+lemma Gets_imp_knows:
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk>  \<Longrightarrow> X \<in> knows B evs"
+apply (case_tac "B = Spy")
+apply (blast dest!: Gets_imp_knows_Spy)
+apply (blast dest!: Gets_imp_knows_agents)
+done
+
+lemma Gets_imp_knows_analz:
+    "\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk>  \<Longrightarrow> X \<in> analz (knows B evs)"
+apply (blast dest: Gets_imp_knows [THEN analz.Inj])
+done
+
+text{*Lemmas for reasoning about predicate "before"*}
+lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs";
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply auto
+done
+
+lemma used_evs_rev: "used evs = used (rev evs)"
+apply (induct_tac "evs")
+apply simp
+apply (induct_tac "a")
+apply (simp add: used_Says_rev)
+apply (simp add: used_Gets_rev)
+apply (simp add: used_Notes_rev)
+done
+
+lemma used_takeWhile_used [rule_format]: 
+      "x : used (takeWhile P X) --> x : used X"
+apply (induct_tac "X")
+apply simp
+apply (induct_tac "a")
+apply (simp_all add: used_Nil)
+apply (blast dest!: initState_into_used)+
+done
+
+lemma set_evs_rev: "set evs = set (rev evs)"
+apply auto
+done
+
+lemma takeWhile_void [rule_format]:
+      "x \<notin> set evs \<longrightarrow> takeWhile (\<lambda>z. z \<noteq> x) evs = evs"
+apply auto
+done
+
+(**** Inductive proofs about bankerb_gets ****)
+
+text{*Forwarding Lemma for reasoning about the encrypted portion of message BK3*}
+lemma BK3_msg_in_parts_knows_Spy:
+     "\<lbrakk>Gets A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs; evs \<in> bankerb_gets \<rbrakk> 
+      \<Longrightarrow> X \<in> parts (knows Spy evs)"
+apply blast
+done
+
+lemma Oops_parts_knows_Spy:
+     "Says Server A (Crypt (shrK A) \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs
+      \<Longrightarrow> K \<in> parts (knows Spy evs)"
+apply blast
+done
+
+
+text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
+lemma Spy_see_shrK [simp]:
+     "evs \<in> bankerb_gets \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast+)
+done
+
+
+lemma Spy_analz_shrK [simp]:
+     "evs \<in> bankerb_gets \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
+by auto
+
+lemma Spy_see_shrK_D [dest!]:
+     "\<lbrakk> Key (shrK A) \<in> parts (knows Spy evs);
+                evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> A:bad"
+by (blast dest: Spy_see_shrK)
+
+lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D,  dest!]
+
+
+text{*Nobody can have used non-existent keys!*}
+lemma new_keys_not_used [simp]:
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerb_gets\<rbrakk>
+     \<Longrightarrow> K \<notin> keysFor (parts (knows Spy evs))"
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)
+txt{*Fake*}
+apply (force dest!: keysFor_parts_insert)
+txt{*BK2, BK3, BK4*}
+apply (force dest!: analz_shrK_Decrypt)+
+done
+
+subsection{* Lemmas concerning the form of items passed in messages *}
+
+text{*Describes the form of K, X and K' when the Server sends this message.*}
+lemma Says_Server_message_form:
+     "\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+         \<in> set evs; evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> K' = shrK A & K \<notin> range shrK &
+          Ticket = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>) &
+          Key K \<notin> used(before
+                  Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+                  on evs) &
+          Tk = CT(before 
+                  Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+                  on evs)"
+apply (unfold before_def)
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct, simp_all)
+txt{*We need this simplification only for Message 2*}
+apply (simp (no_asm) add: takeWhile_tail)
+apply auto
+txt{*Two subcases of Message 2. Subcase: used before*}
+apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
+                   used_takeWhile_used)
+txt{*subcase: CT before*}
+apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
+done
+
+
+text{*If the encrypted message appears then it originated with the Server
+  PROVIDED that A is NOT compromised!
+  This allows A to verify freshness of the session key.
+*}
+lemma Kab_authentic:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
+           \<in> parts (knows Spy evs);
+         A \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+       \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
+             \<in> set evs"
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast)
+done
+
+
+text{*If the TICKET appears then it originated with the Server*}
+text{*FRESHNESS OF THE SESSION KEY to B*}
+lemma ticket_authentic:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (knows Spy evs);
+         B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+       \<Longrightarrow> Says Server A
+            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
+                          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
+           \<in> set evs"
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast)
+done
+
+
+text{*EITHER describes the form of X when the following message is sent,
+  OR     reduces it to the Fake case.
+  Use @{text Says_Server_message_form} if applicable.*}
+lemma Gets_Server_message_form:
+     "\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
+            \<in> set evs;
+         evs \<in> bankerb_gets \<rbrakk>
+ \<Longrightarrow> (K \<notin> range shrK & X = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>))
+          | X \<in> analz (knows Spy evs)"
+apply (case_tac "A \<in> bad")
+apply (force dest!: Gets_imp_knows_Spy [THEN analz.Inj])
+apply (blast dest!: Kab_authentic Says_Server_message_form)
+done
+
+
+text{*Reliability guarantees: honest agents act as we expect*}
+
+lemma BK3_imp_Gets:
+   "\<lbrakk> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;
+      A \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
+    \<Longrightarrow> \<exists> Tk. Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+      \<in> set evs"
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply auto
+done
+
+lemma BK4_imp_Gets:
+   "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
+      B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
+  \<Longrightarrow> \<exists> Tk. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+	            Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply auto
+done
+
+lemma Gets_A_knows_K:
+  "\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
+     evs \<in> bankerb_gets \<rbrakk>
+ \<Longrightarrow> Key K \<in> analz (knows A evs)"
+apply (force dest: Gets_imp_knows_analz)
+done
+
+lemma Gets_B_knows_K:
+  "\<lbrakk> Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+             Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;
+     evs \<in> bankerb_gets \<rbrakk>
+ \<Longrightarrow> Key K \<in> analz (knows B evs)"
+apply (force dest: Gets_imp_knows_analz)
+done
+
+
+(****
+ The following is to prove theorems of the form
+
+  Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow>
+  Key K \<in> analz (knows Spy evs)
+
+ A more general formula must be proved inductively.
+
+****)
+
+
+text{* Session keys are not used to encrypt other session keys *}
+lemma analz_image_freshK [rule_format (no_asm)]:
+     "evs \<in> bankerb_gets \<Longrightarrow>
+   \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
+          (Key K \<in> analz (Key`KK Un (knows Spy evs))) =
+          (K \<in> KK | Key K \<in> analz (knows Spy evs))"
+apply (erule bankerb_gets.induct)
+apply (drule_tac [8] Says_Server_message_form)
+apply (erule_tac [6] Gets_Server_message_form [THEN disjE], analz_freshK, spy_analz, auto) 
+done
+
+
+lemma analz_insert_freshK:
+     "\<lbrakk> evs \<in> bankerb_gets;  KAB \<notin> range shrK \<rbrakk> \<Longrightarrow>
+      (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
+      (K = KAB | Key K \<in> analz (knows Spy evs))"
+by (simp only: analz_image_freshK analz_image_freshK_simps)
+
+
+text{* The session key K uniquely identifies the message *}
+lemma unique_session_keys:
+     "\<lbrakk> Says Server A
+           (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;
+         Says Server A'
+          (Crypt (shrK A') \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs;
+         evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> A=A' & Tk=Tk' & B=B' & X = X'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)
+txt{*BK2: it can't be a new key*}
+apply blast
+done
+
+lemma unique_session_keys_Gets:
+     "\<lbrakk> Gets A
+           (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;
+        Gets A
+          (Crypt (shrK A) \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs;
+        A \<notin> bad; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> Tk=Tk' & B=B' & X = X'"
+apply (blast dest!: Kab_authentic unique_session_keys)
+done
+
+
+lemma Server_Unique:
+     "\<lbrakk> Says Server A
+            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
+        evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> 
+   Unique Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+   on evs"
+apply (erule rev_mp, erule bankerb_gets.induct, simp_all add: Unique_def)
+apply blast
+done
+
+
+
+subsection{*Non-temporal guarantees, explicitly relying on non-occurrence of
+oops events - refined below by temporal guarantees*}
+
+text{*Non temporal treatment of confidentiality*}
+
+text{* Lemma: the session key sent in msg BK2 would be lost by oops
+    if the spy could see it! *}
+lemma lemma_conf [rule_format (no_asm)]:
+     "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+  \<Longrightarrow> Says Server A
+          (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
+                            Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
+         \<in> set evs \<longrightarrow>
+      Key K \<in> analz (knows Spy evs) \<longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<in> set evs"
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Says_Server_message_form)
+apply (frule_tac [6] Gets_Server_message_form [THEN disjE])
+apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*BK2*}
+apply (blast intro: parts_insertI)
+txt{*BK3*}
+apply (case_tac "Aa \<in> bad")
+ prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
+apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)
+txt{*Oops*}
+apply (blast dest: unique_session_keys)
+done
+
+
+text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.*}
+lemma Confidentiality_S:
+     "\<lbrakk> Says Server A
+          (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+apply (frule Says_Server_message_form, assumption)
+apply (blast intro: lemma_conf)
+done
+
+text{*Confidentiality for Alice*}
+lemma Confidentiality_A:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+by (blast dest!: Kab_authentic Confidentiality_S)
+
+text{*Confidentiality for Bob*}
+lemma Confidentiality_B:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
+          \<in> parts (knows Spy evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+by (blast dest!: ticket_authentic Confidentiality_S)
+
+
+text{*Non temporal treatment of authentication*}
+
+text{*Lemmas @{text lemma_A} and @{text lemma_B} in fact are common to both temporal and non-temporal treatments*}
+lemma lemma_A [rule_format]:
+     "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow>
+         Key K \<notin> analz (knows Spy evs) \<longrightarrow>
+         Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
+         \<in> set evs \<longrightarrow>
+          Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
+         Says A B \<lbrace>X, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace>
+             \<in> set evs"
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] Gets_Server_message_form)
+apply (frule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*Fake*}
+apply blast
+txt{*BK2*}
+apply (force dest: Crypt_imp_invKey_keysFor)
+txt{*BK3*}
+apply (blast dest: Kab_authentic unique_session_keys)
+done
+lemma lemma_B [rule_format]:
+     "\<lbrakk> B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs) \<longrightarrow>
+          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
+          \<in> set evs \<longrightarrow>
+          Crypt K (Number Ta) \<in> parts (knows Spy evs) \<longrightarrow>
+          Says B A (Crypt K (Number Ta)) \<in> set evs"
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Oops_parts_knows_Spy)
+apply (frule_tac [6] Gets_Server_message_form)
+apply (drule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)
+apply (simp_all (no_asm_simp) add: all_conj_distrib)
+txt{*Fake*}
+apply blast
+txt{*BK2*} 
+apply (force dest: Crypt_imp_invKey_keysFor)
+txt{*BK4*}
+apply (blast dest: ticket_authentic unique_session_keys
+                   Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad)
+done
+
+
+text{*The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+
+text{*Authentication of A to B*}
+lemma B_authenticates_A_r:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs);
+         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (knows Spy evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+by (blast dest!: ticket_authentic
+          intro!: lemma_A
+          elim!: Confidentiality_S [THEN [2] rev_notE])
+
+text{*Authentication of B to A*}
+lemma A_authenticates_B_r:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (knows Spy evs);
+        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);
+        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
+by (blast dest!: Kab_authentic
+          intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])
+
+lemma B_authenticates_A:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
+         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
+        Key K \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+apply (blast dest!: ticket_authentic intro!: lemma_A)
+done
+
+lemma A_authenticates_B:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
+        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+        Key K \<notin> analz (spies evs);
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
+apply (blast dest!: Kab_authentic intro!: lemma_B)
+done
+
+
+subsection{*Temporal guarantees, relying on a temporal check that insures that
+no oops event occurred. These are available in the sense of goal availability*}
+
+
+text{*Temporal treatment of confidentiality*}
+
+text{* Lemma: the session key sent in msg BK2 would be EXPIRED
+    if the spy could see it! *}
+lemma lemma_conf_temporal [rule_format (no_asm)]:
+     "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+  \<Longrightarrow> Says Server A
+          (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
+                            Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
+         \<in> set evs \<longrightarrow>
+      Key K \<in> analz (knows Spy evs) \<longrightarrow> expiredK Tk evs"
+apply (erule bankerb_gets.induct)
+apply (frule_tac [8] Says_Server_message_form)
+apply (frule_tac [6] Gets_Server_message_form [THEN disjE])
+apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
+txt{*Fake*}
+apply spy_analz
+txt{*BK2*}
+apply (blast intro: parts_insertI less_SucI)
+txt{*BK3*}
+apply (case_tac "Aa \<in> bad")
+ prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
+apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)
+txt{*Oops: PROOF FAILS if unsafe intro below*}
+apply (blast dest: unique_session_keys intro!: less_SucI)
+done
+
+
+text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.*}
+lemma Confidentiality_S_temporal:
+     "\<lbrakk> Says Server A
+          (Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs;
+         \<not> expiredK T evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+apply (frule Says_Server_message_form, assumption)
+apply (blast intro: lemma_conf_temporal)
+done
+
+text{*Confidentiality for Alice*}
+lemma Confidentiality_A_temporal:
+     "\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);
+         \<not> expiredK T evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+by (blast dest!: Kab_authentic Confidentiality_S_temporal)
+
+text{*Confidentiality for Bob*}
+lemma Confidentiality_B_temporal:
+     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
+          \<in> parts (knows Spy evs);
+        \<not> expiredK Tk evs;
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets
+      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+by (blast dest!: ticket_authentic Confidentiality_S_temporal)
+
+
+text{*Temporal treatment of authentication*}
+
+text{*Authentication of A to B*}
+lemma B_authenticates_A_temporal:
+     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs);
+         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
+         \<in> parts (knows Spy evs);
+         \<not> expiredK Tk evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+by (blast dest!: ticket_authentic
+          intro!: lemma_A
+          elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
+
+text{*Authentication of B to A*}
+lemma A_authenticates_B_temporal:
+     "\<lbrakk> Crypt K (Number Ta) \<in> parts (knows Spy evs);
+         Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
+         \<in> parts (knows Spy evs);
+         \<not> expiredK Tk evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
+by (blast dest!: Kab_authentic
+          intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
+
+
+subsection{*Combined guarantees of key distribution and non-injective agreement on the session keys*}
+
+lemma B_authenticates_and_keydist_to_A:
+     "\<lbrakk> Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;
+        Key K \<notin> analz (spies evs);
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+    \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
+                  Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs 
+     \<and>  Key K \<in> analz (knows A evs)"
+apply (blast dest: B_authenticates_A BK3_imp_Gets Gets_A_knows_K)
+done
+
+lemma A_authenticates_and_keydist_to_B:
+     "\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
+        Gets A (Crypt K (Number Ta)) \<in> set evs;
+        Key K \<notin> analz (spies evs);
+        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
+    \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs
+    \<and>   Key K \<in> analz (knows B evs)"
+apply (blast dest: A_authenticates_B BK4_imp_Gets Gets_B_knows_K)
+done
+
+
+
+
+
+end
--- a/src/HOL/Auth/NS_Shared.thy	Wed Feb 01 12:23:14 2006 +0100
+++ b/src/HOL/Auth/NS_Shared.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -1,10 +1,10 @@
 (*  Title:      HOL/Auth/NS_Shared
     ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author:     Lawrence C Paulson and Giampaolo Bella 
     Copyright   1996  University of Cambridge
 *)
 
-header{*The Needham-Schroeder Shared-Key Protocol*}
+header{*Needham-Schroeder Shared-Key Protocol and the Issues Property*}
 
 theory NS_Shared imports Public begin
 
@@ -14,6 +14,17 @@
   Proc. Royal Soc. 426
 *}
 
+constdefs
+
+ (* A is the true creator of X if she has sent X and X never appeared on
+    the trace before this event. Recall that traces grow from head. *)
+  Issues :: "[agent, agent, msg, event list] => bool"
+             ("_ Issues _ with _ on _")
+   "A Issues B with X on evs ==
+      \<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
+      X \<notin> parts (spies (takeWhile (% z. z  \<noteq> Says A B Y) (rev evs)))"
+
+
 consts  ns_shared   :: "event list set"
 inductive "ns_shared"
  intros
@@ -242,11 +253,10 @@
      "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
        Says Server A' (Crypt (shrK A') \<lbrace>NA', Agent B', Key K, X'\<rbrace>) \<in> set evs;
        evs \<in> ns_shared\<rbrakk> \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B' \<and> X = X'"
-apply (erule rev_mp, erule rev_mp, erule ns_shared.induct, simp_all, blast+)
-done
+by (erule rev_mp, erule rev_mp, erule ns_shared.induct, simp_all, blast+)
 
 
-subsubsection{*Crucial secrecy property: Spy does not see the keys sent in msg NS2*}
+subsubsection{*Crucial secrecy property: Spy doesn't see the keys sent in NS2*}
 
 text{*Beware of @{text "[rule_format]"} and the universal quantifier!*}
 lemma secrecy_lemma:
@@ -333,7 +343,9 @@
      Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow>
      Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
      (\<exists>A'. Says A' B X \<in> set evs)"
-apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, analz_mono_contra)
+apply (erule ns_shared.induct, force)
+apply (drule_tac [4] NS3_msg_in_parts_spies)
+apply analz_mono_contra
 apply (simp_all add: ex_disj_distrib, blast)
 txt{*NS2*}
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
@@ -352,7 +364,9 @@
 			    Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
      Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
      Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
-apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, analz_mono_contra, simp_all, blast)
+apply (erule ns_shared.induct, force)
+apply (drule_tac [4] NS3_msg_in_parts_spies)
+apply (analz_mono_contra, simp_all, blast)
 txt{*NS2*}
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
 txt{*NS5*}
@@ -372,4 +386,143 @@
 by (blast intro: B_trusts_NS5_lemma
           dest: B_trusts_NS3 Spy_not_see_encrypted_key)
 
+text{*Unaltered so far wrt original version*}
+
+subsection{*Lemmas for reasoning about predicate "Issues"*}
+
+lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
+          (if A:bad then insert X (spies evs) else spies evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+lemma spies_evs_rev: "spies evs = spies (rev evs)"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a")
+apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
+done
+
+lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono]
+
+lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+done
+
+lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
+
+
+subsection{*Guarantees of non-injective agreement on the session key, and
+of key distribution. They also express forms of freshness of certain messages,
+namely that agents were alive after something happened.*}
+
+lemma B_Issues_A:
+     "\<lbrakk> Says B A (Crypt K (Nonce Nb)) \<in> set evs;
+         Key K \<notin> analz (spies evs);
+         A \<notin> bad;  B \<notin> bad; evs \<in> ns_shared \<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt K (Nonce Nb)) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule ns_shared.induct, analz_mono_contra)
+apply (simp_all)
+txt{*fake*}
+apply blast
+apply (simp_all add: takeWhile_tail)
+txt{*NS3 remains by pure coincidence!*}
+apply (force dest!: A_trusts_NS2 Says_Server_message_form)
+txt{*NS4 would be the non-trivial case can be solved by Nb being used*}
+apply (blast dest: parts_spies_takeWhile_mono [THEN subsetD]
+                   parts_spies_evs_revD2 [THEN subsetD])
+done
+
+text{*Tells A that B was alive after she sent him the session key.  The
+session key must be assumed confidential for this deduction to be meaningful,
+but that assumption can be relaxed by the appropriate argument.
+
+Precisely, the theorem guarantees (to A) key distribution of the session key
+to B. It also guarantees (to A) non-injective agreement of B with A on the
+session key. Both goals are available to A in the sense of Goal Availability.
+*}
+lemma A_authenticates_and_keydist_to_B:
+     "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs);
+       Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
+       Key K \<notin> analz(knows Spy evs);
+       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_shared\<rbrakk>
+      \<Longrightarrow> B Issues A with (Crypt K (Nonce NB)) on evs"
+by (blast intro: A_trusts_NS4_lemma B_Issues_A dest: A_trusts_NS2)
+
+lemma A_trusts_NS5:
+  "\<lbrakk> Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts(spies evs);
+     Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace> \<in> parts(spies evs);
+     Key K \<notin> analz (spies evs);
+     A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk>
+ \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs";
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule ns_shared.induct, analz_mono_contra)
+apply (frule_tac [5] Says_S_message_form)
+apply (simp_all)
+txt{*Fake*}
+apply blast
+txt{*NS2*}
+apply (force dest!: Crypt_imp_keysFor)
+txt{*NS3, much quicker having installed @{term Says_S_message_form} before simplication*}
+apply fastsimp
+txt{*NS5, the most important case, can be solved by unicity*}
+apply (case_tac "Aa \<in> bad")
+apply (force dest!: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
+apply (blast dest: A_trusts_NS2 unique_session_keys)
+done
+
+lemma A_Issues_B:
+     "\<lbrakk> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs;
+        Key K \<notin> analz (spies evs);
+        A \<notin> bad;  B \<notin> bad; evs \<in> ns_shared \<rbrakk>
+    \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) on evs"
+apply (simp (no_asm) add: Issues_def)
+apply (rule exI)
+apply (rule conjI, assumption)
+apply (simp (no_asm))
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule ns_shared.induct, analz_mono_contra)
+apply (simp_all)
+txt{*fake*}
+apply blast
+apply (simp_all add: takeWhile_tail)
+txt{*NS3 remains by pure coincidence!*}
+apply (force dest!: A_trusts_NS2 Says_Server_message_form)
+txt{*NS5 is the non-trivial case and cannot be solved as in @{term B_Issues_A}! because NB is not fresh. We need @{term A_trusts_NS5}, proved for this very purpose*}
+apply (blast dest: A_trusts_NS5 parts_spies_takeWhile_mono [THEN subsetD]
+        parts_spies_evs_revD2 [THEN subsetD])
+done
+
+text{*Tells B that A was alive after B issued NB.
+
+Precisely, the theorem guarantees (to B) key distribution of the session key to A. It also guarantees (to B) non-injective agreement of A with B on the session key. Both goals are available to B in the sense of Goal Availability.
+*}
+lemma B_authenticates_and_keydist_to_A:
+     "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs);
+       Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
+       Key K \<notin> analz (spies evs);
+       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_shared\<rbrakk>
+   \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) on evs"
+by (blast intro: A_Issues_B B_trusts_NS5_lemma dest: B_trusts_NS3)
+
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/OtwayReesBella.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,392 @@
+(*  ID:         $Id$
+    Author:     Giampaolo Bella, Catania University
+*)
+
+header{*Bella's version of the Otway-Rees protocol*}
+
+
+theory OtwayReesBella imports Public begin
+
+text{*Bella's modifications to a version of the Otway-Rees protocol taken from
+the BAN paper only concern message 7. The updated protocol makes the goal of
+key distribution of the session key available to A. Investigating the
+principle of Goal Availability undermines the BAN claim about the original
+protocol, that "this protocol does not make use of Kab as an encryption key,
+so neither principal can know whether the key is known to the other". The
+updated protocol makes no use of the session key to encrypt but informs A that
+B knows it.*}
+
+consts  orb   :: "event list set"
+inductive "orb"
+ intros 
+
+  Nil:  "[]\<in> orb"
+
+  Fake: "\<lbrakk>evsa\<in> orb;  X\<in> synth (analz (knows Spy evsa))\<rbrakk>
+ 	 \<Longrightarrow> Says Spy B X  # evsa \<in> orb"
+
+  Reception: "\<lbrakk>evsr\<in> orb;  Says A B X \<in> set evsr\<rbrakk>
+	      \<Longrightarrow> Gets B X # evsr \<in> orb"
+
+  OR1:  "\<lbrakk>evs1\<in> orb;  Nonce NA \<notin> used evs1\<rbrakk>
+	 \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, 
+		   Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
+	       # evs1 \<in> orb"
+
+  OR2:  "\<lbrakk>evs2\<in> orb;  Nonce NB \<notin> used evs2;
+	   Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
+	\<Longrightarrow> Says B Server 
+		\<lbrace>Nonce M, Agent A, Agent B, X, 
+	   Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+	       # evs2 \<in> orb"
+
+  OR3:  "\<lbrakk>evs3\<in> orb;  Key KAB \<notin> used evs3;
+	  Gets Server 
+	     \<lbrace>Nonce M, Agent A, Agent B, 
+	       Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>, 
+	       Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+	  \<in> set evs3\<rbrakk>
+	\<Longrightarrow> Says Server B \<lbrace>Nonce M,
+		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+				      Nonce NB, Key KAB\<rbrace>\<rbrace>
+	       # evs3 \<in> orb"
+
+  (*B can only check that the message he is bouncing is a ciphertext*)
+  (*Sending M back is omitted*)   
+  OR4:  "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>; 
+	  Says B Server \<lbrace>Nonce M, Agent A, Agent B, X', 
+		Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+	    \<in> set evs4;
+	  Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
+	    \<in> set evs4\<rbrakk>
+	\<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
+
+
+  Oops: "\<lbrakk>evso\<in> orb;  
+	   Says Server B \<lbrace>Nonce M,
+		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+				      Nonce NB, Key KAB\<rbrace>\<rbrace> 
+	     \<in> set evso\<rbrakk>
+ \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso 
+     \<in> orb"
+
+
+
+declare knows_Spy_partsEs [elim]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un  [dest]
+
+
+text{*Fragile proof, with backtracking in the possibility call.*}
+lemma possibility_thm: "\<lbrakk>A \<noteq> Server; B \<noteq> Server; Key K \<notin> used[]\<rbrakk>    
+      \<Longrightarrow>   \<exists> evs \<in> orb.           
+     Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2] orb.Nil
+                    [THEN orb.OR1, THEN orb.Reception,
+                     THEN orb.OR2, THEN orb.Reception,
+                     THEN orb.OR3, THEN orb.Reception, THEN orb.OR4]) 
+apply (possibility, simp add: used_Cons)  
+done
+
+
+lemma Gets_imp_Says :
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
+apply (erule rev_mp)
+apply (erule orb.induct)
+apply auto
+done
+
+lemma Gets_imp_knows_Spy: 
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
+by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+
+declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
+
+lemma Gets_imp_knows:
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk>  \<Longrightarrow> X \<in> knows B evs"
+apply (case_tac "B = Spy")
+apply (blast dest!: Gets_imp_knows_Spy)
+apply (blast dest!: Gets_imp_knows_agents)
+done
+
+
+lemma OR2_analz_knows_Spy: 
+   "\<lbrakk>Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs; evs \<in> orb\<rbrakk>   
+    \<Longrightarrow> X \<in> analz (knows Spy evs)"
+by (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
+
+lemma OR4_parts_knows_Spy: 
+   "\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key Kab\<rbrace>\<rbrace>  \<in> set evs; 
+      evs \<in> orb\<rbrakk>   \<Longrightarrow> X \<in> parts (knows Spy evs)"
+by blast
+
+lemma Oops_parts_knows_Spy: 
+    "Says Server B \<lbrace>Nonce M, Crypt K' \<lbrace>X, Nonce Nb, K\<rbrace>\<rbrace> \<in> set evs  
+     \<Longrightarrow> K \<in> parts (knows Spy evs)"
+by blast
+
+lemmas OR2_parts_knows_Spy =
+    OR2_analz_knows_Spy [THEN analz_into_parts, standard]
+
+ML
+{*
+val Oops_parts_knows_Spy = thm "Oops_parts_knows_Spy"
+val OR4_parts_knows_Spy = thm "OR4_parts_knows_Spy"
+val OR2_parts_knows_Spy = thm "OR2_parts_knows_Spy"
+
+fun parts_explicit_tac i = 
+    forward_tac [Oops_parts_knows_Spy] (i+7) THEN
+    forward_tac [OR4_parts_knows_Spy]  (i+6) THEN
+    forward_tac [OR2_parts_knows_Spy]  (i+4)
+*}
+ 
+method_setup parts_explicit = {*
+    Method.ctxt_args (fn ctxt =>
+        Method.METHOD (fn facts =>
+            parts_explicit_tac 1)) *}
+  "to explicitly state that some message components belong to parts knows Spy"
+
+
+lemma Spy_see_shrK [simp]: 
+    "evs \<in> orb \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
+by (erule orb.induct, parts_explicit, simp_all, blast+)
+
+lemma Spy_analz_shrK [simp]: 
+"evs \<in> orb \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
+by auto
+
+lemma Spy_see_shrK_D [dest!]:
+     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> orb|] ==> A \<in> bad"
+by (blast dest: Spy_see_shrK)
+
+lemma new_keys_not_used [simp]:
+   "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> orb\<rbrakk>  \<Longrightarrow> K \<notin> keysFor (parts (knows Spy evs))"
+apply (erule rev_mp)
+apply (erule orb.induct, parts_explicit, simp_all)
+apply (force dest!: keysFor_parts_insert)
+apply (blast+)
+done
+
+
+
+subsection{* Proofs involving analz *}
+
+text{*Describes the form of K and NA when the Server sends this message.  Also
+  for Oops case.*}
+lemma Says_Server_message_form: 
+"\<lbrakk>Says Server B  \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  
+     evs \<in> orb\<rbrakk>                                            
+ \<Longrightarrow> K \<notin> range shrK & (\<exists> A Na. X=(Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>))"
+by (erule rev_mp, erule orb.induct, simp_all)
+
+lemma Says_Server_imp_Gets: 
+ "\<lbrakk>Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>,
+                                             Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;
+    evs \<in> orb\<rbrakk>
+  \<Longrightarrow>  Gets Server \<lbrace>Nonce M, Agent A, Agent B, 
+                   Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>, 
+               Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+         \<in> set evs"
+by (erule rev_mp, erule orb.induct, simp_all)
+
+
+lemma A_trusts_OR1: 
+"\<lbrakk>Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs);  
+    A \<notin> bad; evs \<in> orb\<rbrakk>                   
+ \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
+apply (blast)
+done
+
+
+lemma B_trusts_OR2:
+ "\<lbrakk>Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>  
+      \<in> parts (knows Spy evs);  B \<notin> bad; evs \<in> orb\<rbrakk>                   
+  \<Longrightarrow> (\<exists> X. Says B Server \<lbrace>Nonce M, Agent A, Agent B, X,  
+              Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
+          \<in> set evs)"
+apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
+apply (blast+)
+done
+
+
+lemma B_trusts_OR3: 
+"\<lbrakk>Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace> \<in> parts (knows Spy evs);  
+   B \<notin> bad; evs \<in> orb\<rbrakk>                   
+\<Longrightarrow> \<exists> M. Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> 
+         \<in> set evs"
+apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
+apply (blast+)
+done
+
+lemma Gets_Server_message_form: 
+"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  
+    evs \<in> orb\<rbrakk>                                              
+ \<Longrightarrow> (K \<notin> range shrK & (\<exists> A Na. X = (Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>)))    
+             | X \<in> analz (knows Spy evs)"
+apply (case_tac "B \<in> bad")
+apply (drule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, 
+                                 THEN analz.Decrypt, THEN analz.Fst])
+prefer 3 apply blast
+prefer 3 apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN 
+                                                 parts.Snd, THEN B_trusts_OR3]
+                             Says_Server_message_form)
+apply simp_all                                    
+done
+
+lemma unique_Na: "\<lbrakk>Says A B  \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs;   
+         Says A B' \<lbrace>Nonce M', Agent A, Agent B', Crypt (shrK A) \<lbrace>Nonce Na, Nonce M', Agent A, Agent B'\<rbrace>\<rbrace> \<in> set evs;  
+    A \<notin> bad; evs \<in> orb\<rbrakk> \<Longrightarrow> B=B' & M=M'"
+by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+)
+
+lemma unique_Nb: "\<lbrakk>Says B Server \<lbrace>Nonce M, Agent A, Agent B, X, Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs;   
+         Says B Server \<lbrace>Nonce M', Agent A', Agent B, X', Crypt (shrK B) \<lbrace>Nonce Nb,Nonce M', Nonce M', Agent A', Agent B\<rbrace>\<rbrace> \<in> set evs;   
+    B \<notin> bad; evs \<in> orb\<rbrakk> \<Longrightarrow>   M=M' & A=A' & X=X'"
+by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+)
+
+lemma analz_image_freshCryptK_lemma:
+"(Crypt K X \<in> analz (Key`nE \<union> H)) \<longrightarrow> (Crypt K X \<in> analz H) \<Longrightarrow>  
+        (Crypt K X \<in> analz (Key`nE \<union> H)) = (Crypt K X \<in> analz H)";
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+ML
+{*
+val analz_image_freshCryptK_lemma = thm "analz_image_freshCryptK_lemma";
+val analz_image_freshK_simps = thms "analz_image_freshK_simps";
+
+val analz_image_freshK_ss = 
+     simpset() delsimps [image_insert, image_Un]
+	       delsimps [imp_disjL]    (*reduces blow-up*)
+	       addsimps thms "analz_image_freshK_simps"
+*}
+
+method_setup analz_freshCryptK = {*
+    Method.no_args
+     (Method.METHOD
+      (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+                          REPEAT_FIRST (rtac analz_image_freshCryptK_lemma),
+                          ALLGOALS (asm_simp_tac analz_image_freshK_ss)])) *}
+  "for proving useful rewrite rule"
+
+
+method_setup disentangle = {*
+    Method.no_args
+     (Method.METHOD
+      (fn facts => REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE] 
+                   ORELSE' hyp_subst_tac))) *}
+  "for eliminating conjunctions, disjunctions and the like"
+
+
+
+lemma analz_image_freshCryptK [rule_format]: 
+"evs \<in> orb \<Longrightarrow>                              
+     Key K \<notin> analz (knows Spy evs) \<longrightarrow>  
+       (\<forall> KK. KK \<subseteq> - (range shrK) \<longrightarrow>                  
+             (Crypt K X \<in> analz (Key`KK \<union> (knows Spy evs))) =   
+             (Crypt K X \<in> analz (knows Spy evs)))"
+apply (erule orb.induct)
+apply (analz_mono_contra)
+apply (frule_tac [7] Gets_Server_message_form)
+apply (frule_tac [9] Says_Server_message_form)
+apply disentangle
+apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN  analz.Snd])
+prefer 8 apply clarify
+apply (analz_freshCryptK, spy_analz, fastsimp)
+done
+
+
+
+lemma analz_insert_freshCryptK: 
+"\<lbrakk>evs \<in> orb;  Key K \<notin> analz (knows Spy evs);  
+         Seskey \<notin> range shrK\<rbrakk> \<Longrightarrow>   
+         (Crypt K X \<in> analz (insert (Key Seskey) (knows Spy evs))) =  
+         (Crypt K X \<in> analz (knows Spy evs))"
+by (simp only: analz_image_freshCryptK analz_image_freshK_simps)
+
+
+lemma analz_hard: 
+"\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B,  
+             Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in>set evs; 
+   Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace> \<in> analz (knows Spy evs);  
+   A \<notin> bad; B \<notin> bad; evs \<in> orb\<rbrakk>                   
+ \<Longrightarrow>  Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule orb.induct)
+apply (frule_tac [7] Gets_Server_message_form)
+apply (frule_tac [9] Says_Server_message_form)
+apply disentangle
+txt{*letting the simplifier solve OR2*}
+apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Snd])
+apply (simp_all (no_asm_simp) add: analz_insert_eq pushes split_ifs)
+apply (spy_analz)
+txt{*OR1*}
+apply blast
+txt{*Oops*}
+prefer 4 apply (blast dest: analz_insert_freshCryptK)
+txt{*OR4 - ii*}
+prefer 3 apply (blast)
+txt{*OR3*}
+(*adding Gets_imp_ and Says_imp_ for efficiency*)
+apply (blast dest: 
+       A_trusts_OR1 unique_Na Key_not_used analz_insert_freshCryptK)
+txt{*OR4 - i *}
+apply clarify
+apply (simp add: pushes split_ifs)
+apply (case_tac "Aaa\<in>bad")
+apply (blast dest: analz_insert_freshCryptK)
+apply clarify
+apply simp
+apply (case_tac "Ba\<in>bad")
+apply (frule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Decrypt, THEN analz.Fst] , assumption)
+apply (simp (no_asm_simp))
+apply clarify
+apply (frule Gets_imp_knows_Spy 
+             [THEN parts.Inj, THEN parts.Snd, THEN B_trusts_OR3],  
+       assumption, assumption, assumption, erule exE)
+apply (frule Says_Server_imp_Gets 
+            [THEN Gets_imp_knows_Spy, THEN parts.Inj, THEN parts.Snd, 
+            THEN parts.Snd, THEN parts.Snd, THEN parts.Fst, THEN A_trusts_OR1],
+       assumption, assumption, assumption, assumption)
+apply (blast dest: Says_Server_imp_Gets B_trusts_OR2 unique_Na unique_Nb)
+done
+
+
+lemma Gets_Server_message_form': 
+"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace>  \<in> set evs;  
+   B \<notin> bad; evs \<in> orb\<rbrakk>                              
+  \<Longrightarrow> K \<notin> range shrK & (\<exists> A Na. X = (Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>))"
+by (blast dest!: B_trusts_OR3 Says_Server_message_form)
+
+
+lemma OR4_imp_Gets: 
+"\<lbrakk>Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs;   
+   B \<notin> bad; evs \<in> orb\<rbrakk>  
+ \<Longrightarrow> (\<exists> Nb. Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>,
+                                             Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
+prefer 3 apply (blast dest: Gets_Server_message_form')
+apply blast+
+done
+
+
+lemma A_keydist_to_B: 
+"\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B,  
+            Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in>set evs; 
+   Gets A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs;    
+   A \<notin> bad; B \<notin> bad; evs \<in> orb\<rbrakk>  
+  \<Longrightarrow> Key K \<in> analz (knows B evs)"
+apply (drule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd], assumption)
+apply (drule analz_hard, assumption, assumption, assumption, assumption)
+apply (drule OR4_imp_Gets, assumption, assumption)
+apply (erule exE)
+(*blast doesn't do because it can't infer that Key (shrK P) \<in> (knows P evs)*)
+apply (fastsimp dest!: Gets_imp_knows [THEN analz.Inj] analz.Decrypt)
+done
+
+
+text{*Other properties as for the original protocol*}
+
+
+end
--- a/src/HOL/Auth/ROOT.ML	Wed Feb 01 12:23:14 2006 +0100
+++ b/src/HOL/Auth/ROOT.ML	Wed Feb 01 15:22:02 2006 +0100
@@ -8,17 +8,22 @@
 
 no_document use_thy "NatPair";
 
-add_path "Guard";
+set timing;
 
-set timing;
+(* Conventional protocols: rely on 
+   conventional Message, Event and Public *)
 
 (*Shared-key protocols*)
 time_use_thy "NS_Shared";
 time_use_thy "Kerberos_BAN";
+time_use_thy "Kerberos_BAN_Gets";
 time_use_thy "KerberosIV";
+time_use_thy "KerberosIV_Gets";
+time_use_thy "KerberosV";
 time_use_thy "OtwayRees";
 time_use_thy "OtwayRees_AN";
 time_use_thy "OtwayRees_Bad";
+time_use_thy "OtwayReesBella";
 time_use_thy "WooLam";
 time_use_thy "Recur";
 time_use_thy "Yahalom";
@@ -32,10 +37,12 @@
 time_use_thy "TLS";
 time_use_thy "CertifiedEmail";
 
+(*Smartcard protocols: rely on conventional Message and on new
+  EventSC and Smartcard *)
+
+with_path "Smartcard" time_use_thy "ShoupRubin";
+with_path "Smartcard" time_use_thy "ShoupRubinBella";
+
 (*Blanqui's "guard" concept: protocol-independent secrecy*)
-time_use_thy "P1";
-time_use_thy "P2";
-time_use_thy "Guard_NS_Public";
-time_use_thy "Guard_OtwayRees";
-time_use_thy "Guard_Yahalom";
-time_use_thy "Proto";
+with_path "Guard" (app time_use_thy)
+  ["P1", "P2", "Guard_NS_Public", "Guard_OtwayRees", "Guard_Yahalom", "Proto"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Smartcard/EventSC.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,447 @@
+header{*Theory of Events for Security Protocols that use smartcards*}
+
+theory EventSC imports "../Message" begin
+
+consts  (*Initial states of agents -- parameter of the construction*)
+  initState :: "agent => msg set"
+
+datatype card = Card agent
+
+text{*Four new events express the traffic between an agent and his card*}
+datatype  
+  event = Says  agent agent msg
+        | Notes agent       msg
+        | Gets  agent       msg
+        | Inputs agent card msg (*Agent sends to card and\<dots>*)
+        | C_Gets card       msg (*\<dots> card receives it*) 
+        | Outpts card agent msg (*Card sends to agent and\<dots>*)
+        | A_Gets agent      msg (*agent receives it*)
+    
+consts 
+ bad     :: "agent set"  (*compromised agents*)
+ knows   :: "agent => event list => msg set" (*agents' knowledge*)
+ stolen    :: "card set" (* stolen smart cards *)
+ cloned  :: "card set"   (* cloned smart cards*)
+ secureM :: "bool"(*assumption of secure means between agents and their cards*)
+
+syntax
+  insecureM :: bool (*certain protocols make no assumption of secure means*)
+translations
+  "insecureM" == "\<not>secureM"
+
+
+text{*Spy has access to his own key for spoof messages, but Server is secure*}
+specification (bad)
+  Spy_in_bad     [iff]: "Spy \<in> bad"
+  Server_not_bad [iff]: "Server \<notin> bad"
+  apply (rule exI [of _ "{Spy}"], simp) done
+
+specification (stolen)
+  (*The server's card is secure by assumption\<dots>*)
+  Card_Server_not_stolen [iff]: "Card Server \<notin> stolen"
+  Card_Spy_not_stolen  	 [iff]: "Card Spy \<notin> stolen"
+  apply blast done
+
+specification (cloned)
+  (*\<dots> the spy's card is secure because she already can use it freely*)
+  Card_Server_not_cloned [iff]: "Card Server \<notin> cloned"
+  Card_Spy_not_cloned  	 [iff]: "Card Spy \<notin> cloned"
+  apply blast done
+
+
+primrec (*This definition is extended over the new events, subject to the 
+          assumption of secure means*)
+  knows_Nil:   "knows A [] = initState A"
+  knows_Cons:  "knows A (ev # evs) =
+	 (case ev of
+	    Says A' B X => 
+                if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
+	  | Notes A' X  => 
+	        if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs) 
+                                             else knows A evs 
+          | Gets A' X   =>
+		if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
+                                     else knows A evs
+          | Inputs A' C X =>
+	      if secureM then
+                if A=A' then insert X (knows A evs) else knows A evs
+	      else
+	        if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
+          | C_Gets C X   => knows A evs
+          | Outpts C A' X =>
+	      if secureM then
+                if A=A' then insert X (knows A evs) else knows A evs
+              else
+	        if A=Spy then insert X (knows A evs) else knows A evs
+          | A_Gets A' X   =>
+		if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
+                                     else knows A evs)"
+
+
+
+consts
+  (*The set of items that might be visible to someone is easily extended 
+    over the new events*)
+  used :: "event list => msg set"
+
+primrec
+  used_Nil:   "used []         = (UN B. parts (initState B))"
+  used_Cons:  "used (ev # evs) =
+	         (case ev of
+		    Says A B X => parts {X} \<union> (used evs)
+		  | Notes A X  => parts {X} \<union> (used evs)
+		  | Gets A X   => used evs
+                  | Inputs A C X  => parts{X} \<union> (used evs) 
+		  | C_Gets C X   => used evs
+                  | Outpts C A X  => parts{X} \<union> (used evs)
+		  | A_Gets A X   => used evs)"
+    --{*@{term Gets} always follows @{term Says} in real protocols. 
+       Likewise, @{term C_Gets} will always have to follow @{term Inputs}
+       and @{term A_Gets} will always have to follow @{term Outpts}*}
+
+
+lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs \<longrightarrow> X \<in> used evs"
+apply (induct_tac evs)
+apply (auto split: event.split) 
+done
+
+lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs \<longrightarrow> X \<in> used evs"
+apply (induct_tac evs)
+apply (auto split: event.split) 
+done
+
+lemma MPair_used [rule_format]:
+     "MPair X Y \<in> used evs \<longrightarrow> X \<in> used evs & Y \<in> used evs"
+apply (induct_tac evs)
+apply (auto split: event.split) 
+done
+
+
+subsection{*Function @{term knows}*}
+
+(*Simplifying   
+ parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
+  This version won't loop with the simplifier.*)
+lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard]
+
+lemma knows_Spy_Says [simp]:
+     "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
+by simp
+
+text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
+      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
+lemma knows_Spy_Notes [simp]:
+     "knows Spy (Notes A X # evs) =  
+          (if A\<in>bad then insert X (knows Spy evs) else knows Spy evs)"
+by simp
+
+lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
+by simp
+
+lemma knows_Spy_Inputs_secureM [simp]: 
+     "secureM \<Longrightarrow> knows Spy (Inputs A C X # evs) =  
+        (if A=Spy then insert X (knows Spy evs) else knows Spy evs)"
+by simp
+
+lemma knows_Spy_Inputs_insecureM [simp]: 
+     "insecureM \<Longrightarrow> knows Spy (Inputs A C X # evs) = insert X (knows Spy evs)"
+by simp
+
+lemma knows_Spy_C_Gets [simp]: "knows Spy (C_Gets C X # evs) = knows Spy evs"
+by simp
+
+lemma knows_Spy_Outpts_secureM [simp]: 
+      "secureM \<Longrightarrow> knows Spy (Outpts C A X # evs) = 
+         (if A=Spy then insert X (knows Spy evs) else knows Spy evs)"
+by simp
+
+lemma knows_Spy_Outpts_insecureM [simp]: 
+      "insecureM \<Longrightarrow> knows Spy (Outpts C A X # evs) = insert X (knows Spy evs)"
+by simp
+
+lemma knows_Spy_A_Gets [simp]: "knows Spy (A_Gets A X # evs) = knows Spy evs"
+by simp
+
+
+
+
+lemma knows_Spy_subset_knows_Spy_Says:
+     "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_Spy_subset_knows_Spy_Notes:
+     "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)"
+by force
+
+lemma knows_Spy_subset_knows_Spy_Gets:
+     "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_Spy_subset_knows_Spy_Inputs: 
+     "knows Spy evs \<subseteq> knows Spy (Inputs A C X # evs)"
+by auto
+
+lemma knows_Spy_equals_knows_Spy_Gets: 
+     "knows Spy evs = knows Spy (C_Gets C X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_Spy_subset_knows_Spy_Outpts: "knows Spy evs \<subseteq> knows Spy (Outpts C A X # evs)"
+by auto
+
+lemma knows_Spy_subset_knows_Spy_A_Gets: "knows Spy evs \<subseteq> knows Spy (A_Gets A X # evs)"
+by (simp add: subset_insertI)
+
+
+
+text{*Spy sees what is sent on the traffic*}
+lemma Says_imp_knows_Spy [rule_format]:
+     "Says A B X \<in> set evs \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+lemma Notes_imp_knows_Spy [rule_format]:
+     "Notes A X \<in> set evs \<longrightarrow> A\<in> bad \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+(*Nothing can be stated on a Gets event*)
+
+lemma Inputs_imp_knows_Spy_secureM [rule_format (no_asm)]: 
+     "Inputs Spy C X \<in> set evs \<longrightarrow> secureM \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+lemma Inputs_imp_knows_Spy_insecureM [rule_format (no_asm)]:
+     "Inputs A C X \<in> set evs \<longrightarrow> insecureM \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+(*Nothing can be stated on a C_Gets event*)
+
+lemma Outpts_imp_knows_Spy_secureM [rule_format (no_asm)]: 
+     "Outpts C Spy X \<in> set evs \<longrightarrow> secureM \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+lemma Outpts_imp_knows_Spy_insecureM [rule_format (no_asm)]:
+     "Outpts C A X \<in> set evs \<longrightarrow> insecureM \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+(*Nothing can be stated on an A_Gets event*)
+
+
+
+text{*Elimination rules: derive contradictions from old Says events containing
+  items known to be fresh*}
+lemmas knows_Spy_partsEs =
+     Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] 
+     parts.Body [THEN revcut_rl, standard]
+
+
+
+subsection{*Knowledge of Agents*}
+
+lemma knows_Says: "knows A (Says A B X # evs) = insert X (knows A evs)"
+by simp
+
+lemma knows_Notes: "knows A (Notes A X # evs) = insert X (knows A evs)"
+by simp
+
+lemma knows_Gets:
+      "A \<noteq> Spy \<longrightarrow> knows A (Gets A X # evs) = insert X (knows A evs)"
+by simp
+
+lemma knows_Inputs: "knows A (Inputs A C X # evs) = insert X (knows A evs)"
+by simp
+
+lemma knows_C_Gets: "knows A (C_Gets C X # evs) = knows A evs"
+by simp
+
+lemma knows_Outpts_secureM: 
+      "secureM \<longrightarrow> knows A (Outpts C A X # evs) = insert X (knows A evs)"
+by simp
+
+lemma knows_Outpts_secureM: 
+      "insecureM \<longrightarrow> knows Spy (Outpts C A X # evs) = insert X (knows Spy evs)"
+by simp
+(*somewhat equivalent to knows_Spy_Outpts_insecureM*)
+
+
+
+
+lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_subset_knows_Inputs: "knows A evs \<subseteq> knows A (Inputs A' C X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_subset_knows_C_Gets: "knows A evs \<subseteq> knows A (C_Gets C X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_subset_knows_Outpts: "knows A evs \<subseteq> knows A (Outpts C A' X # evs)"
+by (simp add: subset_insertI)
+
+lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (A_Gets A' X # evs)"
+by (simp add: subset_insertI)
+
+
+
+text{*Agents know what they say*}
+lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs \<longrightarrow> X \<in> knows A evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+text{*Agents know what they note*}
+lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+text{*Agents know what they receive*}
+lemma Gets_imp_knows_agents [rule_format]:
+     "A \<noteq> Spy \<longrightarrow> Gets A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+(*Agents know what they input to their smart card*)
+lemma Inputs_imp_knows_agents [rule_format (no_asm)]: 
+     "Inputs A (Card A) X \<in> set evs \<longrightarrow> X \<in> knows A evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+(*Nothing to prove about C_Gets*)
+
+(*Agents know what they obtain as output of their smart card,
+  if the means is secure...*)
+lemma Outpts_imp_knows_agents_secureM [rule_format (no_asm)]: 
+     "secureM \<longrightarrow> Outpts (Card A) A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+(*otherwise only the spy knows the outputs*)
+lemma Outpts_imp_knows_agents_insecureM [rule_format (no_asm)]: 
+      "insecureM \<longrightarrow> Outpts (Card A) A X \<in> set evs \<longrightarrow> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+done
+
+(*end lemmas about agents' knowledge*)
+
+
+
+lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs"
+apply (induct_tac "evs", force)  
+apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) 
+done
+
+lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
+
+lemma initState_into_used: "X \<in> parts (initState B) \<Longrightarrow> X \<in> used evs"
+apply (induct_tac "evs")
+apply (simp_all add: parts_insert_knows_A split add: event.split, blast)
+done
+
+lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
+by simp
+
+lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
+by simp
+
+lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
+by simp
+
+lemma used_Inputs [simp]: "used (Inputs A C X # evs) = parts{X} \<union> used evs"
+by simp
+
+lemma used_C_Gets [simp]: "used (C_Gets C X # evs) = used evs"
+by simp
+
+lemma used_Outpts [simp]: "used (Outpts C A X # evs) = parts{X} \<union> used evs"
+by simp
+
+lemma used_A_Gets [simp]: "used (A_Gets A X # evs) = used evs"
+by simp
+
+lemma used_nil_subset: "used [] \<subseteq> used evs"
+apply simp
+apply (blast intro: initState_into_used)
+done
+
+
+
+(*Novel lemmas*)
+lemma Says_parts_used [rule_format (no_asm)]: 
+     "Says A B X \<in> set evs \<longrightarrow> (parts  {X}) \<subseteq> used evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+lemma Notes_parts_used [rule_format (no_asm)]: 
+     "Notes A X \<in> set evs \<longrightarrow> (parts  {X}) \<subseteq> used evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+lemma Outpts_parts_used [rule_format (no_asm)]: 
+     "Outpts C A X \<in> set evs \<longrightarrow> (parts  {X}) \<subseteq> used evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+lemma Inputs_parts_used [rule_format (no_asm)]: 
+     "Inputs A C X \<in> set evs \<longrightarrow> (parts  {X}) \<subseteq> used evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) split add: event.split)
+apply blast
+done
+
+
+
+
+text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
+declare knows_Cons [simp del]
+        used_Nil [simp del] used_Cons [simp del]
+
+
+lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)"
+by (induct e, auto simp: knows_Cons)
+
+lemma initState_subset_knows: "initState A \<subseteq> knows A evs"
+apply (induct_tac evs, simp) 
+apply (blast intro: knows_subset_knows_Cons [THEN subsetD])
+done
+
+
+text{*For proving @{text new_keys_not_used}*}
+lemma keysFor_parts_insert:
+     "\<lbrakk> K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) \<rbrakk>   
+      \<Longrightarrow> K \<in> keysFor (parts (G \<union> H)) \<or> Key (invKey K) \<in> parts H"; 
+by (force 
+    dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
+           analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD]
+    intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD])
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,1396 @@
+(*  ID:         $Id$
+    Author:     Giampaolo Bella, Catania University
+*)
+
+header{*Original Shoup-Rubin protocol*}
+
+theory ShoupRubin imports Smartcard begin
+
+consts
+
+    sesK :: "nat*key => key"
+
+axioms
+     
+   (*sesK is injective on each component*) 
+   inj_sesK [iff]: "(sesK(m,k) = sesK(m',k')) = (m = m' \<and> k = k')"
+
+   (*all long-term keys differ from sesK*)
+   shrK_disj_sesK [iff]: "shrK A \<noteq> sesK(m,pk)"
+   crdK_disj_sesK [iff]: "crdK C \<noteq> sesK(m,pk)"
+   pin_disj_sesK  [iff]: "pin P \<noteq> sesK(m,pk)"
+   pairK_disj_sesK[iff]:"pairK(A,B) \<noteq> sesK(m,pk)"
+
+   (*needed for base case in analz_image_freshK*)
+   Atomic_distrib [iff]: "Atomic`(KEY`K \<union> NONCE`N) =
+                   Atomic`(KEY`K) \<union> Atomic`(NONCE`N)" 
+
+  (*this protocol makes the assumption of secure means
+    between each agent and his smartcard*)
+   shouprubin_assumes_securemeans [iff]: "evs \<in> sr \<Longrightarrow> secureM"
+
+constdefs
+
+  Unique :: "[event, event list] => bool" ("Unique _ on _")
+   "Unique ev on evs == 
+      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
+
+
+consts  sr   :: "event list set"
+inductive "sr"
+  intros 
+
+    Nil:  "[]\<in> sr"
+
+
+
+    Fake: "\<lbrakk> evsF\<in> sr;  X\<in> synth (analz (knows Spy evsF)); 
+             illegalUse(Card B) \<rbrakk>
+          \<Longrightarrow> Says Spy A X # 
+              Inputs Spy (Card B) X # evsF \<in> sr"
+
+(*In general this rule causes the assumption Card B \<notin> cloned
+  in most guarantees for B - starting with confidentiality -
+  otherwise pairK_confidential could not apply*)
+    Forge:
+         "\<lbrakk> evsFo \<in> sr; Nonce Nb \<in> analz (knows Spy evsFo);
+             Key (pairK(A,B)) \<in> knows Spy evsFo \<rbrakk>
+          \<Longrightarrow> Notes Spy (Key (sesK(Nb,pairK(A,B)))) # evsFo \<in> sr"
+
+
+
+   Reception: "\<lbrakk> evsR\<in> sr; Says A B X \<in> set evsR \<rbrakk>
+              \<Longrightarrow> Gets B X # evsR \<in> sr"
+
+
+
+(*A AND THE SERVER *)
+    SR1:  "\<lbrakk> evs1\<in> sr; A \<noteq> Server\<rbrakk>
+          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> 
+                # evs1 \<in> sr"
+
+    SR2:  "\<lbrakk> evs2\<in> sr; 
+             Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Server A \<lbrace>Nonce (Pairkey(A,B)), 
+                           Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>
+                  \<rbrace>
+                # evs2 \<in> sr"
+
+
+
+
+(*A AND HER CARD*)
+(*A cannot decrypt the verifier for she dosn't know shrK A,
+  but the pairkey is recognisable*)
+    SR3:  "\<lbrakk> evs3\<in> sr; legalUse(Card A);
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+             Gets A \<lbrace>Nonce Pk, Certificate\<rbrace> \<in> set evs3 \<rbrakk>
+          \<Longrightarrow> Inputs A (Card A) (Agent A)
+                # evs3 \<in> sr"   (*however A only queries her card 
+if she has previously contacted the server to initiate with some B. 
+Otherwise she would do so even if the Server had not been active. 
+Still, this doesn't and can't mean that the pairkey originated with 
+the server*)
+ 
+(*The card outputs the nonce Na to A*)               
+    SR4:  "\<lbrakk> evs4\<in> sr;  A \<noteq> Server; 
+             Nonce Na \<notin> used evs4; legalUse(Card A);
+             Inputs A (Card A) (Agent A) \<in> set evs4 \<rbrakk> 
+       \<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+              # evs4 \<in> sr"
+
+(*The card can be exploited by the spy*)
+(*because of the assumptions on the card, A is certainly not server nor spy*)
+ SR4Fake: "\<lbrakk> evs4F\<in> sr; Nonce Na \<notin> used evs4F; 
+             illegalUse(Card A);
+             Inputs Spy (Card A) (Agent A) \<in> set evs4F \<rbrakk> 
+      \<Longrightarrow> Outpts (Card A) Spy \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+            # evs4F \<in> sr"
+
+
+
+
+(*A TOWARDS B*)
+    SR5:  "\<lbrakk> evs5\<in> sr; 
+             Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs5;
+             \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> # evs5 \<in> sr"
+(*A must check that the verifier is not a compound message, 
+  otherwise this would also fire after SR7 *)
+
+
+
+
+(*B AND HIS CARD*)
+    SR6:  "\<lbrakk> evs6\<in> sr; legalUse(Card B);
+             Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs6 \<rbrakk>
+          \<Longrightarrow> Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> 
+                # evs6 \<in> sr"
+
+(*B gets back from the card the session key and various verifiers*)
+    SR7:  "\<lbrakk> evs7\<in> sr; 
+             Nonce Nb \<notin> used evs7; legalUse(Card B); B \<noteq> Server;
+             K = sesK(Nb,pairK(A,B));
+             Key K \<notin> used evs7;
+             Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7\<rbrakk>
+    \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Key K,
+                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                            Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
+                # evs7 \<in> sr"
+
+ (*The card can be exploited by the spy*)
+(*because of the assumptions on the card, A is certainly not server nor spy*)
+ SR7Fake:  "\<lbrakk> evs7F\<in> sr; Nonce Nb \<notin> used evs7F; 
+             illegalUse(Card B);
+             K = sesK(Nb,pairK(A,B));
+             Key K \<notin> used evs7F;
+             Inputs Spy (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7F \<rbrakk>
+          \<Longrightarrow> Outpts (Card B) Spy \<lbrace>Nonce Nb, Key K,
+                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                            Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
+                # evs7F \<in> sr"
+
+
+
+
+(*B TOWARDS A*)
+(*having sent an input that mentions A is the only memory B relies on,
+  since the output doesn't mention A - lack of explicitness*) 
+    SR8:  "\<lbrakk> evs8\<in> sr;  
+             Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs8;
+             Outpts (Card B) B \<lbrace>Nonce Nb, Key K, 
+                                 Cert1, Cert2\<rbrace> \<in> set evs8 \<rbrakk>
+          \<Longrightarrow> Says B A \<lbrace>Nonce Nb, Cert1\<rbrace> # evs8 \<in> sr"
+  
+
+
+
+(*A AND HER CARD*)
+(*A cannot check the form of the verifiers - although I can prove the form of
+  Cert2 - and just feeds her card with what she's got*)
+    SR9:  "\<lbrakk> evs9\<in> sr; legalUse(Card A);
+             Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs9;
+             Outpts (Card A) A \<lbrace>Nonce Na, Cert2\<rbrace> \<in> set evs9; 
+             Gets A \<lbrace>Nonce Nb, Cert3\<rbrace> \<in> set evs9;
+             \<forall> p q. Cert2 \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
+          \<Longrightarrow> Inputs A (Card A) 
+                 \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,
+                  Cert1, Cert3, Cert2\<rbrace> 
+                # evs9 \<in> sr"
+
+(*But the card will only give outputs to the inputs of the correct form*)
+    SR10: "\<lbrakk> evs10\<in> sr; legalUse(Card A); A \<noteq> Server;
+             K = sesK(Nb,pairK(A,B));
+             Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, 
+                                 Nonce (Pairkey(A,B)),
+                                 Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), 
+                                                   Agent B\<rbrace>,
+                                 Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                                 Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+               \<in> set evs10 \<rbrakk>
+          \<Longrightarrow> Outpts (Card A) A \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
+                 # evs10 \<in> sr"
+
+(*The card can be exploited by the spy*)
+(*because of the assumptions on the card, A is certainly not server nor spy*)
+SR10Fake: "\<lbrakk> evs10F\<in> sr; 
+             illegalUse(Card A);
+             K = sesK(Nb,pairK(A,B));
+             Inputs Spy (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, 
+                                   Nonce (Pairkey(A,B)),
+                                   Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), 
+                                                    Agent B\<rbrace>,
+                                   Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                                   Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+               \<in> set evs10F \<rbrakk>
+          \<Longrightarrow> Outpts (Card A) Spy \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
+                 # evs10F \<in> sr"
+
+
+
+
+(*A TOWARDS B*)
+(*having initiated with B is the only memory A relies on,
+  since the output doesn't mention B - lack of explicitness*) 
+    SR11: "\<lbrakk> evs11\<in> sr;
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs11;
+             Outpts (Card A) A \<lbrace>Key K, Certificate\<rbrace> \<in> set evs11 \<rbrakk>
+          \<Longrightarrow> Says A B (Certificate) 
+                 # evs11 \<in> sr"
+
+
+
+    (*Both peers may leak by accident the session keys obtained from their
+      cards*)
+    Oops1:
+     "\<lbrakk> evsO1 \<in> sr;
+         Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Certificate, 
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evsO1 \<rbrakk>
+     \<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO1 \<in> sr"
+
+    Oops2:
+     "\<lbrakk> evsO2 \<in> sr;
+         Outpts (Card A) A \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
+           \<in> set evsO2 \<rbrakk>
+    \<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO2 \<in> sr"
+
+
+
+
+
+
+(*To solve Fake case when it doesn't involve analz - used to be condensed
+  into Fake_parts_insert_tac*)
+declare Fake_parts_insert_in_Un  [dest]
+declare analz_into_parts [dest]
+(*declare parts_insertI [intro]*)
+
+
+
+(*General facts about message reception*)
+lemma Gets_imp_Says: 
+       "\<lbrakk> Gets B X \<in> set evs; evs \<in> sr \<rbrakk> \<Longrightarrow> \<exists> A. Says A B X \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Gets_imp_knows_Spy: 
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> sr \<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+done
+
+lemma Gets_imp_knows_Spy_parts_Snd: 
+     "\<lbrakk> Gets B \<lbrace>X, Y\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  \<Longrightarrow> Y \<in> parts (knows Spy evs)"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy parts.Inj parts.Snd)
+done
+
+lemma Gets_imp_knows_Spy_analz_Snd: 
+     "\<lbrakk> Gets B \<lbrace>X, Y\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  \<Longrightarrow> Y \<in> analz (knows Spy evs)"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy analz.Inj analz.Snd)
+done
+
+(*end general facts*)
+
+
+
+(*Begin lemmas on secure means, from Event.ML, proved for shouprubin. They help
+  the simplifier, especially in analz_image_freshK*)
+
+
+lemma Inputs_imp_knows_Spy_secureM_sr: 
+      "\<lbrakk> Inputs Spy C X \<in> set evs; evs \<in> sr \<rbrakk> \<Longrightarrow> X \<in> knows Spy evs"
+apply (simp (no_asm_simp) add: Inputs_imp_knows_Spy_secureM)
+done
+
+lemma knows_Spy_Inputs_secureM_sr_Spy: 
+      "evs \<in>sr \<Longrightarrow> knows Spy (Inputs Spy C X # evs) = insert X (knows Spy evs)"
+apply (simp (no_asm_simp))
+done
+
+lemma knows_Spy_Inputs_secureM_sr: 
+    "\<lbrakk> A \<noteq> Spy; evs \<in>sr \<rbrakk> \<Longrightarrow> knows Spy (Inputs A C X # evs) =  knows Spy evs"
+apply (simp (no_asm_simp))
+done
+
+lemma knows_Spy_Outpts_secureM_sr_Spy: 
+      "evs \<in>sr \<Longrightarrow> knows Spy (Outpts C Spy X # evs) = insert X (knows Spy evs)"
+apply (simp (no_asm_simp))
+done
+
+lemma knows_Spy_Outpts_secureM_sr: 
+     "\<lbrakk> A \<noteq> Spy; evs \<in>sr \<rbrakk> \<Longrightarrow> knows Spy (Outpts C A X # evs) =  knows Spy evs"
+apply (simp (no_asm_simp))
+done
+
+(*End lemmas on secure means for shouprubin*)
+
+
+
+
+(*BEGIN technical lemmas - evolution of forwarding lemmas*)
+
+(*If an honest agent uses a smart card, then the card is his/her own, is
+  not stolen, and the agent has received suitable data to feed the card. 
+  In other words, these are guarantees that an honest agent can only use 
+  his/her own card, and must use it correctly.
+  On the contrary, the spy can "Inputs" any cloned cards also by the Fake rule.
+
+  Instead of Auto_tac, proofs here used to asm-simplify and then force-tac.
+*)
+lemma Inputs_A_Card_3: 
+    "\<lbrakk> Inputs A C (Agent A) \<in> set evs; A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+      (\<exists> Pk Certificate. Gets A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Inputs_B_Card_6: 
+     "\<lbrakk> Inputs B C \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs; B \<noteq> Spy; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and> Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Inputs_A_Card_9: 
+     "\<lbrakk> Inputs A C \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,   
+                                           Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+  \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+      Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs     \<and>  
+      Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace> \<in> set evs        \<and>   
+      Gets A \<lbrace>Nonce Nb, Cert2\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+
+(*The two occurrences of A in the Outpts event don't match SR4Fake, where
+  A cannot be the Spy. Hence the card is legally usable by rule SR4*)
+lemma Outpts_A_Card_4: 
+     "\<lbrakk> Outpts C A \<lbrace>Nonce Na, (Crypt (crdK (Card A)) (Nonce Na))\<rbrace> \<in> set evs;  
+         evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         Inputs A (Card A) (Agent A) \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+
+(*First certificate is made explicit so that a comment similar to the previous
+  applies. This also provides Na to the Inputs event in the conclusion*)
+lemma Outpts_B_Card_7: 
+      "\<lbrakk> Outpts C B \<lbrace>Nonce Nb, Key K,  
+                      Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                      Cert2\<rbrace> \<in> set evs;  
+         evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and>  
+         Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Outpts_A_Card_10: 
+     "\<lbrakk> Outpts C A \<lbrace>Key K, (Crypt (pairK(A,B)) (Nonce Nb))\<rbrace> \<in> set evs; 
+         evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         (\<exists> Na Ver1 Ver2 Ver3.  
+       Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),  
+                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+
+
+(*
+A can't check the form of the certificate, and so cannot associate the sesion 
+key to the other peer! This already shows that the protocol fails to satisfy 
+the principle of goal availability for the goal of key association.
+Similar reasoning below for the goal of confidentiality will be even more
+accessible.
+*)
+lemma Outpts_A_Card_10_imp_Inputs: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Key K, Certificate\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> (\<exists> B Na Nb Ver1 Ver2 Ver3.  
+       Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),  
+                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule sr.induct)
+apply simp_all
+apply blast+
+done
+
+
+
+
+(*Weaker version: if the agent can't check the forms of the verifiers, then
+  the agent must not be the spy so as to solve SR4Fake. The verifier must be
+  recognised as some cyphertex in order to distinguish from case SR7, 
+  concerning B's output, which also begins with a nonce.
+*)
+lemma Outpts_honest_A_Card_4: 
+     "\<lbrakk> Outpts C A \<lbrace>Nonce Na, Crypt K X\<rbrace> \<in>set evs; 
+         A \<noteq> Spy;  evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         Inputs A (Card A) (Agent A) \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+(*alternative formulation of same theorem
+Goal "\<lbrakk> Outpts C A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs;  
+         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>;    
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         Inputs A (Card A) (Agent A) \<in> set evs"
+same proof
+*)
+
+
+lemma Outpts_honest_B_Card_7: 
+     "\<lbrakk> Outpts C B \<lbrace>Nonce Nb, Key K, Cert1, Cert2\<rbrace> \<in> set evs;  
+         B \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and>  
+         (\<exists> A Na. Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Outpts_honest_A_Card_10: 
+     "\<lbrakk> Outpts C A \<lbrace>Key K, Certificate\<rbrace> \<in> set evs;  
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> legalUse (C) \<and> C = (Card A) \<and>  
+         (\<exists> B Na Nb Pk Ver1 Ver2 Ver3.  
+          Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Pk,  
+                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule sr.induct)
+apply simp_all
+apply blast+
+done
+(*-END-*)
+
+
+(*Even weaker versions: if the agent can't check the forms of the verifiers
+  and the agent may be the spy, then we must know what card the agent
+  is getting the output from. 
+*)
+lemma Outpts_which_Card_4: 
+    "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt K X\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+    \<Longrightarrow> Inputs A (Card A) (Agent A) \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply (simp_all (no_asm_simp))
+apply clarify
+done
+
+lemma Outpts_which_Card_7: 
+  "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Cert1, Cert2\<rbrace> \<in> set evs;  
+         evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> \<exists> A Na. Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Outpts_which_Card_10: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Key (sesK(Nb,pairK(A,B))),  
+                             Crypt (pairK(A,B)) (Nonce Nb) \<rbrace> \<in> set evs;  
+         evs \<in> sr \<rbrakk>  
+    \<Longrightarrow> \<exists> Na. Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)), 
+                            Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>,  
+                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                            Crypt (crdK (Card A)) (Nonce Na) \<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+
+(*Lemmas on the form of outputs*)
+
+
+(*A needs to check that the verifier is a cipher for it to come from SR4
+  otherwise it could come from SR7 *)
+lemma Outpts_A_Card_form_4: 
+  "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs;  
+         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> Certificate = (Crypt (crdK (Card A)) (Nonce Na))"
+apply (erule rev_mp, erule sr.induct)
+apply (simp_all (no_asm_simp))
+done
+
+lemma Outpts_B_Card_form_7: 
+   "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Cert1, Cert2\<rbrace> \<in> set evs; 
+         evs \<in> sr \<rbrakk>          
+      \<Longrightarrow> \<exists> A Na.    
+          K = sesK(Nb,pairK(A,B)) \<and>                       
+          Cert1 = (Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>) \<and>  
+          Cert2 = (Crypt (pairK(A,B)) (Nonce Nb))"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Outpts_A_Card_form_10: 
+   "\<lbrakk> Outpts (Card A) A \<lbrace>Key K, Certificate\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> \<exists> B Nb.  
+          K = sesK(Nb,pairK(A,B)) \<and>  
+          Certificate = (Crypt (pairK(A,B)) (Nonce Nb))"
+apply (erule rev_mp, erule sr.induct)
+apply (simp_all (no_asm_simp))
+done
+
+lemma Outpts_A_Card_form_bis: 
+  "\<lbrakk> Outpts (Card A') A' \<lbrace>Key (sesK(Nb,pairK(A,B))), Certificate\<rbrace> \<in> set evs; 
+         evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> A' = A \<and>  
+          Certificate = (Crypt (pairK(A,B)) (Nonce Nb))"
+apply (erule rev_mp, erule sr.induct)
+apply (simp_all (no_asm_simp))
+done
+
+(*\<dots> and Inputs *)
+
+lemma Inputs_A_Card_form_9: 
+     "\<lbrakk> Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,   
+                             Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
+         evs \<in> sr \<rbrakk>  
+  \<Longrightarrow>    Cert3 = Crypt (crdK (Card A)) (Nonce Na)"
+apply (erule rev_mp)
+apply (erule sr.induct)
+apply (simp_all (no_asm_simp))
+(*Fake*)
+apply force
+(*SR9*)
+apply (blast dest!: Outpts_A_Card_form_4)
+done
+(* Pk, Cert1, Cert2 cannot be made explicit because they traversed the network in the clear *)
+
+(*General guarantees on Inputs and Outpts*)
+
+(*for any agents*)
+
+
+lemma Inputs_Card_legalUse: 
+  "\<lbrakk> Inputs A (Card A) X \<in> set evs; evs \<in> sr \<rbrakk> \<Longrightarrow> legalUse(Card A)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Outpts_Card_legalUse: 
+  "\<lbrakk> Outpts (Card A) A X \<in> set evs; evs \<in> sr \<rbrakk> \<Longrightarrow> legalUse(Card A)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+(*for honest agents*)
+
+lemma Inputs_Card: "\<lbrakk> Inputs A C X \<in> set evs; A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> C = (Card A) \<and> legalUse(C)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Outpts_Card: "\<lbrakk> Outpts C A X \<in> set evs; A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> C = (Card A) \<and> legalUse(C)"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+lemma Inputs_Outpts_Card: 
+     "\<lbrakk> Inputs A C X \<in> set evs \<or> Outpts C A Y \<in> set evs;  
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> C = (Card A) \<and> legalUse(Card A)"
+apply (blast dest: Inputs_Card Outpts_Card)
+done
+
+
+(*for the spy - they stress that the model behaves as it is meant to*) 
+
+(*The or version can be also proved directly.
+  It stresses that the spy may use either her own legally usable card or
+  all the illegally usable cards.
+*)
+lemma Inputs_Card_Spy: 
+  "\<lbrakk> Inputs Spy C X \<in> set evs \<or> Outpts C Spy X \<in> set evs; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> C = (Card Spy) \<and> legalUse(Card Spy) \<or>  
+          (\<exists> A. C = (Card A) \<and> illegalUse(Card A))"
+apply (erule rev_mp, erule sr.induct)
+apply auto
+done
+
+
+(*END technical lemmas*)
+
+
+
+
+
+
+(*BEGIN unicity theorems: certain items uniquely identify a smart card's
+                          output*)
+
+(*A's card's first output: the nonce uniquely identifies the rest*)
+lemma Outpts_A_Card_unique_nonce:
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>  
+           \<in> set evs;   
+         Outpts (Card A') A' \<lbrace>Nonce Na, Crypt (crdK (Card A')) (Nonce Na)\<rbrace> 
+           \<in> set evs;   
+         evs \<in> sr \<rbrakk> \<Longrightarrow> A=A'"
+apply (erule rev_mp, erule rev_mp, erule sr.induct, simp_all)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+(*B's card's output: the NONCE uniquely identifies the rest*)
+lemma Outpts_B_Card_unique_nonce: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key SK, Cert1, Cert2\<rbrace> \<in> set evs;   
+         Outpts (Card B') B' \<lbrace>Nonce Nb, Key SK', Cert1', Cert2'\<rbrace> \<in> set evs;   
+         evs \<in> sr \<rbrakk> \<Longrightarrow> B=B' \<and> SK=SK' \<and> Cert1=Cert1' \<and> Cert2=Cert2'"
+apply (erule rev_mp, erule rev_mp, erule sr.induct, simp_all)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+
+(*B's card's output: the SESKEY uniquely identifies the rest*)
+lemma Outpts_B_Card_unique_key: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key SK, Cert1, Cert2\<rbrace> \<in> set evs;   
+         Outpts (Card B') B' \<lbrace>Nonce Nb', Key SK, Cert1', Cert2'\<rbrace> \<in> set evs;   
+         evs \<in> sr \<rbrakk> \<Longrightarrow> B=B' \<and> Nb=Nb' \<and> Cert1=Cert1' \<and> Cert2=Cert2'"
+apply (erule rev_mp, erule rev_mp, erule sr.induct, simp_all)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+lemma Outpts_A_Card_unique_key: "\<lbrakk> Outpts (Card A) A \<lbrace>Key K, V\<rbrace> \<in> set evs;   
+         Outpts (Card A') A' \<lbrace>Key K, V'\<rbrace> \<in> set evs;   
+         evs \<in> sr \<rbrakk> \<Longrightarrow> A=A' \<and> V=V'"
+apply (erule rev_mp, erule rev_mp, erule sr.induct, simp_all)
+apply (blast dest: Outpts_A_Card_form_bis)
+apply blast
+done
+
+
+(*Revised unicity theorems - applies to both steps 4 and 7*)
+lemma Outpts_A_Card_Unique: 
+  "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, rest\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> Unique (Outpts (Card A) A \<lbrace>Nonce Na, rest\<rbrace>) on evs"
+apply (erule rev_mp, erule sr.induct, simp_all add: Unique_def)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+(*can't prove the same on evs10 for it doesn't have a freshness assumption!*)
+
+
+(*END unicity theorems*)
+
+
+(*BEGIN counterguarantees about spy's knowledge*)
+
+(*on nonces*)
+
+lemma Spy_knows_Na: 
+      "\<lbrakk> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> Nonce Na \<in> analz (knows Spy evs)"
+apply (blast dest!: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd])
+done
+
+lemma Spy_knows_Nb: 
+      "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> Nonce Nb \<in> analz (knows Spy evs)"
+apply (blast dest!: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Fst])
+done
+
+
+(*on Pairkey*)
+
+lemma Pairkey_Gets_analz_knows_Spy: 
+      "\<lbrakk> Gets A \<lbrace>Nonce (Pairkey(A,B)), Certificate\<rbrace> \<in> set evs; evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> Nonce (Pairkey(A,B)) \<in> analz (knows Spy evs)"
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
+done
+
+lemma Pairkey_Inputs_imp_Gets: 
+     "\<lbrakk> Inputs A (Card A)             
+           \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),     
+             Cert1, Cert3, Cert2\<rbrace> \<in> set evs;           
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>     
+      \<Longrightarrow> Gets A \<lbrace>Nonce (Pairkey(A,B)), Cert1\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply (simp_all (no_asm_simp))
+apply force
+done
+
+lemma Pairkey_Inputs_analz_knows_Spy: 
+     "\<lbrakk> Inputs A (Card A)             
+           \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),     
+             Cert1, Cert3, Cert2\<rbrace> \<in> set evs;           
+         evs \<in> sr \<rbrakk>     
+     \<Longrightarrow> Nonce (Pairkey(A,B)) \<in> analz (knows Spy evs)"
+apply (case_tac "A = Spy")
+apply (fastsimp dest!: Inputs_imp_knows_Spy_secureM [THEN analz.Inj])
+apply (blast dest!: Pairkey_Inputs_imp_Gets [THEN Pairkey_Gets_analz_knows_Spy])
+done
+
+(* This fails on base case because of XOR properties.
+lemma Pairkey_authentic:
+  "\<lbrakk> Nonce (Pairkey(A,B)) \<in> parts (knows Spy evs);
+     Card A \<notin> cloned; evs \<in> sr \<rbrakk>
+ \<Longrightarrow> \<exists> cert. Says Server A \<lbrace>Nonce (Pairkey(A,B)), Cert\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule sr.induct, simp_all)
+apply clarify
+oops
+
+ 1. \<And>x a b.
+       \<lbrakk>Card A \<notin> cloned; Pairkey (A, B) = Pairkey (a, b); Card a \<in> cloned;
+        Card b \<in> cloned\<rbrakk>
+       \<Longrightarrow> False
+*)
+
+(*END counterguarantees on spy's knowledge*)
+
+
+(*BEGIN rewrite rules for parts operator*)
+
+
+declare shrK_disj_sesK [THEN not_sym, iff] 
+declare pin_disj_sesK [THEN not_sym, iff]
+declare crdK_disj_sesK [THEN not_sym, iff]
+declare pairK_disj_sesK [THEN not_sym, iff]
+
+
+ML
+{*
+val Outpts_B_Card_form_7 = thm "Outpts_B_Card_form_7"
+val Outpts_A_Card_form_4 = thm "Outpts_A_Card_form_4"
+val Outpts_A_Card_form_10 = thm "Outpts_A_Card_form_10"
+val Gets_imp_knows_Spy = thm "Gets_imp_knows_Spy"
+val Outpts_B_Card_form_7 = thm "Outpts_B_Card_form_7"
+val Gets_imp_knows_Spy_parts_Snd = thm "Gets_imp_knows_Spy_parts_Snd"
+val Gets_imp_knows_Spy_analz_Snd = thm "Gets_imp_knows_Spy_analz_Snd"
+
+val prepare_tac = 
+ (*SR8*)   forward_tac [Outpts_B_Card_form_7] 14 THEN
+           eresolve_tac [exE] 15 THEN eresolve_tac [exE] 15 THEN 
+ (*SR9*)   forward_tac [Outpts_A_Card_form_4] 16 THEN 
+ (*SR11*)  forward_tac [Outpts_A_Card_form_10] 21 THEN
+           eresolve_tac [exE] 22 THEN eresolve_tac [exE] 22
+
+val parts_prepare_tac = 
+           prepare_tac THEN
+ (*SR9*)   dresolve_tac [Gets_imp_knows_Spy_parts_Snd] 18 THEN 
+ (*SR9*)   dresolve_tac [Gets_imp_knows_Spy_parts_Snd] 19 THEN 
+ (*Oops1*) dresolve_tac [Outpts_B_Card_form_7] 25    THEN               
+ (*Oops2*) dresolve_tac [Outpts_A_Card_form_10] 27 THEN                
+ (*Base*)  Force_tac 1
+
+val analz_prepare_tac = 
+         prepare_tac THEN
+         dtac (Gets_imp_knows_Spy_analz_Snd) 18 THEN 
+ (*SR9*) dtac (Gets_imp_knows_Spy_analz_Snd) 19 THEN 
+         REPEAT_FIRST (eresolve_tac [asm_rl, conjE] ORELSE' hyp_subst_tac)
+
+*}
+
+method_setup prepare = {*
+    Method.no_args (Method.METHOD (fn facts => prepare_tac)) *}
+  "to launch a few simple facts that'll help the simplifier"
+
+method_setup parts_prepare = {*
+    Method.no_args (Method.METHOD (fn facts => parts_prepare_tac)) *}
+  "additional facts to reason about parts"
+
+method_setup analz_prepare = {*
+    Method.no_args (Method.METHOD (fn facts => analz_prepare_tac)) *}
+  "additional facts to reason about analz"
+
+
+(*Treatment of pins is here for completeness. This protocol doesn't use pins*)
+lemma Spy_parts_keys [simp]: "evs \<in> sr \<Longrightarrow>  
+  (Key (shrK P) \<in> parts (knows Spy evs)) = (Card P \<in> cloned) \<and>  
+  (Key (pin P) \<in> parts (knows Spy evs)) = (P \<in> bad \<or> Card P \<in> cloned) \<and>  
+  (Key (crdK C) \<in> parts (knows Spy evs)) = (C \<in> cloned) \<and>  
+  (Key (pairK(A,B)) \<in> parts (knows Spy evs)) = (Card B \<in> cloned)"
+apply (erule sr.induct)
+apply parts_prepare
+apply simp_all
+apply (blast intro: parts_insertI)
+done
+
+
+(*END rewrite rules for parts operator*)
+
+(*BEGIN rewrite rules for analz operator*)
+
+lemma Spy_analz_shrK[simp]: "evs \<in> sr \<Longrightarrow>  
+  (Key (shrK P) \<in> analz (knows Spy evs)) = (Card P \<in> cloned)" 
+apply (auto dest!: Spy_knows_cloned)
+done
+
+lemma Spy_analz_crdK[simp]: "evs \<in> sr \<Longrightarrow>  
+  (Key (crdK C) \<in> analz (knows Spy evs)) = (C \<in> cloned)"
+apply (auto dest!: Spy_knows_cloned)
+done
+
+lemma Spy_analz_pairK[simp]: "evs \<in> sr \<Longrightarrow>  
+  (Key (pairK(A,B)) \<in> analz (knows Spy evs)) = (Card B \<in> cloned)"
+apply (auto dest!: Spy_knows_cloned)
+done
+
+
+
+(*Because initState contains a set of nonces, this is needed for base case of
+  analz_image_freshK*)
+lemma analz_image_Key_Un_Nonce: "analz (Key`K \<union> Nonce`N) = Key`K \<union> Nonce`N"
+apply auto
+done
+
+ML
+{*
+val knows_Spy_Inputs_secureM_sr_Spy = thm "knows_Spy_Inputs_secureM_sr_Spy"
+val knows_Spy_Outpts_secureM_sr_Spy = thm "knows_Spy_Outpts_secureM_sr_Spy"
+val shouprubin_assumes_securemeans = thm "shouprubin_assumes_securemeans"
+val analz_image_Key_Un_Nonce= thm "analz_image_Key_Un_Nonce"
+*}
+
+method_setup sc_analz_freshK = {*
+    Method.no_args
+     (Method.METHOD
+      (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+                          REPEAT_FIRST (rtac analz_image_freshK_lemma),
+                          ALLGOALS (asm_simp_tac (analz_image_freshK_ss
+                                    addsimps [knows_Spy_Inputs_secureM_sr_Spy,
+                                              knows_Spy_Outpts_secureM_sr_Spy,
+                                              shouprubin_assumes_securemeans, 
+                                              analz_image_Key_Un_Nonce]))])) *}
+    "for proving the Session Key Compromise theorem for smartcard protocols"
+
+
+lemma analz_image_freshK [rule_format]: 
+     "evs \<in> sr \<Longrightarrow>      \<forall> K KK.  
+          (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) =  
+          (K \<in> KK \<or> Key K \<in> analz (knows Spy evs))"
+apply (erule sr.induct)
+apply analz_prepare
+apply sc_analz_freshK
+apply spy_analz
+done
+
+
+lemma analz_insert_freshK: "evs \<in> sr \<Longrightarrow>   
+         Key K \<in> analz (insert (Key K') (knows Spy evs)) =  
+         (K = K' \<or> Key K \<in> analz (knows Spy evs))"
+apply (simp only: analz_image_freshK_simps analz_image_freshK)
+done
+
+(*END rewrite rules for analz operator*)
+
+(*BEGIN authenticity theorems*)
+
+
+
+
+(*Card B \<notin> cloned needed for Fake
+  B \<notin> bad needed for SR7Fake; equivalent to Card B \<notin> stolen
+*)
+
+lemma Na_Nb_certificate_authentic: 
+     "\<lbrakk> Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace> \<in> parts (knows Spy evs);  
+         \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>           
+     \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Key (sesK(Nb,pairK(A,B))),   
+                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply parts_prepare
+apply simp_all
+(*Fake*)
+apply spy_analz
+(*SR7F*)
+apply clarify
+done
+
+(* Card B \<notin> cloned needed for Fake and SR7F
+   B \<noteq> Spy needed for SR7
+   B \<notin> bad - or Card B \<notin> stolen - needed for SR7F
+   Card A \<notin> cloned needed for SR10F
+   A \<notin> bad - or Card A \<notin> stolen - needed for SR10F
+
+   Non-trivial case done by the simplifier.*)
+lemma Nb_certificate_authentic: 
+      "\<lbrakk> Crypt (pairK(A,B)) (Nonce Nb) \<in> parts (knows Spy evs);  
+         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>    
+     \<Longrightarrow> Outpts (Card A) A \<lbrace>Key (sesK(Nb,pairK(A,B))),  
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply parts_prepare
+apply (case_tac [17] "Aa = Spy")
+apply simp_all
+(*Fake*)
+apply spy_analz
+(*SR7F, SR10F*)
+apply clarify+
+done
+
+
+
+(*Discovering the very origin of the Nb certificate... non needed!*)
+(*lemma*)
+lemma Outpts_A_Card_imp_pairK_parts: 
+     "\<lbrakk> Outpts (Card A) A            
+         \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs;  
+         evs \<in> sr \<rbrakk>   
+     \<Longrightarrow> \<exists> Na. Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace> \<in> parts (knows Spy evs)"
+apply (erule rev_mp, erule sr.induct)
+apply parts_prepare
+apply simp_all
+(*Fake*)
+apply (blast dest: parts_insertI)
+(*SR7*)
+apply force
+(*SR7F*)
+apply force
+(*SR8*)
+apply blast
+(*SR10*)
+apply (blast dest: Inputs_imp_knows_Spy_secureM_sr parts.Inj Inputs_A_Card_9 Gets_imp_knows_Spy elim: knows_Spy_partsEs)
+(*SR10F*)
+apply (blast dest: Inputs_imp_knows_Spy_secureM_sr [THEN parts.Inj] 
+                   Inputs_A_Card_9 Gets_imp_knows_Spy 
+             elim: knows_Spy_partsEs)
+done
+
+
+
+lemma Nb_certificate_authentic_bis: 
+     "\<lbrakk> Crypt (pairK(A,B)) (Nonce Nb) \<in> parts (knows Spy evs);  
+         B \<noteq> Spy; \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>    
+     \<Longrightarrow> \<exists> Na. Outpts (Card B) B \<lbrace>Nonce Nb, Key (sesK(Nb,pairK(A,B))),   
+                   Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                   Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply parts_prepare
+apply (simp_all (no_asm_simp))
+(*Fake*)
+apply spy_analz
+(*SR7*)
+apply blast
+(*SR7F*)
+apply blast
+(*SR10*)
+apply (blast dest: Na_Nb_certificate_authentic Inputs_imp_knows_Spy_secureM_sr [THEN parts.Inj] elim: knows_Spy_partsEs)
+(*SR10F*)
+apply (blast dest: Na_Nb_certificate_authentic Inputs_imp_knows_Spy_secureM_sr [THEN parts.Inj] elim: knows_Spy_partsEs)
+(*SR11*)
+apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_imp_pairK_parts)
+done
+
+
+lemma Pairkey_certificate_authentic: 
+    "\<lbrakk> Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace> \<in> parts (knows Spy evs);    
+         Card A \<notin> cloned; evs \<in> sr \<rbrakk>        
+     \<Longrightarrow> Pk = Pairkey(A,B) \<and>              
+         Says Server A \<lbrace>Nonce Pk,  
+                        Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>\<rbrace> 
+           \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply parts_prepare
+apply (simp_all (no_asm_simp))
+(*Fake*)
+apply spy_analz
+done
+
+
+(*Alternatively:  A \<notin> bad; Card A \<notin> cloned; B \<notin> bad; Card B \<notin> cloned;*)
+lemma sesK_authentic: 
+     "\<lbrakk> Key (sesK(Nb,pairK(A,B))) \<in> parts (knows Spy evs);  
+         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>           
+      \<Longrightarrow> Notes Spy \<lbrace>Key (sesK(Nb,pairK(A,B))), Nonce Nb, Agent A, Agent B\<rbrace>  
+           \<in> set evs"
+apply (erule rev_mp, erule sr.induct)
+apply parts_prepare
+apply (simp_all (no_asm_simp))
+(*fake*)
+apply spy_analz
+(*forge*)
+apply (fastsimp dest: analz.Inj)
+(*SR7: used B\<noteq>Spy*)
+(*SR7F*)
+apply clarify
+(*SR10: used A\<noteq>Spy*)
+(*SR10F*)
+apply clarify
+(*Oops*)
+apply simp_all
+done
+
+
+(*END authenticity theorems*)
+
+
+(*BEGIN confidentiality theorems*)
+
+(*If B were bad and his card stolen, they spy could use B's card but would 
+  not obtain this K because B's card only issues new session keys out
+  of new nonces. 
+  If A were bad, then her card could be stolen, hence the spy could feed it
+  with Nb and get this K. Thus, A\<notin>bad can be replaced by Card A \<notin> stolen
+  Hence these are the minimal assumptions:
+        A \<notin> bad; B \<noteq> Spy; Card A \<notin> cloned; Card B \<notin> cloned; 
+         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); Card B \<notin> cloned;
+*)
+
+lemma Confidentiality: 
+     "\<lbrakk> Notes Spy \<lbrace>Key (sesK(Nb,pairK(A,B))), Nonce Nb, Agent A, Agent B\<rbrace>  
+           \<notin> set evs; 
+        A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+        evs \<in> sr \<rbrakk>           
+      \<Longrightarrow> Key (sesK(Nb,pairK(A,B))) \<notin> analz (knows Spy evs)"
+apply (blast intro: sesK_authentic)
+done
+
+lemma Confidentiality_B: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Certificate,          
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs;  
+         Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> \<notin> set evs;  
+         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); Card B \<notin> cloned; 
+         evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+apply (erule rev_mp, erule rev_mp, erule sr.induct)
+apply analz_prepare
+apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs)
+(*Fake*)
+apply spy_analz
+(*Forge*)
+apply (rotate_tac 7)
+apply (drule parts.Inj)
+apply (fastsimp dest: Outpts_B_Card_form_7)
+(*SR7*)
+apply (blast dest!: Outpts_B_Card_form_7)
+(*SR7F*)
+apply clarify
+apply (drule Outpts_parts_used)
+apply simp
+(*faster than
+  by (fast_tac (claset() addDs [Outpts_parts_used] addss (simpset())) 1)
+*)
+(*SR10*)
+apply (fastsimp dest: Outpts_B_Card_form_7)
+(*SR10F - uses assumption Card A not cloned*)
+apply clarify
+apply (drule Outpts_B_Card_form_7, assumption)
+apply simp
+(*Oops1*)
+apply (blast dest!: Outpts_B_Card_form_7)
+(*Oops2*)
+apply (blast dest!: Outpts_B_Card_form_7 Outpts_A_Card_form_10)
+done
+
+(*Confidentiality_A can be is faster to prove in forward style, using
+the authentication theorems. So it is moved below*)
+
+
+(*END confidentiality theorems*)
+
+
+
+(*BEGIN authentication theorems*)
+
+lemma A_authenticates_B: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs;  
+         \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>           
+     \<Longrightarrow> \<exists> Na. 
+            Outpts (Card B) B \<lbrace>Nonce Nb, Key K,   
+                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_form_10 Outpts_A_Card_imp_pairK_parts)
+done
+
+lemma A_authenticates_B_Gets: 
+     "\<lbrakk> Gets A \<lbrace>Nonce Nb, Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>\<rbrace>  
+           \<in> set evs;  
+         \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>           
+     \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Key (sesK(Nb, pairK (A, B))),   
+                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN Na_Nb_certificate_authentic])
+done
+
+
+
+
+lemma B_authenticates_A: 
+     "\<lbrakk> Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
+         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>  
+      \<Longrightarrow> Outpts (Card A) A            
+       \<lbrace>Key (sesK(Nb,pairK(A,B))), Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule sr.induct)
+apply (simp_all (no_asm_simp))
+apply (blast dest: Says_imp_knows_Spy [THEN parts.Inj] Nb_certificate_authentic)
+done
+
+
+(*END authentication theorems*)
+
+lemma Confidentiality_A: "\<lbrakk> Outpts (Card A) A            
+           \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs;  
+         Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> \<notin> set evs;  
+         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>           
+     \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+apply (drule A_authenticates_B)
+prefer 3
+apply (erule exE)
+apply (drule Confidentiality_B)
+apply auto
+done
+
+lemma Outpts_imp_knows_agents_secureM_sr: 
+   "\<lbrakk> Outpts (Card A) A X \<in> set evs; evs \<in> sr \<rbrakk> \<Longrightarrow> X \<in> knows A evs"
+apply (simp (no_asm_simp) add: Outpts_imp_knows_agents_secureM)
+done
+
+
+(*BEGIN key distribution theorems*)
+
+
+(*Alternatively: B \<notin> bad; Card B \<notin> cloned;*)
+lemma A_keydist_to_B: 
+     "\<lbrakk> Outpts (Card A) A            
+           \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs;  
+         \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>           
+     \<Longrightarrow> Key K \<in> analz (knows B evs)"
+apply (drule A_authenticates_B)
+prefer 3
+apply (erule exE)
+apply (rule Outpts_imp_knows_agents_secureM_sr [THEN analz.Inj, THEN analz.Snd, THEN analz.Fst])
+apply assumption+
+done
+
+
+(*Alternatively: A \<notin> bad; B \<notin> bad; Card A \<notin> cloned; Card B \<notin> cloned;*)
+lemma B_keydist_to_A: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Certificate,          
+                             (Crypt (pairK(A,B)) (Nonce Nb))\<rbrace> \<in> set evs;  
+         Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
+         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> Key K \<in> analz (knows A evs)"
+apply (frule B_authenticates_A)
+apply (drule_tac [5] Outpts_B_Card_form_7)
+apply (rule_tac [6] Outpts_imp_knows_agents_secureM_sr [THEN analz.Inj, THEN analz.Fst])
+prefer 6 apply force
+apply assumption+
+done
+
+(*END key distribution theorems*)
+
+
+
+
+
+
+
+
+(*BEGIN further theorems about authenticity of verifiers
+  (useful to agents and cards).                          *)
+
+(*MSG11
+If B receives the verifier of msg11, then the verifier originated with msg7.
+Alternatively: A \<notin> bad; B \<notin> bad; Card A \<notin> cloned; Card B \<notin> cloned;
+*)
+lemma Nb_certificate_authentic_B: 
+     "\<lbrakk> Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
+        B \<noteq> Spy; \<not>illegalUse(Card B); 
+        evs \<in> sr \<rbrakk>  
+    \<Longrightarrow> \<exists> Na. 
+            Outpts (Card B) B \<lbrace>Nonce Nb, Key (sesK(Nb,pairK(A,B))),   
+                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN Nb_certificate_authentic_bis])
+done
+(*Useless to B: B can't check the form of the verifier because he doesn't know
+  pairK(A,B) *)
+
+(*MSG10
+If A obtains the verifier of msg10, then the verifier originated with msg7:
+A_authenticates_B. It is useful to A, who can check the form of the 
+verifier by application of Outpts_A_Card_form_10.
+*)
+
+(*MSG9
+The first verifier verifies the Pairkey to the card: since it's encrypted
+under Ka, it must come from the server (if A's card is not cloned).
+The second verifier verifies both nonces, since it's encrypted under the
+pairK, it must originate with B's card  (if A and B's cards not cloned).
+The third verifier verifies Na: since it's encrytped under the card's key,
+it originated with the card; so the card does not need to save Na
+in the first place and do a comparison now: it just verifies Na through the
+verifier. Three theorems related to these three statements.
+
+Recall that a card can check the form of the verifiers (can decrypt them),
+while an agent in general cannot, if not provided with a suitable theorem.
+*)
+
+(*Card A can't reckon the pairkey - we need to guarantee its integrity!*)
+lemma Pairkey_certificate_authentic_A_Card: 
+     "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+               Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>,  
+               Cert2, Cert3\<rbrace> \<in> set evs; 
+         A \<noteq> Spy; Card A \<notin> cloned; evs \<in> sr \<rbrakk>   
+     \<Longrightarrow> Pk = Pairkey(A,B) \<and>  
+         Says Server A \<lbrace>Nonce (Pairkey(A,B)),  
+                  Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>\<rbrace>   
+           \<in> set evs "
+apply (blast dest: Inputs_A_Card_9 Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd] Pairkey_certificate_authentic)
+done
+(*the second conjunct of the thesis might be regarded as a form of integrity 
+  in the sense of Neuman-Ts'o*)
+
+lemma Na_Nb_certificate_authentic_A_Card: 
+      "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+               Cert1,  
+               Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, Cert3\<rbrace> \<in> set evs; 
+      A \<noteq> Spy; \<not>illegalUse(Card B); evs \<in> sr \<rbrakk> 
+     \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Key (sesK(Nb, pairK (A, B))),    
+                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
+           \<in> set evs "
+apply (blast dest: Inputs_A_Card_9 Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN Na_Nb_certificate_authentic])
+done
+
+lemma Na_authentic_A_Card: 
+     "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+                Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>   
+     \<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace>  
+           \<in> set evs"
+apply (blast dest: Inputs_A_Card_9)
+done
+
+(* The last three theorems for Card A can be put togheter trivially.
+They are separated to highlight the different requirements on agents
+and their cards.*)
+
+
+(*Alternatively:
+  A \<noteq> Spy; B \<notin> bad; Card A \<notin> cloned; Card B \<notin> cloned; evs \<in> sr \<rbrakk> *)
+lemma Inputs_A_Card_9_authentic: 
+  "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+               Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>,  
+               Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, Cert3\<rbrace> \<in> set evs; 
+    A \<noteq> Spy; Card A \<notin> cloned;\<not>illegalUse(Card B); evs \<in> sr \<rbrakk>   
+    \<Longrightarrow>  Says Server A \<lbrace>Nonce Pk, Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>\<rbrace>   
+           \<in> set evs  \<and> 
+         Outpts (Card B) B \<lbrace>Nonce Nb, Key (sesK(Nb, pairK (A, B))),    
+                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
+           \<in> set evs  \<and> 
+         Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace>  
+           \<in> set evs"
+apply (blast dest: Inputs_A_Card_9 Na_Nb_certificate_authentic Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd] Pairkey_certificate_authentic)
+done
+
+(*MSG8
+Nothing to prove because the message is a cleartext that comes from the 
+network*)
+
+(*Other messages: nothing to prove because the verifiers involved are new*)
+
+
+(*END further theorems about authenticity of verifiers*)
+
+
+
+(* BEGIN trivial guarantees on outputs for agents *)
+
+(*MSG4*)
+lemma SR4_imp: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace> 
+           \<in> set evs;  
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>                 
+     \<Longrightarrow> \<exists> Pk V. Gets A \<lbrace>Pk, V\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_A_Card_4 Inputs_A_Card_3)
+done
+(*weak: could strengthen the model adding verifier for the Pairkey to msg3*)
+
+
+(*MSG7*)
+lemma SR7_imp: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Key K,  
+                      Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                      Cert2\<rbrace> \<in> set evs;  
+         B \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_B_Card_7 Inputs_B_Card_6)
+done
+
+(*MSG10*)
+lemma SR10_imp: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
+           \<in> set evs;  
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> \<exists> Cert1 Cert2.  
+                   Gets A \<lbrace>Nonce (Pairkey (A, B)), Cert1\<rbrace> \<in> set evs \<and>  
+                   Gets A \<lbrace>Nonce Nb, Cert2\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_A_Card_10 Inputs_A_Card_9)
+done
+
+
+(*END trivial guarantees on outputs for agents*)
+
+
+
+(*INTEGRITY*)
+lemma Outpts_Server_not_evs: "evs \<in> sr \<Longrightarrow> Outpts (Card Server) P X \<notin> set evs"
+apply (erule sr.induct)
+apply auto
+done
+
+text{*@{term step2_integrity} also is a reliability theorem*}
+lemma Says_Server_message_form: 
+     "\<lbrakk> Says Server A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs;  
+         evs \<in> sr \<rbrakk>                   
+     \<Longrightarrow> \<exists> B. Pk = Nonce (Pairkey(A,B)) \<and>  
+         Certificate = Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>"
+apply (erule rev_mp)
+apply (erule sr.induct)
+apply auto
+apply (blast dest!: Outpts_Server_not_evs)+
+done
+(*cannot be made useful to A in form of a Gets event*)
+
+text{*
+  step4integrity is @{term Outpts_A_Card_form_4}
+
+  step7integrity is @{term Outpts_B_Card_form_7}
+*}
+
+lemma step8_integrity: 
+     "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs;  
+         B \<noteq> Server; B \<noteq> Spy; evs \<in> sr \<rbrakk>                   
+     \<Longrightarrow> \<exists> Cert2 K.   
+          Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Certificate, Cert2\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule sr.induct)
+prefer 18 apply (fastsimp dest: Outpts_A_Card_form_10)
+apply auto
+done
+
+
+text{*  step9integrity is @{term Inputs_A_Card_form_9}
+
+        step10integrity is @{term Outpts_A_Card_form_10}.
+*}
+
+lemma step11_integrity: 
+     "\<lbrakk> Says A B (Certificate) \<in> set evs; 
+         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>;  
+         A \<noteq> Spy; evs \<in> sr \<rbrakk>  
+     \<Longrightarrow> \<exists> K.  
+            Outpts (Card A) A \<lbrace>Key K, Certificate\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule sr.induct)
+apply auto
+done
+
+
+
+
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,1390 @@
+(*  ID:         $Id$
+    Author:     Giampaolo Bella, Catania University
+*)
+
+header{*Bella's modification of the Shoup-Rubin protocol*}
+
+theory ShoupRubinBella imports Smartcard begin
+
+text{*The modifications are that message 7 now mentions A, while message 10
+now mentions Nb and B. The lack of explicitness of the original version was
+discovered by investigating adherence to the principle of Goal
+Availability. Only the updated version makes the goals of confidentiality,
+authentication and key distribution available to both peers.*}
+
+consts
+
+    sesK :: "nat*key => key"
+
+axioms
+     
+   (*sesK is injective on each component*) 
+   inj_sesK [iff]: "(sesK(m,k) = sesK(m',k')) = (m = m' \<and> k = k')"
+
+   (*all long-term keys differ from sesK*)
+   shrK_disj_sesK [iff]: "shrK A \<noteq> sesK(m,pk)"
+   crdK_disj_sesK [iff]: "crdK C \<noteq> sesK(m,pk)"
+   pin_disj_sesK  [iff]: "pin P \<noteq> sesK(m,pk)"
+   pairK_disj_sesK[iff]: "pairK(A,B) \<noteq> sesK(m,pk)"
+
+   (*needed for base case in analz_image_freshK*)
+   Atomic_distrib [iff]: "Atomic`(KEY`K \<union> NONCE`N) =
+                   Atomic`(KEY`K) \<union> Atomic`(NONCE`N)" 
+
+  (*this protocol makes the assumption of secure means
+    between each agent and his smartcard*)
+   shouprubin_assumes_securemeans [iff]: "evs \<in> srb \<Longrightarrow> secureM"
+
+constdefs
+
+  Unique :: "[event, event list] => bool" ("Unique _ on _")
+   "Unique ev on evs == 
+      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
+
+
+consts  srb   :: "event list set"
+inductive "srb"
+  intros 
+
+    Nil:  "[]\<in> srb"
+
+
+
+    Fake: "\<lbrakk> evsF \<in> srb;  X \<in> synth (analz (knows Spy evsF)); 
+             illegalUse(Card B) \<rbrakk>
+          \<Longrightarrow> Says Spy A X # 
+              Inputs Spy (Card B) X # evsF \<in> srb"
+
+(*In general this rule causes the assumption Card B \<notin> cloned
+  in most guarantees for B - starting with confidentiality -
+  otherwise pairK_confidential could not apply*)
+    Forge:
+         "\<lbrakk> evsFo \<in> srb; Nonce Nb \<in> analz (knows Spy evsFo);
+             Key (pairK(A,B)) \<in> knows Spy evsFo \<rbrakk>
+          \<Longrightarrow> Notes Spy (Key (sesK(Nb,pairK(A,B)))) # evsFo \<in> srb"
+
+
+
+   Reception: "\<lbrakk> evsrb\<in> srb; Says A B X \<in> set evsrb \<rbrakk>
+              \<Longrightarrow> Gets B X # evsrb \<in> srb"
+
+
+
+(*A AND THE SERVER*)
+    SR_U1:  "\<lbrakk> evs1 \<in> srb; A \<noteq> Server \<rbrakk>
+          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> 
+                # evs1 \<in> srb"
+
+    SR_U2:  "\<lbrakk> evs2 \<in> srb; 
+             Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Server A \<lbrace>Nonce (Pairkey(A,B)), 
+                           Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>
+                  \<rbrace>
+                # evs2 \<in> srb"
+
+
+
+
+(*A AND HER CARD*)
+(*A cannot decrypt the verifier for she dosn't know shrK A,
+  but the pairkey is recognisable*)
+    SR_U3:  "\<lbrakk> evs3 \<in> srb; legalUse(Card A);
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+             Gets A \<lbrace>Nonce Pk, Certificate\<rbrace> \<in> set evs3 \<rbrakk>
+          \<Longrightarrow> Inputs A (Card A) (Agent A)
+                # evs3 \<in> srb"   (*however A only queries her card 
+if she has previously contacted the server to initiate with some B. 
+Otherwise she would do so even if the Server had not been active. 
+Still, this doesn't and can't mean that the pairkey originated with 
+the server*)
+ 
+(*The card outputs the nonce Na to A*)               
+    SR_U4:  "\<lbrakk> evs4 \<in> srb; 
+             Nonce Na \<notin> used evs4; legalUse(Card A); A \<noteq> Server;
+             Inputs A (Card A) (Agent A) \<in> set evs4 \<rbrakk> 
+       \<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+              # evs4 \<in> srb"
+
+(*The card can be exploited by the spy*)
+(*because of the assumptions on the card, A is certainly not server nor spy*)
+ SR_U4Fake: "\<lbrakk> evs4F \<in> srb; Nonce Na \<notin> used evs4F; 
+             illegalUse(Card A);
+             Inputs Spy (Card A) (Agent A) \<in> set evs4F \<rbrakk> 
+      \<Longrightarrow> Outpts (Card A) Spy \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+            # evs4F \<in> srb"
+
+
+
+
+(*A TOWARDS B*)
+    SR_U5:  "\<lbrakk> evs5 \<in> srb; 
+             Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs5;
+             \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> # evs5 \<in> srb"
+(*A must check that the verifier is not a compound message, 
+  otherwise this would also fire after SR_U7 *)
+
+
+
+
+(*B AND HIS CARD*)
+    SR_U6:  "\<lbrakk> evs6 \<in> srb; legalUse(Card B);
+             Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs6 \<rbrakk>
+          \<Longrightarrow> Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> 
+                # evs6 \<in> srb"
+(*B gets back from the card the session key and various verifiers*)
+    SR_U7:  "\<lbrakk> evs7 \<in> srb; 
+             Nonce Nb \<notin> used evs7; legalUse(Card B); B \<noteq> Server;
+             K = sesK(Nb,pairK(A,B));
+             Key K \<notin> used evs7;
+             Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7\<rbrakk>
+    \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,
+                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                            Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
+                # evs7 \<in> srb"
+(*The card can be exploited by the spy*)
+(*because of the assumptions on the card, A is certainly not server nor spy*)
+ SR_U7Fake:  "\<lbrakk> evs7F \<in> srb; Nonce Nb \<notin> used evs7F; 
+             illegalUse(Card B);
+             K = sesK(Nb,pairK(A,B));
+             Key K \<notin> used evs7F;
+             Inputs Spy (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7F \<rbrakk>
+          \<Longrightarrow> Outpts (Card B) Spy \<lbrace>Nonce Nb, Agent A, Key K,
+                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                            Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
+                # evs7F \<in> srb"
+
+
+
+
+(*B TOWARDS A*)
+(*having sent an input that mentions A is the only memory B relies on,
+  since the output doesn't mention A - lack of explicitness*) 
+    SR_U8:  "\<lbrakk> evs8 \<in> srb;  
+             Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs8;
+             Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, 
+                                 Cert1, Cert2\<rbrace> \<in> set evs8 \<rbrakk>
+          \<Longrightarrow> Says B A \<lbrace>Nonce Nb, Cert1\<rbrace> # evs8 \<in> srb"
+  
+
+
+
+(*A AND HER CARD*)
+(*A cannot check the form of the verifiers - although I can prove the form of
+  Cert2 - and just feeds her card with what she's got*)
+    SR_U9:  "\<lbrakk> evs9 \<in> srb; legalUse(Card A);
+             Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs9;
+             Outpts (Card A) A \<lbrace>Nonce Na, Cert2\<rbrace> \<in> set evs9; 
+             Gets A \<lbrace>Nonce Nb, Cert3\<rbrace> \<in> set evs9;
+             \<forall> p q. Cert2 \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
+          \<Longrightarrow> Inputs A (Card A) 
+                 \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,
+                  Cert1, Cert3, Cert2\<rbrace> 
+                # evs9 \<in> srb"
+(*But the card will only give outputs to the inputs of the correct form*)
+    SR_U10: "\<lbrakk> evs10 \<in> srb; legalUse(Card A); A \<noteq> Server;
+             K = sesK(Nb,pairK(A,B));
+             Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, 
+                                 Nonce (Pairkey(A,B)),
+                                 Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), 
+                                                   Agent B\<rbrace>,
+                                 Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                                 Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+               \<in> set evs10 \<rbrakk>
+          \<Longrightarrow> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, 
+                                 Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
+                 # evs10 \<in> srb"
+(*The card can be exploited by the spy*)
+(*because of the assumptions on the card, A is certainly not server nor spy*)
+SR_U10Fake: "\<lbrakk> evs10F \<in> srb; 
+             illegalUse(Card A);
+             K = sesK(Nb,pairK(A,B));
+             Inputs Spy (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, 
+                                   Nonce (Pairkey(A,B)),
+                                   Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), 
+                                                    Agent B\<rbrace>,
+                                   Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                                   Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+               \<in> set evs10F \<rbrakk>
+          \<Longrightarrow> Outpts (Card A) Spy \<lbrace>Agent B, Nonce Nb, 
+                                   Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
+                 # evs10F \<in> srb"
+
+
+
+
+(*A TOWARDS B*)
+(*having initiated with B is the only memory A relies on,
+  since the output doesn't mention B - lack of explicitness*) 
+    SR_U11: "\<lbrakk> evs11 \<in> srb;
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs11;
+             Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
+               \<in> set evs11 \<rbrakk>
+          \<Longrightarrow> Says A B (Certificate) 
+                 # evs11 \<in> srb"
+
+
+
+(*Both peers may leak by accident the session keys obtained from their
+  cards*)
+    Oops1:
+     "\<lbrakk> evsO1 \<in> srb;
+         Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
+           \<in> set evsO1 \<rbrakk>
+     \<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO1 \<in> srb"
+
+    Oops2:
+     "\<lbrakk> evsO2 \<in> srb;
+         Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
+           \<in> set evsO2 \<rbrakk>
+    \<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO2 \<in> srb"
+
+
+
+
+
+
+(*To solve Fake case when it doesn't involve analz - used to be condensed
+  into Fake_parts_insert_tac*)
+declare Fake_parts_insert_in_Un  [dest]
+declare analz_into_parts [dest]
+(*declare parts_insertI [intro]*)
+
+
+
+(*General facts about message reception*)
+lemma Gets_imp_Says: 
+       "\<lbrakk> Gets B X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> \<exists> A. Says A B X \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Gets_imp_knows_Spy: 
+     "\<lbrakk> Gets B X \<in> set evs; evs \<in> srb \<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+done
+
+lemma Gets_imp_knows_Spy_parts_Snd: 
+     "\<lbrakk> Gets B \<lbrace>X, Y\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  \<Longrightarrow> Y \<in> parts (knows Spy evs)"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy parts.Inj parts.Snd)
+done
+
+lemma Gets_imp_knows_Spy_analz_Snd: 
+     "\<lbrakk> Gets B \<lbrace>X, Y\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  \<Longrightarrow> Y \<in> analz (knows Spy evs)"
+apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy analz.Inj analz.Snd)
+done
+
+(*end general facts*)
+
+
+
+(*Begin lemmas on secure means, from Event.ML, proved for shouprubin. They help
+  the simplifier, especially in analz_image_freshK*)
+
+
+lemma Inputs_imp_knows_Spy_secureM_srb: 
+      "\<lbrakk> Inputs Spy C X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> X \<in> knows Spy evs"
+apply (simp (no_asm_simp) add: Inputs_imp_knows_Spy_secureM)
+done
+
+lemma knows_Spy_Inputs_secureM_srb_Spy: 
+      "evs \<in>srb \<Longrightarrow> knows Spy (Inputs Spy C X # evs) = insert X (knows Spy evs)"
+apply (simp (no_asm_simp))
+done
+
+lemma knows_Spy_Inputs_secureM_srb: 
+    "\<lbrakk> A \<noteq> Spy; evs \<in>srb \<rbrakk> \<Longrightarrow> knows Spy (Inputs A C X # evs) =  knows Spy evs"
+apply (simp (no_asm_simp))
+done
+
+lemma knows_Spy_Outpts_secureM_srb_Spy: 
+      "evs \<in>srb \<Longrightarrow> knows Spy (Outpts C Spy X # evs) = insert X (knows Spy evs)"
+apply (simp (no_asm_simp))
+done
+
+lemma knows_Spy_Outpts_secureM_srb: 
+     "\<lbrakk> A \<noteq> Spy; evs \<in>srb \<rbrakk> \<Longrightarrow> knows Spy (Outpts C A X # evs) =  knows Spy evs"
+apply (simp (no_asm_simp))
+done
+
+(*End lemmas on secure means for shouprubin*)
+
+
+
+
+(*BEGIN technical lemmas - evolution of forwarding lemmas*)
+
+(*If an honest agent uses a smart card, then the card is his/her own, is
+  not stolen, and the agent has received suitable data to feed the card. 
+  In other words, these are guarantees that an honest agent can only use 
+  his/her own card, and must use it correctly.
+  On the contrary, the spy can "Inputs" any cloned cards also by the Fake rule.
+
+  Instead of Auto_tac, proofs here used to asm-simplify and then force-tac.
+*)
+lemma Inputs_A_Card_3: 
+    "\<lbrakk> Inputs A C (Agent A) \<in> set evs; A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+      (\<exists> Pk Certificate. Gets A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Inputs_B_Card_6: 
+     "\<lbrakk> Inputs B C \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs; B \<noteq> Spy; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and> Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Inputs_A_Card_9: 
+     "\<lbrakk> Inputs A C \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,   
+                                           Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+  \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+      Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs     \<and>  
+      Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace> \<in> set evs        \<and>   
+      Gets A \<lbrace>Nonce Nb, Cert2\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+
+(*The two occurrences of A in the Outpts event don't match SR_U4Fake, where
+  A cannot be the Spy. Hence the card is legally usable by rule SR_U4*)
+lemma Outpts_A_Card_4: 
+     "\<lbrakk> Outpts C A \<lbrace>Nonce Na, (Crypt (crdK (Card A)) (Nonce Na))\<rbrace> \<in> set evs;  
+         evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         Inputs A (Card A) (Agent A) \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+
+(*First certificate is made explicit so that a comment similar to the previous
+  applies. This also provides Na to the Inputs event in the conclusion*)
+lemma Outpts_B_Card_7: 
+      "\<lbrakk> Outpts C B \<lbrace>Nonce Nb, Agent A, Key K,  
+                      Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                      Cert2\<rbrace> \<in> set evs;  
+         evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and>  
+         Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Outpts_A_Card_10: 
+     "\<lbrakk> Outpts C A \<lbrace>Agent B, Nonce Nb, 
+                    Key K, (Crypt (pairK(A,B)) (Nonce Nb))\<rbrace> \<in> set evs; 
+         evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         (\<exists> Na Ver1 Ver2 Ver3.  
+       Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),  
+                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+
+
+(*
+Contrarily to original version, A doesn't need to check the form of the 
+certificate to learn that her peer is B. The goal is available to A.
+*)
+lemma Outpts_A_Card_10_imp_Inputs: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
+          \<in> set evs; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> (\<exists> Na Ver1 Ver2 Ver3.  
+       Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),  
+                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule srb.induct)
+apply simp_all
+apply blast+
+done
+
+
+
+
+(*Weaker version: if the agent can't check the forms of the verifiers, then
+  the agent must not be the spy so as to solve SR_U4Fake. The verifier must be
+  recognised as some cyphertex in order to distinguish from case SR_U7, 
+  concerning B's output, which also begins with a nonce.
+*)
+lemma Outpts_honest_A_Card_4: 
+     "\<lbrakk> Outpts C A \<lbrace>Nonce Na, Crypt K X\<rbrace> \<in>set evs; 
+         A \<noteq> Spy;  evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         Inputs A (Card A) (Agent A) \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+(*alternative formulation of same theorem
+Goal "\<lbrakk> Outpts C A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs;  
+         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>;    
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
+         Inputs A (Card A) (Agent A) \<in> set evs"
+same proof
+*)
+
+
+lemma Outpts_honest_B_Card_7: 
+    "\<lbrakk> Outpts C B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> \<in> set evs;  
+       B \<noteq> Spy; evs \<in> srb \<rbrakk>  
+   \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and>  
+       (\<exists> Na. Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Outpts_honest_A_Card_10: 
+     "\<lbrakk> Outpts C A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs;  
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> legalUse (C) \<and> C = (Card A) \<and>  
+         (\<exists> Na Pk Ver1 Ver2 Ver3.  
+          Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Pk,  
+                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
+apply (erule rev_mp, erule srb.induct)
+apply simp_all
+apply blast+
+done
+(*-END-*)
+
+
+(*Even weaker versions: if the agent can't check the forms of the verifiers
+  and the agent may be the spy, then we must know what card the agent
+  is getting the output from. 
+*)
+lemma Outpts_which_Card_4: 
+    "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt K X\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
+    \<Longrightarrow> Inputs A (Card A) (Agent A) \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply (simp_all (no_asm_simp))
+apply clarify
+done
+
+lemma Outpts_which_Card_7: 
+  "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
+       \<in> set evs;  evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> \<exists> Na. Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+(*This goal is now available - in the sense of Goal Availability*)
+lemma Outpts_which_Card_10: 
+    "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate \<rbrace> \<in> set evs;
+       evs \<in> srb \<rbrakk>  
+    \<Longrightarrow> \<exists> Na. Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)), 
+                            Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>,  
+                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                            Crypt (crdK (Card A)) (Nonce Na) \<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+
+(*Lemmas on the form of outputs*)
+
+
+(*A needs to check that the verifier is a cipher for it to come from SR_U4
+  otherwise it could come from SR_U7 *)
+lemma Outpts_A_Card_form_4: 
+  "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs;  
+         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> Certificate = (Crypt (crdK (Card A)) (Nonce Na))"
+apply (erule rev_mp, erule srb.induct)
+apply (simp_all (no_asm_simp))
+done
+
+lemma Outpts_B_Card_form_7: 
+   "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
+        \<in> set evs; evs \<in> srb \<rbrakk>          
+      \<Longrightarrow> \<exists> Na.    
+          K = sesK(Nb,pairK(A,B)) \<and>                       
+          Cert1 = (Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>) \<and>  
+          Cert2 = (Crypt (pairK(A,B)) (Nonce Nb))"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Outpts_A_Card_form_10: 
+   "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
+        \<in> set evs; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> K = sesK(Nb,pairK(A,B)) \<and>  
+          Certificate = (Crypt (pairK(A,B)) (Nonce Nb))"
+apply (erule rev_mp, erule srb.induct)
+apply (simp_all (no_asm_simp))
+done
+
+lemma Outpts_A_Card_form_bis: 
+  "\<lbrakk> Outpts (Card A') A' \<lbrace>Agent B', Nonce Nb', Key (sesK(Nb,pairK(A,B))), 
+     Certificate\<rbrace> \<in> set evs; 
+         evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> A' = A \<and> B' = B \<and> Nb = Nb' \<and>  
+          Certificate = (Crypt (pairK(A,B)) (Nonce Nb))"
+apply (erule rev_mp, erule srb.induct)
+apply (simp_all (no_asm_simp))
+done
+
+(*\<dots> and Inputs *)
+
+lemma Inputs_A_Card_form_9: 
+
+     "\<lbrakk> Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,   
+                             Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
+         evs \<in> srb \<rbrakk>  
+  \<Longrightarrow>    Cert3 = Crypt (crdK (Card A)) (Nonce Na)"
+apply (erule rev_mp)
+apply (erule srb.induct)
+apply (simp_all (no_asm_simp))
+(*Fake*)
+apply force
+(*SR_U9*)
+apply (blast dest!: Outpts_A_Card_form_4)
+done
+(* Pk, Cert1, Cert2 cannot be made explicit because they traversed the network in the clear *)
+
+
+
+(*General guarantees on Inputs and Outpts*)
+
+(*for any agents*)
+
+
+lemma Inputs_Card_legalUse: 
+  "\<lbrakk> Inputs A (Card A) X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> legalUse(Card A)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Outpts_Card_legalUse: 
+  "\<lbrakk> Outpts (Card A) A X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> legalUse(Card A)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+(*for honest agents*)
+
+lemma Inputs_Card: "\<lbrakk> Inputs A C X \<in> set evs; A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> C = (Card A) \<and> legalUse(C)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Outpts_Card: "\<lbrakk> Outpts C A X \<in> set evs; A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> C = (Card A) \<and> legalUse(C)"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+lemma Inputs_Outpts_Card: 
+     "\<lbrakk> Inputs A C X \<in> set evs \<or> Outpts C A Y \<in> set evs;  
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> C = (Card A) \<and> legalUse(Card A)"
+apply (blast dest: Inputs_Card Outpts_Card)
+done
+
+
+(*for the spy - they stress that the model behaves as it is meant to*) 
+
+(*The or version can be also proved directly.
+  It stresses that the spy may use either her own legally usable card or
+  all the illegally usable cards.
+*)
+lemma Inputs_Card_Spy: 
+  "\<lbrakk> Inputs Spy C X \<in> set evs \<or> Outpts C Spy X \<in> set evs; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> C = (Card Spy) \<and> legalUse(Card Spy) \<or>  
+          (\<exists> A. C = (Card A) \<and> illegalUse(Card A))"
+apply (erule rev_mp, erule srb.induct)
+apply auto
+done
+
+
+(*END technical lemmas*)
+
+
+
+
+
+
+(*BEGIN unicity theorems: certain items uniquely identify a smart card's
+                          output*)
+
+(*A's card's first output: the nonce uniquely identifies the rest*)
+lemma Outpts_A_Card_unique_nonce:
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>  
+           \<in> set evs;   
+         Outpts (Card A') A' \<lbrace>Nonce Na, Crypt (crdK (Card A')) (Nonce Na)\<rbrace> 
+           \<in> set evs;   
+         evs \<in> srb \<rbrakk> \<Longrightarrow> A=A'"
+apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+(*B's card's output: the NONCE uniquely identifies the rest*)
+lemma Outpts_B_Card_unique_nonce: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key SK, Cert1, Cert2\<rbrace> \<in> set evs;   
+      Outpts (Card B') B' \<lbrace>Nonce Nb, Agent A', Key SK', Cert1', Cert2'\<rbrace> \<in> set evs;
+       evs \<in> srb \<rbrakk> \<Longrightarrow> B=B' \<and> A=A' \<and> SK=SK' \<and> Cert1=Cert1' \<and> Cert2=Cert2'"
+apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+
+(*B's card's output: the SESKEY uniquely identifies the rest*)
+lemma Outpts_B_Card_unique_key: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key SK, Cert1, Cert2\<rbrace> \<in> set evs;   
+      Outpts (Card B') B' \<lbrace>Nonce Nb', Agent A', Key SK, Cert1', Cert2'\<rbrace> \<in> set evs; 
+       evs \<in> srb \<rbrakk> \<Longrightarrow> B=B' \<and> A=A' \<and> Nb=Nb' \<and> Cert1=Cert1' \<and> Cert2=Cert2'"
+apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+lemma Outpts_A_Card_unique_key: 
+   "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, V\<rbrace> \<in> set evs;   
+      Outpts (Card A') A' \<lbrace>Agent B', Nonce Nb', Key K, V'\<rbrace> \<in> set evs;   
+         evs \<in> srb \<rbrakk> \<Longrightarrow> A=A' \<and> B=B' \<and> Nb=Nb' \<and> V=V'"
+apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
+apply (blast dest: Outpts_A_Card_form_bis)
+apply blast
+done
+
+
+(*Revised unicity theorem - applies to both steps 4 and 7*)
+lemma Outpts_A_Card_Unique: 
+  "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, rest\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> Unique (Outpts (Card A) A \<lbrace>Nonce Na, rest\<rbrace>) on evs"
+apply (erule rev_mp, erule srb.induct, simp_all add: Unique_def)
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+apply (fastsimp dest: Outpts_parts_used)
+apply blast
+done
+
+(*can't prove the same on evs10 for it doesn't have a freshness assumption!*)
+
+
+(*END unicity theorems*)
+
+
+(*BEGIN counterguarantees about spy's knowledge*)
+
+(*on nonces*)
+
+lemma Spy_knows_Na: 
+      "\<lbrakk> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> Nonce Na \<in> analz (knows Spy evs)"
+apply (blast dest!: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd])
+done
+
+lemma Spy_knows_Nb: 
+      "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> Nonce Nb \<in> analz (knows Spy evs)"
+apply (blast dest!: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Fst])
+done
+
+
+(*on Pairkey*)
+
+lemma Pairkey_Gets_analz_knows_Spy: 
+      "\<lbrakk> Gets A \<lbrace>Nonce (Pairkey(A,B)), Certificate\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> Nonce (Pairkey(A,B)) \<in> analz (knows Spy evs)"
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
+done
+
+lemma Pairkey_Inputs_imp_Gets: 
+     "\<lbrakk> Inputs A (Card A)             
+           \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),     
+             Cert1, Cert3, Cert2\<rbrace> \<in> set evs;           
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>     
+      \<Longrightarrow> Gets A \<lbrace>Nonce (Pairkey(A,B)), Cert1\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply (simp_all (no_asm_simp))
+apply force
+done
+
+lemma Pairkey_Inputs_analz_knows_Spy: 
+     "\<lbrakk> Inputs A (Card A)             
+           \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),     
+             Cert1, Cert3, Cert2\<rbrace> \<in> set evs;           
+         evs \<in> srb \<rbrakk>     
+     \<Longrightarrow> Nonce (Pairkey(A,B)) \<in> analz (knows Spy evs)"
+apply (case_tac "A = Spy")
+apply (fastsimp dest!: Inputs_imp_knows_Spy_secureM [THEN analz.Inj])
+apply (blast dest!: Pairkey_Inputs_imp_Gets [THEN Pairkey_Gets_analz_knows_Spy])
+done
+
+(* This fails on base case because of XOR properties.
+lemma Pairkey_authentic:
+  "\<lbrakk> Nonce (Pairkey(A,B)) \<in> parts (knows Spy evs);
+     Card A \<notin> cloned; evs \<in> sr \<rbrakk>
+ \<Longrightarrow> \<exists> cert. Says Server A \<lbrace>Nonce (Pairkey(A,B)), Cert\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule sr.induct, simp_all)
+apply clarify
+oops
+
+ 1. \<And>x a b.
+       \<lbrakk>Card A \<notin> cloned; Pairkey (A, B) = Pairkey (a, b); Card a \<in> cloned;
+        Card b \<in> cloned\<rbrakk>
+       \<Longrightarrow> False
+*)
+
+
+(*END counterguarantees on spy's knowledge*)
+
+
+(*BEGIN rewrite rules for parts operator*)
+
+declare shrK_disj_sesK [THEN not_sym, iff] 
+declare pin_disj_sesK [THEN not_sym, iff]
+declare crdK_disj_sesK [THEN not_sym, iff]
+declare pairK_disj_sesK [THEN not_sym, iff]
+
+
+ML
+{*
+val Outpts_B_Card_form_7 = thm "Outpts_B_Card_form_7"
+val Outpts_A_Card_form_4 = thm "Outpts_A_Card_form_4"
+val Outpts_A_Card_form_10 = thm "Outpts_A_Card_form_10"
+val Gets_imp_knows_Spy = thm "Gets_imp_knows_Spy"
+val Outpts_B_Card_form_7 = thm "Outpts_B_Card_form_7"
+val Gets_imp_knows_Spy_parts_Snd = thm "Gets_imp_knows_Spy_parts_Snd"
+val Gets_imp_knows_Spy_analz_Snd = thm "Gets_imp_knows_Spy_analz_Snd"
+
+val prepare_tac = 
+ (*SR_U8*)   forward_tac [Outpts_B_Card_form_7] 14 THEN
+ (*SR_U8*)   Clarify_tac 15 THEN
+ (*SR_U9*)   forward_tac [Outpts_A_Card_form_4] 16 THEN 
+ (*SR_U11*)  forward_tac [Outpts_A_Card_form_10] 21 
+
+val parts_prepare_tac = 
+           prepare_tac THEN
+ (*SR_U9*)   dresolve_tac [Gets_imp_knows_Spy_parts_Snd] 18 THEN 
+ (*SR_U9*)   dresolve_tac [Gets_imp_knows_Spy_parts_Snd] 19 THEN 
+ (*Oops1*) dresolve_tac [Outpts_B_Card_form_7] 25    THEN               
+ (*Oops2*) dresolve_tac [Outpts_A_Card_form_10] 27 THEN                
+ (*Base*)  Force_tac 1
+
+val analz_prepare_tac = 
+         prepare_tac THEN
+         dtac (Gets_imp_knows_Spy_analz_Snd) 18 THEN 
+ (*SR_U9*) dtac (Gets_imp_knows_Spy_analz_Snd) 19 THEN 
+         REPEAT_FIRST (eresolve_tac [asm_rl, conjE] ORELSE' hyp_subst_tac)
+
+*}
+
+method_setup prepare = {*
+    Method.no_args (Method.METHOD (fn facts => prepare_tac)) *}
+  "to launch a few simple facts that'll help the simplifier"
+
+method_setup parts_prepare = {*
+    Method.no_args (Method.METHOD (fn facts => parts_prepare_tac)) *}
+  "additional facts to reason about parts"
+
+method_setup analz_prepare = {*
+    Method.no_args (Method.METHOD (fn facts => analz_prepare_tac)) *}
+  "additional facts to reason about analz"
+
+
+
+lemma Spy_parts_keys [simp]: "evs \<in> srb \<Longrightarrow>  
+  (Key (shrK P) \<in> parts (knows Spy evs)) = (Card P \<in> cloned) \<and>  
+  (Key (pin P) \<in> parts (knows Spy evs)) = (P \<in> bad \<or> Card P \<in> cloned) \<and>  
+  (Key (crdK C) \<in> parts (knows Spy evs)) = (C \<in> cloned) \<and>  
+  (Key (pairK(A,B)) \<in> parts (knows Spy evs)) = (Card B \<in> cloned)"
+apply (erule srb.induct)
+apply parts_prepare
+apply simp_all
+apply (blast intro: parts_insertI)
+done
+
+(*END rewrite rules for parts operator*)
+
+(*BEGIN rewrite rules for analz operator*)
+
+
+lemma Spy_analz_shrK[simp]: "evs \<in> srb \<Longrightarrow>  
+  (Key (shrK P) \<in> analz (knows Spy evs)) = (Card P \<in> cloned)"
+apply (auto dest!: Spy_knows_cloned)
+done
+
+lemma Spy_analz_crdK[simp]: "evs \<in> srb \<Longrightarrow>  
+  (Key (crdK C) \<in> analz (knows Spy evs)) = (C \<in> cloned)"
+apply (auto dest!: Spy_knows_cloned)
+done
+
+lemma Spy_analz_pairK[simp]: "evs \<in> srb \<Longrightarrow>  
+  (Key (pairK(A,B)) \<in> analz (knows Spy evs)) = (Card B \<in> cloned)"
+apply (auto dest!: Spy_knows_cloned)
+done
+
+
+
+
+(*Because initState contains a set of nonces, this is needed for base case of
+  analz_image_freshK*)
+lemma analz_image_Key_Un_Nonce: "analz (Key`K \<union> Nonce`N) = Key`K \<union> Nonce`N"
+apply auto
+done
+
+ML
+{*
+val knows_Spy_Inputs_secureM_srb_Spy = thm "knows_Spy_Inputs_secureM_srb_Spy"
+val knows_Spy_Outpts_secureM_srb_Spy = thm "knows_Spy_Outpts_secureM_srb_Spy"
+val shouprubin_assumes_securemeans = thm "shouprubin_assumes_securemeans"
+val analz_image_Key_Un_Nonce= thm "analz_image_Key_Un_Nonce"
+*}
+
+method_setup sc_analz_freshK = {*
+    Method.no_args
+     (Method.METHOD
+      (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+                          REPEAT_FIRST (rtac analz_image_freshK_lemma),
+                          ALLGOALS (asm_simp_tac (analz_image_freshK_ss
+                                    addsimps [knows_Spy_Inputs_secureM_srb_Spy,
+                                              knows_Spy_Outpts_secureM_srb_Spy,
+                                              shouprubin_assumes_securemeans, 
+                                              analz_image_Key_Un_Nonce]))])) *}
+    "for proving the Session Key Compromise theorem for smartcard protocols"
+
+
+lemma analz_image_freshK [rule_format]: 
+     "evs \<in> srb \<Longrightarrow>      \<forall> K KK.  
+          (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) =  
+          (K \<in> KK \<or> Key K \<in> analz (knows Spy evs))"
+apply (erule srb.induct)
+apply analz_prepare
+apply sc_analz_freshK
+apply spy_analz
+done
+
+
+lemma analz_insert_freshK: "evs \<in> srb \<Longrightarrow>   
+         Key K \<in> analz (insert (Key K') (knows Spy evs)) =  
+         (K = K' \<or> Key K \<in> analz (knows Spy evs))"
+apply (simp only: analz_image_freshK_simps analz_image_freshK)
+done
+
+(*END rewrite rules for analz operator*)
+
+(*BEGIN authenticity theorems*)
+
+
+
+
+lemma Na_Nb_certificate_authentic: 
+     "\<lbrakk> Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace> \<in> parts (knows Spy evs);  
+         \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>           
+     \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb,pairK(A,B))),   
+                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply parts_prepare
+apply simp_all
+(*Fake*)
+apply spy_analz
+(*SR_U7F*)
+apply clarify
+(*SR_U8*)
+apply clarify
+done
+
+lemma Nb_certificate_authentic:
+      "\<lbrakk> Crypt (pairK(A,B)) (Nonce Nb) \<in> parts (knows Spy evs);  
+         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>    
+     \<Longrightarrow> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key (sesK(Nb,pairK(A,B))),  
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply parts_prepare
+apply (case_tac [17] "Aa = Spy")
+apply simp_all
+(*Fake*)
+apply spy_analz
+(*SR_U7F, SR_U10F*)
+apply clarify+
+done
+
+
+
+(*Discovering the very origin of the Nb certificate...*)
+lemma Outpts_A_Card_imp_pairK_parts: 
+     "\<lbrakk> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, 
+                    Key K, Certificate\<rbrace> \<in> set evs;  
+        evs \<in> srb \<rbrakk>   
+    \<Longrightarrow> \<exists> Na. Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace> \<in> parts (knows Spy evs)"
+apply (erule rev_mp, erule srb.induct)
+apply parts_prepare
+apply simp_all
+(*Fake*)
+apply (blast dest: parts_insertI)
+(*SR_U7*)
+apply force
+(*SR_U7F*)
+apply force
+(*SR_U8*)
+apply blast
+(*SR_U10*)
+apply (blast dest: Inputs_imp_knows_Spy_secureM_srb parts.Inj Inputs_A_Card_9 Gets_imp_knows_Spy elim: knows_Spy_partsEs)
+(*SR_U10F*)
+apply (blast dest: Inputs_imp_knows_Spy_secureM_srb [THEN parts.Inj] 
+                   Inputs_A_Card_9 Gets_imp_knows_Spy 
+             elim: knows_Spy_partsEs)
+done
+
+               
+lemma Nb_certificate_authentic_bis: 
+     "\<lbrakk> Crypt (pairK(A,B)) (Nonce Nb) \<in> parts (knows Spy evs);  
+         B \<noteq> Spy; \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>    
+ \<Longrightarrow> \<exists> Na. Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb,pairK(A,B))),   
+                   Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                   Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply parts_prepare
+apply (simp_all (no_asm_simp))
+(*Fake*)
+apply spy_analz
+(*SR_U7*)
+apply blast
+(*SR_U7F*)
+apply blast
+(*SR_U8*)
+apply force
+(*SR_U10*)
+apply (blast dest: Na_Nb_certificate_authentic Inputs_imp_knows_Spy_secureM_srb [THEN parts.Inj] elim: knows_Spy_partsEs)
+(*SR_U10F*)
+apply (blast dest: Na_Nb_certificate_authentic Inputs_imp_knows_Spy_secureM_srb [THEN parts.Inj] elim: knows_Spy_partsEs)
+(*SR_U11*)
+apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_imp_pairK_parts)
+done
+
+
+lemma Pairkey_certificate_authentic: 
+    "\<lbrakk> Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace> \<in> parts (knows Spy evs);    
+         Card A \<notin> cloned; evs \<in> srb \<rbrakk>        
+     \<Longrightarrow> Pk = Pairkey(A,B) \<and>              
+         Says Server A \<lbrace>Nonce Pk,  
+                        Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>\<rbrace> 
+           \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply parts_prepare
+apply (simp_all (no_asm_simp))
+(*Fake*)
+apply spy_analz
+(*SR_U8*)
+apply force
+done
+
+
+lemma sesK_authentic: 
+     "\<lbrakk> Key (sesK(Nb,pairK(A,B))) \<in> parts (knows Spy evs);  
+         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>           
+      \<Longrightarrow> Notes Spy \<lbrace>Key (sesK(Nb,pairK(A,B))), Nonce Nb, Agent A, Agent B\<rbrace>  
+           \<in> set evs"
+apply (erule rev_mp, erule srb.induct)
+apply parts_prepare
+apply (simp_all)
+(*fake*)
+apply spy_analz
+(*forge*)
+apply (fastsimp dest: analz.Inj)
+(*SR_U7: used B\<noteq>Spy*)
+(*SR_U7F*)
+apply clarify
+(*SR_U10: used A\<noteq>Spy*)
+(*SR_U10F*)
+apply clarify
+done
+
+
+(*END authenticity theorems*)
+
+
+(*BEGIN confidentiality theorems*)
+
+
+lemma Confidentiality: 
+     "\<lbrakk> Notes Spy \<lbrace>Key (sesK(Nb,pairK(A,B))), Nonce Nb, Agent A, Agent B\<rbrace>  
+           \<notin> set evs; 
+        A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+        evs \<in> srb \<rbrakk>           
+      \<Longrightarrow> Key (sesK(Nb,pairK(A,B))) \<notin> analz (knows Spy evs)"
+apply (blast intro: sesK_authentic)
+done
+
+lemma Confidentiality_B: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
+          \<in> set evs;  
+        Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> \<notin> set evs;  
+        A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); Card B \<notin> cloned; 
+        evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+apply (erule rev_mp, erule rev_mp, erule srb.induct)
+apply analz_prepare
+apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs)
+(*Fake*)
+apply spy_analz
+(*Forge*)
+apply (rotate_tac 7)
+apply (drule parts.Inj)
+apply (fastsimp dest: Outpts_B_Card_form_7)
+(*SR_U7*)
+apply (blast dest!: Outpts_B_Card_form_7)
+(*SR_U7F*)
+apply clarify
+apply (drule Outpts_parts_used)
+apply simp
+(*faster than
+  by (fast_tac (claset() addDs [Outpts_parts_used] addss (simpset())) 1)
+*)
+(*SR_U10*)
+apply (fastsimp dest: Outpts_B_Card_form_7)
+(*SR_U10F - uses assumption Card A not cloned*)
+apply clarify
+apply (drule Outpts_B_Card_form_7, assumption)
+apply simp
+(*Oops1*)
+apply (blast dest!: Outpts_B_Card_form_7)
+(*Oops2*)
+apply (blast dest!: Outpts_B_Card_form_7 Outpts_A_Card_form_10)
+done
+
+
+(*END confidentiality theorems*)
+
+
+(*BEGIN authentication theorems*)
+
+lemma A_authenticates_B: 
+     "\<lbrakk> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs;
+        \<not>illegalUse(Card B); 
+        evs \<in> srb \<rbrakk>           
+ \<Longrightarrow> \<exists> Na. Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,   
+                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_form_10 Outpts_A_Card_imp_pairK_parts)
+done
+
+lemma A_authenticates_B_Gets: 
+     "\<lbrakk> Gets A \<lbrace>Nonce Nb, Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>\<rbrace>  
+           \<in> set evs;  
+         \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>           
+    \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb, pairK (A, B))),   
+                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN Na_Nb_certificate_authentic])
+done
+
+
+lemma A_authenticates_B_bis: 
+     "\<lbrakk> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, Key K, Cert2\<rbrace> \<in> set evs;  
+        \<not>illegalUse(Card B); 
+        evs \<in> srb \<rbrakk>           
+ \<Longrightarrow> \<exists> Cert1. Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
+                \<in> set evs"
+apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_form_10 Outpts_A_Card_imp_pairK_parts)
+done
+
+
+
+
+
+
+lemma B_authenticates_A: 
+     "\<lbrakk> Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
+         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, 
+         Key (sesK(Nb,pairK(A,B))), Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule srb.induct)
+apply (simp_all (no_asm_simp))
+apply (blast dest: Says_imp_knows_Spy [THEN parts.Inj] Nb_certificate_authentic)
+done
+
+
+lemma B_authenticates_A_bis: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> \<in> set evs;
+        Gets B (Cert2) \<in> set evs;  
+         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>  
+      \<Longrightarrow> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, Key K, Cert2\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_B_Card_form_7 B_authenticates_A)
+done
+
+
+(*END authentication theorems*)
+
+
+lemma Confidentiality_A: 
+      "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, 
+                       Key K, Certificate\<rbrace> \<in> set evs;  
+         Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> \<notin> set evs;  
+         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>           
+     \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
+apply (drule A_authenticates_B)
+prefer 3
+apply (erule exE)
+apply (drule Confidentiality_B)
+apply auto
+done
+
+
+lemma Outpts_imp_knows_agents_secureM_srb: 
+   "\<lbrakk> Outpts (Card A) A X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> X \<in> knows A evs"
+apply (simp (no_asm_simp) add: Outpts_imp_knows_agents_secureM)
+done
+
+
+(*BEGIN key distribution theorems*)
+lemma A_keydist_to_B: 
+   "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs;  
+         \<not>illegalUse(Card B); 
+         evs \<in> srb \<rbrakk>           
+     \<Longrightarrow> Key K \<in> analz (knows B evs)"
+apply (drule A_authenticates_B)
+prefer 3
+apply (erule exE)
+apply (rule Outpts_imp_knows_agents_secureM_srb [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
+apply assumption+
+done
+
+
+lemma B_keydist_to_A: 
+"\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> \<in> set evs;  
+   Gets B (Cert2) \<in> set evs;  
+   B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
+   evs \<in> srb \<rbrakk>  
+ \<Longrightarrow> Key K \<in> analz (knows A evs)"
+apply (frule Outpts_B_Card_form_7)
+apply assumption apply simp
+apply (frule B_authenticates_A)
+apply (rule_tac [5] Outpts_imp_knows_agents_secureM_srb [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
+apply simp+
+done
+
+(*END key distribution theorems*)
+
+
+
+
+(*BEGIN further theorems about authenticity of verifiers - useful to cards,
+  and somewhat to agents *)
+
+(*MSG11
+If B receives the verifier of msg11, then the verifier originated with msg7.
+This is clearly not available to B: B can't check the form of the verifier because he doesn't know pairK(A,B)
+*)
+lemma Nb_certificate_authentic_B: 
+     "\<lbrakk> Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
+        B \<noteq> Spy; \<not>illegalUse(Card B); 
+        evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> \<exists> Na. 
+            Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb,pairK(A,B))),   
+                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
+                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
+apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN Nb_certificate_authentic_bis])
+done
+
+(*MSG10
+If A obtains the verifier of msg10, then the verifier originated with msg7:
+A_authenticates_B. It is useful to A, who can check the form of the 
+verifier by application of Outpts_A_Card_form_10.
+*)
+
+(*MSG9
+The first verifier verifies the Pairkey to the card: since it's encrypted
+under Ka, it must come from the server (if A's card is not cloned).
+The second verifier verifies both nonces, since it's encrypted under the
+pairK, it must originate with B's card  (if A and B's cards not cloned).
+The third verifier verifies Na: since it's encrytped under the card's key,
+it originated with the card; so the card does not need to save Na
+in the first place and do a comparison now: it just verifies Na through the
+verifier. Three theorems related to these three statements.
+
+Recall that a card can check the form of the verifiers (can decrypt them),
+while an agent in general cannot, if not provided with a suitable theorem.
+*)
+
+(*Card A can't reckon the pairkey - we need to guarantee its integrity!*)
+lemma Pairkey_certificate_authentic_A_Card: 
+     "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+               Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>,  
+               Cert2, Cert3\<rbrace> \<in> set evs; 
+         A \<noteq> Spy; Card A \<notin> cloned; evs \<in> srb \<rbrakk>   
+     \<Longrightarrow> Pk = Pairkey(A,B) \<and>  
+         Says Server A \<lbrace>Nonce (Pairkey(A,B)),  
+                  Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>\<rbrace>   
+           \<in> set evs "
+apply (blast dest: Inputs_A_Card_9 Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd] Pairkey_certificate_authentic)
+done
+(*the second conjunct of the thesis might be regarded as a form of integrity 
+  in the sense of Neuman-Ts'o*)
+
+lemma Na_Nb_certificate_authentic_A_Card: 
+      "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+          Cert1, Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, Cert3\<rbrace> \<in> set evs; 
+      A \<noteq> Spy; \<not>illegalUse(Card B); evs \<in> srb \<rbrakk> 
+   \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb, pairK (A, B))),    
+                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
+           \<in> set evs "
+apply (frule Inputs_A_Card_9)
+apply assumption+
+apply (blast dest: Inputs_A_Card_9 Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN Na_Nb_certificate_authentic])
+done
+
+lemma Na_authentic_A_Card: 
+     "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+                Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>   
+     \<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace>  
+           \<in> set evs"
+apply (blast dest: Inputs_A_Card_9)
+done
+
+(* These three theorems for Card A can be put together trivially.
+They are separated to highlight the different requirements on agents
+and their cards.*)
+
+
+lemma Inputs_A_Card_9_authentic: 
+  "\<lbrakk> Inputs A (Card A)   
+             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
+               Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>,  
+               Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, Cert3\<rbrace> \<in> set evs; 
+    A \<noteq> Spy; Card A \<notin> cloned; \<not>illegalUse(Card B); evs \<in> srb \<rbrakk>   
+    \<Longrightarrow>  Says Server A \<lbrace>Nonce Pk, Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>\<rbrace>   
+           \<in> set evs  \<and> 
+       Outpts (Card B) B \<lbrace>Nonce Nb, Agent A,  Key (sesK(Nb, pairK (A, B))),    
+                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
+           \<in> set evs  \<and> 
+         Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace>  
+           \<in> set evs"
+apply (blast dest: Inputs_A_Card_9 Na_Nb_certificate_authentic Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd] Pairkey_certificate_authentic)
+done
+
+
+(*MSG8
+Nothing to prove because the message is a cleartext that comes from the 
+network*)
+
+(*Other messages: nothing to prove because the verifiers involved are new*)
+
+(*END further theorems about authenticity of verifiers*)
+
+
+
+(* BEGIN trivial guarantees on outputs for agents *)
+
+(*MSG4*)
+lemma SR_U4_imp: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace> 
+           \<in> set evs;  
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>                 
+     \<Longrightarrow> \<exists> Pk V. Gets A \<lbrace>Pk, V\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_A_Card_4 Inputs_A_Card_3)
+done
+(*weak: could strengthen the model adding verifier for the Pairkey to msg3*)
+
+
+(*MSG7*)
+lemma SR_U7_imp: 
+     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,  
+                      Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
+                      Cert2\<rbrace> \<in> set evs;  
+         B \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_B_Card_7 Inputs_B_Card_6)
+done
+
+(*MSG10*)
+lemma SR_U10_imp: 
+     "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, 
+                           Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
+         \<in> set evs;  
+        A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> \<exists> Cert1 Cert2.  
+                   Gets A \<lbrace>Nonce (Pairkey (A, B)), Cert1\<rbrace> \<in> set evs \<and>  
+                   Gets A \<lbrace>Nonce Nb, Cert2\<rbrace> \<in> set evs"
+apply (blast dest: Outpts_A_Card_10 Inputs_A_Card_9)
+done
+
+
+(*END trivial guarantees on outputs for agents*)
+
+
+
+(*INTEGRITY*)
+lemma Outpts_Server_not_evs: 
+      "evs \<in> srb \<Longrightarrow> Outpts (Card Server) P X \<notin> set evs"
+apply (erule srb.induct)
+apply auto
+done
+
+text{*@{term step2_integrity} also is a reliability theorem*}
+lemma Says_Server_message_form: 
+     "\<lbrakk> Says Server A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs;  
+         evs \<in> srb \<rbrakk>                   
+     \<Longrightarrow> \<exists> B. Pk = Nonce (Pairkey(A,B)) \<and>  
+         Certificate = Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>"
+apply (erule rev_mp)
+apply (erule srb.induct)
+apply auto
+apply (blast dest!: Outpts_Server_not_evs)+
+done
+(*cannot be made useful to A in form of a Gets event*)
+
+text{*
+  step4integrity is @{term Outpts_A_Card_form_4}
+
+  step7integrity is @{term Outpts_B_Card_form_7}
+*}
+
+lemma step8_integrity: 
+     "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs;  
+         B \<noteq> Server; B \<noteq> Spy; evs \<in> srb \<rbrakk>                   
+     \<Longrightarrow> \<exists> Cert2 K.   
+    Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Certificate, Cert2\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule srb.induct)
+prefer 18 apply (fastsimp dest: Outpts_A_Card_form_10)
+apply auto
+done
+
+
+text{*  step9integrity is @{term Inputs_A_Card_form_9}
+        step10integrity is @{term Outpts_A_Card_form_10}.
+*}
+
+
+lemma step11_integrity: 
+     "\<lbrakk> Says A B (Certificate) \<in> set evs; 
+         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>;  
+         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
+     \<Longrightarrow> \<exists> K Nb.  
+      Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs"
+apply (erule rev_mp)
+apply (erule srb.induct)
+apply auto
+done
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Wed Feb 01 15:22:02 2006 +0100
@@ -0,0 +1,474 @@
+(*  ID:         $Id$
+    Author:     Giampaolo Bella, Catania University
+*)
+
+header{*Theory of smartcards*}
+
+theory Smartcard imports EventSC begin
+
+text{*  
+As smartcards handle long-term (symmetric) keys, this theoy extends and 
+supersedes theory Private.thy
+
+An agent is bad if she reveals her PIN to the spy, not the shared key that
+is embedded in her card. An agent's being bad implies nothing about her 
+smartcard, which independently may be stolen or cloned.
+*}
+
+consts
+  shrK    :: "agent => key"  (*long-term keys saved in smart cards*)
+  crdK    :: "card  => key"  (*smart cards' symmetric keys*)
+  pin     :: "agent => key"  (*pin to activate the smart cards*)
+
+  (*Mostly for Shoup-Rubin*)
+  Pairkey :: "agent * agent => nat"
+  pairK   :: "agent * agent => key"
+
+axioms
+  inj_shrK: "inj shrK"  --{*No two smartcards store the same key*}
+  inj_crdK: "inj crdK"  --{*Nor do two cards*}
+  inj_pin : "inj pin"   --{*Nor do two agents have the same pin*}
+
+  (*pairK is injective on each component, if we assume encryption to be a PRF
+    or at least collision free *)
+  inj_pairK    [iff]: "(pairK(A,B) = pairK(A',B')) = (A = A' & B = B')"
+  comm_Pairkey [iff]: "Pairkey(A,B) = Pairkey(B,A)"
+
+  (*long-term keys differ from each other*)
+  pairK_disj_crdK [iff]: "pairK(A,B) \<noteq> crdK C"
+  pairK_disj_shrK [iff]: "pairK(A,B) \<noteq> shrK P"
+  pairK_disj_pin [iff]:  "pairK(A,B) \<noteq> pin P"
+  shrK_disj_crdK [iff]:  "shrK P \<noteq> crdK C"
+  shrK_disj_pin [iff]:  "shrK P \<noteq> pin Q"
+  crdK_disj_pin [iff]:   "crdK C \<noteq> pin P"
+
+
+text{*All keys are symmetric*}
+defs  all_symmetric_def: "all_symmetric == True"
+
+lemma isSym_keys: "K \<in> symKeys"	
+by (simp add: symKeys_def all_symmetric_def invKey_symmetric) 
+
+
+constdefs
+  legalUse :: "card => bool" ("legalUse (_)")
+  "legalUse C == C \<notin> stolen"
+
+consts  
+  illegalUse :: "card  => bool"
+primrec
+  illegalUse_def: 
+  "illegalUse (Card A) = ( (Card A \<in> stolen \<and> A \<in> bad)  \<or>  Card A \<in> cloned )"
+
+
+text{*initState must be defined with care*}
+primrec
+(*Server knows all long-term keys; adding cards' keys may be redundant but
+  helps prove crdK_in_initState and crdK_in_used to distinguish cards' keys
+  from fresh (session) keys*)
+  initState_Server:  "initState Server = 
+        (Key`(range shrK \<union> range crdK \<union> range pin \<union> range pairK)) \<union> 
+        (Nonce`(range Pairkey))"
+
+(*Other agents know only their own*)
+  initState_Friend:  "initState (Friend i) = {Key (pin (Friend i))}"
+
+(*Spy knows bad agents' pins, cloned cards' keys, pairKs, and Pairkeys *)
+  initState_Spy: "initState Spy  = 
+                 (Key`((pin`bad) \<union> (pin `{A. Card A \<in> cloned}) \<union> 
+                                      (shrK`{A. Card A \<in> cloned}) \<union> 
+                        (crdK`cloned) \<union> 
+                        (pairK`{(X,A). Card A \<in> cloned})))
+           \<union> (Nonce`(Pairkey`{(A,B). Card A \<in> cloned & Card B \<in> cloned}))"
+
+
+text{*Still relying on axioms*}
+axioms
+  Key_supply_ax:  "finite KK \<Longrightarrow> \<exists> K. K \<notin> KK & Key K \<notin> used evs"
+
+  (*Needed because of Spy's knowledge of Pairkeys*)
+  Nonce_supply_ax: "finite NN \<Longrightarrow> \<exists> N. N \<notin> NN & Nonce N \<notin> used evs"
+
+
+
+
+
+
+
+subsection{*Basic properties of shrK*}
+
+(*Injectiveness: Agents' long-term keys are distinct.*)
+declare inj_shrK [THEN inj_eq, iff]
+declare inj_crdK [THEN inj_eq, iff]
+declare inj_pin  [THEN inj_eq, iff]
+
+lemma invKey_K [simp]: "invKey K = K"
+apply (insert isSym_keys)
+apply (simp add: symKeys_def) 
+done
+
+
+lemma analz_Decrypt' [dest]:
+     "\<lbrakk> Crypt K X \<in> analz H;  Key K  \<in> analz H \<rbrakk> \<Longrightarrow> X \<in> analz H"
+by auto
+
+text{*Now cancel the @{text dest} attribute given to
+ @{text analz.Decrypt} in its declaration.*}
+declare analz.Decrypt [rule del]
+
+text{*Rewrites should not refer to  @{term "initState(Friend i)"} because
+  that expression is not in normal form.*}
+
+text{*Added to extend initstate with set of nonces*}
+lemma parts_image_Nonce [simp]: "parts (Nonce`N) = Nonce`N"
+apply auto
+apply (erule parts.induct)
+apply auto
+done
+
+lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
+apply (unfold keysFor_def)
+apply (induct_tac "C", auto)
+done
+
+(*Specialized to shared-key model: no @{term invKey}*)
+lemma keysFor_parts_insert:
+     "\<lbrakk> K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) \<rbrakk> 
+     \<Longrightarrow> K \<in> keysFor (parts (G \<union> H)) | Key K \<in> parts H";
+by (force dest: EventSC.keysFor_parts_insert)  
+
+lemma Crypt_imp_keysFor: "Crypt K X \<in> H \<Longrightarrow> K \<in> keysFor H"
+by (drule Crypt_imp_invKey_keysFor, simp)
+
+
+subsection{*Function "knows"*}
+
+(*Spy knows the pins of bad agents!*)
+lemma Spy_knows_bad [intro!]: "A \<in> bad \<Longrightarrow> Key (pin A) \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
+done
+
+(*Spy knows the long-term keys of cloned cards!*)
+lemma Spy_knows_cloned [intro!]: 
+     "Card A \<in> cloned \<Longrightarrow>  Key (crdK (Card A)) \<in> knows Spy evs &   
+                           Key (shrK A) \<in> knows Spy evs &  
+                           Key (pin A)  \<in> knows Spy evs &  
+                          (\<forall> B. Key (pairK(B,A)) \<in> knows Spy evs)"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
+done
+
+lemma Spy_knows_cloned1 [intro!]: "C \<in> cloned \<Longrightarrow> Key (crdK C) \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
+done
+
+lemma Spy_knows_cloned2 [intro!]: "\<lbrakk> Card A \<in> cloned; Card B \<in> cloned \<rbrakk>  
+   \<Longrightarrow> Nonce (Pairkey(A,B))\<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
+done
+
+(*Spy only knows pins of bad agents!*)
+lemma Spy_knows_Spy_bad [intro!]: "A\<in> bad \<Longrightarrow> Key (pin A) \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
+done
+
+
+(*For case analysis on whether or not an agent is compromised*)
+lemma Crypt_Spy_analz_bad: 
+  "\<lbrakk> Crypt (pin A) X \<in> analz (knows Spy evs);  A\<in>bad \<rbrakk>   
+      \<Longrightarrow> X \<in> analz (knows Spy evs)"
+apply (force dest!: analz.Decrypt)
+done
+
+(** Fresh keys never clash with other keys **)
+
+lemma shrK_in_initState [iff]: "Key (shrK A) \<in> initState Server"
+apply (induct_tac "A")
+apply auto
+done
+
+lemma shrK_in_used [iff]: "Key (shrK A) \<in> used evs"
+apply (rule initState_into_used)
+apply blast
+done
+
+lemma crdK_in_initState [iff]: "Key (crdK A) \<in> initState Server"
+apply (induct_tac "A")
+apply auto
+done
+
+lemma crdK_in_used [iff]: "Key (crdK A) \<in> used evs"
+apply (rule initState_into_used)
+apply blast
+done
+
+lemma pin_in_initState [iff]: "Key (pin A) \<in> initState A"
+apply (induct_tac "A")
+apply auto
+done
+
+lemma pin_in_used [iff]: "Key (pin A) \<in> used evs"
+apply (rule initState_into_used)
+apply blast
+done
+
+lemma pairK_in_initState [iff]: "Key (pairK X) \<in> initState Server"
+apply (induct_tac "X")
+apply auto
+done
+
+lemma pairK_in_used [iff]: "Key (pairK X) \<in> used evs"
+apply (rule initState_into_used)
+apply blast
+done
+
+
+
+(*Used in parts_induct_tac and analz_Fake_tac to distinguish session keys
+  from long-term shared keys*)
+lemma Key_not_used [simp]: "Key K \<notin> used evs \<Longrightarrow> K \<notin> range shrK"
+by blast
+
+lemma shrK_neq [simp]: "Key K \<notin> used evs \<Longrightarrow> shrK B \<noteq> K"
+by blast
+
+lemma crdK_not_used [simp]: "Key K \<notin> used evs \<Longrightarrow> K \<notin> range crdK"
+apply clarify
+done
+
+lemma crdK_neq [simp]: "Key K \<notin> used evs \<Longrightarrow> crdK C \<noteq> K"
+apply clarify
+done
+
+lemma pin_not_used [simp]: "Key K \<notin> used evs \<Longrightarrow> K \<notin> range pin"
+apply clarify
+done
+
+lemma pin_neq [simp]: "Key K \<notin> used evs \<Longrightarrow> pin A \<noteq> K"
+apply clarify
+done
+
+lemma pairK_not_used [simp]: "Key K \<notin> used evs \<Longrightarrow> K \<notin> range pairK"
+apply clarify
+done
+
+lemma pairK_neq [simp]: "Key K \<notin> used evs \<Longrightarrow> pairK(A,B) \<noteq> K"
+apply clarify
+done
+
+declare shrK_neq [THEN not_sym, simp]
+declare crdK_neq [THEN not_sym, simp]
+declare pin_neq [THEN not_sym, simp]
+declare pairK_neq [THEN not_sym, simp]
+
+
+subsection{*Fresh nonces*}
+
+lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState (Friend i))"
+by auto
+
+
+(*This lemma no longer holds of smartcard protocols, where the cards can store
+  nonces.
+
+lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []"
+apply (simp (no_asm) add: used_Nil)
+done
+
+So, we must use old-style supply fresh nonce theorems relying on the appropriate axiom*)
+
+
+subsection{*Supply fresh nonces for possibility theorems.*}
+
+
+lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs"
+apply (rule Finites.emptyI [THEN Nonce_supply_ax, THEN exE], blast)
+done
+
+lemma Nonce_supply2: 
+  "\<exists>N N'. Nonce N \<notin> used evs & Nonce N' \<notin> used evs' & N \<noteq> N'"
+apply (cut_tac evs = evs in Finites.emptyI [THEN Nonce_supply_ax])
+apply (erule exE)
+apply (cut_tac evs = evs' in Finites.emptyI [THEN Finites.insertI, THEN Nonce_supply_ax]) 
+apply auto
+done
+
+
+lemma Nonce_supply3: "\<exists>N N' N''. Nonce N \<notin> used evs & Nonce N' \<notin> used evs' &  
+                    Nonce N'' \<notin> used evs'' & N \<noteq> N' & N' \<noteq> N'' & N \<noteq> N''"
+apply (cut_tac evs = evs in Finites.emptyI [THEN Nonce_supply_ax])
+apply (erule exE)
+apply (cut_tac evs = evs' and a1 = N in Finites.emptyI [THEN Finites.insertI, THEN Nonce_supply_ax]) 
+apply (erule exE)
+apply (cut_tac evs = evs'' and a1 = Na and a2 = N in Finites.emptyI [THEN Finites.insertI, THEN Finites.insertI, THEN Nonce_supply_ax]) 
+apply blast
+done
+
+lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs"
+apply (rule Finites.emptyI [THEN Nonce_supply_ax, THEN exE])
+apply (rule someI, blast)
+done
+
+
+
+text{*Unlike the corresponding property of nonces, we cannot prove
+    @{term "finite KK \<Longrightarrow> \<exists>K. K \<notin> KK & Key K \<notin> used evs"}.
+    We have infinitely many agents and there is nothing to stop their
+    long-term keys from exhausting all the natural numbers.  Instead,
+    possibility theorems must assume the existence of a few keys.*}
+
+
+subsection{*Tactics for possibility theorems*}
+
+ML
+{*
+val inj_shrK      = thm "inj_shrK";
+val isSym_keys    = thm "isSym_keys";
+val Nonce_supply = thm "Nonce_supply";
+val invKey_K = thm "invKey_K";
+val analz_Decrypt' = thm "analz_Decrypt'";
+val keysFor_parts_initState = thm "keysFor_parts_initState";
+val keysFor_parts_insert = thm "keysFor_parts_insert";
+val Crypt_imp_keysFor = thm "Crypt_imp_keysFor";
+val Spy_knows_Spy_bad = thm "Spy_knows_Spy_bad";
+val Crypt_Spy_analz_bad = thm "Crypt_Spy_analz_bad";
+val shrK_in_initState = thm "shrK_in_initState";
+val shrK_in_used = thm "shrK_in_used";
+val Key_not_used = thm "Key_not_used";
+val shrK_neq = thm "shrK_neq";
+val Nonce_notin_initState = thm "Nonce_notin_initState";
+val Nonce_supply1 = thm "Nonce_supply1";
+val Nonce_supply2 = thm "Nonce_supply2";
+val Nonce_supply3 = thm "Nonce_supply3";
+val Nonce_supply = thm "Nonce_supply";
+val used_Says = thm "used_Says";
+val used_Gets = thm "used_Gets";
+val used_Notes = thm "used_Notes";
+val used_Inputs = thm "used_Inputs";
+val used_C_Gets = thm "used_C_Gets";
+val used_Outpts = thm "used_Outpts";
+val used_A_Gets = thm "used_A_Gets";
+*}
+
+
+ML
+{*
+(*Omitting used_Says makes the tactic much faster: it leaves expressions
+    such as  Nonce ?N \<notin> used evs that match Nonce_supply*)
+fun gen_possibility_tac ss state = state |>
+   (REPEAT 
+    (ALLGOALS (simp_tac (ss delsimps [used_Says, used_Notes, used_Gets,
+                         used_Inputs, used_C_Gets, used_Outpts, used_A_Gets] 
+                         setSolver safe_solver))
+     THEN
+     REPEAT_FIRST (eq_assume_tac ORELSE' 
+                   resolve_tac [refl, conjI, Nonce_supply])))
+
+(*Tactic for possibility theorems (ML script version)*)
+fun possibility_tac state = gen_possibility_tac (simpset()) state
+
+(*For harder protocols (such as Recur) where we have to set up some
+  nonces and keys initially*)
+fun basic_possibility_tac st = st |>
+    REPEAT 
+    (ALLGOALS (asm_simp_tac (simpset() setSolver safe_solver))
+     THEN
+     REPEAT_FIRST (resolve_tac [refl, conjI]))
+*}
+
+subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
+
+lemma subset_Compl_range_shrK: "A \<subseteq> - (range shrK) \<Longrightarrow> shrK x \<notin> A"
+by blast
+
+lemma subset_Compl_range_crdK: "A \<subseteq> - (range crdK) \<Longrightarrow> crdK x \<notin> A"
+apply blast
+done
+
+lemma subset_Compl_range_pin: "A \<subseteq> - (range pin) \<Longrightarrow> pin x \<notin> A"
+apply blast
+done
+
+lemma subset_Compl_range_pairK: "A \<subseteq> - (range pairK) \<Longrightarrow> pairK x \<notin> A"
+apply blast
+done
+lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \<union> H"
+by blast
+
+lemma insert_Key_image: "insert (Key K) (Key`KK \<union> C) = Key`(insert K KK) \<union> C"
+by blast
+
+(** Reverse the normal simplification of "image" to build up (not break down)
+    the set of keys.  Use analz_insert_eq with (Un_upper2 RS analz_mono) to
+    erase occurrences of forwarded message components (X). **)
+
+lemmas analz_image_freshK_simps =
+       simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
+       disj_comms 
+       image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset
+       analz_insert_eq Un_upper2 [THEN analz_mono, THEN [2] rev_subsetD]
+       insert_Key_singleton subset_Compl_range_shrK subset_Compl_range_crdK
+       subset_Compl_range_pin subset_Compl_range_pairK
+       Key_not_used insert_Key_image Un_assoc [THEN sym]
+
+(*Lemma for the trivial direction of the if-and-only-if*)
+lemma analz_image_freshK_lemma:
+     "(Key K \<in> analz (Key`nE \<union> H)) \<longrightarrow> (K \<in> nE | Key K \<in> analz H)  \<Longrightarrow>  
+         (Key K \<in> analz (Key`nE \<union> H)) = (K \<in> nE | Key K \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+ML
+{*
+val analz_image_freshK_lemma = thm "analz_image_freshK_lemma";
+
+val analz_image_freshK_ss = 
+     simpset() delsimps [image_insert, image_Un]
+	       delsimps [imp_disjL]    (*reduces blow-up*)
+	       addsimps thms "analz_image_freshK_simps"
+*}
+
+
+
+(*Lets blast_tac perform this step without needing the simplifier*)
+lemma invKey_shrK_iff [iff]:
+     "(Key (invKey K) \<in> X) = (Key K \<in> X)"
+by auto
+
+(*Specialized methods*)
+
+method_setup analz_freshK = {*
+    Method.no_args
+     (Method.METHOD
+      (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+                          REPEAT_FIRST (rtac analz_image_freshK_lemma),
+                          ALLGOALS (asm_simp_tac analz_image_freshK_ss)])) *}
+    "for proving the Session Key Compromise theorem"
+
+method_setup possibility = {*
+    Method.ctxt_args (fn ctxt =>
+        Method.METHOD (fn facts =>
+            gen_possibility_tac (Simplifier.get_local_simpset ctxt))) *}
+    "for proving possibility theorems"
+
+lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)"
+by (induct e, auto simp: knows_Cons)
+
+(*Needed for actual protocols that will follow*)
+declare shrK_disj_crdK[THEN not_sym, iff]
+declare shrK_disj_pin[THEN not_sym, iff]
+declare pairK_disj_shrK[THEN not_sym, iff]
+declare pairK_disj_crdK[THEN not_sym, iff]
+declare pairK_disj_pin[THEN not_sym, iff]
+declare crdK_disj_pin[THEN not_sym, iff]
+
+declare legalUse_def [iff] illegalUse_def [iff]
+
+
+
+
+
+end
--- a/src/HOL/IsaMakefile	Wed Feb 01 12:23:14 2006 +0100
+++ b/src/HOL/IsaMakefile	Wed Feb 01 15:22:02 2006 +0100
@@ -398,10 +398,11 @@
 $(LOG)/HOL-Auth.gz: $(OUT)/HOL Library/NatPair.thy \
   Auth/CertifiedEmail.thy Auth/Event.thy \
   Auth/Message.thy Auth/NS_Public.thy Auth/NS_Public_Bad.thy \
-  Auth/NS_Shared.thy Auth/OtwayRees.thy Auth/OtwayRees_AN.thy \
+  Auth/NS_Shared.thy Auth/OtwayRees.thy Auth/OtwayReesBella.thy Auth/OtwayRees_AN.thy \
   Auth/OtwayRees_Bad.thy Auth/Public.thy Auth/ROOT.ML \
   Auth/Recur.thy Auth/Shared.thy \
-  Auth/TLS.thy Auth/WooLam.thy Auth/Kerberos_BAN.thy Auth/KerberosIV.thy \
+  Auth/TLS.thy Auth/WooLam.thy Auth/Kerberos_BAN.thy Auth/Kerberos_BAN_Gets.thy \
+  Auth/KerberosIV.thy Auth/KerberosIV_Gets.thy Auth/KerberosV.thy \
   Auth/Yahalom.thy Auth/Yahalom2.thy Auth/Yahalom_Bad.thy \
   Auth/ZhouGollmann.thy \
   Auth/Guard/Analz.thy Auth/Guard/Extensions.thy Auth/Guard/GuardK.thy \
@@ -410,6 +411,8 @@
   Auth/Guard/Guard_NS_Public.thy Auth/Guard/Guard_OtwayRees.thy \
   Auth/Guard/P1.thy Auth/Guard/P2.thy \
   Auth/Guard/Proto.thy Auth/Guard/Guard_Yahalom.thy\
+  Auth/Smartcard/EventSC.thy Auth/Smartcard/ShoupRubinBella.thy\
+  Auth/Smartcard/ShoupRubin.thy Auth/Smartcard/Smartcard.thy\
   Auth/document/root.tex 
 	@$(ISATOOL) usedir -g true $(OUT)/HOL Auth