# HG changeset patch # User paulson # Date 1477315865 -3600 # Node ID 464420ba7f74e27f5aaf960a5b7921ab8e69e597 # Parent 8a0fe5469ba08e00d729cef19c57f7408716f902 "subgoal" examples diff -r 8a0fe5469ba0 -r 464420ba7f74 src/HOL/Auth/Yahalom.thy --- a/src/HOL/Auth/Yahalom.thy Sun Oct 23 16:44:17 2016 +0200 +++ b/src/HOL/Auth/Yahalom.thy Mon Oct 24 14:31:05 2016 +0100 @@ -22,32 +22,32 @@ (*The spy MAY say anything he CAN say. We do not expect him to invent new nonces here, but he can also use NS1. Common to all similar protocols.*) - | Fake: "[| evsf \ yahalom; X \ synth (analz (knows Spy evsf)) |] - ==> Says Spy B X # evsf \ yahalom" + | Fake: "\evsf \ yahalom; X \ synth (analz (knows Spy evsf))\ + \ Says Spy B X # evsf \ yahalom" (*A message that has been sent can be received by the intended recipient.*) - | Reception: "[| evsr \ yahalom; Says A B X \ set evsr |] - ==> Gets B X # evsr \ yahalom" + | Reception: "\evsr \ yahalom; Says A B X \ set evsr\ + \ Gets B X # evsr \ yahalom" (*Alice initiates a protocol run*) - | YM1: "[| evs1 \ yahalom; Nonce NA \ used evs1 |] - ==> Says A B \Agent A, Nonce NA\ # evs1 \ yahalom" + | YM1: "\evs1 \ yahalom; Nonce NA \ used evs1\ + \ Says A B \Agent A, Nonce NA\ # evs1 \ yahalom" (*Bob's response to Alice's message.*) - | YM2: "[| evs2 \ yahalom; Nonce NB \ used evs2; - Gets B \Agent A, Nonce NA\ \ set evs2 |] - ==> Says B Server + | YM2: "\evs2 \ yahalom; Nonce NB \ used evs2; + Gets B \Agent A, Nonce NA\ \ set evs2\ + \ Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, Nonce NB\\ # evs2 \ yahalom" (*The Server receives Bob's message. He responds by sending a new session key to Alice, with a packet for forwarding to Bob.*) - | YM3: "[| evs3 \ yahalom; Key KAB \ used evs3; KAB \ symKeys; + | YM3: "\evs3 \ yahalom; Key KAB \ used evs3; KAB \ symKeys; Gets Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, Nonce NB\\ - \ set evs3 |] - ==> Says Server A + \ set evs3\ + \ Says Server A \Crypt (shrK A) \Agent B, Key KAB, Nonce NA, Nonce NB\, Crypt (shrK B) \Agent A, Key KAB\\ # evs3 \ yahalom" @@ -57,20 +57,20 @@ uses the new session key to send Bob his Nonce. The premise @{term "A \ Server"} is needed for \Says_Server_not_range\. Alice can check that K is symmetric by its length.\ - "[| evs4 \ yahalom; A \ Server; K \ symKeys; + "\evs4 \ yahalom; A \ Server; K \ symKeys; Gets A \Crypt(shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, X\ \ set evs4; - Says A B \Agent A, Nonce NA\ \ set evs4 |] - ==> Says A B \X, Crypt K (Nonce NB)\ # evs4 \ yahalom" + Says A B \Agent A, Nonce NA\ \ set evs4\ + \ Says A B \X, Crypt K (Nonce NB)\ # evs4 \ yahalom" (*This message models possible leaks of session keys. The Nonces identify the protocol run. Quoting Server here ensures they are correct.*) - | Oops: "[| evso \ yahalom; + | Oops: "\evso \ yahalom; Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, - X\ \ set evso |] - ==> Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ yahalom" + X\ \ set evso\ + \ Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ yahalom" definition KeyWithNonce :: "[key, nat, event list] => bool" where @@ -86,8 +86,8 @@ declare analz_into_parts [dest] text\A "possibility property": there are traces that reach the end\ -lemma "[| A \ Server; K \ symKeys; Key K \ used [] |] - ==> \X NB. \evs \ yahalom. +lemma "\A \ Server; K \ symKeys; Key K \ used []\ + \ \X NB. \evs \ yahalom. Says A B \X, Crypt K (Nonce NB)\ \ set evs" apply (intro exI bexI) apply (rule_tac [2] yahalom.Nil @@ -102,12 +102,12 @@ subsection\Regularity Lemmas for Yahalom\ lemma Gets_imp_Says: - "[| Gets B X \ set evs; evs \ yahalom |] ==> \A. Says A B X \ set evs" + "\Gets B X \ set evs; evs \ yahalom\ \ \A. Says A B X \ set evs" by (erule rev_mp, erule yahalom.induct, auto) text\Must be proved separately for each protocol\ lemma Gets_imp_knows_Spy: - "[| Gets B X \ set evs; evs \ yahalom |] ==> X \ knows Spy evs" + "\Gets B X \ set evs; evs \ yahalom\ \ X \ knows Spy evs" by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) lemmas Gets_imp_analz_Spy = Gets_imp_knows_Spy [THEN analz.Inj] @@ -116,8 +116,8 @@ text\Lets us treat YM4 using a similar argument as for the Fake case.\ lemma YM4_analz_knows_Spy: - "[| Gets A \Crypt (shrK A) Y, X\ \ set evs; evs \ yahalom |] - ==> X \ analz (knows Spy evs)" + "\Gets A \Crypt (shrK A) Y, X\ \ set evs; evs \ yahalom\ + \ X \ analz (knows Spy evs)" by blast lemmas YM4_parts_knows_Spy = @@ -126,7 +126,7 @@ text\For Oops\ lemma YM4_Key_parts_knows_Spy: "Says Server A \Crypt (shrK A) \B,K,NA,NB\, X\ \ set evs - ==> K \ parts (knows Spy evs)" + \ K \ parts (knows Spy evs)" by (metis parts.Body parts.Fst parts.Snd Says_imp_knows_Spy parts.Inj) text\Theorems of the form @{term "X \ parts (knows Spy evs)"} imply @@ -134,23 +134,23 @@ text\Spy never sees a good agent's shared key!\ lemma Spy_see_shrK [simp]: - "evs \ yahalom ==> (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" + "evs \ yahalom \ (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" by (erule yahalom.induct, force, drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+) lemma Spy_analz_shrK [simp]: - "evs \ yahalom ==> (Key (shrK A) \ analz (knows Spy evs)) = (A \ bad)" + "evs \ yahalom \ (Key (shrK A) \ analz (knows Spy evs)) = (A \ bad)" by auto lemma Spy_see_shrK_D [dest!]: - "[|Key (shrK A) \ parts (knows Spy evs); evs \ yahalom|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ yahalom\ \ A \ bad" by (blast dest: Spy_see_shrK) text\Nobody can have used non-existent keys! Needed to apply \analz_insert_Key\\ lemma new_keys_not_used [simp]: - "[|Key K \ used evs; K \ symKeys; evs \ yahalom|] - ==> K \ keysFor (parts (spies evs))" + "\Key K \ used evs; K \ symKeys; evs \ yahalom\ + \ K \ keysFor (parts (spies evs))" apply (erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy, simp_all) @@ -162,17 +162,17 @@ text\Earlier, all protocol proofs declared this theorem. But only a few proofs need it, e.g. Yahalom and Kerberos IV.\ lemma new_keys_not_analzd: - "[|K \ symKeys; evs \ yahalom; Key K \ used evs|] - ==> K \ keysFor (analz (knows Spy evs))" + "\K \ symKeys; evs \ yahalom; Key K \ used evs\ + \ K \ keysFor (analz (knows Spy evs))" by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD]) text\Describes the form of K when the Server sends this message. Useful for Oops as well as main secrecy property.\ lemma Says_Server_not_range [simp]: - "[| Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, X\ - \ set evs; evs \ yahalom |] - ==> K \ range shrK" + "\Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, X\ + \ set evs; evs \ yahalom\ + \ K \ range shrK" by (erule rev_mp, erule yahalom.induct, simp_all) @@ -181,7 +181,7 @@ (**** The following is to prove theorems of the form - Key K \ analz (insert (Key KAB) (knows Spy evs)) ==> + Key K \ analz (insert (Key KAB) (knows Spy evs)) \ Key K \ analz (knows Spy evs) A more general formula must be proved inductively. @@ -190,7 +190,7 @@ text\Session keys are not used to encrypt other session keys\ lemma analz_image_freshK [rule_format]: - "evs \ yahalom ==> + "evs \ yahalom \ \K KK. KK <= - (range shrK) --> (Key K \ analz (Key`KK Un (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" @@ -201,7 +201,7 @@ done lemma analz_insert_freshK: - "[| evs \ yahalom; KAB \ range shrK |] ==> + "\evs \ yahalom; KAB \ range shrK\ \ (Key K \ analz (insert (Key KAB) (knows Spy evs))) = (K = KAB | Key K \ analz (knows Spy evs))" by (simp only: analz_image_freshK analz_image_freshK_simps) @@ -209,12 +209,12 @@ text\The Key K uniquely identifies the Server's message.\ lemma unique_session_keys: - "[| Says Server A + "\Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, X\ \ set evs; Says Server A' \Crypt (shrK A') \Agent B', Key K, na', nb'\, X'\ \ set evs; - evs \ yahalom |] - ==> A=A' & B=B' & na=na' & nb=nb'" + evs \ yahalom\ + \ A=A' & B=B' & na=na' & nb=nb'" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, simp_all) txt\YM3, by freshness, and YM4\ @@ -224,8 +224,8 @@ text\Crucial secrecy property: Spy does not see the keys sent in msg YM3\ lemma secrecy_lemma: - "[| A \ bad; B \ bad; evs \ yahalom |] - ==> Says Server A + "\A \ bad; B \ bad; evs \ yahalom\ + \ Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, Crypt (shrK B) \Agent A, Key K\\ \ set evs --> @@ -233,19 +233,21 @@ Key K \ analz (knows Spy evs)" apply (erule yahalom.induct, force, drule_tac [6] YM4_analz_knows_Spy) -apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz) \\Fake\ -apply (blast dest: unique_session_keys)+ \\YM3, Oops\ +apply (simp_all add: pushes analz_insert_eq analz_insert_freshK) + subgoal --\Fake\ by spy_analz + subgoal --\YM3\ by blast + subgoal --\Oops\ by (blast dest: unique_session_keys) done text\Final version\ lemma Spy_not_see_encrypted_key: - "[| Says Server A + "\Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, Crypt (shrK B) \Agent A, Key K\\ \ set evs; Notes Spy \na, nb, Key K\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ analz (knows Spy evs)" by (blast dest: secrecy_lemma) @@ -253,9 +255,9 @@ text\If the encrypted message appears then it originated with the Server\ lemma A_trusts_YM3: - "[| Crypt (shrK A) \Agent B, Key K, na, nb\ \ parts (knows Spy evs); - A \ bad; evs \ yahalom |] - ==> Says Server A + "\Crypt (shrK A) \Agent B, Key K, na, nb\ \ parts (knows Spy evs); + A \ bad; evs \ yahalom\ + \ Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, Crypt (shrK B) \Agent A, Key K\\ \ set evs" @@ -269,10 +271,10 @@ text\The obvious combination of \A_trusts_YM3\ with \Spy_not_see_encrypted_key\\ lemma A_gets_good_key: - "[| Crypt (shrK A) \Agent B, Key K, na, nb\ \ parts (knows Spy evs); + "\Crypt (shrK A) \Agent B, Key K, na, nb\ \ parts (knows Spy evs); Notes Spy \na, nb, Key K\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ analz (knows Spy evs)" by (metis A_trusts_YM3 secrecy_lemma) @@ -281,9 +283,9 @@ text\B knows, by the first part of A's message, that the Server distributed the key for A and B. But this part says nothing about nonces.\ lemma B_trusts_YM4_shrK: - "[| Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs); - B \ bad; evs \ yahalom |] - ==> \NA NB. Says Server A + "\Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs); + B \ bad; evs \ yahalom\ + \ \NA NB. Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, Crypt (shrK B) \Agent A, Key K\\ @@ -302,19 +304,18 @@ induction formula.\ lemma B_trusts_YM4_newK [rule_format]: - "[|Crypt K (Nonce NB) \ parts (knows Spy evs); - Nonce NB \ analz (knows Spy evs); evs \ yahalom|] - ==> \A B NA. Says Server A + "\Crypt K (Nonce NB) \ parts (knows Spy evs); + Nonce NB \ analz (knows Spy evs); evs \ yahalom\ + \ \A B NA. Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, Crypt (shrK B) \Agent A, Key K\\ \ set evs" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) -apply (analz_mono_contra, simp_all) -txt\Fake, YM3\ -apply blast -apply blast + apply (analz_mono_contra, simp_all) + subgoal --\Fake\ by blast + subgoal --\YM3\ by blast txt\YM4. A is uncompromised because NB is secure A's certificate guarantees the existence of the Server message\ apply (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad @@ -330,7 +331,7 @@ lemma KeyWithNonceI: "Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB\, X\ - \ set evs ==> KeyWithNonce K NB evs" + \ set evs \ KeyWithNonce K NB evs" by (unfold KeyWithNonce_def, blast) lemma KeyWithNonce_Says [simp]: @@ -352,16 +353,16 @@ text\A fresh key cannot be associated with any nonce (with respect to a given trace).\ lemma fresh_not_KeyWithNonce: - "Key K \ used evs ==> ~ KeyWithNonce K NB evs" + "Key K \ used evs \ ~ KeyWithNonce K NB evs" by (unfold KeyWithNonce_def, blast) text\The Server message associates K with NB' and therefore not with any other nonce NB.\ lemma Says_Server_KeyWithNonce: - "[| Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB'\, X\ + "\Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB'\, X\ \ set evs; - NB \ NB'; evs \ yahalom |] - ==> ~ KeyWithNonce K NB evs" + NB \ NB'; evs \ yahalom\ + \ ~ KeyWithNonce K NB evs" by (unfold KeyWithNonce_def, blast dest: unique_session_keys) @@ -373,12 +374,12 @@ text\As with \analz_image_freshK\, we take some pains to express the property as a logical equivalence so that the simplifier can apply it.\ lemma Nonce_secrecy_lemma: - "P --> (X \ analz (G Un H)) --> (X \ analz H) ==> + "P --> (X \ analz (G Un H)) --> (X \ analz H) \ P --> (X \ analz (G Un H)) = (X \ analz H)" by (blast intro: analz_mono [THEN subsetD]) lemma Nonce_secrecy: - "evs \ yahalom ==> + "evs \ yahalom \ (\KK. KK <= - (range shrK) --> (\K \ KK. K \ symKeys --> ~ KeyWithNonce K NB evs) --> (Nonce NB \ analz (Key`KK Un (knows Spy evs))) = @@ -396,19 +397,12 @@ @{term Says_Server_KeyWithNonce}, we get @{prop "~ KeyWithNonce K NB evs"}; then simplification can apply the induction hypothesis with @{term "KK = {K}"}.\ -txt\Fake\ -apply spy_analz -txt\YM2\ -apply blast -txt\YM3\ -apply blast -txt\YM4\ -apply (erule_tac V = "\KK. P KK" for P in thin_rl, clarify) -txt\If @{prop "A \ bad"} then @{term NBa} is known, therefore - @{prop "NBa \ NB"}. Previous two steps make the next step - faster.\ -apply (metis A_trusts_YM3 Gets_imp_analz_Spy Gets_imp_knows_Spy KeyWithNonce_def - Spy_analz_shrK analz.Fst analz.Snd analz_shrK_Decrypt parts.Fst parts.Inj) + subgoal --\Fake\ by spy_analz + subgoal --\YM2\ by blast + subgoal --\YM3\ by blast + subgoal --\YM4: If @{prop "A \ bad"} then @{term NBa} is known, therefore @{prop "NBa \ NB"}.\ + by (metis A_trusts_YM3 Gets_imp_analz_Spy Gets_imp_knows_Spy KeyWithNonce_def + Spy_analz_shrK analz.Fst analz.Snd analz_shrK_Decrypt parts.Fst parts.Inj) done @@ -416,11 +410,11 @@ it was distributed with that key. The more general form above is required for the induction to carry through.\ lemma single_Nonce_secrecy: - "[| Says Server A + "\Says Server A \Crypt (shrK A) \Agent B, Key KAB, na, Nonce NB'\, X\ \ set evs; - NB \ NB'; KAB \ range shrK; evs \ yahalom |] - ==> (Nonce NB \ analz (insert (Key KAB) (knows Spy evs))) = + NB \ NB'; KAB \ range shrK; evs \ yahalom\ + \ (Nonce NB \ analz (insert (Key KAB) (knows Spy evs))) = (Nonce NB \ analz (knows Spy evs))" by (simp_all del: image_insert image_Un imp_disjL add: analz_image_freshK_simps split_ifs @@ -430,10 +424,10 @@ subsubsection\The Nonce NB uniquely identifies B's message.\ lemma unique_NB: - "[| Crypt (shrK B) \Agent A, Nonce NA, nb\ \ parts (knows Spy evs); + "\Crypt (shrK B) \Agent A, Nonce NA, nb\ \ parts (knows Spy evs); Crypt (shrK B') \Agent A', Nonce NA', nb\ \ parts (knows Spy evs); - evs \ yahalom; B \ bad; B' \ bad |] - ==> NA' = NA & A' = A & B' = B" + evs \ yahalom; B \ bad; B' \ bad\ + \ NA' = NA & A' = A & B' = B" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy, simp_all) @@ -445,12 +439,12 @@ text\Variant useful for proving secrecy of NB. Because nb is assumed to be secret, we no longer must assume B, B' not bad.\ lemma Says_unique_NB: - "[| Says C S \X, Crypt (shrK B) \Agent A, Nonce NA, nb\\ + "\Says C S \X, Crypt (shrK B) \Agent A, Nonce NA, nb\\ \ set evs; Gets S' \X', Crypt (shrK B') \Agent A', Nonce NA', nb\\ \ set evs; - nb \ analz (knows Spy evs); evs \ yahalom |] - ==> NA' = NA & A' = A & B' = B" + nb \ analz (knows Spy evs); evs \ yahalom\ + \ NA' = NA & A' = A & B' = B" by (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad dest: Says_imp_spies unique_NB parts.Inj analz.Inj) @@ -458,9 +452,9 @@ subsubsection\A nonce value is never used both as NA and as NB\ lemma no_nonce_YM1_YM2: - "[|Crypt (shrK B') \Agent A', Nonce NB, nb'\ \ parts(knows Spy evs); - Nonce NB \ analz (knows Spy evs); evs \ yahalom|] - ==> Crypt (shrK B) \Agent A, na, Nonce NB\ \ parts(knows Spy evs)" + "\Crypt (shrK B') \Agent A', Nonce NB, nb'\ \ parts(knows Spy evs); + Nonce NB \ analz (knows Spy evs); evs \ yahalom\ + \ Crypt (shrK B) \Agent A, na, Nonce NB\ \ parts(knows Spy evs)" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) @@ -471,64 +465,60 @@ text\The Server sends YM3 only in response to YM2.\ lemma Says_Server_imp_YM2: - "[| Says Server A \Crypt (shrK A) \Agent B, k, na, nb\, X\ \ set evs; - evs \ yahalom |] - ==> Gets Server \Agent B, Crypt (shrK B) \Agent A, na, nb\\ + "\Says Server A \Crypt (shrK A) \Agent B, k, na, nb\, X\ \ set evs; + evs \ yahalom\ + \ Gets Server \Agent B, Crypt (shrK B) \Agent A, na, nb\\ \ set evs" by (erule rev_mp, erule yahalom.induct, auto) text\A vital theorem for B, that nonce NB remains secure from the Spy.\ -lemma Spy_not_see_NB : - "[| Says B Server +theorem Spy_not_see_NB : + "\Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, Nonce NB\\ \ set evs; (\k. Notes Spy \Nonce NA, Nonce NB, k\ \ set evs); - A \ bad; B \ bad; evs \ yahalom |] - ==> Nonce NB \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Nonce NB \ analz (knows Spy evs)" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_analz_knows_Spy) apply (simp_all add: split_ifs pushes new_keys_not_analzd analz_insert_eq analz_insert_freshK) -txt\Fake\ -apply spy_analz -txt\YM1: NB=NA is impossible anyway, but NA is secret because it is fresh!\ -apply blast -txt\YM2\ -apply blast -txt\Prove YM3 by showing that no NB can also be an NA\ -apply (blast dest!: no_nonce_YM1_YM2 dest: Gets_imp_Says Says_unique_NB) -txt\LEVEL 7: YM4 and Oops remain\ -apply (clarify, simp add: all_conj_distrib) -txt\YM4: key K is visible to Spy, contradicting session key secrecy theorem\ -txt\Case analysis on Aa:bad; PROOF FAILED problems - use \Says_unique_NB\ to identify message components: @{term "Aa=A"}, @{term "Ba=B"}\ -apply (blast dest!: Says_unique_NB analz_shrK_Decrypt - parts.Inj [THEN parts.Fst, THEN A_trusts_YM3] - dest: Gets_imp_Says Says_imp_spies Says_Server_imp_YM2 - Spy_not_see_encrypted_key) -txt\Oops case: if the nonce is betrayed now, show that the Oops event is - covered by the quantified Oops assumption.\ -apply (clarify, simp add: all_conj_distrib) -apply (frule Says_Server_imp_YM2, assumption) -apply (metis Gets_imp_Says Says_Server_not_range Says_unique_NB no_nonce_YM1_YM2 parts.Snd single_Nonce_secrecy spies_partsEs(1)) + subgoal --\Fake\ by spy_analz + subgoal --\YM1: NB=NA is impossible anyway, but NA is secret because it is fresh!\ by blast + subgoal --\YM2\ by blast + subgoal --\YM3: because no NB can also be an NA\ + by (blast dest!: no_nonce_YM1_YM2 dest: Gets_imp_Says Says_unique_NB) + subgoal --\YM4: key K is visible to Spy, contradicting session key secrecy theorem\ + --\Case analysis on whether Aa is bad; + use \Says_unique_NB\ to identify message components: @{term "Aa=A"}, @{term "Ba=B"}\ + apply clarify + apply (blast dest!: Says_unique_NB analz_shrK_Decrypt + parts.Inj [THEN parts.Fst, THEN A_trusts_YM3] + dest: Gets_imp_Says Says_imp_spies Says_Server_imp_YM2 + Spy_not_see_encrypted_key) + done + subgoal --\Oops case: if the nonce is betrayed now, show that the Oops event is + covered by the quantified Oops assumption.\ + apply clarsimp + apply (metis Says_Server_imp_YM2 Gets_imp_Says Says_Server_not_range Says_unique_NB no_nonce_YM1_YM2 parts.Snd single_Nonce_secrecy spies_partsEs(1)) + done done - text\B's session key guarantee from YM4. The two certificates contribute to a single conclusion about the Server's message. Note that the "Notes Spy" assumption must quantify over \\\ POSSIBLE keys instead of our particular K. If this run is broken and the spy substitutes a certificate containing an old key, B has no means of telling.\ lemma B_trusts_YM4: - "[| Gets B \Crypt (shrK B) \Agent A, Key K\, + "\Gets B \Crypt (shrK B) \Agent A, Key K\, Crypt K (Nonce NB)\ \ set evs; Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, Nonce NB\\ \ set evs; \k. Notes Spy \Nonce NA, Nonce NB, k\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Says Server A + A \ bad; B \ bad; evs \ yahalom\ + \ Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, Crypt (shrK B) \Agent A, Key K\\ @@ -541,14 +531,14 @@ text\The obvious combination of \B_trusts_YM4\ with \Spy_not_see_encrypted_key\\ lemma B_gets_good_key: - "[| Gets B \Crypt (shrK B) \Agent A, Key K\, + "\Gets B \Crypt (shrK B) \Agent A, Key K\, Crypt K (Nonce NB)\ \ set evs; Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, Nonce NB\\ \ set evs; \k. Notes Spy \Nonce NA, Nonce NB, k\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ analz (knows Spy evs)" by (metis B_trusts_YM4 Spy_not_see_encrypted_key) @@ -556,9 +546,9 @@ text\The encryption in message YM2 tells us it cannot be faked.\ lemma B_Said_YM2 [rule_format]: - "[|Crypt (shrK B) \Agent A, Nonce NA, nb\ \ parts (knows Spy evs); - evs \ yahalom|] - ==> B \ bad --> + "\Crypt (shrK B) \Agent A, Nonce NA, nb\ \ parts (knows Spy evs); + evs \ yahalom\ + \ B \ bad --> Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, nb\\ \ set evs" apply (erule rev_mp, erule yahalom.induct, force, @@ -569,9 +559,9 @@ text\If the server sends YM3 then B sent YM2\ lemma YM3_auth_B_to_A_lemma: - "[|Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, nb\, X\ - \ set evs; evs \ yahalom|] - ==> B \ bad --> + "\Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, nb\, X\ + \ set evs; evs \ yahalom\ + \ B \ bad --> Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, nb\\ \ set evs" apply (erule rev_mp, erule yahalom.induct, simp_all) @@ -580,11 +570,11 @@ done text\If A receives YM3 then B has used nonce NA (and therefore is alive)\ -lemma YM3_auth_B_to_A: - "[| Gets A \Crypt (shrK A) \Agent B, Key K, Nonce NA, nb\, X\ +theorem YM3_auth_B_to_A: + "\Gets A \Crypt (shrK A) \Agent B, Key K, Nonce NA, nb\, X\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, nb\\ + A \ bad; B \ bad; evs \ yahalom\ + \ Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, nb\\ \ set evs" by (metis A_trusts_YM3 Gets_imp_analz_Spy YM3_auth_B_to_A_lemma analz.Fst not_parts_not_analz) @@ -596,9 +586,9 @@ text\Assuming the session key is secure, if both certificates are present then A has said NB. We can't be sure about the rest of A's message, but only NB matters for freshness.\ -lemma A_Said_YM3_lemma [rule_format]: +theorem A_Said_YM3_lemma [rule_format]: "evs \ yahalom - ==> Key K \ analz (knows Spy evs) --> + \ Key K \ analz (knows Spy evs) --> Crypt K (Nonce NB) \ parts (knows Spy evs) --> Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs) --> B \ bad --> @@ -606,31 +596,27 @@ apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) apply (analz_mono_contra, simp_all) -txt\Fake\ -apply blast -txt\YM3: by \new_keys_not_used\, the message - @{term "Crypt K (Nonce NB)"} could not exist\ -apply (force dest!: Crypt_imp_keysFor) -txt\YM4: was @{term "Crypt K (Nonce NB)"} the very last message? - If not, use the induction hypothesis\ -apply (simp add: ex_disj_distrib) -txt\yes: apply unicity of session keys\ -apply (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK - Crypt_Spy_analz_bad + subgoal --\Fake\ by blast + subgoal --\YM3 because the message @{term "Crypt K (Nonce NB)"} could not exist\ + by (force dest!: Crypt_imp_keysFor) + subgoal --\YM4: was @{term "Crypt K (Nonce NB)"} the very last message? If not, use the induction hypothesis, + otherwise by unicity of session keys\ + by (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK Crypt_Spy_analz_bad dest: Says_imp_knows_Spy [THEN parts.Inj] unique_session_keys) done text\If B receives YM4 then A has used nonce NB (and therefore is alive). Moreover, A associates K with NB (thus is talking about the same run). Other premises guarantee secrecy of K.\ -lemma YM4_imp_A_Said_YM3 [rule_format]: - "[| Gets B \Crypt (shrK B) \Agent A, Key K\, +theorem YM4_imp_A_Said_YM3 [rule_format]: + "\Gets B \Crypt (shrK B) \Agent A, Key K\, Crypt K (Nonce NB)\ \ set evs; Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, Nonce NB\\ \ set evs; (\NA k. Notes Spy \Nonce NA, Nonce NB, k\ \ set evs); - A \ bad; B \ bad; evs \ yahalom |] - ==> \X. Says A B \X, Crypt K (Nonce NB)\ \ set evs" + A \ bad; B \ bad; evs \ yahalom\ + \ \X. Says A B \X, Crypt K (Nonce NB)\ \ set evs" by (metis A_Said_YM3_lemma B_gets_good_key Gets_imp_analz_Spy YM4_parts_knows_Spy analz.Fst not_parts_not_analz) + end diff -r 8a0fe5469ba0 -r 464420ba7f74 src/HOL/Auth/Yahalom2.thy --- a/src/HOL/Auth/Yahalom2.thy Sun Oct 23 16:44:17 2016 +0200 +++ b/src/HOL/Auth/Yahalom2.thy Mon Oct 24 14:31:05 2016 +0100 @@ -26,33 +26,33 @@ (*The spy MAY say anything he CAN say. We do not expect him to invent new nonces here, but he can also use NS1. Common to all similar protocols.*) - | Fake: "[| evsf \ yahalom; X \ synth (analz (knows Spy evsf)) |] - ==> Says Spy B X # evsf \ yahalom" + | Fake: "\evsf \ yahalom; X \ synth (analz (knows Spy evsf))\ + \ Says Spy B X # evsf \ yahalom" (*A message that has been sent can be received by the intended recipient.*) - | Reception: "[| evsr \ yahalom; Says A B X \ set evsr |] - ==> Gets B X # evsr \ yahalom" + | Reception: "\evsr \ yahalom; Says A B X \ set evsr\ + \ Gets B X # evsr \ yahalom" (*Alice initiates a protocol run*) - | YM1: "[| evs1 \ yahalom; Nonce NA \ used evs1 |] - ==> Says A B \Agent A, Nonce NA\ # evs1 \ yahalom" + | YM1: "\evs1 \ yahalom; Nonce NA \ used evs1\ + \ Says A B \Agent A, Nonce NA\ # evs1 \ yahalom" (*Bob's response to Alice's message.*) - | YM2: "[| evs2 \ yahalom; Nonce NB \ used evs2; - Gets B \Agent A, Nonce NA\ \ set evs2 |] - ==> Says B Server + | YM2: "\evs2 \ yahalom; Nonce NB \ used evs2; + Gets B \Agent A, Nonce NA\ \ set evs2\ + \ Says B Server \Agent B, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ # evs2 \ yahalom" (*The Server receives Bob's message. He responds by sending a new session key to Alice, with a certificate for forwarding to Bob. Both agents are quoted in the 2nd certificate to prevent attacks!*) - | YM3: "[| evs3 \ yahalom; Key KAB \ used evs3; + | YM3: "\evs3 \ yahalom; Key KAB \ used evs3; Gets Server \Agent B, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ - \ set evs3 |] - ==> Says Server A + \ set evs3\ + \ Says Server A \Nonce NB, Crypt (shrK A) \Agent B, Key KAB, Nonce NA\, Crypt (shrK B) \Agent A, Agent B, Key KAB, Nonce NB\\ @@ -60,20 +60,20 @@ (*Alice receives the Server's (?) message, checks her Nonce, and uses the new session key to send Bob his Nonce.*) - | YM4: "[| evs4 \ yahalom; + | YM4: "\evs4 \ yahalom; Gets A \Nonce NB, Crypt (shrK A) \Agent B, Key K, Nonce NA\, X\ \ set evs4; - Says A B \Agent A, Nonce NA\ \ set evs4 |] - ==> Says A B \X, Crypt K (Nonce NB)\ # evs4 \ yahalom" + Says A B \Agent A, Nonce NA\ \ set evs4\ + \ Says A B \X, Crypt K (Nonce NB)\ # evs4 \ yahalom" (*This message models possible leaks of session keys. The nonces identify the protocol run. Quoting Server here ensures they are correct. *) - | Oops: "[| evso \ yahalom; + | Oops: "\evso \ yahalom; Says Server A \Nonce NB, Crypt (shrK A) \Agent B, Key K, Nonce NA\, - X\ \ set evso |] - ==> Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ yahalom" + X\ \ set evso\ + \ Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ yahalom" declare Says_imp_knows_Spy [THEN analz.Inj, dest] @@ -83,7 +83,7 @@ text\A "possibility property": there are traces that reach the end\ lemma "Key K \ used [] - ==> \X NB. \evs \ yahalom. + \ \X NB. \evs \ yahalom. Says A B \X, Crypt K (Nonce NB)\ \ set evs" apply (intro exI bexI) apply (rule_tac [2] yahalom.Nil @@ -95,12 +95,12 @@ done lemma Gets_imp_Says: - "[| Gets B X \ set evs; evs \ yahalom |] ==> \A. Says A B X \ set evs" + "\Gets B X \ set evs; evs \ yahalom\ \ \A. Says A B X \ set evs" by (erule rev_mp, erule yahalom.induct, auto) text\Must be proved separately for each protocol\ lemma Gets_imp_knows_Spy: - "[| Gets B X \ set evs; evs \ yahalom |] ==> X \ knows Spy evs" + "\Gets B X \ set evs; evs \ yahalom\ \ X \ knows Spy evs" by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) declare Gets_imp_knows_Spy [THEN analz.Inj, dest] @@ -111,8 +111,8 @@ text\Result for reasoning about the encrypted portion of messages. Lets us treat YM4 using a similar argument as for the Fake case.\ lemma YM4_analz_knows_Spy: - "[| Gets A \NB, Crypt (shrK A) Y, X\ \ set evs; evs \ yahalom |] - ==> X \ analz (knows Spy evs)" + "\Gets A \NB, Crypt (shrK A) Y, X\ \ set evs; evs \ yahalom\ + \ X \ analz (knows Spy evs)" by blast lemmas YM4_parts_knows_Spy = @@ -124,49 +124,45 @@ text\Spy never sees a good agent's shared key!\ lemma Spy_see_shrK [simp]: - "evs \ yahalom ==> (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" + "evs \ yahalom \ (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" by (erule yahalom.induct, force, drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+) lemma Spy_analz_shrK [simp]: - "evs \ yahalom ==> (Key (shrK A) \ analz (knows Spy evs)) = (A \ bad)" + "evs \ yahalom \ (Key (shrK A) \ analz (knows Spy evs)) = (A \ bad)" by auto lemma Spy_see_shrK_D [dest!]: - "[|Key (shrK A) \ parts (knows Spy evs); evs \ yahalom|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ yahalom\ \ A \ bad" by (blast dest: Spy_see_shrK) text\Nobody can have used non-existent keys! Needed to apply \analz_insert_Key\\ lemma new_keys_not_used [simp]: - "[|Key K \ used evs; K \ symKeys; evs \ yahalom|] - ==> K \ keysFor (parts (spies evs))" + "\Key K \ used evs; K \ symKeys; evs \ yahalom\ + \ K \ keysFor (parts (spies evs))" apply (erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy, simp_all) -txt\Fake\ -apply (force dest!: keysFor_parts_insert) -txt\YM3\ -apply blast -txt\YM4\ -apply auto -apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj]) +subgoal --\Fake\ by (force dest!: keysFor_parts_insert) +subgoal --\YM3 \by blast +subgoal --\YM4\ by (fastforce dest!: Gets_imp_knows_Spy [THEN parts.Inj]) done text\Describes the form of K when the Server sends this message. Useful for Oops as well as main secrecy property.\ lemma Says_Server_message_form: - "[| Says Server A \nb', Crypt (shrK A) \Agent B, Key K, na\, X\ - \ set evs; evs \ yahalom |] - ==> K \ range shrK" + "\Says Server A \nb', Crypt (shrK A) \Agent B, Key K, na\, X\ + \ set evs; evs \ yahalom\ + \ K \ range shrK" by (erule rev_mp, erule yahalom.induct, simp_all) (**** The following is to prove theorems of the form - Key K \ analz (insert (Key KAB) (knows Spy evs)) ==> + Key K \ analz (insert (Key KAB) (knows Spy evs)) \ Key K \ analz (knows Spy evs) A more general formula must be proved inductively. @@ -175,7 +171,7 @@ (** Session keys are not used to encrypt other session keys **) lemma analz_image_freshK [rule_format]: - "evs \ yahalom ==> + "evs \ yahalom \ \K KK. KK <= - (range shrK) --> (Key K \ analz (Key`KK Un (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" @@ -185,7 +181,7 @@ done lemma analz_insert_freshK: - "[| evs \ yahalom; KAB \ range shrK |] ==> + "\evs \ yahalom; KAB \ range shrK\ \ (Key K \ analz (insert (Key KAB) (knows Spy evs))) = (K = KAB | Key K \ analz (knows Spy evs))" by (simp only: analz_image_freshK analz_image_freshK_simps) @@ -193,12 +189,12 @@ text\The Key K uniquely identifies the Server's message\ lemma unique_session_keys: - "[| Says Server A + "\Says Server A \nb, Crypt (shrK A) \Agent B, Key K, na\, X\ \ set evs; Says Server A' \nb', Crypt (shrK A') \Agent B', Key K, na'\, X'\ \ set evs; - evs \ yahalom |] - ==> A=A' & B=B' & na=na' & nb=nb'" + evs \ yahalom\ + \ A=A' & B=B' & na=na' & nb=nb'" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, simp_all) txt\YM3, by freshness\ @@ -209,8 +205,8 @@ subsection\Crucial Secrecy Property: Spy Does Not See Key @{term KAB}\ lemma secrecy_lemma: - "[| A \ bad; B \ bad; evs \ yahalom |] - ==> Says Server A + "\A \ bad; B \ bad; evs \ yahalom\ + \ Says Server A \nb, Crypt (shrK A) \Agent B, Key K, na\, Crypt (shrK B) \Agent A, Agent B, Key K, nb\\ \ set evs --> @@ -225,13 +221,13 @@ text\Final version\ lemma Spy_not_see_encrypted_key: - "[| Says Server A + "\Says Server A \nb, Crypt (shrK A) \Agent B, Key K, na\, Crypt (shrK B) \Agent A, Agent B, Key K, nb\\ \ set evs; Notes Spy \na, nb, Key K\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ analz (knows Spy evs)" by (blast dest: secrecy_lemma Says_Server_message_form) @@ -244,13 +240,13 @@ other than Fake are trivial, while Fake requires @{term "Key K \ analz (knows Spy evs)"}.\ lemma Spy_not_know_encrypted_key: - "[| Says Server A + "\Says Server A \nb, Crypt (shrK A) \Agent B, Key K, na\, Crypt (shrK B) \Agent A, Agent B, Key K, nb\\ \ set evs; Notes Spy \na, nb, Key K\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ knows Spy evs" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ knows Spy evs" by (blast dest: Spy_not_see_encrypted_key) @@ -259,9 +255,9 @@ text\If the encrypted message appears then it originated with the Server. May now apply \Spy_not_see_encrypted_key\, subject to its conditions.\ lemma A_trusts_YM3: - "[| Crypt (shrK A) \Agent B, Key K, na\ \ parts (knows Spy evs); - A \ bad; evs \ yahalom |] - ==> \nb. Says Server A + "\Crypt (shrK A) \Agent B, Key K, na\ \ parts (knows Spy evs); + A \ bad; evs \ yahalom\ + \ \nb. Says Server A \nb, Crypt (shrK A) \Agent B, Key K, na\, Crypt (shrK B) \Agent A, Agent B, Key K, nb\\ \ set evs" @@ -275,10 +271,10 @@ text\The obvious combination of \A_trusts_YM3\ with \Spy_not_see_encrypted_key\\ theorem A_gets_good_key: - "[| Crypt (shrK A) \Agent B, Key K, na\ \ parts (knows Spy evs); + "\Crypt (shrK A) \Agent B, Key K, na\ \ parts (knows Spy evs); \nb. Notes Spy \na, nb, Key K\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ analz (knows Spy evs)" by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key) @@ -287,10 +283,10 @@ text\B knows, by the first part of A's message, that the Server distributed the key for A and B, and has associated it with NB.\ lemma B_trusts_YM4_shrK: - "[| Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\ + "\Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\ \ parts (knows Spy evs); - B \ bad; evs \ yahalom |] - ==> \NA. Says Server A + B \ bad; evs \ yahalom\ + \ \NA. Says Server A \Nonce NB, Crypt (shrK A) \Agent B, Key K, Nonce NA\, Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\\ @@ -309,10 +305,10 @@ text\What can B deduce from receipt of YM4? Stronger and simpler than Yahalom because we do not have to show that NB is secret.\ lemma B_trusts_YM4: - "[| Gets B \Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\, X\ + "\Gets B \Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\, X\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> \NA. Says Server A + A \ bad; B \ bad; evs \ yahalom\ + \ \NA. Says Server A \Nonce NB, Crypt (shrK A) \Agent B, Key K, Nonce NA\, Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\\ @@ -323,11 +319,11 @@ text\The obvious combination of \B_trusts_YM4\ with \Spy_not_see_encrypted_key\\ theorem B_gets_good_key: - "[| Gets B \Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\, X\ + "\Gets B \Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\, X\ \ set evs; \na. Notes Spy \na, Nonce NB, Key K\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ yahalom\ + \ Key K \ analz (knows Spy evs)" by (blast dest!: B_trusts_YM4 Spy_not_see_encrypted_key) @@ -335,9 +331,9 @@ text\The encryption in message YM2 tells us it cannot be faked.\ lemma B_Said_YM2: - "[| Crypt (shrK B) \Agent A, Nonce NA\ \ parts (knows Spy evs); - B \ bad; evs \ yahalom |] - ==> \NB. Says B Server \Agent B, Nonce NB, + "\Crypt (shrK B) \Agent A, Nonce NA\ \ parts (knows Spy evs); + B \ bad; evs \ yahalom\ + \ \NB. Says B Server \Agent B, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ \ set evs" apply (erule rev_mp) @@ -350,10 +346,10 @@ text\If the server sends YM3 then B sent YM2, perhaps with a different NB\ lemma YM3_auth_B_to_A_lemma: - "[| Says Server A \nb, Crypt (shrK A) \Agent B, Key K, Nonce NA\, X\ + "\Says Server A \nb, Crypt (shrK A) \Agent B, Key K, Nonce NA\, X\ \ set evs; - B \ bad; evs \ yahalom |] - ==> \nb'. Says B Server \Agent B, nb', + B \ bad; evs \ yahalom\ + \ \nb'. Says B Server \Agent B, nb', Crypt (shrK B) \Agent A, Nonce NA\\ \ set evs" apply (erule rev_mp) @@ -364,10 +360,10 @@ text\If A receives YM3 then B has used nonce NA (and therefore is alive)\ theorem YM3_auth_B_to_A: - "[| Gets A \nb, Crypt (shrK A) \Agent B, Key K, Nonce NA\, X\ + "\Gets A \nb, Crypt (shrK A) \Agent B, Key K, Nonce NA\, X\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> \nb'. Says B Server + A \ bad; B \ bad; evs \ yahalom\ + \ \nb'. Says B Server \Agent B, nb', Crypt (shrK B) \Agent A, Nonce NA\\ \ set evs" by (blast dest!: A_trusts_YM3 YM3_auth_B_to_A_lemma) @@ -385,16 +381,16 @@ text\This lemma allows a use of \unique_session_keys\ in the next proof, which otherwise is extremely slow.\ lemma secure_unique_session_keys: - "[| Crypt (shrK A) \Agent B, Key K, na\ \ analz (spies evs); + "\Crypt (shrK A) \Agent B, Key K, na\ \ analz (spies evs); Crypt (shrK A') \Agent B', Key K, na'\ \ analz (spies evs); - Key K \ analz (knows Spy evs); evs \ yahalom |] - ==> A=A' & B=B'" + Key K \ analz (knows Spy evs); evs \ yahalom\ + \ A=A' & B=B'" by (blast dest!: A_trusts_YM3 dest: unique_session_keys Crypt_Spy_analz_bad) lemma Auth_A_to_B_lemma [rule_format]: "evs \ yahalom - ==> Key K \ analz (knows Spy evs) --> + \ Key K \ analz (knows Spy evs) --> K \ symKeys --> Crypt K (Nonce NB) \ parts (knows Spy evs) --> Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\ @@ -404,14 +400,12 @@ apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) apply (analz_mono_contra, simp_all) -txt\Fake\ -apply blast -txt\YM3: by \new_keys_not_used\, the message - @{term "Crypt K (Nonce NB)"} could not exist\ -apply (force dest!: Crypt_imp_keysFor) -txt\YM4: was @{term "Crypt K (Nonce NB)"} the very last message? If so, - apply unicity of session keys; if not, use the induction hypothesis\ -apply (blast dest!: B_trusts_YM4_shrK dest: secure_unique_session_keys) + subgoal --\Fake\ by blast + subgoal --\YM3 because the message @{term "Crypt K (Nonce NB)"} could not exist\ + by (force dest!: Crypt_imp_keysFor) + subgoal --\YM4: was @{term "Crypt K (Nonce NB)"} the very last message? If not, use the induction hypothesis, + otherwise by unicity of session keys\ + by (blast dest!: B_trusts_YM4_shrK dest: secure_unique_session_keys) done @@ -419,11 +413,11 @@ Moreover, A associates K with NB (thus is talking about the same run). Other premises guarantee secrecy of K.\ theorem YM4_imp_A_Said_YM3 [rule_format]: - "[| Gets B \Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\, + "\Gets B \Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\, Crypt K (Nonce NB)\ \ set evs; (\NA. Notes Spy \Nonce NA, Nonce NB, Key K\ \ set evs); - K \ symKeys; A \ bad; B \ bad; evs \ yahalom |] - ==> \X. Says A B \X, Crypt K (Nonce NB)\ \ set evs" + K \ symKeys; A \ bad; B \ bad; evs \ yahalom\ + \ \X. Says A B \X, Crypt K (Nonce NB)\ \ set evs" by (blast intro: Auth_A_to_B_lemma dest: Spy_not_see_encrypted_key B_trusts_YM4_shrK)