# HG changeset patch # User paulson # Date 1665677996 -3600 # Node ID 2317e9a192390b003faa707403c7c6d053e3377b # Parent 71bf371a97844d3cd85fd6793faeb80645b70dee# Parent 61640505795152a7652e776be5fdc76f4745ef88 merged diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/CertifiedEmail.thy --- a/src/HOL/Auth/CertifiedEmail.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/CertifiedEmail.thy Thu Oct 13 17:19:56 2022 +0100 @@ -36,33 +36,33 @@ | Fake: \ \The Spy may say anything he can say. The sender field is correct, but agents don't use that information.\ - "[| evsf \ certified_mail; X \ synth(analz(spies evsf))|] - ==> Says Spy B X # evsf \ certified_mail" + "\evsf \ certified_mail; X \ synth(analz(spies evsf))\ + \ Says Spy B X # evsf \ certified_mail" | FakeSSL: \ \The Spy may open SSL sessions with TTP, who is the only agent equipped with the necessary credentials to serve as an SSL server.\ - "[| evsfssl \ certified_mail; X \ synth(analz(spies evsfssl))|] - ==> Notes TTP \Agent Spy, Agent TTP, X\ # evsfssl \ certified_mail" + "\evsfssl \ certified_mail; X \ synth(analz(spies evsfssl))\ + \ Notes TTP \Agent Spy, Agent TTP, X\ # evsfssl \ certified_mail" | CM1: \ \The sender approaches the recipient. The message is a number.\ - "[|evs1 \ certified_mail; + "\evs1 \ certified_mail; Key K \ used evs1; K \ symKeys; Nonce q \ used evs1; hs = Hash\Number cleartext, Nonce q, response S R q, Crypt K (Number m)\; - S2TTP = Crypt(pubEK TTP) \Agent S, Number BothAuth, Key K, Agent R, hs\|] - ==> Says S R \Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, + S2TTP = Crypt(pubEK TTP) \Agent S, Number BothAuth, Key K, Agent R, hs\\ + \ Says S R \Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, Number cleartext, Nonce q, S2TTP\ # evs1 \ certified_mail" | CM2: \ \The recipient records \<^term>\S2TTP\ while transmitting it and her password to \<^term>\TTP\ over an SSL channel.\ - "[|evs2 \ certified_mail; + "\evs2 \ certified_mail; Gets R \Agent S, Agent TTP, em, Number BothAuth, Number cleartext, Nonce q, S2TTP\ \ set evs2; TTP \ R; - hr = Hash \Number cleartext, Nonce q, response S R q, em\ |] - ==> + hr = Hash \Number cleartext, Nonce q, response S R q, em\\ + \ Notes TTP \Agent R, Agent TTP, S2TTP, Key(RPwd R), hr\ # evs2 \ certified_mail" @@ -71,26 +71,26 @@ the client (\<^term>\R\), but \<^term>\TTP\ accepts the message only if the given password is that of the claimed sender, \<^term>\R\. He replies over the established SSL channel.\ - "[|evs3 \ certified_mail; + "\evs3 \ certified_mail; Notes TTP \Agent R, Agent TTP, S2TTP, Key(RPwd R), hr\ \ set evs3; S2TTP = Crypt (pubEK TTP) \Agent S, Number BothAuth, Key k, Agent R, hs\; - TTP \ R; hs = hr; k \ symKeys|] - ==> + TTP \ R; hs = hr; k \ symKeys\ + \ Notes R \Agent TTP, Agent R, Key k, hr\ # Gets S (Crypt (priSK TTP) S2TTP) # Says TTP S (Crypt (priSK TTP) S2TTP) # evs3 \ certified_mail" | Reception: - "[|evsr \ certified_mail; Says A B X \ set evsr|] - ==> Gets B X#evsr \ certified_mail" + "\evsr \ certified_mail; Says A B X \ set evsr\ + \ Gets B X#evsr \ certified_mail" declare Says_imp_knows_Spy [THEN analz.Inj, dest] declare analz_into_parts [dest] (*A "possibility property": there are traces that reach the end*) -lemma "[| Key K \ used []; K \ symKeys |] ==> +lemma "\Key K \ used []; K \ symKeys\ \ \S2TTP. \evs \ certified_mail. Says TTP S (Crypt (priSK TTP) S2TTP) \ set evs" apply (intro exI bexI) @@ -103,23 +103,23 @@ lemma Gets_imp_Says: - "[| Gets B X \ set evs; evs \ certified_mail |] ==> \A. Says A B X \ set evs" + "\Gets B X \ set evs; evs \ certified_mail\ \ \A. Says A B X \ set evs" apply (erule rev_mp) apply (erule certified_mail.induct, auto) done lemma Gets_imp_parts_knows_Spy: - "[|Gets A X \ set evs; evs \ certified_mail|] ==> X \ parts(spies evs)" + "\Gets A X \ set evs; evs \ certified_mail\ \ X \ parts(spies evs)" apply (drule Gets_imp_Says, simp) apply (blast dest: Says_imp_knows_Spy parts.Inj) done lemma CM2_S2TTP_analz_knows_Spy: - "[|Gets R \Agent A, Agent B, em, Number AO, Number cleartext, + "\Gets R \Agent A, Agent B, em, Number AO, Number cleartext, Nonce q, S2TTP\ \ set evs; - evs \ certified_mail|] - ==> S2TTP \ analz(spies evs)" + evs \ certified_mail\ + \ S2TTP \ analz(spies evs)" apply (drule Gets_imp_Says, simp) apply (blast dest: Says_imp_knows_Spy analz.Inj) done @@ -141,15 +141,15 @@ the fakessl rule allows Spy to spoof the sender's name. Maybe can strengthen the second disjunct with \<^term>\R\Spy\.\ lemma hr_form: - "[|Notes TTP \Agent R, Agent TTP, S2TTP, pwd, hr\ \ set evs; - evs \ certified_mail|] - ==> hr \ synth (analz (spies evs)) | + "\Notes TTP \Agent R, Agent TTP, S2TTP, pwd, hr\ \ set evs; + evs \ certified_mail\ + \ hr \ synth (analz (spies evs)) | (\clt q S em. hr = Hash \Number clt, Nonce q, response S R q, em\)" by (blast intro: hr_form_lemma) lemma Spy_dont_know_private_keys [dest!]: - "[|Key (privateKey b A) \ parts (spies evs); evs \ certified_mail|] - ==> A \ bad" + "\Key (privateKey b A) \ parts (spies evs); evs \ certified_mail\ + \ A \ bad" apply (erule rev_mp) apply (erule certified_mail.induct, simp_all) txt\Fake\ @@ -166,15 +166,15 @@ lemma Spy_know_private_keys_iff [simp]: "evs \ certified_mail - ==> (Key (privateKey b A) \ parts (spies evs)) = (A \ bad)" + \ (Key (privateKey b A) \ parts (spies evs)) = (A \ bad)" by blast lemma Spy_dont_know_TTPKey_parts [simp]: - "evs \ certified_mail ==> Key (privateKey b TTP) \ parts(spies evs)" + "evs \ certified_mail \ Key (privateKey b TTP) \ parts(spies evs)" by simp lemma Spy_dont_know_TTPKey_analz [simp]: - "evs \ certified_mail ==> Key (privateKey b TTP) \ analz(spies evs)" + "evs \ certified_mail \ Key (privateKey b TTP) \ analz(spies evs)" by auto text\Thus, prove any goal that assumes that \<^term>\Spy\ knows a private key @@ -183,11 +183,11 @@ lemma CM3_k_parts_knows_Spy: - "[| evs \ certified_mail; + "\evs \ certified_mail; Notes TTP \Agent A, Agent TTP, Crypt (pubEK TTP) \Agent S, Number AO, Key K, - Agent R, hs\, Key (RPwd R), hs\ \ set evs|] - ==> Key K \ parts(spies evs)" + Agent R, hs\, Key (RPwd R), hs\ \ set evs\ + \ Key K \ parts(spies evs)" apply (rotate_tac 1) apply (erule rev_mp) apply (erule certified_mail.induct, simp_all) @@ -201,7 +201,7 @@ done lemma Spy_dont_know_RPwd [rule_format]: - "evs \ certified_mail ==> Key (RPwd A) \ parts(spies evs) \ A \ bad" + "evs \ certified_mail \ Key (RPwd A) \ parts(spies evs) \ A \ bad" apply (erule certified_mail.induct, simp_all) txt\Fake\ apply (blast dest: Fake_parts_insert_in_Un) @@ -218,17 +218,17 @@ lemma Spy_know_RPwd_iff [simp]: - "evs \ certified_mail ==> (Key (RPwd A) \ parts(spies evs)) = (A\bad)" + "evs \ certified_mail \ (Key (RPwd A) \ parts(spies evs)) = (A\bad)" by (auto simp add: Spy_dont_know_RPwd) lemma Spy_analz_RPwd_iff [simp]: - "evs \ certified_mail ==> (Key (RPwd A) \ analz(spies evs)) = (A\bad)" + "evs \ certified_mail \ (Key (RPwd A) \ analz(spies evs)) = (A\bad)" by (metis Spy_know_RPwd_iff Spy_spies_bad_shrK analz.Inj analz_into_parts) text\Unused, but a guarantee of sorts\ theorem CertAutenticity: - "[|Crypt (priSK TTP) X \ parts (spies evs); evs \ certified_mail|] - ==> \A. Says TTP A (Crypt (priSK TTP) X) \ set evs" + "\Crypt (priSK TTP) X \ parts (spies evs); evs \ certified_mail\ + \ \A. Says TTP A (Crypt (priSK TTP) X) \ set evs" apply (erule rev_mp) apply (erule certified_mail.induct, simp_all) txt\Fake\ @@ -246,7 +246,7 @@ subsection\Proving Confidentiality Results\ lemma analz_image_freshK [rule_format]: - "evs \ certified_mail ==> + "evs \ certified_mail \ \K KK. invKey (pubEK TTP) \ KK \ (Key K \ analz (Key`KK \ (spies evs))) = (K \ KK | Key K \ analz (spies evs))" @@ -264,7 +264,7 @@ lemma analz_insert_freshK: - "[| evs \ certified_mail; KAB \ invKey (pubEK TTP) |] ==> + "\evs \ certified_mail; KAB \ invKey (pubEK TTP)\ \ (Key K \ analz (insert (Key KAB) (spies evs))) = (K = KAB | Key K \ analz (spies evs))" by (simp only: analz_image_freshK analz_image_freshK_simps) @@ -273,14 +273,14 @@ provided \<^term>\K\ is secure. Proof is surprisingly hard.\ lemma Notes_SSL_imp_used: - "[|Notes B \Agent A, Agent B, X\ \ set evs|] ==> X \ used evs" + "\Notes B \Agent A, Agent B, X\ \ set evs\ \ X \ used evs" by (blast dest!: Notes_imp_used) (*The weaker version, replacing "used evs" by "parts (spies evs)", isn't inductive: message 3 case can't be proved *) lemma S2TTP_sender_lemma [rule_format]: - "evs \ certified_mail ==> + "evs \ certified_mail \ Key K \ analz (spies evs) \ (\AO. Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\ \ used evs \ @@ -310,10 +310,10 @@ done lemma S2TTP_sender: - "[|Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\ \ used evs; + "\Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\ \ used evs; Key K \ analz (spies evs); - evs \ certified_mail|] - ==> \m ctxt q. + evs \ certified_mail\ + \ \m ctxt q. hs = Hash\Number ctxt, Nonce q, response S R q, Crypt K (Number m)\ \ Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, @@ -325,8 +325,8 @@ text\Nobody can have used non-existent keys!\ lemma new_keys_not_used [simp]: - "[|Key K \ used evs; K \ symKeys; evs \ certified_mail|] - ==> K \ keysFor (parts (spies evs))" + "\Key K \ used evs; K \ symKeys; evs \ certified_mail\ + \ K \ keysFor (parts (spies evs))" apply (erule rev_mp) apply (erule certified_mail.induct, simp_all) txt\Fake\ @@ -344,7 +344,7 @@ theorem for ciphertexts of the form \<^term>\Crypt K (Number m)\, where \<^term>\K\ is secure.\ lemma Key_unique_lemma [rule_format]: - "evs \ certified_mail ==> + "evs \ certified_mail \ Key K \ analz (spies evs) \ (\m cleartext q hs. Says S R @@ -368,7 +368,7 @@ text\The key determines the sender, recipient and protocol options.\ lemma Key_unique: - "[|Says S R + "\Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number cleartext, Nonce q, Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\\ @@ -379,8 +379,8 @@ Crypt (pubEK TTP) \Agent S', Number AO', Key K, Agent R', hs'\\ \ set evs; Key K \ analz (spies evs); - evs \ certified_mail|] - ==> R' = R \ S' = S \ AO' = AO \ hs' = hs" + evs \ certified_mail\ + \ R' = R \ S' = S \ AO' = AO \ hs' = hs" by (rule Key_unique_lemma, assumption+) @@ -390,13 +390,13 @@ If Spy gets the key then \<^term>\R\ is bad and \<^term>\S\ moreover gets his return receipt (and therefore has no grounds for complaint).\ theorem S_fairness_bad_R: - "[|Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, + "\Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number cleartext, Nonce q, S2TTP\ \ set evs; S2TTP = Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\; Key K \ analz (spies evs); evs \ certified_mail; - S\Spy|] - ==> R \ bad \ Gets S (Crypt (priSK TTP) S2TTP) \ set evs" + S\Spy\ + \ R \ bad \ Gets S (Crypt (priSK TTP) S2TTP) \ set evs" apply (erule rev_mp) apply (erule ssubst) apply (erule rev_mp) @@ -418,24 +418,24 @@ text\Confidentially for the symmetric key\ theorem Spy_not_see_encrypted_key: - "[|Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, + "\Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number cleartext, Nonce q, S2TTP\ \ set evs; S2TTP = Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\; evs \ certified_mail; - S\Spy; R \ bad|] - ==> Key K \ analz(spies evs)" + S\Spy; R \ bad\ + \ Key K \ analz(spies evs)" by (blast dest: S_fairness_bad_R) text\Agent \<^term>\R\, who may be the Spy, doesn't receive the key until \<^term>\S\ has access to the return receipt.\ theorem S_guarantee: - "[|Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, + "\Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number cleartext, Nonce q, S2TTP\ \ set evs; S2TTP = Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\; Notes R \Agent TTP, Agent R, Key K, hs\ \ set evs; - S\Spy; evs \ certified_mail|] - ==> Gets S (Crypt (priSK TTP) S2TTP) \ set evs" + S\Spy; evs \ certified_mail\ + \ Gets S (Crypt (priSK TTP) S2TTP) \ set evs" apply (erule rev_mp) apply (erule ssubst) apply (erule rev_mp) @@ -451,13 +451,13 @@ then \<^term>\R\ receives the necessary key. This result is also important to \<^term>\S\, as it confirms the validity of the return receipt.\ theorem RR_validity: - "[|Crypt (priSK TTP) S2TTP \ used evs; + "\Crypt (priSK TTP) S2TTP \ used evs; S2TTP = Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, Hash \Number cleartext, Nonce q, r, em\\; hr = Hash \Number cleartext, Nonce q, r, em\; - R\Spy; evs \ certified_mail|] - ==> Notes R \Agent TTP, Agent R, Key K, hr\ \ set evs" + R\Spy; evs \ certified_mail\ + \ Notes R \Agent TTP, Agent R, Key K, hr\ \ set evs" apply (erule rev_mp) apply (erule ssubst) apply (erule ssubst) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Event.thy --- a/src/HOL/Auth/Event.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Event.thy Thu Oct 13 17:19:56 2022 +0100 @@ -186,7 +186,7 @@ text\What agents DIFFERENT FROM Spy know was either said, or noted, or got, or known initially\ lemma knows_imp_Says_Gets_Notes_initState [rule_format]: - "[| X \ knows A evs; A \ Spy |] ==> \ B. + "\X \ knows A evs; A \ Spy\ \ \ B. Says A B X \ set evs \ Gets A X \ set evs \ Notes A X \ set evs \ X \ initState A" apply (erule rev_mp) apply (induct_tac "evs") @@ -258,8 +258,8 @@ text\For proving \new_keys_not_used\\ lemma keysFor_parts_insert: - "[| K \ keysFor (parts (insert X G)); X \ synth (analz H) |] - ==> K \ keysFor (parts (G \ H)) | Key (invKey K) \ parts H" + "\K \ keysFor (parts (insert X G)); X \ synth (analz H)\ + \ K \ keysFor (parts (G \ H)) | Key (invKey K) \ 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] diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Analz.thy --- a/src/HOL/Auth/Guard/Analz.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Analz.thy Thu Oct 13 17:19:56 2022 +0100 @@ -16,9 +16,9 @@ pparts :: "msg set => msg set" for H :: "msg set" where - Inj [intro]: "[| X \ H; is_MPair X |] ==> X \ pparts H" -| Fst [dest]: "[| \X,Y\ \ pparts H; is_MPair X |] ==> X \ pparts H" -| Snd [dest]: "[| \X,Y\ \ pparts H; is_MPair Y |] ==> Y \ pparts H" + Inj [intro]: "\X \ H; is_MPair X\ \ X \ pparts H" +| Fst [dest]: "\\X,Y\ \ pparts H; is_MPair X\ \ X \ pparts H" +| Snd [dest]: "\\X,Y\ \ pparts H; is_MPair Y\ \ Y \ pparts H" subsection\basic facts about \<^term>\pparts\\ @@ -46,7 +46,7 @@ lemma pparts_insertI [intro]: "X \ pparts H \ X \ pparts (insert Y H)" by (erule pparts.induct, auto) -lemma pparts_sub: "[| X \ pparts G; G \ H |] ==> X \ pparts H" +lemma pparts_sub: "\X \ pparts G; G \ H\ \ X \ pparts H" by (erule pparts.induct, auto) lemma pparts_insert2 [iff]: "pparts (insert X (insert Y H)) @@ -100,8 +100,8 @@ subsection\facts about \<^term>\pparts\ and \<^term>\parts\\ -lemma pparts_no_Nonce [dest]: "[| X \ pparts {Y}; Nonce n \ parts {Y} |] -==> Nonce n \ parts {X}" +lemma pparts_no_Nonce [dest]: "\X \ pparts {Y}; Nonce n \ parts {Y}\ +\ Nonce n \ parts {X}" by (erule pparts.induct, simp_all) subsection\facts about \<^term>\pparts\ and \<^term>\analz\\ @@ -109,7 +109,7 @@ lemma pparts_analz: "X \ pparts H \ X \ analz H" by (erule pparts.induct, auto) -lemma pparts_analz_sub: "[| X \ pparts G; G \ H |] ==> X \ analz H" +lemma pparts_analz_sub: "\X \ pparts G; G \ H\ \ X \ analz H" by (auto dest: pparts_sub pparts_analz) subsection\messages that contribute to analz\ @@ -118,9 +118,9 @@ kparts :: "msg set => msg set" for H :: "msg set" where - Inj [intro]: "[| X \ H; not_MPair X |] ==> X \ kparts H" -| Fst [intro]: "[| \X,Y\ \ pparts H; not_MPair X |] ==> X \ kparts H" -| Snd [intro]: "[| \X,Y\ \ pparts H; not_MPair Y |] ==> Y \ kparts H" + Inj [intro]: "\X \ H; not_MPair X\ \ X \ kparts H" +| Fst [intro]: "\\X,Y\ \ pparts H; not_MPair X\ \ X \ kparts H" +| Snd [intro]: "\\X,Y\ \ pparts H; not_MPair Y\ \ Y \ kparts H" subsection\basic facts about \<^term>\kparts\\ @@ -172,7 +172,7 @@ X \ kparts H \ X \ kparts {Z}" by (erule kparts.induct, (blast dest: pparts_insert)+) -lemma kparts_sub: "[| X \ kparts G; G \ H |] ==> X \ kparts H" +lemma kparts_sub: "\X \ kparts G; G \ H\ \ X \ kparts H" by (erule kparts.induct, auto dest: pparts_sub) lemma kparts_Un [iff]: "kparts (G \ H) = kparts G \ kparts H" @@ -197,8 +197,8 @@ subsection\facts about \<^term>\kparts\ and \<^term>\parts\\ -lemma kparts_no_Nonce [dest]: "[| X \ kparts {Y}; Nonce n \ parts {Y} |] -==> Nonce n \ parts {X}" +lemma kparts_no_Nonce [dest]: "\X \ kparts {Y}; Nonce n \ parts {Y}\ +\ Nonce n \ parts {X}" by (erule kparts.induct, auto) lemma kparts_parts: "X \ kparts H \ X \ parts H" @@ -208,8 +208,8 @@ by (erule parts.induct, auto dest: kparts_parts intro: parts.Fst parts.Snd parts.Body) -lemma Crypt_kparts_Nonce_parts [dest]: "[| Crypt K Y \ kparts {Z}; -Nonce n \ parts {Y} |] ==> Nonce n \ parts {Z}" +lemma Crypt_kparts_Nonce_parts [dest]: "\Crypt K Y \ kparts {Z}; +Nonce n \ parts {Y}\ \ Nonce n \ parts {Z}" by auto subsection\facts about \<^term>\kparts\ and \<^term>\analz\\ @@ -217,7 +217,7 @@ lemma kparts_analz: "X \ kparts H \ X \ analz H" by (erule kparts.induct, auto dest: pparts_analz) -lemma kparts_analz_sub: "[| X \ kparts G; G \ H |] ==> X \ analz H" +lemma kparts_analz_sub: "\X \ kparts G; G \ H\ \ X \ analz H" by (erule kparts.induct, auto dest: pparts_analz_sub) lemma analz_kparts [rule_format,dest]: "X \ analz H \ @@ -234,16 +234,16 @@ \ Nonce n \ kparts {Y} \ Nonce n \ analz G" by (erule synth.induct, auto) -lemma kparts_insert_synth: "[| Y \ parts (insert X G); X \ synth (analz G); -Nonce n \ kparts {Y}; Nonce n \ analz G |] ==> Y \ parts G" +lemma kparts_insert_synth: "\Y \ parts (insert X G); X \ synth (analz G); +Nonce n \ kparts {Y}; Nonce n \ analz G\ \ Y \ parts G" apply (drule parts_insert_substD, clarify) apply (drule in_sub, drule_tac X=Y in parts_sub, simp) apply (auto dest: Nonce_kparts_synth) done lemma Crypt_insert_synth: - "[| Crypt K Y \ parts (insert X G); X \ synth (analz G); Nonce n \ kparts {Y}; Nonce n \ analz G |] - ==> Crypt K Y \ parts G" + "\Crypt K Y \ parts (insert X G); X \ synth (analz G); Nonce n \ kparts {Y}; Nonce n \ analz G\ + \ Crypt K Y \ parts G" by (metis Fake_parts_insert_in_Un Nonce_kparts_synth UnE analz_conj_parts synth_simps(5)) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Extensions.thy --- a/src/HOL/Auth/Guard/Extensions.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Extensions.thy Thu Oct 13 17:19:56 2022 +0100 @@ -11,7 +11,7 @@ subsection\Extensions to Theory \Set\\ -lemma eq: "[| \x. x\A \ x\B; \x. x\B \ x\A |] ==> A=B" +lemma eq: "\\x. x\A \ x\B; \x. x\B \ x\A\ \ A=B" by auto lemma insert_Un: "P ({x} \ A) \ P (insert x A)" @@ -80,7 +80,7 @@ not_MPair :: "msg => bool" where "not_MPair X == ~ is_MPair X" -lemma is_MPairE: "[| is_MPair X ==> P; not_MPair X ==> P |] ==> P" +lemma is_MPairE: "\is_MPair X \ P; not_MPair X \ P\ \ P" by auto declare is_MPair_def [simp del] @@ -109,7 +109,7 @@ definition usekeys :: "msg set => key set" where "usekeys G \ {K. \Y. Crypt K Y \ G}" -lemma finite_keysFor [intro]: "finite G ==> finite (keysFor G)" +lemma finite_keysFor [intro]: "finite G \ finite (keysFor G)" apply (simp add: keysFor_def) apply (rule finite_imageI) apply (induct G rule: finite_induct) @@ -122,14 +122,14 @@ subsubsection\lemmas on parts\ -lemma parts_sub: "[| X \ parts G; G \ H |] ==> X \ parts H" +lemma parts_sub: "\X \ parts G; G \ H\ \ X \ parts H" by (auto dest: parts_mono) lemma parts_Diff [dest]: "X \ parts (G - H) \ X \ parts G" by (erule parts_sub, auto) -lemma parts_Diff_notin: "[| Y \ H; Nonce n \ parts (H - {Y}) |] -==> Nonce n \ parts H" +lemma parts_Diff_notin: "\Y \ H; Nonce n \ parts (H - {Y})\ +\ Nonce n \ parts H" by simp lemmas parts_insert_substI = parts_insert [THEN ssubst] @@ -142,32 +142,32 @@ apply (erule finite_induct, simp) by (rule parts_insert_substI, simp) -lemma parts_parts: "[| X \ parts {Y}; Y \ parts G |] ==> X \ parts G" +lemma parts_parts: "\X \ parts {Y}; Y \ parts G\ \ X \ parts G" by (frule parts_cut, auto) -lemma parts_parts_parts: "[| X \ parts {Y}; Y \ parts {Z}; Z \ parts G |] ==> X \ parts G" +lemma parts_parts_parts: "\X \ parts {Y}; Y \ parts {Z}; Z \ parts G\ \ X \ parts G" by (auto dest: parts_parts) -lemma parts_parts_Crypt: "[| Crypt K X \ parts G; Nonce n \ parts {X} |] -==> Nonce n \ parts G" +lemma parts_parts_Crypt: "\Crypt K X \ parts G; Nonce n \ parts {X}\ +\ Nonce n \ parts G" by (blast intro: parts.Body dest: parts_parts) subsubsection\lemmas on synth\ -lemma synth_sub: "[| X \ synth G; G \ H |] ==> X \ synth H" +lemma synth_sub: "\X \ synth G; G \ H\ \ X \ synth H" by (auto dest: synth_mono) -lemma Crypt_synth [rule_format]: "[| X \ synth G; Key K \ G |] ==> +lemma Crypt_synth [rule_format]: "\X \ synth G; Key K \ G\ \ Crypt K Y \ parts {X} \ Crypt K Y \ parts G" by (erule synth.induct, auto dest: parts_sub) subsubsection\lemmas on analz\ -lemma analz_UnI1 [intro]: "X \ analz G ==> X \ analz (G \ H)" +lemma analz_UnI1 [intro]: "X \ analz G \ X \ analz (G \ H)" by (subgoal_tac "G <= G Un H") (blast dest: analz_mono)+ -lemma analz_sub: "[| X \ analz G; G \ H |] ==> X \ analz H" +lemma analz_sub: "\X \ analz G; G \ H\ \ X \ analz H" by (auto dest: analz_mono) lemma analz_Diff [dest]: "X \ analz (G - H) \ X \ analz G" @@ -175,16 +175,16 @@ lemmas in_analz_subset_cong = analz_subset_cong [THEN subsetD] -lemma analz_eq: "A=A' ==> analz A = analz A'" +lemma analz_eq: "A=A' \ analz A = analz A'" by auto lemmas insert_commute_substI = insert_commute [THEN ssubst] lemma analz_insertD: - "[| Crypt K Y \ H; Key (invKey K) \ H |] ==> analz (insert Y H) = analz H" + "\Crypt K Y \ H; Key (invKey K) \ H\ \ analz (insert Y H) = analz H" by (blast intro: analz.Decrypt analz_insert_eq) -lemma must_decrypt [rule_format,dest]: "[| X \ analz H; has_no_pair H |] ==> +lemma must_decrypt [rule_format,dest]: "\X \ analz H; has_no_pair H\ \ X \ H \ (\K Y. Crypt K Y \ H \ Key (invKey K) \ H)" by (erule analz.induct, auto) @@ -205,8 +205,8 @@ lemmas in_analz_subset_parts = analz_subset_parts [THEN subsetD] -lemma Crypt_synth_insert: "[| Crypt K X \ parts (insert Y H); -Y \ synth (analz H); Key K \ analz H |] ==> Crypt K X \ parts H" +lemma Crypt_synth_insert: "\Crypt K X \ parts (insert Y H); +Y \ synth (analz H); Key K \ analz H\ \ Crypt K X \ parts H" apply (drule parts_insert_substD, clarify) apply (frule in_sub) apply (frule parts_mono) @@ -230,10 +230,10 @@ definition keyset :: "msg set => bool" where "keyset G \ \X. X \ G \ (\K. X = Key K)" -lemma keyset_in [dest]: "[| keyset G; X \ G |] ==> \K. X = Key K" +lemma keyset_in [dest]: "\keyset G; X \ G\ \ \K. X = Key K" by (auto simp: keyset_def) -lemma MPair_notin_keyset [simp]: "keyset G ==> \X,Y\ \ G" +lemma MPair_notin_keyset [simp]: "keyset G \ \X,Y\ \ G" by auto lemma Crypt_notin_keyset [simp]: "keyset G \ Crypt K X \ G" @@ -242,7 +242,7 @@ lemma Nonce_notin_keyset [simp]: "keyset G \ Nonce n \ G" by auto -lemma parts_keyset [simp]: "keyset G ==> parts G = G" +lemma parts_keyset [simp]: "keyset G \ parts G = G" by (auto, erule parts.induct, auto) subsubsection\keys a priori necessary for decrypting the messages of G\ @@ -253,7 +253,7 @@ lemma keyset_keysfor [iff]: "keyset (keysfor G)" by (simp add: keyset_def keysfor_def, blast) -lemma keyset_Diff_keysfor [simp]: "keyset H ==> keyset (H - keysfor G)" +lemma keyset_Diff_keysfor [simp]: "keyset H \ keyset (H - keysfor G)" by (auto simp: keyset_def) lemma keysfor_Crypt: "Crypt K X \ parts G \ Key (invKey K) \ keysfor G" @@ -262,12 +262,12 @@ lemma no_key_no_Crypt: "Key K \ keysfor G \ Crypt (invKey K) X \ parts G" by (auto dest: keysfor_Crypt) -lemma finite_keysfor [intro]: "finite G ==> finite (keysfor G)" +lemma finite_keysfor [intro]: "finite G \ finite (keysfor G)" by (auto simp: keysfor_def intro: finite_UN_I) subsubsection\only the keys necessary for G are useful in analz\ -lemma analz_keyset: "keyset H ==> +lemma analz_keyset: "keyset H \ analz (G Un H) = H - keysfor G Un (analz (G Un (H Int keysfor G)))" apply (rule eq) apply (erule analz.induct, blast) @@ -297,37 +297,37 @@ "Gets_correct p == \evs B X. evs \ p \ Gets B X \ set evs \ (\A. Says A B X \ set evs)" -lemma Gets_correct_Says: "[| Gets_correct p; Gets B X # evs \ p |] -==> \A. Says A B X \ set evs" +lemma Gets_correct_Says: "\Gets_correct p; Gets B X # evs \ p\ +\ \A. Says A B X \ set evs" apply (simp add: Gets_correct_def) by (drule_tac x="Gets B X # evs" in spec, auto) definition one_step :: "event list set => bool" where "one_step p == \evs ev. ev#evs \ p \ evs \ p" -lemma one_step_Cons [dest]: "[| one_step p; ev#evs \ p |] ==> evs \ p" -by (unfold one_step_def, blast) +lemma one_step_Cons [dest]: "\one_step p; ev#evs \ p\ \ evs \ p" + unfolding one_step_def by blast -lemma one_step_app: "[| evs@evs' \ p; one_step p; [] \ p |] ==> evs' \ p" +lemma one_step_app: "\evs@evs' \ p; one_step p; [] \ p\ \ evs' \ p" by (induct evs, auto) -lemma trunc: "[| evs @ evs' \ p; one_step p |] ==> evs' \ p" +lemma trunc: "\evs @ evs' \ p; one_step p\ \ evs' \ p" by (induct evs, auto) definition has_only_Says :: "event list set => bool" where "has_only_Says p \ \evs ev. evs \ p \ ev \ set evs \ (\A B X. ev = Says A B X)" -lemma has_only_SaysD: "[| ev \ set evs; evs \ p; has_only_Says p |] -==> \A B X. ev = Says A B X" -by (unfold has_only_Says_def, blast) +lemma has_only_SaysD: "\ev \ set evs; evs \ p; has_only_Says p\ +\ \A B X. ev = Says A B X" + unfolding has_only_Says_def by blast -lemma in_has_only_Says [dest]: "[| has_only_Says p; evs \ p; ev \ set evs |] -==> \A B X. ev = Says A B X" +lemma in_has_only_Says [dest]: "\has_only_Says p; evs \ p; ev \ set evs\ +\ \A B X. ev = Says A B X" by (auto simp: has_only_Says_def) lemma has_only_Says_imp_Gets_correct [simp]: "has_only_Says p -==> Gets_correct p" +\ Gets_correct p" by (auto simp: has_only_Says_def Gets_correct_def) subsubsection\lemma on knows\ @@ -335,8 +335,8 @@ lemma Says_imp_spies2: "Says A B \X,Y\ \ set evs \ Y \ parts (spies evs)" by (drule Says_imp_spies, drule parts.Inj, drule parts.Snd, simp) -lemma Says_not_parts: "[| Says A B X \ set evs; Y \ parts (spies evs) |] -==> Y \ parts {X}" +lemma Says_not_parts: "\Says A B X \ set evs; Y \ parts (spies evs)\ +\ Y \ parts {X}" by (auto dest: Says_imp_spies parts_parts) subsubsection\knows without initState\ @@ -390,8 +390,8 @@ lemmas knows_Cons_substI = knows_Cons [THEN ssubst] lemmas knows_Cons_substD = knows_Cons [THEN sym, THEN ssubst] -lemma knows'_sub_spies': "[| evs \ p; has_only_Says p; one_step p |] -==> knows' A evs \ spies' evs" +lemma knows'_sub_spies': "\evs \ p; has_only_Says p; one_step p\ +\ knows' A evs \ spies' evs" by (induct evs, auto split: event.splits) subsubsection\knows' is finite\ @@ -466,12 +466,12 @@ lemma finite_knows_max' [iff]: "finite (knows_max' A evs)" by (induct evs, auto split: event.split) -lemma knows_max'_sub_spies': "[| evs \ p; has_only_Says p; one_step p |] -==> knows_max' A evs \ spies' evs" +lemma knows_max'_sub_spies': "\evs \ p; has_only_Says p; one_step p\ +\ knows_max' A evs \ spies' evs" by (induct evs, auto split: event.splits) -lemma knows_max'_in_spies' [dest]: "[| evs \ p; X \ knows_max' A evs; -has_only_Says p; one_step p |] ==> X \ spies' evs" +lemma knows_max'_in_spies' [dest]: "\evs \ p; X \ knows_max' A evs; +has_only_Says p; one_step p\ \ X \ spies' evs" by (rule knows_max'_sub_spies' [THEN subsetD], auto) lemma knows_max'_app: "knows_max' A (evs @ evs') @@ -537,11 +537,11 @@ lemma used_appIR: "X \ used evs \ X \ used (evs @ evs')" by (erule used_sub_app [THEN subsetD]) -lemma used_parts: "[| X \ parts {Y}; Y \ used evs |] ==> X \ used evs" +lemma used_parts: "\X \ parts {Y}; Y \ used evs\ \ X \ used evs" apply (auto simp: used_decomp dest: used'_parts) by (auto simp: init_def used_Nil dest: parts_trans) -lemma parts_Says_used: "[| Says A B X \ set evs; Y \ parts {X} |] ==> Y \ used evs" +lemma parts_Says_used: "\Says A B X \ set evs; Y \ parts {X}\ \ Y \ used evs" by (induct evs, simp_all, safe, auto intro: used_ConsI) lemma parts_used_app: "X \ parts {Y} \ X \ used (evs @ Says A B Y # evs')" @@ -561,16 +561,16 @@ lemma not_used_not_spied: "X \ used evs \ X \ parts (spies evs)" by (induct evs, auto simp: used_Nil) -lemma not_used_not_parts: "[| Y \ used evs; Says A B X \ set evs |] -==> Y \ parts {X}" +lemma not_used_not_parts: "\Y \ used evs; Says A B X \ set evs\ +\ Y \ parts {X}" by (induct evs, auto intro: used_ConsI) -lemma not_used_parts_false: "[| X \ used evs; Y \ parts (spies evs) |] -==> X \ parts {Y}" +lemma not_used_parts_false: "\X \ used evs; Y \ parts (spies evs)\ +\ X \ parts {Y}" by (auto dest: parts_parts) -lemma known_used [rule_format]: "[| evs \ p; Gets_correct p; one_step p |] -==> X \ parts (knows A evs) \ X \ used evs" +lemma known_used [rule_format]: "\evs \ p; Gets_correct p; one_step p\ +\ X \ parts (knows A evs) \ X \ used evs" apply (case_tac "A=Spy", blast) apply (induct evs) apply (simp add: used.simps, blast) @@ -583,8 +583,8 @@ apply (drule_tac B=A and X=msg and evs=evs in Gets_correct_Says) by (auto dest: Says_imp_used intro: used_ConsI) -lemma known_max_used [rule_format]: "[| evs \ p; Gets_correct p; one_step p |] -==> X \ parts (knows_max A evs) \ X \ used evs" +lemma known_max_used [rule_format]: "\evs \ p; Gets_correct p; one_step p\ +\ X \ parts (knows_max A evs) \ X \ used evs" apply (case_tac "A=Spy") apply force apply (induct evs) @@ -597,22 +597,22 @@ apply (drule_tac B=A and X=msg and evs=evs in Gets_correct_Says) by (auto simp: knows_max'_Cons dest: Says_imp_used intro: used_ConsI) -lemma not_used_not_known: "[| evs \ p; X \ used evs; -Gets_correct p; one_step p |] ==> X \ parts (knows A evs)" +lemma not_used_not_known: "\evs \ p; X \ used evs; +Gets_correct p; one_step p\ \ X \ parts (knows A evs)" by (case_tac "A=Spy", auto dest: not_used_not_spied known_used) -lemma not_used_not_known_max: "[| evs \ p; X \ used evs; -Gets_correct p; one_step p |] ==> X \ parts (knows_max A evs)" +lemma not_used_not_known_max: "\evs \ p; X \ used evs; +Gets_correct p; one_step p\ \ X \ parts (knows_max A evs)" by (case_tac "A=Spy", auto dest: not_used_not_spied known_max_used) subsubsection\a nonce or key in a message cannot equal a fresh nonce or key\ -lemma Nonce_neq [dest]: "[| Nonce n' \ used evs; -Says A B X \ set evs; Nonce n \ parts {X} |] ==> n \ n'" +lemma Nonce_neq [dest]: "\Nonce n' \ used evs; +Says A B X \ set evs; Nonce n \ parts {X}\ \ n \ n'" by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub) -lemma Key_neq [dest]: "[| Key n' \ used evs; -Says A B X \ set evs; Key n \ parts {X} |] ==> n ~= n'" +lemma Key_neq [dest]: "\Key n' \ used evs; +Says A B X \ set evs; Key n \ parts {X}\ \ n ~= n'" by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub) subsubsection\message of an event\ @@ -623,7 +623,7 @@ | "msg (Gets A X) = X" | "msg (Notes A X) = X" -lemma used_sub_parts_used: "X \ used (ev # evs) ==> X \ parts {msg ev} \ used evs" +lemma used_sub_parts_used: "X \ used (ev # evs) \ X \ parts {msg ev} \ used evs" by (induct ev, auto) end diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Guard.thy --- a/src/HOL/Auth/Guard/Guard.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Guard.thy Thu Oct 13 17:19:56 2022 +0100 @@ -19,7 +19,7 @@ No_Nonce [intro]: "Nonce n \ parts {X} \ X \ guard n Ks" | Guard_Nonce [intro]: "invKey K \ Ks \ Crypt K X \ guard n Ks" | Crypt [intro]: "X \ guard n Ks \ Crypt K X \ guard n Ks" -| Pair [intro]: "[| X \ guard n Ks; Y \ guard n Ks |] ==> \X,Y\ \ guard n Ks" +| Pair [intro]: "\X \ guard n Ks; Y \ guard n Ks\ \ \X,Y\ \ guard n Ks" subsection\basic facts about \<^term>\guard\\ @@ -38,7 +38,7 @@ lemma Nonce_notin_guard_iff [iff]: "Nonce n \ guard n Ks" by (auto dest: Nonce_notin_guard) -lemma guard_has_Crypt [rule_format]: "X \ guard n Ks ==> Nonce n \ parts {X} +lemma guard_has_Crypt [rule_format]: "X \ guard n Ks \ Nonce n \ parts {X} \ (\K Y. Crypt K Y \ kparts {X} \ Nonce n \ parts {Y})" by (erule guard.induct, auto) @@ -55,7 +55,7 @@ Y \ kparts {X} \ Y \ guard n Ks" by (erule guard.induct, auto) -lemma guard_Crypt: "[| Crypt K Y \ guard n Ks; K \ invKey`Ks |] ==> Y \ guard n Ks" +lemma guard_Crypt: "\Crypt K Y \ guard n Ks; K \ invKey`Ks\ \ Y \ guard n Ks" by (ind_cases "Crypt K Y \ guard n Ks") (auto intro!: image_eqI) lemma guard_MPair [iff]: "(\X,Y\ \ guard n Ks) = (X \ guard n Ks \ Y \ guard n Ks)" @@ -65,7 +65,7 @@ Crypt K Y \ kparts {X} \ Nonce n \ kparts {Y} \ Y \ guard n Ks" by (erule guard.induct, auto dest: guard_kparts) -lemma guard_extand: "[| X \ guard n Ks; Ks \ Ks' |] ==> X \ guard n Ks'" +lemma guard_extand: "\X \ guard n Ks; Ks \ Ks'\ \ X \ guard n Ks'" by (erule guard.induct, auto) subsection\guarded sets\ @@ -86,15 +86,15 @@ lemma Nonce_notin_kparts [simplified]: "Guard n Ks H \ Nonce n \ kparts H" by (auto simp: Guard_def dest: in_kparts Nonce_notin_kparts_msg) -lemma Guard_must_decrypt: "[| Guard n Ks H; Nonce n \ analz H |] ==> +lemma Guard_must_decrypt: "\Guard n Ks H; Nonce n \ analz H\ \ \K Y. Crypt K Y \ kparts H \ Key (invKey K) \ kparts H" apply (drule_tac P="\G. Nonce n \ G" in analz_pparts_kparts_substD, simp) by (drule must_decrypt, auto dest: Nonce_notin_kparts) -lemma Guard_kparts [intro]: "Guard n Ks H ==> Guard n Ks (kparts H)" +lemma Guard_kparts [intro]: "Guard n Ks H \ Guard n Ks (kparts H)" by (auto simp: Guard_def dest: in_kparts guard_kparts) -lemma Guard_mono: "[| Guard n Ks H; G <= H |] ==> Guard n Ks G" +lemma Guard_mono: "\Guard n Ks H; G <= H\ \ Guard n Ks G" by (auto simp: Guard_def) lemma Guard_insert [iff]: "Guard n Ks (insert X H) @@ -104,54 +104,54 @@ lemma Guard_Un [iff]: "Guard n Ks (G Un H) = (Guard n Ks G & Guard n Ks H)" by (auto simp: Guard_def) -lemma Guard_synth [intro]: "Guard n Ks G ==> Guard n Ks (synth G)" +lemma Guard_synth [intro]: "Guard n Ks G \ Guard n Ks (synth G)" by (auto simp: Guard_def, erule synth.induct, auto) -lemma Guard_analz [intro]: "[| Guard n Ks G; \K. K \ Ks \ Key K \ analz G |] -==> Guard n Ks (analz G)" +lemma Guard_analz [intro]: "\Guard n Ks G; \K. K \ Ks \ Key K \ analz G\ +\ Guard n Ks (analz G)" apply (auto simp: Guard_def) apply (erule analz.induct, auto) by (ind_cases "Crypt K Xa \ guard n Ks" for K Xa, auto) -lemma in_Guard [dest]: "[| X \ G; Guard n Ks G |] ==> X \ guard n Ks" +lemma in_Guard [dest]: "\X \ G; Guard n Ks G\ \ X \ guard n Ks" by (auto simp: Guard_def) -lemma in_synth_Guard: "[| X \ synth G; Guard n Ks G |] ==> X \ guard n Ks" +lemma in_synth_Guard: "\X \ synth G; Guard n Ks G\ \ X \ guard n Ks" by (drule Guard_synth, auto) -lemma in_analz_Guard: "[| X \ analz G; Guard n Ks G; -\K. K \ Ks \ Key K \ analz G |] ==> X \ guard n Ks" +lemma in_analz_Guard: "\X \ analz G; Guard n Ks G; +\K. K \ Ks \ Key K \ analz G\ \ X \ guard n Ks" by (drule Guard_analz, auto) -lemma Guard_keyset [simp]: "keyset G ==> Guard n Ks G" +lemma Guard_keyset [simp]: "keyset G \ Guard n Ks G" by (auto simp: Guard_def) -lemma Guard_Un_keyset: "[| Guard n Ks G; keyset H |] ==> Guard n Ks (G \ H)" +lemma Guard_Un_keyset: "\Guard n Ks G; keyset H\ \ Guard n Ks (G \ H)" by auto -lemma in_Guard_kparts: "[| X \ G; Guard n Ks G; Y \ kparts {X} |] ==> Y \ guard n Ks" +lemma in_Guard_kparts: "\X \ G; Guard n Ks G; Y \ kparts {X}\ \ Y \ guard n Ks" by blast -lemma in_Guard_kparts_neq: "[| X \ G; Guard n Ks G; Nonce n' \ kparts {X} |] -==> n \ n'" +lemma in_Guard_kparts_neq: "\X \ G; Guard n Ks G; Nonce n' \ kparts {X}\ +\ n \ n'" by (blast dest: in_Guard_kparts) -lemma in_Guard_kparts_Crypt: "[| X \ G; Guard n Ks G; is_MPair X; -Crypt K Y \ kparts {X}; Nonce n \ kparts {Y} |] ==> invKey K \ Ks" +lemma in_Guard_kparts_Crypt: "\X \ G; Guard n Ks G; is_MPair X; +Crypt K Y \ kparts {X}; Nonce n \ kparts {Y}\ \ invKey K \ Ks" apply (drule in_Guard, simp) apply (frule guard_not_guard, simp+) apply (drule guard_kparts, simp) by (ind_cases "Crypt K Y \ guard n Ks", auto) -lemma Guard_extand: "[| Guard n Ks G; Ks \ Ks' |] ==> Guard n Ks' G" +lemma Guard_extand: "\Guard n Ks G; Ks \ Ks'\ \ Guard n Ks' G" by (auto simp: Guard_def dest: guard_extand) -lemma guard_invKey [rule_format]: "[| X \ guard n Ks; Nonce n \ kparts {Y} |] ==> +lemma guard_invKey [rule_format]: "\X \ guard n Ks; Nonce n \ kparts {Y}\ \ Crypt K Y \ kparts {X} \ invKey K \ Ks" by (erule guard.induct, auto) -lemma Crypt_guard_invKey [rule_format]: "[| Crypt K Y \ guard n Ks; -Nonce n \ kparts {Y} |] ==> invKey K \ Ks" +lemma Crypt_guard_invKey [rule_format]: "\Crypt K Y \ guard n Ks; +Nonce n \ kparts {Y}\ \ invKey K \ Ks" by (auto dest: guard_invKey) subsection\set obtained by decrypting a message\ @@ -160,14 +160,14 @@ decrypt :: "msg set => key => msg => msg set" where "decrypt H K Y == insert Y (H - {Crypt K Y})" -lemma analz_decrypt: "[| Crypt K Y \ H; Key (invKey K) \ H; Nonce n \ analz H |] -==> Nonce n \ analz (decrypt H K Y)" +lemma analz_decrypt: "\Crypt K Y \ H; Key (invKey K) \ H; Nonce n \ analz H\ +\ Nonce n \ analz (decrypt H K Y)" apply (drule_tac P="\H. Nonce n \ analz H" in ssubst [OF insert_Diff]) apply assumption apply (simp only: analz_Crypt_if, simp) done -lemma parts_decrypt: "[| Crypt K Y \ H; X \ parts (decrypt H K Y) |] ==> X \ parts H" +lemma parts_decrypt: "\Crypt K Y \ H; X \ parts (decrypt H K Y)\ \ X \ parts H" by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body) subsection\number of Crypt's in a message\ @@ -195,12 +195,12 @@ lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'" by (induct l, auto) -lemma mem_cnb_minus: "x \ set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)" +lemma mem_cnb_minus: "x \ set l \ cnb l = crypt_nb x + (cnb l - crypt_nb x)" by (induct l) auto lemmas mem_cnb_minus_substI = mem_cnb_minus [THEN ssubst] -lemma cnb_minus [simp]: "x \ set l ==> cnb (remove l x) = cnb l - crypt_nb x" +lemma cnb_minus [simp]: "x \ set l \ cnb (remove l x) = cnb l - crypt_nb x" apply (induct l, auto) apply (erule_tac l=l and x=x in mem_cnb_minus_substI) apply simp @@ -286,20 +286,20 @@ apply (rule_tac B="set l'" in subset_trans, rule set_remove, blast) by (rule kparts_set) -lemma Guard_invKey_finite: "[| Nonce n \ analz G; Guard n Ks G; finite G |] -==> \K. K \ Ks \ Key K \ analz G" +lemma Guard_invKey_finite: "\Nonce n \ analz G; Guard n Ks G; finite G\ +\ \K. K \ Ks \ Key K \ analz G" apply (drule finite_list, clarify) by (rule Guard_invKey_by_list, auto) -lemma Guard_invKey: "[| Nonce n \ analz G; Guard n Ks G |] -==> \K. K \ Ks \ Key K \ analz G" +lemma Guard_invKey: "\Nonce n \ analz G; Guard n Ks G\ +\ \K. K \ Ks \ Key K \ analz G" by (auto dest: analz_needs_only_finite Guard_invKey_finite) subsection\if the analyse of a finite guarded set and a (possibly infinite) set of keys gives n then it must also gives Ks\ -lemma Guard_invKey_keyset: "[| Nonce n \ analz (G \ H); Guard n Ks G; finite G; -keyset H |] ==> \K. K \ Ks \ Key K \ analz (G \ H)" +lemma Guard_invKey_keyset: "\Nonce n \ analz (G \ H); Guard n Ks G; finite G; +keyset H\ \ \K. K \ Ks \ Key K \ analz (G \ H)" apply (frule_tac P="\G. Nonce n \ G" and G=G in analz_keyset_substD, simp_all) apply (drule_tac G="G Un (H Int keysfor G)" in Guard_invKey_finite) by (auto simp: Guard_def intro: analz_sub) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/GuardK.thy --- a/src/HOL/Auth/Guard/GuardK.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/GuardK.thy Thu Oct 13 17:19:56 2022 +0100 @@ -24,9 +24,9 @@ for n :: nat and Ks :: "key set" where No_Key [intro]: "Key n \ parts {X} \ X \ guardK n Ks" -| Guard_Key [intro]: "invKey K \ Ks ==> Crypt K X \ guardK n Ks" +| Guard_Key [intro]: "invKey K \ Ks \ Crypt K X \ guardK n Ks" | Crypt [intro]: "X \ guardK n Ks \ Crypt K X \ guardK n Ks" -| Pair [intro]: "[| X \ guardK n Ks; Y \ guardK n Ks |] ==> \X,Y\ \ guardK n Ks" +| Pair [intro]: "\X \ guardK n Ks; Y \ guardK n Ks\ \ \X,Y\ \ guardK n Ks" subsection\basic facts about \<^term>\guardK\\ @@ -62,7 +62,7 @@ Y \ kparts {X} \ Y \ guardK n Ks" by (erule guardK.induct, auto dest: kparts_parts parts_sub) -lemma guardK_Crypt: "[| Crypt K Y \ guardK n Ks; K \ invKey`Ks |] ==> Y \ guardK n Ks" +lemma guardK_Crypt: "\Crypt K Y \ guardK n Ks; K \ invKey`Ks\ \ Y \ guardK n Ks" by (ind_cases "Crypt K Y \ guardK n Ks") (auto intro!: image_eqI) lemma guardK_MPair [iff]: "(\X,Y\ \ guardK n Ks) @@ -73,8 +73,8 @@ Crypt K Y \ kparts {X} \ Key n \ kparts {Y} \ Y \ guardK n Ks" by (erule guardK.induct, auto dest: guardK_kparts) -lemma guardK_extand: "[| X \ guardK n Ks; Ks \ Ks'; -[| K \ Ks'; K \ Ks |] ==> Key K \ parts {X} |] ==> X \ guardK n Ks'" +lemma guardK_extand: "\X \ guardK n Ks; Ks \ Ks'; +\K \ Ks'; K \ Ks\ \ Key K \ parts {X}\ \ X \ guardK n Ks'" by (erule guardK.induct, auto) subsection\guarded sets\ @@ -90,7 +90,7 @@ lemma Key_notin_kparts [simplified]: "GuardK n Ks H \ Key n \ kparts H" by (auto simp: GuardK_def dest: in_kparts Key_notin_kparts_msg) -lemma GuardK_must_decrypt: "[| GuardK n Ks H; Key n \ analz H |] ==> +lemma GuardK_must_decrypt: "\GuardK n Ks H; Key n \ analz H\ \ \K Y. Crypt K Y \ kparts H \ Key (invKey K) \ kparts H" apply (drule_tac P="\G. Key n \ G" in analz_pparts_kparts_substD, simp) by (drule must_decrypt, auto dest: Key_notin_kparts) @@ -98,7 +98,7 @@ lemma GuardK_kparts [intro]: "GuardK n Ks H \ GuardK n Ks (kparts H)" by (auto simp: GuardK_def dest: in_kparts guardK_kparts) -lemma GuardK_mono: "[| GuardK n Ks H; G \ H |] ==> GuardK n Ks G" +lemma GuardK_mono: "\GuardK n Ks H; G \ H\ \ GuardK n Ks G" by (auto simp: GuardK_def) lemma GuardK_insert [iff]: "GuardK n Ks (insert X H) @@ -111,45 +111,45 @@ lemma GuardK_synth [intro]: "GuardK n Ks G \ GuardK n Ks (synth G)" by (auto simp: GuardK_def, erule synth.induct, auto) -lemma GuardK_analz [intro]: "[| GuardK n Ks G; \K. K \ Ks \ Key K \ analz G |] -==> GuardK n Ks (analz G)" +lemma GuardK_analz [intro]: "\GuardK n Ks G; \K. K \ Ks \ Key K \ analz G\ +\ GuardK n Ks (analz G)" apply (auto simp: GuardK_def) apply (erule analz.induct, auto) by (ind_cases "Crypt K Xa \ guardK n Ks" for K Xa, auto) -lemma in_GuardK [dest]: "[| X \ G; GuardK n Ks G |] ==> X \ guardK n Ks" +lemma in_GuardK [dest]: "\X \ G; GuardK n Ks G\ \ X \ guardK n Ks" by (auto simp: GuardK_def) -lemma in_synth_GuardK: "[| X \ synth G; GuardK n Ks G |] ==> X \ guardK n Ks" +lemma in_synth_GuardK: "\X \ synth G; GuardK n Ks G\ \ X \ guardK n Ks" by (drule GuardK_synth, auto) -lemma in_analz_GuardK: "[| X \ analz G; GuardK n Ks G; -\K. K \ Ks \ Key K \ analz G |] ==> X \ guardK n Ks" +lemma in_analz_GuardK: "\X \ analz G; GuardK n Ks G; +\K. K \ Ks \ Key K \ analz G\ \ X \ guardK n Ks" by (drule GuardK_analz, auto) -lemma GuardK_keyset [simp]: "[| keyset G; Key n \ G |] ==> GuardK n Ks G" +lemma GuardK_keyset [simp]: "\keyset G; Key n \ G\ \ GuardK n Ks G" by (simp only: GuardK_def, clarify, drule keyset_in, auto) -lemma GuardK_Un_keyset: "[| GuardK n Ks G; keyset H; Key n \ H |] -==> GuardK n Ks (G Un H)" +lemma GuardK_Un_keyset: "\GuardK n Ks G; keyset H; Key n \ H\ +\ GuardK n Ks (G Un H)" by auto -lemma in_GuardK_kparts: "[| X \ G; GuardK n Ks G; Y \ kparts {X} |] ==> Y \ guardK n Ks" +lemma in_GuardK_kparts: "\X \ G; GuardK n Ks G; Y \ kparts {X}\ \ Y \ guardK n Ks" by blast -lemma in_GuardK_kparts_neq: "[| X \ G; GuardK n Ks G; Key n' \ kparts {X} |] -==> n \ n'" +lemma in_GuardK_kparts_neq: "\X \ G; GuardK n Ks G; Key n' \ kparts {X}\ +\ n \ n'" by (blast dest: in_GuardK_kparts) -lemma in_GuardK_kparts_Crypt: "[| X \ G; GuardK n Ks G; is_MPair X; -Crypt K Y \ kparts {X}; Key n \ kparts {Y} |] ==> invKey K \ Ks" +lemma in_GuardK_kparts_Crypt: "\X \ G; GuardK n Ks G; is_MPair X; +Crypt K Y \ kparts {X}; Key n \ kparts {Y}\ \ invKey K \ Ks" apply (drule in_GuardK, simp) apply (frule guardK_not_guardK, simp+) apply (drule guardK_kparts, simp) by (ind_cases "Crypt K Y \ guardK n Ks", auto) -lemma GuardK_extand: "[| GuardK n Ks G; Ks \ Ks'; -[| K \ Ks'; K \ Ks |] ==> Key K \ parts G |] ==> GuardK n Ks' G" +lemma GuardK_extand: "\GuardK n Ks G; Ks \ Ks'; +\K \ Ks'; K \ Ks\ \ Key K \ parts G\ \ GuardK n Ks' G" by (auto simp: GuardK_def dest: guardK_extand parts_sub) subsection\set obtained by decrypting a message\ @@ -158,14 +158,14 @@ decrypt :: "msg set \ key \ msg \ msg set" where "decrypt H K Y \ insert Y (H - {Crypt K Y})" -lemma analz_decrypt: "[| Crypt K Y \ H; Key (invKey K) \ H; Key n \ analz H |] -==> Key n \ analz (decrypt H K Y)" +lemma analz_decrypt: "\Crypt K Y \ H; Key (invKey K) \ H; Key n \ analz H\ +\ Key n \ analz (decrypt H K Y)" apply (drule_tac P="\H. Key n \ analz H" in ssubst [OF insert_Diff]) apply assumption apply (simp only: analz_Crypt_if, simp) done -lemma parts_decrypt: "[| Crypt K Y \ H; X \ parts (decrypt H K Y) |] ==> X \ parts H" +lemma parts_decrypt: "\Crypt K Y \ H; X \ parts (decrypt H K Y)\ \ X \ parts H" by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body) subsection\number of Crypt's in a message\ @@ -191,12 +191,12 @@ lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'" by (induct l, auto) -lemma mem_cnb_minus: "x \ set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)" +lemma mem_cnb_minus: "x \ set l \ cnb l = crypt_nb x + (cnb l - crypt_nb x)" by (induct l, auto) lemmas mem_cnb_minus_substI = mem_cnb_minus [THEN ssubst] -lemma cnb_minus [simp]: "x \ set l ==> cnb (remove l x) = cnb l - crypt_nb x" +lemma cnb_minus [simp]: "x \ set l \ cnb (remove l x) = cnb l - crypt_nb x" apply (induct l, auto) by (erule_tac l=l and x=x in mem_cnb_minus_substI, simp) @@ -280,20 +280,20 @@ apply (rule_tac B="set l'" in subset_trans, rule set_remove, blast) by (rule kparts_set) -lemma GuardK_invKey_finite: "[| Key n \ analz G; GuardK n Ks G; finite G |] -==> \K. K \ Ks \ Key K \ analz G" +lemma GuardK_invKey_finite: "\Key n \ analz G; GuardK n Ks G; finite G\ +\ \K. K \ Ks \ Key K \ analz G" apply (drule finite_list, clarify) by (rule GuardK_invKey_by_list, auto) -lemma GuardK_invKey: "[| Key n \ analz G; GuardK n Ks G |] -==> \K. K \ Ks \ Key K \ analz G" +lemma GuardK_invKey: "\Key n \ analz G; GuardK n Ks G\ +\ \K. K \ Ks \ Key K \ analz G" by (auto dest: analz_needs_only_finite GuardK_invKey_finite) text\if the analyse of a finite guarded set and a (possibly infinite) set of keys gives n then it must also gives Ks\ -lemma GuardK_invKey_keyset: "[| Key n \ analz (G \ H); GuardK n Ks G; finite G; -keyset H; Key n \ H |] ==> \K. K \ Ks \ Key K \ analz (G \ H)" +lemma GuardK_invKey_keyset: "\Key n \ analz (G \ H); GuardK n Ks G; finite G; +keyset H; Key n \ H\ \ \K. K \ Ks \ Key K \ analz (G \ H)" apply (frule_tac P="\G. Key n \ G" and G=G in analz_keyset_substD, simp_all) apply (drule_tac G="G Un (H Int keysfor G)" in GuardK_invKey_finite) apply (auto simp: GuardK_def intro: analz_sub) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Guard_NS_Public.thy --- a/src/HOL/Auth/Guard/Guard_NS_Public.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Guard_NS_Public.thy Thu Oct 13 17:19:56 2022 +0100 @@ -39,14 +39,14 @@ Nil: "[] \ nsp" -| Fake: "[| evs \ nsp; X \ synth (analz (spies evs)) |] ==> Says Spy B X # evs \ nsp" +| Fake: "\evs \ nsp; X \ synth (analz (spies evs))\ \ Says Spy B X # evs \ nsp" -| NS1: "[| evs1 \ nsp; Nonce NA \ used evs1 |] ==> ns1 A B NA # evs1 \ nsp" +| NS1: "\evs1 \ nsp; Nonce NA \ used evs1\ \ ns1 A B NA # evs1 \ nsp" -| NS2: "[| evs2 \ nsp; Nonce NB \ used evs2; ns1' A' A B NA \ set evs2 |] ==> +| NS2: "\evs2 \ nsp; Nonce NB \ used evs2; ns1' A' A B NA \ set evs2\ \ ns2 B A NA NB # evs2 \ nsp" -| NS3: "\A B B' NA NB evs3. [| evs3 \ nsp; ns1 A B NA \ set evs3; ns2' B' B A NA NB \ set evs3 |] ==> +| NS3: "\A B B' NA NB evs3. \evs3 \ nsp; ns1 A B NA \ set evs3; ns2' B' B A NA NB \ set evs3\ \ ns3 A B NB # evs3 \ nsp" subsection\declarations for tactics\ @@ -64,7 +64,7 @@ by (auto simp: Gets_correct_def dest: nsp_has_no_Gets) lemma nsp_is_one_step [iff]: "one_step nsp" -by (unfold one_step_def, clarify, ind_cases "ev#evs \ nsp" for ev evs, auto) + unfolding one_step_def by (clarify, ind_cases "ev#evs \ nsp" for ev evs, auto) lemma nsp_has_only_Says' [rule_format]: "evs \ nsp \ ev \ set evs \ (\A B X. ev=Says A B X)" @@ -94,9 +94,9 @@ by (blast intro: analz_insertI)+ lemma no_Nonce_NS1_NS2' [rule_format]: -"[| Crypt (pubK B') \Nonce NA', Nonce NA, Agent A'\ \ parts (spies evs); -Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs); evs \ nsp |] -==> Nonce NA \ analz (spies evs)" +"\Crypt (pubK B') \Nonce NA', Nonce NA, Agent A'\ \ parts (spies evs); +Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs); evs \ nsp\ +\ Nonce NA \ analz (spies evs)" by (rule no_Nonce_NS1_NS2, auto) lemma NB_is_uniq [rule_format]: "evs \ nsp \ @@ -108,7 +108,7 @@ subsection\guardedness of NA\ -lemma ns1_imp_Guard [rule_format]: "[| evs \ nsp; A \ bad; B \ bad |] ==> +lemma ns1_imp_Guard [rule_format]: "\evs \ nsp; A \ bad; B \ bad\ \ ns1 A B NA \ set evs \ Guard NA {priK A,priK B} (spies evs)" apply (erule nsp.induct) (* Nil *) @@ -135,7 +135,7 @@ subsection\guardedness of NB\ -lemma ns2_imp_Guard [rule_format]: "[| evs \ nsp; A \ bad; B \ bad |] ==> +lemma ns2_imp_Guard [rule_format]: "\evs \ nsp; A \ bad; B \ bad\ \ ns2 B A NA NB \ set evs \ Guard NB {priK A,priK B} (spies evs)" apply (erule nsp.induct) (* Nil *) @@ -165,13 +165,13 @@ subsection\Agents' Authentication\ -lemma B_trusts_NS1: "[| evs \ nsp; A \ bad; B \ bad |] ==> +lemma B_trusts_NS1: "\evs \ nsp; A \ bad; B \ bad\ \ Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs) \ Nonce NA \ analz (spies evs) \ ns1 A B NA \ set evs" apply (erule nsp.induct, simp_all) by (blast intro: analz_insertI)+ -lemma A_trusts_NS2: "[| evs \ nsp; A \ bad; B \ bad |] ==> ns1 A B NA \ set evs +lemma A_trusts_NS2: "\evs \ nsp; A \ bad; B \ bad\ \ ns1 A B NA \ set evs \ Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\ \ parts (spies evs) \ ns2 B A NA NB \ set evs" apply (erule nsp.induct, simp_all, safe) @@ -182,7 +182,7 @@ apply (frule_tac B=B in ns1_imp_Guard, simp+) by (drule Guard_Nonce_analz, simp+, blast+) -lemma B_trusts_NS3: "[| evs \ nsp; A \ bad; B \ bad |] ==> ns2 B A NA NB \ set evs +lemma B_trusts_NS3: "\evs \ nsp; A \ bad; B \ bad\ \ ns2 B A NA NB \ set evs \ Crypt (pubK B) (Nonce NB) \ parts (spies evs) \ ns3 A B NB \ set evs" apply (erule nsp.induct, simp_all, safe) apply (frule_tac B=B in ns2_imp_Guard, simp+) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Guard_OtwayRees.thy --- a/src/HOL/Auth/Guard/Guard_OtwayRees.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Guard_OtwayRees.thy Thu Oct 13 17:19:56 2022 +0100 @@ -61,18 +61,18 @@ Nil: "[] \ or" -| Fake: "[| evs \ or; X \ synth (analz (spies evs)) |] ==> Says Spy B X # evs \ or" +| Fake: "\evs \ or; X \ synth (analz (spies evs))\ \ Says Spy B X # evs \ or" -| OR1: "[| evs1 \ or; Nonce NA \ used evs1 |] ==> or1 A B NA # evs1 \ or" +| OR1: "\evs1 \ or; Nonce NA \ used evs1\ \ or1 A B NA # evs1 \ or" -| OR2: "[| evs2 \ or; or1' A' A B NA X \ set evs2; Nonce NB \ used evs2 |] - ==> or2 A B NA NB X # evs2 \ or" +| OR2: "\evs2 \ or; or1' A' A B NA X \ set evs2; Nonce NB \ used evs2\ + \ or2 A B NA NB X # evs2 \ or" -| OR3: "[| evs3 \ or; or2' B' A B NA NB \ set evs3; Key K \ used evs3 |] - ==> or3 A B NA NB K # evs3 \ or" +| OR3: "\evs3 \ or; or2' B' A B NA NB \ set evs3; Key K \ used evs3\ + \ or3 A B NA NB K # evs3 \ or" -| OR4: "[| evs4 \ or; or2 A B NA NB X \ set evs4; or3' S Y A B NA NB K \ set evs4 |] - ==> or4 A B NA X # evs4 \ or" +| OR4: "\evs4 \ or; or2 A B NA NB X \ set evs4; or3' S Y A B NA NB K \ set evs4\ + \ or4 A B NA X # evs4 \ or" subsection\declarations for tactics\ @@ -89,7 +89,7 @@ by (auto simp: Gets_correct_def dest: or_has_no_Gets) lemma or_is_one_step [iff]: "one_step or" -by (unfold one_step_def, clarify, ind_cases "ev#evs \ or" for ev evs, auto) + unfolding one_step_def by (clarify, ind_cases "ev#evs \ or" for ev evs, auto) lemma or_has_only_Says' [rule_format]: "evs \ or \ ev \ set evs \ (\A B X. ev=Says A B X)" @@ -119,7 +119,7 @@ subsection\guardedness of KAB\ -lemma Guard_KAB [rule_format]: "[| evs \ or; A \ bad; B \ bad |] ==> +lemma Guard_KAB [rule_format]: "\evs \ or; A \ bad; B \ bad\ \ or3 A B NA NB K \ set evs \ GuardK K {shrK A,shrK B} (spies evs)" apply (erule or.induct) (* Nil *) @@ -140,7 +140,7 @@ subsection\guardedness of NB\ -lemma Guard_NB [rule_format]: "[| evs \ or; B \ bad |] ==> +lemma Guard_NB [rule_format]: "\evs \ or; B \ bad\ \ or2 A B NA NB X \ set evs \ Guard NB {shrK B} (spies evs)" apply (erule or.induct) (* Nil *) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Guard_Public.thy --- a/src/HOL/Auth/Guard/Guard_Public.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Guard_Public.thy Thu Oct 13 17:19:56 2022 +0100 @@ -54,7 +54,7 @@ definition priK_set :: "key set => bool" where "priK_set Ks \ \K. K \ Ks \ (\A. K = priK A)" -lemma in_priK_set: "[| priK_set Ks; K \ Ks |] ==> \A. K = priK A" +lemma in_priK_set: "\priK_set Ks; K \ Ks\ \ \A. K = priK A" by (simp add: priK_set_def) lemma priK_set1 [iff]: "priK_set {priK A}" @@ -68,13 +68,13 @@ definition good :: "key set => bool" where "good Ks == \K. K \ Ks \ agt K \ bad" -lemma in_good: "[| good Ks; K \ Ks |] ==> agt K \ bad" +lemma in_good: "\good Ks; K \ Ks\ \ agt K \ bad" by (simp add: good_def) lemma good1 [simp]: "A \ bad \ good {priK A}" by (simp add: good_def) -lemma good2 [simp]: "[| A \ bad; B \ bad |] ==> good {priK A, priK B}" +lemma good2 [simp]: "\A \ bad; B \ bad\ \ good {priK A, priK B}" by (simp add: good_def) subsubsection\greatest nonce used in a trace, 0 if there is no nonce\ @@ -122,23 +122,23 @@ by (induct B, auto simp: Guard_def initState.simps) lemma Guard_knows_max': "Guard n Ks (knows_max' C evs) -==> Guard n Ks (knows_max C evs)" +\ Guard n Ks (knows_max C evs)" by (simp add: knows_max_def) lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \ used evs \ Guard n Ks (spies evs)" by (auto simp: Guard_def dest: not_used_not_known parts_sub) -lemma Nonce_not_used_Guard [dest]: "[| evs \ p; Nonce n \ used evs; -Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" +lemma Nonce_not_used_Guard [dest]: "\evs \ p; Nonce n \ used evs; +Gets_correct p; one_step p\ \ Guard n Ks (knows (Friend C) evs)" by (auto simp: Guard_def dest: known_used parts_trans) -lemma Nonce_not_used_Guard_max [dest]: "[| evs \ p; Nonce n \ used evs; -Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" +lemma Nonce_not_used_Guard_max [dest]: "\evs \ p; Nonce n \ used evs; +Gets_correct p; one_step p\ \ Guard n Ks (knows_max (Friend C) evs)" by (auto simp: Guard_def dest: known_max_used parts_trans) -lemma Nonce_not_used_Guard_max' [dest]: "[| evs \ p; Nonce n \ used evs; -Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" +lemma Nonce_not_used_Guard_max' [dest]: "\evs \ p; Nonce n \ used evs; +Gets_correct p; one_step p\ \ Guard n Ks (knows_max' (Friend C) evs)" apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) by (auto simp: knows_max_def) @@ -147,16 +147,16 @@ definition regular :: "event list set \ bool" where "regular p \ \evs A. evs \ p \ (Key (priK A) \ parts (spies evs)) = (A \ bad)" -lemma priK_parts_iff_bad [simp]: "[| evs \ p; regular p |] ==> +lemma priK_parts_iff_bad [simp]: "\evs \ p; regular p\ \ (Key (priK A) \ parts (spies evs)) = (A \ bad)" by (auto simp: regular_def) -lemma priK_analz_iff_bad [simp]: "[| evs \ p; regular p |] ==> +lemma priK_analz_iff_bad [simp]: "\evs \ p; regular p\ \ (Key (priK A) \ analz (spies evs)) = (A \ bad)" by auto -lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \ p; -priK_set Ks; good Ks; regular p |] ==> Nonce n \ analz (spies evs)" +lemma Guard_Nonce_analz: "\Guard n Ks (spies evs); evs \ p; +priK_set Ks; good Ks; regular p\ \ Nonce n \ analz (spies evs)" apply (clarify, simp only: knows_decomp) apply (drule Guard_invKey_keyset, simp+, safe) apply (drule in_good, simp) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Guard_Shared.thy --- a/src/HOL/Auth/Guard/Guard_Shared.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Guard_Shared.thy Thu Oct 13 17:19:56 2022 +0100 @@ -49,7 +49,7 @@ definition shrK_set :: "key set => bool" where "shrK_set Ks \ \K. K \ Ks \ (\A. K = shrK A)" -lemma in_shrK_set: "[| shrK_set Ks; K \ Ks |] ==> \A. K = shrK A" +lemma in_shrK_set: "\shrK_set Ks; K \ Ks\ \ \A. K = shrK A" by (simp add: shrK_set_def) lemma shrK_set1 [iff]: "shrK_set {shrK A}" @@ -63,13 +63,13 @@ definition good :: "key set \ bool" where "good Ks \ \K. K \ Ks \ agt K \ bad" -lemma in_good: "[| good Ks; K \ Ks |] ==> agt K \ bad" +lemma in_good: "\good Ks; K \ Ks\ \ agt K \ bad" by (simp add: good_def) lemma good1 [simp]: "A \ bad \ good {shrK A}" by (simp add: good_def) -lemma good2 [simp]: "[| A \ bad; B \ bad |] ==> good {shrK A, shrK B}" +lemma good2 [simp]: "\A \ bad; B \ bad\ \ good {shrK A, shrK B}" by (simp add: good_def) @@ -100,23 +100,23 @@ by (induct B, auto simp: Guard_def initState.simps) lemma Guard_knows_max': "Guard n Ks (knows_max' C evs) -==> Guard n Ks (knows_max C evs)" +\ Guard n Ks (knows_max C evs)" by (simp add: knows_max_def) lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \ used evs \ Guard n Ks (spies evs)" by (auto simp: Guard_def dest: not_used_not_known parts_sub) -lemma Nonce_not_used_Guard [dest]: "[| evs \ p; Nonce n \ used evs; -Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" +lemma Nonce_not_used_Guard [dest]: "\evs \ p; Nonce n \ used evs; +Gets_correct p; one_step p\ \ Guard n Ks (knows (Friend C) evs)" by (auto simp: Guard_def dest: known_used parts_trans) -lemma Nonce_not_used_Guard_max [dest]: "[| evs \ p; Nonce n \ used evs; -Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" +lemma Nonce_not_used_Guard_max [dest]: "\evs \ p; Nonce n \ used evs; +Gets_correct p; one_step p\ \ Guard n Ks (knows_max (Friend C) evs)" by (auto simp: Guard_def dest: known_max_used parts_trans) -lemma Nonce_not_used_Guard_max' [dest]: "[| evs \ p; Nonce n \ used evs; -Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" +lemma Nonce_not_used_Guard_max' [dest]: "\evs \ p; Nonce n \ used evs; +Gets_correct p; one_step p\ \ Guard n Ks (knows_max' (Friend C) evs)" apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) by (auto simp: knows_max_def) @@ -125,24 +125,24 @@ lemma GuardK_init [simp]: "n \ range shrK \ GuardK n Ks (initState B)" by (induct B, auto simp: GuardK_def initState.simps) -lemma GuardK_knows_max': "[| GuardK n A (knows_max' C evs); n \ range shrK |] -==> GuardK n A (knows_max C evs)" +lemma GuardK_knows_max': "\GuardK n A (knows_max' C evs); n \ range shrK\ +\ GuardK n A (knows_max C evs)" by (simp add: knows_max_def) lemma Key_not_used_GuardK_spies [dest]: "Key n \ used evs \ GuardK n A (spies evs)" by (auto simp: GuardK_def dest: not_used_not_known parts_sub) -lemma Key_not_used_GuardK [dest]: "[| evs \ p; Key n \ used evs; -Gets_correct p; one_step p |] ==> GuardK n A (knows (Friend C) evs)" +lemma Key_not_used_GuardK [dest]: "\evs \ p; Key n \ used evs; +Gets_correct p; one_step p\ \ GuardK n A (knows (Friend C) evs)" by (auto simp: GuardK_def dest: known_used parts_trans) -lemma Key_not_used_GuardK_max [dest]: "[| evs \ p; Key n \ used evs; -Gets_correct p; one_step p |] ==> GuardK n A (knows_max (Friend C) evs)" +lemma Key_not_used_GuardK_max [dest]: "\evs \ p; Key n \ used evs; +Gets_correct p; one_step p\ \ GuardK n A (knows_max (Friend C) evs)" by (auto simp: GuardK_def dest: known_max_used parts_trans) -lemma Key_not_used_GuardK_max' [dest]: "[| evs \ p; Key n \ used evs; -Gets_correct p; one_step p |] ==> GuardK n A (knows_max' (Friend C) evs)" +lemma Key_not_used_GuardK_max' [dest]: "\evs \ p; Key n \ used evs; +Gets_correct p; one_step p\ \ GuardK n A (knows_max' (Friend C) evs)" apply (rule_tac H="knows_max (Friend C) evs" in GuardK_mono) by (auto simp: knows_max_def) @@ -151,16 +151,16 @@ definition regular :: "event list set => bool" where "regular p \ \evs A. evs \ p \ (Key (shrK A) \ parts (spies evs)) = (A \ bad)" -lemma shrK_parts_iff_bad [simp]: "[| evs \ p; regular p |] ==> +lemma shrK_parts_iff_bad [simp]: "\evs \ p; regular p\ \ (Key (shrK A) \ parts (spies evs)) = (A \ bad)" by (auto simp: regular_def) -lemma shrK_analz_iff_bad [simp]: "[| evs \ p; regular p |] ==> +lemma shrK_analz_iff_bad [simp]: "\evs \ p; regular p\ \ (Key (shrK A) \ analz (spies evs)) = (A \ bad)" by auto -lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \ p; -shrK_set Ks; good Ks; regular p |] ==> Nonce n \ analz (spies evs)" +lemma Guard_Nonce_analz: "\Guard n Ks (spies evs); evs \ p; +shrK_set Ks; good Ks; regular p\ \ Nonce n \ analz (spies evs)" apply (clarify, simp only: knows_decomp) apply (drule Guard_invKey_keyset, simp+, safe) apply (drule in_good, simp) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Guard_Yahalom.thy --- a/src/HOL/Auth/Guard/Guard_Yahalom.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy Thu Oct 13 17:19:56 2022 +0100 @@ -52,18 +52,18 @@ Nil: "[] \ ya" -| Fake: "[| evs \ ya; X \ synth (analz (spies evs)) |] ==> Says Spy B X # evs \ ya" +| Fake: "\evs \ ya; X \ synth (analz (spies evs))\ \ Says Spy B X # evs \ ya" -| YA1: "[| evs1 \ ya; Nonce NA \ used evs1 |] ==> ya1 A B NA # evs1 \ ya" +| YA1: "\evs1 \ ya; Nonce NA \ used evs1\ \ ya1 A B NA # evs1 \ ya" -| YA2: "[| evs2 \ ya; ya1' A' A B NA \ set evs2; Nonce NB \ used evs2 |] - ==> ya2 A B NA NB # evs2 \ ya" +| YA2: "\evs2 \ ya; ya1' A' A B NA \ set evs2; Nonce NB \ used evs2\ + \ ya2 A B NA NB # evs2 \ ya" -| YA3: "[| evs3 \ ya; ya2' B' A B NA NB \ set evs3; Key K \ used evs3 |] - ==> ya3 A B NA NB K # evs3 \ ya" +| YA3: "\evs3 \ ya; ya2' B' A B NA NB \ set evs3; Key K \ used evs3\ + \ ya3 A B NA NB K # evs3 \ ya" -| YA4: "[| evs4 \ ya; ya1 A B NA \ set evs4; ya3' S Y A B NA NB K \ set evs4 |] - ==> ya4 A B K NB Y # evs4 \ ya" +| YA4: "\evs4 \ ya; ya1 A B NA \ set evs4; ya3' S Y A B NA NB K \ set evs4\ + \ ya4 A B K NB Y # evs4 \ ya" subsection\declarations for tactics\ @@ -80,7 +80,7 @@ by (auto simp: Gets_correct_def dest: ya_has_no_Gets) lemma ya_is_one_step [iff]: "one_step ya" -by (unfold one_step_def, clarify, ind_cases "ev#evs \ ya" for ev evs, auto) + unfolding one_step_def by (clarify, ind_cases "ev#evs \ ya" for ev evs, auto) lemma ya_has_only_Says' [rule_format]: "evs \ ya \ ev \ set evs \ (\A B X. ev=Says A B X)" @@ -96,7 +96,7 @@ subsection\guardedness of KAB\ -lemma Guard_KAB [rule_format]: "[| evs \ ya; A \ bad; B \ bad |] ==> +lemma Guard_KAB [rule_format]: "\evs \ ya; A \ bad; B \ bad\ \ ya3 A B NA NB K \ set evs \ GuardK K {shrK A,shrK B} (spies evs)" apply (erule ya.induct) (* Nil *) @@ -127,18 +127,18 @@ subsection\ya2' implies ya1'\ lemma ya2'_parts_imp_ya1'_parts [rule_format]: - "[| evs \ ya; B \ bad |] ==> + "\evs \ ya; B \ bad\ \ Ciph B \Agent A, Nonce NA, Nonce NB\ \ parts (spies evs) \ \Agent A, Nonce NA\ \ spies evs" by (erule ya.induct, auto dest: Says_imp_spies intro: parts_parts) -lemma ya2'_imp_ya1'_parts: "[| ya2' B' A B NA NB \ set evs; evs \ ya; B \ bad |] -==> \Agent A, Nonce NA\ \ spies evs" +lemma ya2'_imp_ya1'_parts: "\ya2' B' A B NA NB \ set evs; evs \ ya; B \ bad\ +\ \Agent A, Nonce NA\ \ spies evs" by (blast dest: Says_imp_spies ya2'_parts_imp_ya1'_parts) subsection\uniqueness of NB\ -lemma NB_is_uniq_in_ya2'_parts [rule_format]: "[| evs \ ya; B \ bad; B' \ bad |] ==> +lemma NB_is_uniq_in_ya2'_parts [rule_format]: "\evs \ ya; B \ bad; B' \ bad\ \ Ciph B \Agent A, Nonce NA, Nonce NB\ \ parts (spies evs) \ Ciph B' \Agent A', Nonce NA', Nonce NB\ \ parts (spies evs) \ A=A' \ B=B' \ NA=NA'" @@ -148,14 +148,14 @@ apply (drule not_used_parts_false, simp+)+ by (drule Says_not_parts, simp+)+ -lemma NB_is_uniq_in_ya2': "[| ya2' C A B NA NB \ set evs; -ya2' C' A' B' NA' NB \ set evs; evs \ ya; B \ bad; B' \ bad |] -==> A=A' \ B=B' \ NA=NA'" +lemma NB_is_uniq_in_ya2': "\ya2' C A B NA NB \ set evs; +ya2' C' A' B' NA' NB \ set evs; evs \ ya; B \ bad; B' \ bad\ +\ A=A' \ B=B' \ NA=NA'" by (drule NB_is_uniq_in_ya2'_parts, auto dest: Says_imp_spies) subsection\ya3' implies ya2'\ -lemma ya3'_parts_imp_ya2'_parts [rule_format]: "[| evs \ ya; A \ bad |] ==> +lemma ya3'_parts_imp_ya2'_parts [rule_format]: "\evs \ ya; A \ bad\ \ Ciph A \Agent B, Key K, Nonce NA, Nonce NB\ \ parts (spies evs) \ Ciph B \Agent A, Nonce NA, Nonce NB\ \ parts (spies evs)" apply (erule ya.induct, simp_all) @@ -163,7 +163,7 @@ apply (blast intro: parts_sub, blast) by (auto dest: Says_imp_spies parts_parts) -lemma ya3'_parts_imp_ya2' [rule_format]: "[| evs \ ya; A \ bad |] ==> +lemma ya3'_parts_imp_ya2' [rule_format]: "\evs \ ya; A \ bad\ \ Ciph A \Agent B, Key K, Nonce NA, Nonce NB\ \ parts (spies evs) \ (\B'. ya2' B' A B NA NB \ set evs)" apply (erule ya.induct, simp_all, safe) @@ -173,21 +173,21 @@ apply blast by (auto dest: Says_imp_spies2 parts_parts) -lemma ya3'_imp_ya2': "[| ya3' S Y A B NA NB K \ set evs; evs \ ya; A \ bad |] -==> (\B'. ya2' B' A B NA NB \ set evs)" +lemma ya3'_imp_ya2': "\ya3' S Y A B NA NB K \ set evs; evs \ ya; A \ bad\ +\ (\B'. ya2' B' A B NA NB \ set evs)" by (drule ya3'_parts_imp_ya2', auto dest: Says_imp_spies) subsection\ya3' implies ya3\ -lemma ya3'_parts_imp_ya3 [rule_format]: "[| evs \ ya; A \ bad |] ==> +lemma ya3'_parts_imp_ya3 [rule_format]: "\evs \ ya; A \ bad\ \ Ciph A \Agent B, Key K, Nonce NA, Nonce NB\ \ parts(spies evs) \ ya3 A B NA NB K \ set evs" apply (erule ya.induct, simp_all, safe) apply (drule Crypt_synth_insert, simp+) by (blast dest: Says_imp_spies2 parts_parts) -lemma ya3'_imp_ya3: "[| ya3' S Y A B NA NB K \ set evs; evs \ ya; A \ bad |] -==> ya3 A B NA NB K \ set evs" +lemma ya3'_imp_ya3: "\ya3' S Y A B NA NB K \ set evs; evs \ ya; A \ bad\ +\ ya3 A B NA NB K \ set evs" by (blast dest: Says_imp_spies ya3'_parts_imp_ya3) subsection\guardedness of NB\ @@ -195,7 +195,7 @@ definition ya_keys :: "agent \ agent \ nat \ nat \ event list \ key set" where "ya_keys A B NA NB evs \ {shrK A,shrK B} \ {K. ya3 A B NA NB K \ set evs}" -lemma Guard_NB [rule_format]: "[| evs \ ya; A \ bad; B \ bad |] ==> +lemma Guard_NB [rule_format]: "\evs \ ya; A \ bad; B \ bad\ \ ya2 A B NA NB \ set evs \ Guard NB (ya_keys A B NA NB evs) (spies evs)" apply (erule ya.induct) (* Nil *) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/List_Msg.thy --- a/src/HOL/Auth/Guard/List_Msg.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/List_Msg.thy Thu Oct 13 17:19:56 2022 +0100 @@ -17,8 +17,8 @@ subsubsection\induction principle\ -lemma lmsg_induct: "[| !!x. not_MPair x ==> P x; !!x l. P l ==> P (cons x l) |] -==> P l" +lemma lmsg_induct: "\!!x. not_MPair x \ P x; !!x l. P l \ P (cons x l)\ +\ P l" by (induct l) auto subsubsection\head\ @@ -52,7 +52,7 @@ "del (x, cons y l) = (if x=y then l else cons y (del (x,l)))" | "del (x, other) = other" -lemma notin_del [simp]: "~ isin (x,l) ==> del (x,l) = l" +lemma notin_del [simp]: "~ isin (x,l) \ del (x,l) = l" by (induct l) auto lemma isin_del [rule_format]: "isin (y, del (x,l)) --> isin (y,l)" @@ -81,7 +81,7 @@ "ith (cons x l, 0) = x" | "ith (other, i) = other" -lemma ith_head: "0 < len l ==> ith (l,0) = head l" +lemma ith_head: "0 < len l \ ith (l,0) = head l" by (cases l) auto subsubsection\insertion\ @@ -114,14 +114,14 @@ inductive_set agl :: "msg set" where Nil[intro]: "nil \ agl" -| Cons[intro]: "[| A \ agent; I \ agl |] ==> cons (Agent A) I \ agl" +| Cons[intro]: "\A \ agent; I \ agl\ \ cons (Agent A) I \ agl" subsubsection\basic facts about agent lists\ lemma del_in_agl [intro]: "I \ agl \ del (a,I) \ agl" by (erule agl.induct, auto) -lemma app_in_agl [intro]: "[| I \ agl; J \ agl |] ==> app (I,J) \ agl" +lemma app_in_agl [intro]: "\I \ agl; J \ agl\ \ app (I,J) \ agl" by (erule agl.induct, auto) lemma no_Key_in_agl: "I \ agl \ Key K \ parts {I}" @@ -130,18 +130,18 @@ lemma no_Nonce_in_agl: "I \ agl \ Nonce n \ parts {I}" by (erule agl.induct, auto) -lemma no_Key_in_appdel: "[| I \ agl; J \ agl |] ==> +lemma no_Key_in_appdel: "\I \ agl; J \ agl\ \ Key K \ parts {app (J, del (Agent B, I))}" by (rule no_Key_in_agl, auto) -lemma no_Nonce_in_appdel: "[| I \ agl; J \ agl |] ==> +lemma no_Nonce_in_appdel: "\I \ agl; J \ agl\ \ Nonce n \ parts {app (J, del (Agent B, I))}" by (rule no_Nonce_in_agl, auto) lemma no_Crypt_in_agl: "I \ agl \ Crypt K X \ parts {I}" by (erule agl.induct, auto) -lemma no_Crypt_in_appdel: "[| I \ agl; J \ agl |] ==> +lemma no_Crypt_in_appdel: "\I \ agl; J \ agl\ \ Crypt K X \ parts {app (J, del (Agent B,I))}" by (rule no_Crypt_in_agl, auto) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/P1.thy --- a/src/HOL/Auth/Guard/P1.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/P1.thy Thu Oct 13 17:19:56 2022 +0100 @@ -122,7 +122,7 @@ lemma prom_inj [dest]: "prom B ofr A r I L J C = prom B' ofr' A' r' I' L' J' C' -==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" +\ B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" by (auto simp: prom_def) lemma Nonce_in_prom [iff]: "Nonce ofr \ parts {prom B ofr A r I L J C}" @@ -133,7 +133,7 @@ "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' -==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" +\ B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" by (auto simp: pro_def dest: prom_inj) subsubsection\protocol\ @@ -143,13 +143,13 @@ Nil: "[] \ p1" -| Fake: "[| evsf \ p1; X \ synth (analz (spies evsf)) |] ==> Says Spy B X # evsf \ p1" +| Fake: "\evsf \ p1; X \ synth (analz (spies evsf))\ \ Says Spy B X # evsf \ p1" -| Request: "[| evsr \ p1; Nonce n \ used evsr; I \ agl |] ==> req A r n I B # evsr \ p1" +| Request: "\evsr \ p1; Nonce n \ used evsr; I \ agl\ \ req A r n I B # evsr \ p1" -| Propose: "[| evsp \ p1; Says A' B \Agent A,Number r,I,cons M L\ \ set evsp; +| Propose: "\evsp \ p1; Says A' B \Agent A,Number r,I,cons M L\ \ set evsp; I \ agl; J \ agl; isin (Agent C, app (J, del (Agent B, I))); - Nonce ofr \ used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp \ p1" + Nonce ofr \ used evsp\ \ pro B ofr A r I (cons M L) J C # evsp \ p1" subsubsection\Composition of Traces\ @@ -354,7 +354,7 @@ by (auto simp: Gets_correct_def dest: p1_has_no_Gets) lemma p1_is_one_step [iff]: "one_step p1" -by (unfold one_step_def, clarify, ind_cases "ev#evs \ p1" for ev evs, auto) + unfolding one_step_def by (clarify, ind_cases "ev#evs \ p1" for ev evs, auto) lemma p1_has_only_Says' [rule_format]: "evs \ p1 \ ev \ set evs \ (\A B X. ev=Says A B X)" @@ -373,8 +373,8 @@ subsubsection\private keys are safe\ lemma priK_parts_Friend_imp_bad [rule_format,dest]: - "[| evs \ p1; Friend B \ A |] - ==> (Key (priK A) \ parts (knows (Friend B) evs)) \ (A \ bad)" + "\evs \ p1; Friend B \ A\ + \ (Key (priK A) \ parts (knows (Friend B) evs)) \ (A \ bad)" apply (erule p1.induct) apply (simp_all add: initState.simps knows.simps pro_def prom_def req_def reqm_def anchor_def chain_def sign_def) @@ -384,12 +384,12 @@ done lemma priK_analz_Friend_imp_bad [rule_format,dest]: - "[| evs \ p1; Friend B \ A |] -==> (Key (priK A) \ analz (knows (Friend B) evs)) \ (A \ bad)" + "\evs \ p1; Friend B \ A\ +\ (Key (priK A) \ analz (knows (Friend B) evs)) \ (A \ bad)" by auto -lemma priK_notin_knows_max_Friend: "[| evs \ p1; A \ bad; A \ Friend C |] -==> Key (priK A) \ analz (knows_max (Friend C) evs)" +lemma priK_notin_knows_max_Friend: "\evs \ p1; A \ bad; A \ Friend C\ +\ Key (priK A) \ analz (knows_max (Friend C) evs)" apply (rule not_parts_not_analz, simp add: knows_max_def, safe) apply (drule_tac H="spies' evs" in parts_sub) apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) @@ -401,16 +401,16 @@ lemma agl_guard [intro]: "I \ agl \ I \ guard n Ks" by (erule agl.induct, auto) -lemma Says_to_knows_max'_guard: "[| Says A' C \A'',r,I,L\ \ set evs; -Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" +lemma Says_to_knows_max'_guard: "\Says A' C \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs)\ \ L \ guard n Ks" by (auto dest: Says_to_knows_max') -lemma Says_from_knows_max'_guard: "[| Says C A' \A'',r,I,L\ \ set evs; -Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" +lemma Says_from_knows_max'_guard: "\Says C A' \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs)\ \ L \ guard n Ks" by (auto dest: Says_from_knows_max') -lemma Says_Nonce_not_used_guard: "[| Says A' B \A'',r,I,L\ \ set evs; -Nonce n \ used evs |] ==> L \ guard n Ks" +lemma Says_Nonce_not_used_guard: "\Says A' B \A'',r,I,L\ \ set evs; +Nonce n \ used evs\ \ L \ guard n Ks" by (drule not_used_not_parts, auto) subsubsection\guardedness of messages\ @@ -432,16 +432,16 @@ lemma reqm_guard [intro]: "I \ agl \ reqm A r n' I B \ guard n {priK A}" by (case_tac "n'=n", auto simp: reqm_def) -lemma reqm_guard_Nonce_neq [intro]: "[| n \ n'; I \ agl |] -==> reqm A' r n' I B \ guard n {priK A}" +lemma reqm_guard_Nonce_neq [intro]: "\n \ n'; I \ agl\ +\ reqm A' r n' I B \ guard n {priK A}" by (auto simp: reqm_def) -lemma prom_guard [intro]: "[| I \ agl; J \ agl; L \ guard n {priK A} |] -==> prom B ofr A r I L J C \ guard n {priK A}" +lemma prom_guard [intro]: "\I \ agl; J \ agl; L \ guard n {priK A}\ +\ prom B ofr A r I L J C \ guard n {priK A}" by (auto simp: prom_def) -lemma prom_guard_Nonce_neq [intro]: "[| n \ ofr; I \ agl; J \ agl; -L \ guard n {priK A} |] ==> prom B ofr A' r I L J C \ guard n {priK A}" +lemma prom_guard_Nonce_neq [intro]: "\n \ ofr; I \ agl; J \ agl; +L \ guard n {priK A}\ \ prom B ofr A' r I L J C \ guard n {priK A}" by (auto simp: prom_def) subsubsection\Nonce uniqueness\ @@ -452,25 +452,25 @@ lemma uniq_Nonce_in_anchor [dest]: "Nonce k \ parts {anchor A n B} \ k=n" by (auto simp: anchor_def chain_def sign_def) -lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k \ parts {reqm A r n I B}; -I \ agl |] ==> k=n" +lemma uniq_Nonce_in_reqm [dest]: "\Nonce k \ parts {reqm A r n I B}; +I \ agl\ \ k=n" by (auto simp: reqm_def dest: no_Nonce_in_agl) -lemma uniq_Nonce_in_prom [dest]: "[| Nonce k \ parts {prom B ofr A r I L J C}; -I \ agl; J \ agl; Nonce k \ parts {L} |] ==> k=ofr" +lemma uniq_Nonce_in_prom [dest]: "\Nonce k \ parts {prom B ofr A r I L J C}; +I \ agl; J \ agl; Nonce k \ parts {L}\ \ k=ofr" by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) subsubsection\requests are guarded\ -lemma req_imp_Guard [rule_format]: "[| evs \ p1; A \ bad |] ==> +lemma req_imp_Guard [rule_format]: "\evs \ p1; A \ bad\ \ req A r n I B \ set evs \ Guard n {priK A} (spies evs)" apply (erule p1.induct, simp) apply (simp add: req_def knows.simps, safe) apply (erule in_synth_Guard, erule Guard_analz, simp) by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) -lemma req_imp_Guard_Friend: "[| evs \ p1; A \ bad; req A r n I B \ set evs |] -==> Guard n {priK A} (knows_max (Friend C) evs)" +lemma req_imp_Guard_Friend: "\evs \ p1; A \ bad; req A r n I B \ set evs\ +\ Guard n {priK A} (knows_max (Friend C) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) apply (rule req_imp_Guard, simp+) @@ -480,7 +480,7 @@ subsubsection\propositions are guarded\ -lemma pro_imp_Guard [rule_format]: "[| evs \ p1; B \ bad; A \ bad |] ==> +lemma pro_imp_Guard [rule_format]: "\evs \ p1; B \ bad; A \ bad\ \ pro B ofr A r I (cons M L) J C \ set evs \ Guard ofr {priK A} (spies evs)" supply [[simproc del: defined_all]] apply (erule p1.induct) (* +3 subgoals *) @@ -518,9 +518,9 @@ apply (simp add: pro_def) by (blast dest: Says_imp_knows_Spy) -lemma pro_imp_Guard_Friend: "[| evs \ p1; B \ bad; A \ bad; -pro B ofr A r I (cons M L) J C \ set evs |] -==> Guard ofr {priK A} (knows_max (Friend D) evs)" +lemma pro_imp_Guard_Friend: "\evs \ p1; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs\ +\ Guard ofr {priK A} (knows_max (Friend D) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) apply (rule pro_imp_Guard, simp+) @@ -531,23 +531,23 @@ subsubsection\data confidentiality: no one other than the originator can decrypt the offers\ -lemma Nonce_req_notin_spies: "[| evs \ p1; req A r n I B \ set evs; A \ bad |] -==> Nonce n \ analz (spies evs)" +lemma Nonce_req_notin_spies: "\evs \ p1; req A r n I B \ set evs; A \ bad\ +\ Nonce n \ analz (spies evs)" by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_req_notin_knows_max_Friend: "[| evs \ p1; req A r n I B \ set evs; -A \ bad; A \ Friend C |] ==> Nonce n \ analz (knows_max (Friend C) evs)" +lemma Nonce_req_notin_knows_max_Friend: "\evs \ p1; req A r n I B \ set evs; +A \ bad; A \ Friend C\ \ Nonce n \ analz (knows_max (Friend C) evs)" apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) -lemma Nonce_pro_notin_spies: "[| evs \ p1; B \ bad; A \ bad; -pro B ofr A r I (cons M L) J C \ set evs |] ==> Nonce ofr \ analz (spies evs)" +lemma Nonce_pro_notin_spies: "\evs \ p1; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs\ \ Nonce ofr \ analz (spies evs)" by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_pro_notin_knows_max_Friend: "[| evs \ p1; B \ bad; A \ bad; -A \ Friend D; pro B ofr A r I (cons M L) J C \ set evs |] -==> Nonce ofr \ analz (knows_max (Friend D) evs)" +lemma Nonce_pro_notin_knows_max_Friend: "\evs \ p1; B \ bad; A \ bad; +A \ Friend D; pro B ofr A r I (cons M L) J C \ set evs\ +\ Nonce ofr \ analz (knows_max (Friend D) evs)" apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) @@ -555,15 +555,15 @@ subsubsection\non repudiability: an offer signed by B has been sent by B\ -lemma Crypt_reqm: "[| Crypt (priK A) X \ parts {reqm A' r n I B}; I \ agl |] ==> A=A'" +lemma Crypt_reqm: "\Crypt (priK A) X \ parts {reqm A' r n I B}; I \ agl\ \ A=A'" by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) -lemma Crypt_prom: "[| Crypt (priK A) X \ parts {prom B ofr A' r I L J C}; -I \ agl; J \ agl |] ==> A=B \ Crypt (priK A) X \ parts {L}" +lemma Crypt_prom: "\Crypt (priK A) X \ parts {prom B ofr A' r I L J C}; +I \ agl; J \ agl\ \ A=B \ Crypt (priK A) X \ parts {L}" apply (simp add: prom_def anchor_def chain_def sign_def) by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) -lemma Crypt_safeness: "[| evs \ p1; A \ bad |] ==> Crypt (priK A) X \ parts (spies evs) +lemma Crypt_safeness: "\evs \ p1; A \ bad\ \ Crypt (priK A) X \ parts (spies evs) \ (\B Y. Says A B Y \ set evs \ Crypt (priK A) X \ parts {Y})" apply (erule p1.induct) (* Nil *) @@ -593,7 +593,7 @@ apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) by auto -lemma Crypt_Hash_imp_sign: "[| evs \ p1; A \ bad |] ==> +lemma Crypt_Hash_imp_sign: "\evs \ p1; A \ bad\ \ Crypt (priK A) (Hash X) \ parts (spies evs) \ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (erule p1.induct) @@ -630,7 +630,7 @@ apply (blast del: MPair_parts)+ done -lemma sign_safeness: "[| evs \ p1; A \ bad |] ==> sign A X \ parts (spies evs) +lemma sign_safeness: "\evs \ p1; A \ bad\ \ sign A X \ parts (spies evs) \ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (clarify, simp add: sign_def, frule parts.Snd) apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/P2.thy --- a/src/HOL/Auth/Guard/P2.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/P2.thy Thu Oct 13 17:19:56 2022 +0100 @@ -102,7 +102,7 @@ app (J, del (Agent B, I)), cons (chain B ofr A L C) L\" lemma prom_inj [dest]: "prom B ofr A r I L J C = prom B' ofr' A' r' I' L' J' C' -==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" +\ B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" by (auto simp: prom_def) lemma Nonce_in_prom [iff]: "Nonce ofr \ parts {prom B ofr A r I L J C}" @@ -113,7 +113,7 @@ "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' -==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" +\ B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" by (auto simp: pro_def dest: prom_inj) subsubsection\protocol\ @@ -123,13 +123,13 @@ Nil: "[] \ p2" -| Fake: "[| evsf \ p2; X \ synth (analz (spies evsf)) |] ==> Says Spy B X # evsf \ p2" +| Fake: "\evsf \ p2; X \ synth (analz (spies evsf))\ \ Says Spy B X # evsf \ p2" -| Request: "[| evsr \ p2; Nonce n \ used evsr; I \ agl |] ==> req A r n I B # evsr \ p2" +| Request: "\evsr \ p2; Nonce n \ used evsr; I \ agl\ \ req A r n I B # evsr \ p2" -| Propose: "[| evsp \ p2; Says A' B \Agent A,Number r,I,cons M L\ \ set evsp; +| Propose: "\evsp \ p2; Says A' B \Agent A,Number r,I,cons M L\ \ set evsp; I \ agl; J \ agl; isin (Agent C, app (J, del (Agent B, I))); - Nonce ofr \ used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp \ p2" + Nonce ofr \ used evsp\ \ pro B ofr A r I (cons M L) J C # evsp \ p2" subsubsection\valid offer lists\ @@ -265,7 +265,7 @@ by (auto simp: Gets_correct_def dest: p2_has_no_Gets) lemma p2_is_one_step [iff]: "one_step p2" -by (unfold one_step_def, clarify, ind_cases "ev#evs \ p2" for ev evs, auto) + unfolding one_step_def by (clarify, ind_cases "ev#evs \ p2" for ev evs, auto) lemma p2_has_only_Says' [rule_format]: "evs \ p2 \ ev \ set evs \ (\A B X. ev=Says A B X)" @@ -284,8 +284,8 @@ subsection\private keys are safe\ lemma priK_parts_Friend_imp_bad [rule_format,dest]: - "[| evs \ p2; Friend B \ A |] - ==> (Key (priK A) \ parts (knows (Friend B) evs)) \ (A \ bad)" + "\evs \ p2; Friend B \ A\ + \ (Key (priK A) \ parts (knows (Friend B) evs)) \ (A \ bad)" apply (erule p2.induct) apply (simp_all add: initState.simps knows.simps pro_def prom_def req_def reqm_def anchor_def chain_def sign_def) @@ -295,13 +295,13 @@ done lemma priK_analz_Friend_imp_bad [rule_format,dest]: - "[| evs \ p2; Friend B \ A |] -==> (Key (priK A) \ analz (knows (Friend B) evs)) \ (A \ bad)" + "\evs \ p2; Friend B \ A\ +\ (Key (priK A) \ analz (knows (Friend B) evs)) \ (A \ bad)" by auto lemma priK_notin_knows_max_Friend: - "[| evs \ p2; A \ bad; A \ Friend C |] - ==> Key (priK A) \ analz (knows_max (Friend C) evs)" + "\evs \ p2; A \ bad; A \ Friend C\ + \ Key (priK A) \ analz (knows_max (Friend C) evs)" apply (rule not_parts_not_analz, simp add: knows_max_def, safe) apply (drule_tac H="spies' evs" in parts_sub) apply (rule_tac p=p2 in knows_max'_sub_spies', simp+) @@ -313,16 +313,16 @@ lemma agl_guard [intro]: "I \ agl \ I \ guard n Ks" by (erule agl.induct, auto) -lemma Says_to_knows_max'_guard: "[| Says A' C \A'',r,I,L\ \ set evs; -Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" +lemma Says_to_knows_max'_guard: "\Says A' C \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs)\ \ L \ guard n Ks" by (auto dest: Says_to_knows_max') -lemma Says_from_knows_max'_guard: "[| Says C A' \A'',r,I,L\ \ set evs; -Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" +lemma Says_from_knows_max'_guard: "\Says C A' \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs)\ \ L \ guard n Ks" by (auto dest: Says_from_knows_max') -lemma Says_Nonce_not_used_guard: "[| Says A' B \A'',r,I,L\ \ set evs; -Nonce n \ used evs |] ==> L \ guard n Ks" +lemma Says_Nonce_not_used_guard: "\Says A' B \A'',r,I,L\ \ set evs; +Nonce n \ used evs\ \ L \ guard n Ks" by (drule not_used_not_parts, auto) subsection\guardedness of messages\ @@ -344,16 +344,16 @@ lemma reqm_guard [intro]: "I \ agl \ reqm A r n' I B \ guard n {priK A}" by (case_tac "n'=n", auto simp: reqm_def) -lemma reqm_guard_Nonce_neq [intro]: "[| n \ n'; I \ agl |] -==> reqm A' r n' I B \ guard n {priK A}" +lemma reqm_guard_Nonce_neq [intro]: "\n \ n'; I \ agl\ +\ reqm A' r n' I B \ guard n {priK A}" by (auto simp: reqm_def) -lemma prom_guard [intro]: "[| I \ agl; J \ agl; L \ guard n {priK A} |] -==> prom B ofr A r I L J C \ guard n {priK A}" +lemma prom_guard [intro]: "\I \ agl; J \ agl; L \ guard n {priK A}\ +\ prom B ofr A r I L J C \ guard n {priK A}" by (auto simp: prom_def) -lemma prom_guard_Nonce_neq [intro]: "[| n \ ofr; I \ agl; J \ agl; -L \ guard n {priK A} |] ==> prom B ofr A' r I L J C \ guard n {priK A}" +lemma prom_guard_Nonce_neq [intro]: "\n \ ofr; I \ agl; J \ agl; +L \ guard n {priK A}\ \ prom B ofr A' r I L J C \ guard n {priK A}" by (auto simp: prom_def) subsection\Nonce uniqueness\ @@ -364,25 +364,25 @@ lemma uniq_Nonce_in_anchor [dest]: "Nonce k \ parts {anchor A n B} \ k=n" by (auto simp: anchor_def chain_def sign_def) -lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k \ parts {reqm A r n I B}; -I \ agl |] ==> k=n" +lemma uniq_Nonce_in_reqm [dest]: "\Nonce k \ parts {reqm A r n I B}; +I \ agl\ \ k=n" by (auto simp: reqm_def dest: no_Nonce_in_agl) -lemma uniq_Nonce_in_prom [dest]: "[| Nonce k \ parts {prom B ofr A r I L J C}; -I \ agl; J \ agl; Nonce k \ parts {L} |] ==> k=ofr" +lemma uniq_Nonce_in_prom [dest]: "\Nonce k \ parts {prom B ofr A r I L J C}; +I \ agl; J \ agl; Nonce k \ parts {L}\ \ k=ofr" by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) subsection\requests are guarded\ -lemma req_imp_Guard [rule_format]: "[| evs \ p2; A \ bad |] ==> +lemma req_imp_Guard [rule_format]: "\evs \ p2; A \ bad\ \ req A r n I B \ set evs \ Guard n {priK A} (spies evs)" apply (erule p2.induct, simp) apply (simp add: req_def knows.simps, safe) apply (erule in_synth_Guard, erule Guard_analz, simp) by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) -lemma req_imp_Guard_Friend: "[| evs \ p2; A \ bad; req A r n I B \ set evs |] -==> Guard n {priK A} (knows_max (Friend C) evs)" +lemma req_imp_Guard_Friend: "\evs \ p2; A \ bad; req A r n I B \ set evs\ +\ Guard n {priK A} (knows_max (Friend C) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) apply (rule req_imp_Guard, simp+) @@ -392,7 +392,7 @@ subsection\propositions are guarded\ -lemma pro_imp_Guard [rule_format]: "[| evs \ p2; B \ bad; A \ bad |] ==> +lemma pro_imp_Guard [rule_format]: "\evs \ p2; B \ bad; A \ bad\ \ pro B ofr A r I (cons M L) J C \ set evs \ Guard ofr {priK A} (spies evs)" supply [[simproc del: defined_all]] apply (erule p2.induct) (* +3 subgoals *) @@ -430,9 +430,9 @@ apply (simp add: pro_def) by (blast dest: Says_imp_knows_Spy) -lemma pro_imp_Guard_Friend: "[| evs \ p2; B \ bad; A \ bad; -pro B ofr A r I (cons M L) J C \ set evs |] -==> Guard ofr {priK A} (knows_max (Friend D) evs)" +lemma pro_imp_Guard_Friend: "\evs \ p2; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs\ +\ Guard ofr {priK A} (knows_max (Friend D) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) apply (rule pro_imp_Guard, simp+) @@ -443,23 +443,23 @@ subsection\data confidentiality: no one other than the originator can decrypt the offers\ -lemma Nonce_req_notin_spies: "[| evs \ p2; req A r n I B \ set evs; A \ bad |] -==> Nonce n \ analz (spies evs)" +lemma Nonce_req_notin_spies: "\evs \ p2; req A r n I B \ set evs; A \ bad\ +\ Nonce n \ analz (spies evs)" by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_req_notin_knows_max_Friend: "[| evs \ p2; req A r n I B \ set evs; -A \ bad; A \ Friend C |] ==> Nonce n \ analz (knows_max (Friend C) evs)" +lemma Nonce_req_notin_knows_max_Friend: "\evs \ p2; req A r n I B \ set evs; +A \ bad; A \ Friend C\ \ Nonce n \ analz (knows_max (Friend C) evs)" apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) -lemma Nonce_pro_notin_spies: "[| evs \ p2; B \ bad; A \ bad; -pro B ofr A r I (cons M L) J C \ set evs |] ==> Nonce ofr \ analz (spies evs)" +lemma Nonce_pro_notin_spies: "\evs \ p2; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs\ \ Nonce ofr \ analz (spies evs)" by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_pro_notin_knows_max_Friend: "[| evs \ p2; B \ bad; A \ bad; -A \ Friend D; pro B ofr A r I (cons M L) J C \ set evs |] -==> Nonce ofr \ analz (knows_max (Friend D) evs)" +lemma Nonce_pro_notin_knows_max_Friend: "\evs \ p2; B \ bad; A \ bad; +A \ Friend D; pro B ofr A r I (cons M L) J C \ set evs\ +\ Nonce ofr \ analz (knows_max (Friend D) evs)" apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) @@ -467,27 +467,27 @@ subsection\forward privacy: only the originator can know the identity of the shops\ -lemma forward_privacy_Spy: "[| evs \ p2; B \ bad; A \ bad; -pro B ofr A r I (cons M L) J C \ set evs |] -==> sign B (Nonce ofr) \ analz (spies evs)" +lemma forward_privacy_Spy: "\evs \ p2; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs\ +\ sign B (Nonce ofr) \ analz (spies evs)" by (auto simp:sign_def dest: Nonce_pro_notin_spies) -lemma forward_privacy_Friend: "[| evs \ p2; B \ bad; A \ bad; A \ Friend D; -pro B ofr A r I (cons M L) J C \ set evs |] -==> sign B (Nonce ofr) \ analz (knows_max (Friend D) evs)" +lemma forward_privacy_Friend: "\evs \ p2; B \ bad; A \ bad; A \ Friend D; +pro B ofr A r I (cons M L) J C \ set evs\ +\ sign B (Nonce ofr) \ analz (knows_max (Friend D) evs)" by (auto simp:sign_def dest:Nonce_pro_notin_knows_max_Friend ) subsection\non repudiability: an offer signed by B has been sent by B\ -lemma Crypt_reqm: "[| Crypt (priK A) X \ parts {reqm A' r n I B}; I \ agl |] ==> A=A'" +lemma Crypt_reqm: "\Crypt (priK A) X \ parts {reqm A' r n I B}; I \ agl\ \ A=A'" by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) -lemma Crypt_prom: "[| Crypt (priK A) X \ parts {prom B ofr A' r I L J C}; -I \ agl; J \ agl |] ==> A=B | Crypt (priK A) X \ parts {L}" +lemma Crypt_prom: "\Crypt (priK A) X \ parts {prom B ofr A' r I L J C}; +I \ agl; J \ agl\ \ A=B | Crypt (priK A) X \ parts {L}" apply (simp add: prom_def anchor_def chain_def sign_def) by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) -lemma Crypt_safeness: "[| evs \ p2; A \ bad |] ==> Crypt (priK A) X \ parts (spies evs) +lemma Crypt_safeness: "\evs \ p2; A \ bad\ \ Crypt (priK A) X \ parts (spies evs) \ (\B Y. Says A B Y \ set evs & Crypt (priK A) X \ parts {Y})" apply (erule p2.induct) (* Nil *) @@ -517,7 +517,7 @@ apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) by auto -lemma Crypt_Hash_imp_sign: "[| evs \ p2; A \ bad |] ==> +lemma Crypt_Hash_imp_sign: "\evs \ p2; A \ bad\ \ Crypt (priK A) (Hash X) \ parts (spies evs) \ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (erule p2.induct) @@ -554,7 +554,7 @@ apply (blast del: MPair_parts)+ done -lemma sign_safeness: "[| evs \ p2; A \ bad |] ==> sign A X \ parts (spies evs) +lemma sign_safeness: "\evs \ p2; A \ bad\ \ sign A X \ parts (spies evs) \ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (clarify, simp add: sign_def, frule parts.Snd) apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Guard/Proto.thy --- a/src/HOL/Auth/Guard/Proto.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Guard/Proto.thy Thu Oct 13 17:19:56 2022 +0100 @@ -52,8 +52,8 @@ (\k. Nonce k \ parts {X} \ nonce s k = n)" by (induct X, simp_all, blast) -lemma wdef_Nonce: "[| Nonce n \ parts {apm s X}; R \ p; msg' R = X; wdef p; -Nonce n \ parts (apm s `(msg `(fst R))) |] ==> +lemma wdef_Nonce: "\Nonce n \ parts {apm s X}; R \ p; msg' R = X; wdef p; +Nonce n \ parts (apm s `(msg `(fst R)))\ \ (\k. Nonce k \ parts {X} \ nonce s k = n)" apply (erule Nonce_apm, unfold wdef_def) apply (drule_tac x=R in spec, drule_tac x=k in spec, clarsimp) @@ -103,10 +103,10 @@ Nil [intro]: "[] \ tr p" -| Fake [intro]: "[| evsf \ tr p; X \ synth (analz (spies evsf)) |] - ==> Says Spy B X # evsf \ tr p" +| Fake [intro]: "\evsf \ tr p; X \ synth (analz (spies evsf))\ + \ Says Spy B X # evsf \ tr p" -| Proto [intro]: "[| evs \ tr p; R \ p; ok evs R s |] ==> ap' s R # evs \ tr p" +| Proto [intro]: "\evs \ tr p; R \ p; ok evs R s\ \ ap' s R # evs \ tr p" subsection\general properties\ @@ -117,29 +117,29 @@ definition has_only_Says' :: "proto => bool" where "has_only_Says' p \ \R. R \ p \ is_Says (snd R)" -lemma has_only_Says'D: "[| R \ p; has_only_Says' p |] -==> (\A B X. snd R = Says A B X)" +lemma has_only_Says'D: "\R \ p; has_only_Says' p\ +\ (\A B X. snd R = Says A B X)" by (unfold has_only_Says'_def is_Says_def, blast) -lemma has_only_Says_tr [simp]: "has_only_Says' p ==> has_only_Says (tr p)" -apply (unfold has_only_Says_def) +lemma has_only_Says_tr [simp]: "has_only_Says' p \ has_only_Says (tr p)" +unfolding has_only_Says_def apply (rule allI, rule allI, rule impI) apply (erule tr.induct) apply (auto simp: has_only_Says'_def ok_def) by (drule_tac x=a in spec, auto simp: is_Says_def) -lemma has_only_Says'_in_trD: "[| has_only_Says' p; list @ ev # evs1 \ tr p |] -==> (\A B X. ev = Says A B X)" +lemma has_only_Says'_in_trD: "\has_only_Says' p; list @ ev # evs1 \ tr p\ +\ (\A B X. ev = Says A B X)" by (drule has_only_Says_tr, auto) -lemma ok_not_used: "[| Nonce n \ used evs; ok evs R s; -\x. x \ fst R \ is_Says x |] ==> Nonce n \ parts (apm s `(msg `(fst R)))" +lemma ok_not_used: "\Nonce n \ used evs; ok evs R s; +\x. x \ fst R \ is_Says x\ \ Nonce n \ parts (apm s `(msg `(fst R)))" apply (unfold ok_def, clarsimp) apply (drule_tac x=x in spec, drule_tac x=x in spec) by (auto simp: is_Says_def dest: Says_imp_spies not_used_not_spied parts_parts) -lemma ok_is_Says: "[| evs' @ ev # evs \ tr p; ok evs R s; has_only_Says' p; -R \ p; x \ fst R |] ==> is_Says x" +lemma ok_is_Says: "\evs' @ ev # evs \ tr p; ok evs R s; has_only_Says' p; +R \ p; x \ fst R\ \ is_Says x" apply (unfold ok_def is_Says_def, clarify) apply (drule_tac x=x in spec, simp) apply (subgoal_tac "one_step (tr p)") @@ -164,27 +164,27 @@ lemma freshD: "fresh p R s n Ks evs \ (\evs1 evs2. evs = evs2 @ ap' s R # evs1 \ Nonce n \ used evs1 \ R \ p \ ok evs1 R s \ Nonce n \ parts {apm' s R} \ apm' s R \ guard n Ks)" -by (unfold fresh_def, blast) + unfolding fresh_def by blast -lemma freshI [intro]: "[| Nonce n \ used evs1; R \ p; Nonce n \ parts {apm' s R}; -ok evs1 R s; apm' s R \ guard n Ks |] -==> fresh p R s n Ks (list @ ap' s R # evs1)" -by (unfold fresh_def, blast) +lemma freshI [intro]: "\Nonce n \ used evs1; R \ p; Nonce n \ parts {apm' s R}; +ok evs1 R s; apm' s R \ guard n Ks\ +\ fresh p R s n Ks (list @ ap' s R # evs1)" + unfolding fresh_def by blast -lemma freshI': "[| Nonce n \ used evs1; (l,r) \ p; -Nonce n \ parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r) \ guard n Ks |] -==> fresh p (l,r) s n Ks (evs2 @ ap s r # evs1)" +lemma freshI': "\Nonce n \ used evs1; (l,r) \ p; +Nonce n \ parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r) \ guard n Ks\ +\ fresh p (l,r) s n Ks (evs2 @ ap s r # evs1)" by (drule freshI, simp+) -lemma fresh_used: "[| fresh p R' s' n Ks evs; has_only_Says' p |] -==> Nonce n \ used evs" +lemma fresh_used: "\fresh p R' s' n Ks evs; has_only_Says' p\ +\ Nonce n \ used evs" apply (unfold fresh_def, clarify) apply (drule has_only_Says'D) by (auto intro: parts_used_app) -lemma fresh_newn: "[| evs' @ ap' s R # evs \ tr p; wdef p; has_only_Says' p; -Nonce n \ used evs; R \ p; ok evs R s; Nonce n \ parts {apm' s R} |] -==> \k. k \ newn R \ nonce s k = n" +lemma fresh_newn: "\evs' @ ap' s R # evs \ tr p; wdef p; has_only_Says' p; +Nonce n \ used evs; R \ p; ok evs R s; Nonce n \ parts {apm' s R}\ +\ \k. k \ newn R \ nonce s k = n" apply (drule wdef_Nonce, simp+) apply (frule ok_not_used, simp+) apply (clarify, erule ok_is_Says, simp+) @@ -193,15 +193,15 @@ apply (drule ok_not_used, simp+) by (clarify, erule ok_is_Says, simp_all) -lemma fresh_rule: "[| evs' @ ev # evs \ tr p; wdef p; Nonce n \ used evs; -Nonce n \ parts {msg ev} |] ==> \R s. R \ p \ ap' s R = ev" +lemma fresh_rule: "\evs' @ ev # evs \ tr p; wdef p; Nonce n \ used evs; +Nonce n \ parts {msg ev}\ \ \R s. R \ p \ ap' s R = ev" apply (drule trunc, simp, ind_cases "ev # evs \ tr p", simp) by (drule_tac x=X in in_sub, drule parts_sub, simp, simp, blast+) -lemma fresh_ruleD: "[| fresh p R' s' n Ks evs; keys R' s' n evs \ Ks; wdef p; +lemma fresh_ruleD: "\fresh p R' s' n Ks evs; keys R' s' n evs \ Ks; wdef p; has_only_Says' p; evs \ tr p; \R k s. nonce s k = n \ Nonce n \ used evs \ R \ p \ k \ newn R \ Nonce n \ parts {apm' s R} \ apm' s R \ guard n Ks \ -apm' s R \ parts (spies evs) \ keys R s n evs \ Ks \ P |] ==> P" +apm' s R \ parts (spies evs) \ keys R s n evs \ Ks \ P\ \ P" apply (frule fresh_used, simp) apply (unfold fresh_def, clarify) apply (drule_tac x=R' in spec) @@ -219,13 +219,13 @@ definition safe :: "key set \ msg set \ bool" where "safe Ks G \ \K. K \ Ks \ Key K \ analz G" -lemma safeD [dest]: "[| safe Ks G; K \ Ks |] ==> Key K \ analz G" -by (unfold safe_def, blast) +lemma safeD [dest]: "\safe Ks G; K \ Ks\ \ Key K \ analz G" + unfolding safe_def by blast -lemma safe_insert: "safe Ks (insert X G) ==> safe Ks G" -by (unfold safe_def, blast) +lemma safe_insert: "safe Ks (insert X G) \ safe Ks G" + unfolding safe_def by blast -lemma Guard_safe: "[| Guard n Ks G; safe Ks G |] ==> Nonce n \ analz G" +lemma Guard_safe: "\Guard n Ks G; safe Ks G\ \ Nonce n \ analz G" by (blast dest: Guard_invKey) subsection\guardedness preservation\ @@ -235,14 +235,14 @@ Guard n Ks (spies evs) \ safe Ks (spies evs) \ fresh p R' s' n Ks evs \ keys R' s' n evs \ Ks \ R \ p \ ok evs R s \ apm' s R \ guard n Ks)" -lemma preservD: "[| preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); +lemma preservD: "\preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); safe Ks (spies evs); fresh p R' s' n Ks evs; R \ p; ok evs R s; -keys R' s' n evs \ Ks |] ==> apm' s R \ guard n Ks" -by (unfold preserv_def, blast) +keys R' s' n evs \ Ks\ \ apm' s R \ guard n Ks" + unfolding preserv_def by blast -lemma preservD': "[| preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); +lemma preservD': "\preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); safe Ks (spies evs); fresh p R' s' n Ks evs; (l,Says A B X) \ p; -ok evs (l,Says A B X) s; keys R' s' n evs \ Ks |] ==> apm s X \ guard n Ks" +ok evs (l,Says A B X) s; keys R' s' n evs \ Ks\ \ apm s X \ guard n Ks" by (drule preservD, simp+) subsection\monotonic keyfun\ @@ -251,14 +251,14 @@ "monoton p keys \ \R' s' n ev evs. ev # evs \ tr p \ keys R' s' n evs \ keys R' s' n (ev # evs)" -lemma monotonD [dest]: "[| keys R' s' n (ev # evs) \ Ks; monoton p keys; -ev # evs \ tr p |] ==> keys R' s' n evs \ Ks" -by (unfold monoton_def, blast) +lemma monotonD [dest]: "\keys R' s' n (ev # evs) \ Ks; monoton p keys; +ev # evs \ tr p\ \ keys R' s' n evs \ Ks" + unfolding monoton_def by blast subsection\guardedness theorem\ -lemma Guard_tr [rule_format]: "[| evs \ tr p; has_only_Says' p; -preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy) |] ==> +lemma Guard_tr [rule_format]: "\evs \ tr p; has_only_Says' p; +preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy)\ \ safe Ks (spies evs) \ fresh p R' s' n Ks evs \ keys R' s' n evs \ Ks \ Guard n Ks (spies evs)" apply (erule tr.induct) @@ -297,18 +297,18 @@ subsection\useful properties for guardedness\ -lemma newn_neq_used: "[| Nonce n \ used evs; ok evs R s; k \ newn R |] -==> n \ nonce s k" +lemma newn_neq_used: "\Nonce n \ used evs; ok evs R s; k \ newn R\ +\ n \ nonce s k" by (auto simp: ok_def) -lemma ok_Guard: "[| ok evs R s; Guard n Ks (spies evs); x \ fst R; is_Says x |] -==> apm s (msg x) \ parts (spies evs) \ apm s (msg x) \ guard n Ks" +lemma ok_Guard: "\ok evs R s; Guard n Ks (spies evs); x \ fst R; is_Says x\ +\ apm s (msg x) \ parts (spies evs) \ apm s (msg x) \ guard n Ks" apply (unfold ok_def is_Says_def, clarify) apply (drule_tac x="Says A B X" in spec, simp) by (drule Says_imp_spies, auto intro: parts_parts) -lemma ok_parts_not_new: "[| Y \ parts (spies evs); Nonce (nonce s n) \ parts {Y}; -ok evs R s |] ==> n \ newn R" +lemma ok_parts_not_new: "\Y \ parts (spies evs); Nonce (nonce s n) \ parts {Y}; +ok evs R s\ \ n \ newn R" by (auto simp: ok_def dest: not_used_not_spied parts_parts) subsection\unicity\ @@ -322,19 +322,19 @@ secret R n s Ks \ parts (spies evs) \ secret R' n' s' Ks \ parts (spies evs) \ secret R n s Ks = secret R' n' s' Ks" -lemma uniqD: "[| uniq p secret; evs \ tr p; R \ p; R' \ p; n \ newn R; n' \ newn R'; +lemma uniqD: "\uniq p secret; evs \ tr p; R \ p; R' \ p; n \ newn R; n' \ newn R'; nonce s n = nonce s' n'; Nonce (nonce s n) \ analz (spies evs); Nonce (nonce s n) \ parts {apm' s R}; Nonce (nonce s n) \ parts {apm' s' R'}; secret R n s Ks \ parts (spies evs); secret R' n' s' Ks \ parts (spies evs); -apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks |] ==> +apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks\ \ secret R n s Ks = secret R' n' s' Ks" -by (unfold uniq_def, blast) + unfolding uniq_def by blast definition ord :: "proto \ (rule \ rule \ bool) \ bool" where "ord p inff \ \R R'. R \ p \ R' \ p \ \ inff R R' \ inff R' R" -lemma ordD: "[| ord p inff; \ inff R R'; R \ p; R' \ p |] ==> inff R' R" -by (unfold ord_def, blast) +lemma ordD: "\ord p inff; \ inff R R'; R \ p; R' \ p\ \ inff R' R" + unfolding ord_def by blast definition uniq' :: "proto \ (rule \ rule \ bool) \ secfun \ bool" where "uniq' p inff secret \ \evs R R' n n' Ks s s'. R \ p \ R' \ p \ @@ -345,16 +345,16 @@ secret R n s Ks \ parts (spies evs) \ secret R' n' s' Ks \ parts (spies evs) \ secret R n s Ks = secret R' n' s' Ks" -lemma uniq'D: "[| uniq' p inff secret; evs \ tr p; inff R R'; R \ p; R' \ p; n \ newn R; +lemma uniq'D: "\uniq' p inff secret; evs \ tr p; inff R R'; R \ p; R' \ p; n \ newn R; n' \ newn R'; nonce s n = nonce s' n'; Nonce (nonce s n) \ analz (spies evs); Nonce (nonce s n) \ parts {apm' s R}; Nonce (nonce s n) \ parts {apm' s' R'}; secret R n s Ks \ parts (spies evs); secret R' n' s' Ks \ parts (spies evs); -apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks |] ==> +apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks\ \ secret R n s Ks = secret R' n' s' Ks" by (unfold uniq'_def, blast) -lemma uniq'_imp_uniq: "[| uniq' p inff secret; ord p inff |] ==> uniq p secret" -apply (unfold uniq_def) +lemma uniq'_imp_uniq: "\uniq' p inff secret; ord p inff\ \ uniq p secret" +unfolding uniq_def apply (rule allI)+ apply (case_tac "inff R R'") apply (blast dest: uniq'D) @@ -441,8 +441,8 @@ subsection\guardedness for NSL\ -lemma "uniq ns secret ==> preserv ns keys n Ks" -apply (unfold preserv_def) +lemma "uniq ns secret \ preserv ns keys n Ks" +unfolding preserv_def apply (rule allI)+ apply (rule impI, rule impI, rule impI, rule impI, rule impI) apply (erule fresh_ruleD, simp, simp, simp, simp) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/KerberosIV.thy --- a/src/HOL/Auth/KerberosIV.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/KerberosIV.thy Thu Oct 13 17:19:56 2022 +0100 @@ -295,7 +295,7 @@ subsection\Lemmas about \<^term>\authKeys\\ lemma authKeys_empty: "authKeys [] = {}" -apply (unfold authKeys_def) +unfolding authKeys_def apply (simp (no_asm)) done @@ -304,27 +304,27 @@ ev \ Says Kas A (Crypt (shrK A) \akey, Agent Peer, Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, akey, Ta\)\)) \ authKeys (ev # evs) = authKeys evs" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeys_insert: "authKeys (Says Kas A (Crypt (shrK A) \Key K, Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K, Number Ta\)\) # evs) = insert K (authKeys evs)" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeys_simp: "K \ authKeys (Says Kas A (Crypt (shrK A) \Key K', Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K', Number Ta\)\) # evs) \ K = K' | K \ authKeys evs" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeysI: "Says Kas A (Crypt (shrK A) \Key K, Agent Tgs, Number Ta, (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key K, Number Ta\)\) \ set evs \ K \ authKeys evs" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeys_used: "K \ authKeys evs \ Key K \ used evs" by (simp add: authKeys_def, blast) @@ -478,7 +478,7 @@ Ta = CT (before Says Kas A (Crypt K \Key authK, Agent Peer, Number Ta, authTicket\) on evs)" -apply (unfold before_def) +unfolding before_def apply (erule rev_mp) apply (erule kerbIV.induct) apply (simp_all (no_asm) add: authKeys_def authKeys_insert, blast, blast) @@ -550,7 +550,7 @@ Ts = CT(before Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, servTicket\) on evs) " -apply (unfold before_def) +unfolding before_def apply (erule rev_mp) apply (erule kerbIV.induct) apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast) @@ -1014,7 +1014,7 @@ lemma AKcryptSKI: "\ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, X \) \ set evs; evs \ kerbIV \ \ AKcryptSK authK servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (blast dest: Says_Tgs_message_form) done @@ -1033,7 +1033,7 @@ lemma Auth_fresh_not_AKcryptSK: "\ Key authK \ used evs; evs \ kerbIV \ \ \ AKcryptSK authK servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (erule rev_mp) apply (erule kerbIV.induct) apply (frule_tac [7] K5_msg_in_parts_spies) @@ -1044,7 +1044,7 @@ (with respect to a given trace). *) lemma Serv_fresh_not_AKcryptSK: "Key servK \ used evs \ \ AKcryptSK authK servK evs" -by (unfold AKcryptSK_def, blast) + unfolding AKcryptSK_def by blast lemma authK_not_AKcryptSK: "\ Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, tk\ @@ -1080,7 +1080,7 @@ text\Long term keys are not issued as servKeys\ lemma shrK_not_AKcryptSK: "evs \ kerbIV \ \ AKcryptSK K (shrK A) evs" -apply (unfold AKcryptSK_def) +unfolding 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) @@ -1093,7 +1093,7 @@ \ set evs; authK' \ authK; evs \ kerbIV \ \ \ AKcryptSK authK' servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (blast dest: unique_servKeys) done diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/KerberosIV_Gets.thy --- a/src/HOL/Auth/KerberosIV_Gets.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/KerberosIV_Gets.thy Thu Oct 13 17:19:56 2022 +0100 @@ -270,27 +270,27 @@ ev \ Says Kas A (Crypt (shrK A) \akey, Agent Peer, Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, akey, Ta\)\)) \ authKeys (ev # evs) = authKeys evs" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeys_insert: "authKeys (Says Kas A (Crypt (shrK A) \Key K, Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K, Number Ta\)\) # evs) = insert K (authKeys evs)" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeys_simp: "K \ authKeys (Says Kas A (Crypt (shrK A) \Key K', Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K', Number Ta\)\) # evs) \ K = K' | K \ authKeys evs" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeysI: "Says Kas A (Crypt (shrK A) \Key K, Agent Tgs, Number Ta, (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key K, Number Ta\)\) \ set evs \ K \ authKeys evs" -by (unfold authKeys_def, auto) + unfolding authKeys_def by auto lemma authKeys_used: "K \ authKeys evs \ Key K \ used evs" by (simp add: authKeys_def, blast) @@ -876,7 +876,7 @@ lemma AKcryptSKI: "\ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, X \) \ set evs; evs \ kerbIV_gets \ \ AKcryptSK authK servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (blast dest: Says_Tgs_message_form) done @@ -894,7 +894,7 @@ lemma Auth_fresh_not_AKcryptSK: "\ Key authK \ used evs; evs \ kerbIV_gets \ \ \ AKcryptSK authK servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (erule rev_mp) apply (erule kerbIV_gets.induct) apply (frule_tac [8] Gets_ticket_parts) @@ -905,7 +905,7 @@ (with respect to a given trace). *) lemma Serv_fresh_not_AKcryptSK: "Key servK \ used evs \ \ AKcryptSK authK servK evs" -by (unfold AKcryptSK_def, blast) + unfolding AKcryptSK_def by blast lemma authK_not_AKcryptSK: "\ Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, tk\ @@ -948,7 +948,7 @@ text\Long term keys are not issued as servKeys\ lemma shrK_not_AKcryptSK: "evs \ kerbIV_gets \ \ AKcryptSK K (shrK A) evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (erule kerbIV_gets.induct) apply (frule_tac [8] Gets_ticket_parts) by (frule_tac [6] Gets_ticket_parts, auto) @@ -960,7 +960,7 @@ \ set evs; authK' \ authK; evs \ kerbIV_gets \ \ \ AKcryptSK authK' servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def by (blast dest: unique_servKeys) text\Equivalently\ diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/KerberosV.thy --- a/src/HOL/Auth/KerberosV.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/KerberosV.thy Thu Oct 13 17:19:56 2022 +0100 @@ -778,7 +778,7 @@ lemma Auth_fresh_not_AKcryptSK: "\ Key authK \ used evs; evs \ kerbV \ \ \ AKcryptSK authK servK evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (erule rev_mp) apply (erule kerbV.induct) apply (frule_tac [7] Says_ticket_parts) @@ -821,7 +821,7 @@ text\Long term keys are not issued as servKeys\ lemma shrK_not_AKcryptSK: "evs \ kerbV \ \ AKcryptSK K (shrK A) evs" -apply (unfold AKcryptSK_def) +unfolding AKcryptSK_def apply (erule kerbV.induct) apply (frule_tac [7] Says_ticket_parts) apply (frule_tac [5] Says_ticket_parts, auto) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Kerberos_BAN.thy --- a/src/HOL/Auth/Kerberos_BAN.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Kerberos_BAN.thy Thu Oct 13 17:19:56 2022 +0100 @@ -295,7 +295,7 @@ Tk = CT(before Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) on evs)" -apply (unfold before_def) +unfolding before_def apply (erule rev_mp) apply (erule bankerberos.induct, simp_all add: takeWhile_tail) apply auto diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Kerberos_BAN_Gets.thy --- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Thu Oct 13 17:19:56 2022 +0100 @@ -283,7 +283,7 @@ Tk = CT(before Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) on evs)" -apply (unfold before_def) +unfolding before_def apply (erule rev_mp) apply (erule bankerb_gets.induct, simp_all) txt\We need this simplification only for Message 2\ diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Message.thy Thu Oct 13 17:19:56 2022 +0100 @@ -14,7 +14,7 @@ (*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*) lemma [simp] : "A \ (B \ A) = B \ A" -by blast + by blast type_synonym key = nat @@ -71,87 +71,87 @@ parts :: "msg set \ msg set" for H :: "msg set" where - Inj [intro]: "X \ H ==> X \ parts H" - | Fst: "\X,Y\ \ parts H ==> X \ parts H" - | Snd: "\X,Y\ \ parts H ==> Y \ parts H" - | Body: "Crypt K X \ parts H ==> X \ parts H" + Inj [intro]: "X \ H \ X \ parts H" + | Fst: "\X,Y\ \ parts H \ X \ parts H" + | Snd: "\X,Y\ \ parts H \ Y \ parts H" + | Body: "Crypt K X \ parts H \ X \ parts H" text\Monotonicity\ -lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" -apply auto -apply (erule parts.induct) -apply (blast dest: parts.Fst parts.Snd parts.Body)+ -done +lemma parts_mono_aux: "\G \ H; X \ parts G\ \ X \ parts H" + by (erule parts.induct) (auto dest: parts.Fst parts.Snd parts.Body) + +lemma parts_mono: "G \ H \ parts(G) \ parts(H)" + using parts_mono_aux by blast text\Equations hold because constructors are injective.\ lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x\A)" -by auto + by auto lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x\A)" -by auto + by auto lemma Nonce_Key_image_eq [simp]: "(Nonce x \ Key`A)" -by auto + by auto subsubsection\Inverse of keys\ lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')" -by (metis invKey) + by (metis invKey) subsection\keysFor operator\ lemma keysFor_empty [simp]: "keysFor {} = {}" -by (unfold keysFor_def, blast) + unfolding keysFor_def by blast lemma keysFor_Un [simp]: "keysFor (H \ H') = keysFor H \ keysFor H'" -by (unfold keysFor_def, blast) + unfolding keysFor_def by blast lemma keysFor_UN [simp]: "keysFor (\i\A. H i) = (\i\A. keysFor (H i))" -by (unfold keysFor_def, blast) + unfolding keysFor_def by blast text\Monotonicity\ -lemma keysFor_mono: "G \ H ==> keysFor(G) \ keysFor(H)" -by (unfold keysFor_def, blast) +lemma keysFor_mono: "G \ H \ keysFor(G) \ keysFor(H)" + unfolding keysFor_def by blast lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_insert_MPair [simp]: "keysFor (insert \X,Y\ H) = keysFor H" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_insert_Crypt [simp]: "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}" -by (unfold keysFor_def, auto) + unfolding keysFor_def by auto -lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H ==> invKey K \ keysFor H" -by (unfold keysFor_def, blast) +lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H \ invKey K \ keysFor H" + unfolding keysFor_def by blast subsection\Inductive relation "parts"\ lemma MPair_parts: - "[| \X,Y\ \ parts H; - [| X \ parts H; Y \ parts H |] ==> P |] ==> P" -by (blast dest: parts.Fst parts.Snd) + "\\X,Y\ \ parts H; + \X \ parts H; Y \ parts H\ \ P\ \ P" + by (blast dest: parts.Fst parts.Snd) declare MPair_parts [elim!] parts.Body [dest!] text\NB These two rules are UNSAFE in the formal sense, as they discard the @@ -160,52 +160,53 @@ The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\ lemma parts_increasing: "H \ parts(H)" -by blast + by blast lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD] +lemma parts_empty_aux: "X \ parts{} \ False" + by (induction rule: parts.induct) (blast+) + lemma parts_empty [simp]: "parts{} = {}" -apply safe -apply (erule parts.induct, blast+) -done + using parts_empty_aux by blast -lemma parts_emptyE [elim!]: "X\ parts{} ==> P" -by simp +lemma parts_emptyE [elim!]: "X\ parts{} \ P" + by simp text\WARNING: loops if H = {Y}, therefore must not be repeated!\ -lemma parts_singleton: "X\ parts H ==> \Y\H. X\ parts {Y}" -by (erule parts.induct, fast+) +lemma parts_singleton: "X \ parts H \ \Y\H. X \ parts {Y}" + by (erule parts.induct, fast+) subsubsection\Unions\ lemma parts_Un_subset1: "parts(G) \ parts(H) \ parts(G \ H)" -by (intro Un_least parts_mono Un_upper1 Un_upper2) + by (intro Un_least parts_mono Un_upper1 Un_upper2) lemma parts_Un_subset2: "parts(G \ H) \ parts(G) \ parts(H)" -apply (rule subsetI) -apply (erule parts.induct, blast+) -done + apply (rule subsetI) + apply (erule parts.induct, blast+) + done lemma parts_Un [simp]: "parts(G \ H) = parts(G) \ parts(H)" -by (intro equalityI parts_Un_subset1 parts_Un_subset2) + by (intro equalityI parts_Un_subset1 parts_Un_subset2) lemma parts_insert: "parts (insert X H) = parts {X} \ parts H" -by (metis insert_is_Un parts_Un) + by (metis insert_is_Un parts_Un) text\TWO inserts to avoid looping. This rewrite is better than nothing. But its behaviour can be strange.\ lemma parts_insert2: - "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" -by (metis Un_commute Un_empty_right Un_insert_right insert_is_Un parts_Un) + "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" + by (metis Un_commute Un_empty_right Un_insert_right insert_is_Un parts_Un) lemma parts_UN_subset1: "(\x\A. parts(H x)) \ parts(\x\A. H x)" -by (intro UN_least parts_mono UN_upper) + by (intro UN_least parts_mono UN_upper) lemma parts_UN_subset2: "parts(\x\A. H x) \ (\x\A. parts(H x))" -apply (rule subsetI) -apply (erule parts.induct, blast+) -done + apply (rule subsetI) + apply (erule parts.induct, blast+) + done lemma parts_UN [simp]: "parts (\x\A. H x) = (\x\A. parts (H x))" @@ -214,7 +215,7 @@ lemma parts_image [simp]: "parts (f ` A) = (\x\A. parts {f x})" apply auto - apply (metis (mono_tags, opaque_lifting) image_iff parts_singleton) + apply (metis (mono_tags, opaque_lifting) image_iff parts_singleton) apply (metis empty_subsetI image_eqI insert_absorb insert_subset parts_mono) done @@ -229,29 +230,29 @@ lemma parts_insert_subset: "insert X (parts H) \ parts(insert X H)" -by (blast intro: parts_mono [THEN [2] rev_subsetD]) + by (blast intro: parts_mono [THEN [2] rev_subsetD]) subsubsection\Idempotence and transitivity\ -lemma parts_partsD [dest!]: "X\ parts (parts H) ==> X\ parts H" -by (erule parts.induct, blast+) +lemma parts_partsD [dest!]: "X\ parts (parts H) \ X\ parts H" + by (erule parts.induct, blast+) lemma parts_idem [simp]: "parts (parts H) = parts H" -by blast + by blast lemma parts_subset_iff [simp]: "(parts G \ parts H) = (G \ parts H)" -by (metis parts_idem parts_increasing parts_mono subset_trans) + by (metis parts_idem parts_increasing parts_mono subset_trans) -lemma parts_trans: "[| X\ parts G; G \ parts H |] ==> X\ parts H" -by (metis parts_subset_iff subsetD) +lemma parts_trans: "\X\ parts G; G \ parts H\ \ X\ parts H" + by (metis parts_subset_iff subsetD) text\Cut\ lemma parts_cut: - "[| Y\ parts (insert X G); X\ parts H |] ==> Y\ parts (G \ H)" -by (blast intro: parts_trans) + "\Y\ parts (insert X G); X\ parts H\ \ Y\ parts (G \ H)" + by (blast intro: parts_trans) -lemma parts_cut_eq [simp]: "X\ parts H ==> parts (insert X H) = parts H" -by (metis insert_absorb parts_idem parts_insert) +lemma parts_cut_eq [simp]: "X\ parts H \ parts (insert X H) = parts H" + by (metis insert_absorb parts_idem parts_insert) subsubsection\Rewrite rules for pulling out atomic messages\ @@ -260,65 +261,65 @@ lemma parts_insert_Agent [simp]: - "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done + "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)" + apply (rule parts_insert_eq_I) + apply (erule parts.induct, auto) + done lemma parts_insert_Nonce [simp]: - "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done + "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)" + apply (rule parts_insert_eq_I) + apply (erule parts.induct, auto) + done lemma parts_insert_Number [simp]: - "parts (insert (Number N) H) = insert (Number N) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done + "parts (insert (Number N) H) = insert (Number N) (parts H)" + apply (rule parts_insert_eq_I) + apply (erule parts.induct, auto) + done lemma parts_insert_Key [simp]: - "parts (insert (Key K) H) = insert (Key K) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done + "parts (insert (Key K) H) = insert (Key K) (parts H)" + apply (rule parts_insert_eq_I) + apply (erule parts.induct, auto) + done lemma parts_insert_Hash [simp]: - "parts (insert (Hash X) H) = insert (Hash X) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done + "parts (insert (Hash X) H) = insert (Hash X) (parts H)" + apply (rule parts_insert_eq_I) + apply (erule parts.induct, auto) + done lemma parts_insert_Crypt [simp]: - "parts (insert (Crypt K X) H) = insert (Crypt K X) (parts (insert X H))" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct, auto) -apply (blast intro: parts.Body) -done + "parts (insert (Crypt K X) H) = insert (Crypt K X) (parts (insert X H))" + apply (rule equalityI) + apply (rule subsetI) + apply (erule parts.induct, auto) + apply (blast intro: parts.Body) + done lemma parts_insert_MPair [simp]: - "parts (insert \X,Y\ H) = + "parts (insert \X,Y\ H) = insert \X,Y\ (parts (insert X (insert Y H)))" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct, auto) -apply (blast intro: parts.Fst parts.Snd)+ -done + apply (rule equalityI) + apply (rule subsetI) + apply (erule parts.induct, auto) + apply (blast intro: parts.Fst parts.Snd)+ + done lemma parts_image_Key [simp]: "parts (Key`N) = Key`N" -by auto + by auto text\In any message, there is an upper bound N on its greatest nonce.\ lemma msg_Nonce_supply: "\N. \n. N\n \ Nonce n \ parts {msg}" proof (induct msg) case (Nonce n) - show ?case - by simp (metis Suc_n_not_le_n) + show ?case + by simp (metis Suc_n_not_le_n) next case (MPair X Y) - then show ?case \ \metis works out the necessary sum itself!\ - by (simp add: parts_insert2) (metis le_trans nat_le_linear) + then show ?case \ \metis works out the necessary sum itself!\ + by (simp add: parts_insert2) (metis le_trans nat_le_linear) qed auto subsection\Inductive relation "analz"\ @@ -331,230 +332,224 @@ analz :: "msg set \ msg set" for H :: "msg set" where - Inj [intro,simp]: "X \ H ==> X \ analz H" - | Fst: "\X,Y\ \ analz H ==> X \ analz H" - | Snd: "\X,Y\ \ analz H ==> Y \ analz H" + Inj [intro,simp]: "X \ H \ X \ analz H" + | Fst: "\X,Y\ \ analz H \ X \ analz H" + | Snd: "\X,Y\ \ analz H \ Y \ analz H" | Decrypt [dest]: - "\Crypt K X \ analz H; Key(invKey K) \ analz H\ \ X \ analz H" + "\Crypt K X \ analz H; Key(invKey K) \ analz H\ \ X \ analz H" text\Monotonicity; Lemma 1 of Lowe's paper\ -lemma analz_mono: "G\H ==> analz(G) \ analz(H)" -apply auto -apply (erule analz.induct) -apply (auto dest: analz.Fst analz.Snd) -done +lemma analz_mono_aux: "\G \ H; X \ analz G\ \ X \ analz H" + by (erule analz.induct) (auto dest: analz.Fst analz.Snd) + +lemma analz_mono: "G\H \ analz(G) \ analz(H)" + using analz_mono_aux by blast text\Making it safe speeds up proofs\ lemma MPair_analz [elim!]: - "[| \X,Y\ \ analz H; - [| X \ analz H; Y \ analz H |] ==> P - |] ==> P" -by (blast dest: analz.Fst analz.Snd) + "\\X,Y\ \ analz H; + \X \ analz H; Y \ analz H\ \ P\ \ P" + by (blast dest: analz.Fst analz.Snd) lemma analz_increasing: "H \ analz(H)" -by blast + by blast + +lemma analz_into_parts: "X \ analz H \ X \ parts H" + by (erule analz.induct) auto lemma analz_subset_parts: "analz H \ parts H" -apply (rule subsetI) -apply (erule analz.induct, blast+) -done + using analz_into_parts by blast -lemmas analz_into_parts = analz_subset_parts [THEN subsetD] +lemma analz_parts [simp]: "analz (parts H) = parts H" + using analz_subset_parts by blast lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD] lemma parts_analz [simp]: "parts (analz H) = parts H" -by (metis analz_increasing analz_subset_parts equalityI parts_mono parts_subset_iff) - -lemma analz_parts [simp]: "analz (parts H) = parts H" -apply auto -apply (erule analz.induct, auto) -done + by (metis analz_increasing analz_subset_parts parts_idem parts_mono subset_antisym) lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD] subsubsection\General equational properties\ lemma analz_empty [simp]: "analz{} = {}" -apply safe -apply (erule analz.induct, blast+) -done + using analz_parts by fastforce text\Converse fails: we can analz more from the union than from the separate parts, as a key in one might decrypt a message in the other\ lemma analz_Un: "analz(G) \ analz(H) \ analz(G \ H)" -by (intro Un_least analz_mono Un_upper1 Un_upper2) + by (intro Un_least analz_mono Un_upper1 Un_upper2) lemma analz_insert: "insert X (analz H) \ analz(insert X H)" -by (blast intro: analz_mono [THEN [2] rev_subsetD]) + by (blast intro: analz_mono [THEN [2] rev_subsetD]) subsubsection\Rewrite rules for pulling out atomic messages\ lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert] lemma analz_insert_Agent [simp]: - "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done + "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)" + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) + done lemma analz_insert_Nonce [simp]: - "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done + "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)" + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) + done lemma analz_insert_Number [simp]: - "analz (insert (Number N) H) = insert (Number N) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done + "analz (insert (Number N) H) = insert (Number N) (analz H)" + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) + done lemma analz_insert_Hash [simp]: - "analz (insert (Hash X) H) = insert (Hash X) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done + "analz (insert (Hash X) H) = insert (Hash X) (analz H)" + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) + done text\Can only pull out Keys if they are not needed to decrypt the rest\ lemma analz_insert_Key [simp]: - "K \ keysFor (analz H) ==> + "K \ keysFor (analz H) \ analz (insert (Key K) H) = insert (Key K) (analz H)" -apply (unfold keysFor_def) -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done + unfolding keysFor_def + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) + done lemma analz_insert_MPair [simp]: - "analz (insert \X,Y\ H) = + "analz (insert \X,Y\ H) = insert \X,Y\ (analz (insert X (insert Y H)))" -apply (rule equalityI) -apply (rule subsetI) -apply (erule analz.induct, auto) -apply (erule analz.induct) -apply (blast intro: analz.Fst analz.Snd)+ -done + apply (rule equalityI) + apply (rule subsetI) + apply (erule analz.induct, auto) + apply (erule analz.induct) + apply (blast intro: analz.Fst analz.Snd)+ + done text\Can pull out enCrypted message if the Key is not known\ lemma analz_insert_Crypt: - "Key (invKey K) \ analz H - ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) + "Key (invKey K) \ analz H + \ analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) -done + done -lemma lemma1: "Key (invKey K) \ analz H ==> +lemma lemma1: "Key (invKey K) \ analz H \ analz (insert (Crypt K X) H) \ insert (Crypt K X) (analz (insert X H))" -apply (rule subsetI) -apply (erule_tac x = x in analz.induct, auto) -done + apply (rule subsetI) + apply (erule_tac x = x in analz.induct, auto) + done -lemma lemma2: "Key (invKey K) \ analz H ==> +lemma lemma2: "Key (invKey K) \ analz H \ insert (Crypt K X) (analz (insert X H)) \ analz (insert (Crypt K X) H)" -apply auto -apply (erule_tac x = x in analz.induct, auto) -apply (blast intro: analz_insertI analz.Decrypt) -done + apply auto + apply (erule_tac x = x in analz.induct, auto) + apply (blast intro: analz_insertI analz.Decrypt) + done lemma analz_insert_Decrypt: - "Key (invKey K) \ analz H ==> + "Key (invKey K) \ analz H \ analz (insert (Crypt K X) H) = insert (Crypt K X) (analz (insert X H))" -by (intro equalityI lemma1 lemma2) + by (intro equalityI lemma1 lemma2) text\Case analysis: either the message is secure, or it is not! Effective, but can cause subgoals to blow up! Use with \if_split\; apparently \split_tac\ does not cope with patterns such as \<^term>\analz (insert (Crypt K X) H)\\ lemma analz_Crypt_if [simp]: - "analz (insert (Crypt K X) H) = + "analz (insert (Crypt K X) H) = (if (Key (invKey K) \ analz H) then insert (Crypt K X) (analz (insert X H)) else insert (Crypt K X) (analz H))" -by (simp add: analz_insert_Crypt analz_insert_Decrypt) + by (simp add: analz_insert_Crypt analz_insert_Decrypt) text\This rule supposes "for the sake of argument" that we have the key.\ lemma analz_insert_Crypt_subset: - "analz (insert (Crypt K X) H) \ + "analz (insert (Crypt K X) H) \ insert (Crypt K X) (analz (insert X H))" -apply (rule subsetI) -apply (erule analz.induct, auto) -done + apply (rule subsetI) + apply (erule analz.induct, auto) + done lemma analz_image_Key [simp]: "analz (Key`N) = Key`N" -apply auto -apply (erule analz.induct, auto) -done + apply auto + apply (erule analz.induct, auto) + done subsubsection\Idempotence and transitivity\ -lemma analz_analzD [dest!]: "X\ analz (analz H) ==> X\ analz H" -by (erule analz.induct, blast+) +lemma analz_analzD [dest!]: "X\ analz (analz H) \ X\ analz H" + by (erule analz.induct, blast+) lemma analz_idem [simp]: "analz (analz H) = analz H" -by blast + by blast lemma analz_subset_iff [simp]: "(analz G \ analz H) = (G \ analz H)" -by (metis analz_idem analz_increasing analz_mono subset_trans) + by (metis analz_idem analz_increasing analz_mono subset_trans) -lemma analz_trans: "[| X\ analz G; G \ analz H |] ==> X\ analz H" -by (drule analz_mono, blast) +lemma analz_trans: "\X\ analz G; G \ analz H\ \ X\ analz H" + by (drule analz_mono, blast) text\Cut; Lemma 2 of Lowe\ -lemma analz_cut: "[| Y\ analz (insert X H); X\ analz H |] ==> Y\ analz H" -by (erule analz_trans, blast) +lemma analz_cut: "\Y\ analz (insert X H); X\ analz H\ \ Y\ analz H" + by (erule analz_trans, blast) (*Cut can be proved easily by induction on - "Y: analz (insert X H) ==> X: analz H \ Y: analz H" + "Y: analz (insert X H) \ X: analz H \ Y: analz H" *) text\This rewrite rule helps in the simplification of messages that involve the forwarding of unknown components (X). Without it, removing occurrences of X can be very complicated.\ -lemma analz_insert_eq: "X\ analz H ==> analz (insert X H) = analz H" -by (metis analz_cut analz_insert_eq_I insert_absorb) +lemma analz_insert_eq: "X\ analz H \ analz (insert X H) = analz H" + by (metis analz_cut analz_insert_eq_I insert_absorb) text\A congruence rule for "analz"\ lemma analz_subset_cong: - "[| analz G \ analz G'; analz H \ analz H' |] - ==> analz (G \ H) \ analz (G' \ H')" -by (metis Un_mono analz_Un analz_subset_iff subset_trans) + "\analz G \ analz G'; analz H \ analz H'\ + \ analz (G \ H) \ analz (G' \ H')" + by (metis Un_mono analz_Un analz_subset_iff subset_trans) lemma analz_cong: - "[| analz G = analz G'; analz H = analz H' |] - ==> analz (G \ H) = analz (G' \ H')" -by (intro equalityI analz_subset_cong, simp_all) + "\analz G = analz G'; analz H = analz H'\ + \ analz (G \ H) = analz (G' \ H')" + by (intro equalityI analz_subset_cong, simp_all) lemma analz_insert_cong: - "analz H = analz H' ==> analz(insert X H) = analz(insert X H')" -by (force simp only: insert_def intro!: analz_cong) + "analz H = analz H' \ analz(insert X H) = analz(insert X H')" + by (force simp only: insert_def intro!: analz_cong) text\If there are no pairs or encryptions then analz does nothing\ lemma analz_trivial: - "[| \X Y. \X,Y\ \ H; \X K. Crypt K X \ H |] ==> analz H = H" -apply safe -apply (erule analz.induct, blast+) -done + "\\X Y. \X,Y\ \ H; \X K. Crypt K X \ H\ \ analz H = H" + apply safe + apply (erule analz.induct, blast+) + done text\These two are obsolete (with a single Spy) but cost little to prove...\ lemma analz_UN_analz_lemma: - "X\ analz (\i\A. analz (H i)) ==> X\ analz (\i\A. H i)" -apply (erule analz.induct) -apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ -done + "X\ analz (\i\A. analz (H i)) \ X\ analz (\i\A. H i)" + apply (erule analz.induct) + apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ + done lemma analz_UN_analz [simp]: "analz (\i\A. analz (H i)) = analz (\i\A. H i)" -by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) + by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) subsection\Inductive relation "synth"\ @@ -568,147 +563,142 @@ synth :: "msg set => msg set" for H :: "msg set" where - Inj [intro]: "X \ H ==> X \ synth H" + Inj [intro]: "X \ H \ X \ synth H" | Agent [intro]: "Agent agt \ synth H" | Number [intro]: "Number n \ synth H" - | Hash [intro]: "X \ synth H ==> Hash X \ synth H" - | MPair [intro]: "[|X \ synth H; Y \ synth H|] ==> \X,Y\ \ synth H" - | Crypt [intro]: "[|X \ synth H; Key(K) \ H|] ==> Crypt K X \ synth H" + | Hash [intro]: "X \ synth H \ Hash X \ synth H" + | MPair [intro]: "\X \ synth H; Y \ synth H\ \ \X,Y\ \ synth H" + | Crypt [intro]: "\X \ synth H; Key(K) \ H\ \ Crypt K X \ synth H" text\Monotonicity\ -lemma synth_mono: "G\H ==> synth(G) \ synth(H)" +lemma synth_mono: "G\H \ synth(G) \ synth(H)" by (auto, erule synth.induct, auto) text\NO \Agent_synth\, as any Agent name can be synthesized. The same holds for \<^term>\Number\\ inductive_simps synth_simps [iff]: - "Nonce n \ synth H" - "Key K \ synth H" - "Hash X \ synth H" - "\X,Y\ \ synth H" - "Crypt K X \ synth H" + "Nonce n \ synth H" + "Key K \ synth H" + "Hash X \ synth H" + "\X,Y\ \ synth H" + "Crypt K X \ synth H" lemma synth_increasing: "H \ synth(H)" -by blast + by blast subsubsection\Unions\ text\Converse fails: we can synth more from the union than from the separate parts, building a compound message using elements of each.\ lemma synth_Un: "synth(G) \ synth(H) \ synth(G \ H)" -by (intro Un_least synth_mono Un_upper1 Un_upper2) + by (intro Un_least synth_mono Un_upper1 Un_upper2) lemma synth_insert: "insert X (synth H) \ synth(insert X H)" -by (blast intro: synth_mono [THEN [2] rev_subsetD]) + by (blast intro: synth_mono [THEN [2] rev_subsetD]) subsubsection\Idempotence and transitivity\ -lemma synth_synthD [dest!]: "X\ synth (synth H) ==> X\ synth H" -by (erule synth.induct, auto) +lemma synth_synthD [dest!]: "X\ synth (synth H) \ X\ synth H" + by (erule synth.induct, auto) lemma synth_idem: "synth (synth H) = synth H" -by blast + by blast lemma synth_subset_iff [simp]: "(synth G \ synth H) = (G \ synth H)" -by (metis subset_trans synth_idem synth_increasing synth_mono) + by (metis subset_trans synth_idem synth_increasing synth_mono) -lemma synth_trans: "[| X\ synth G; G \ synth H |] ==> X\ synth H" -by (drule synth_mono, blast) +lemma synth_trans: "\X\ synth G; G \ synth H\ \ X\ synth H" + by (drule synth_mono, blast) text\Cut; Lemma 2 of Lowe\ -lemma synth_cut: "[| Y\ synth (insert X H); X\ synth H |] ==> Y\ synth H" -by (erule synth_trans, blast) +lemma synth_cut: "\Y\ synth (insert X H); X\ synth H\ \ Y\ synth H" + by (erule synth_trans, blast) lemma Crypt_synth_eq [simp]: - "Key K \ H ==> (Crypt K X \ synth H) = (Crypt K X \ H)" -by blast + "Key K \ H \ (Crypt K X \ synth H) = (Crypt K X \ H)" + by blast lemma keysFor_synth [simp]: - "keysFor (synth H) = keysFor H \ invKey`{K. Key K \ H}" -by (unfold keysFor_def, blast) + "keysFor (synth H) = keysFor H \ invKey`{K. Key K \ H}" + unfolding keysFor_def by blast subsubsection\Combinations of parts, analz and synth\ lemma parts_synth [simp]: "parts (synth H) = parts H \ synth H" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct) -apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD] - parts.Fst parts.Snd parts.Body)+ -done + apply (rule equalityI) + apply (rule subsetI) + apply (erule parts.induct) + apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD] + parts.Fst parts.Snd parts.Body)+ + done lemma analz_analz_Un [simp]: "analz (analz G \ H) = analz (G \ H)" -apply (intro equalityI analz_subset_cong)+ -apply simp_all -done + apply (intro equalityI analz_subset_cong)+ + apply simp_all + done lemma analz_synth_Un [simp]: "analz (synth G \ H) = analz (G \ H) \ synth G" -apply (rule equalityI) -apply (rule subsetI) -apply (erule analz.induct) -prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD]) -apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+ -done + apply (rule equalityI) + apply (rule subsetI) + apply (erule analz.induct) + prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD]) + apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+ + done lemma analz_synth [simp]: "analz (synth H) = analz H \ synth H" -by (metis Un_empty_right analz_synth_Un) + by (metis Un_empty_right analz_synth_Un) subsubsection\For reasoning about the Fake rule in traces\ -lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" -by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono) +lemma parts_insert_subset_Un: "X\ G \ parts(insert X H) \ parts G \ parts H" + by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono) text\More specifically for Fake. See also \Fake_parts_sing\ below\ lemma Fake_parts_insert: - "X \ synth (analz H) ==> + "X \ synth (analz H) \ parts (insert X H) \ synth (analz H) \ parts H" -by (metis Un_commute analz_increasing insert_subset parts_analz parts_mono - parts_synth synth_mono synth_subset_iff) + by (metis Un_commute analz_increasing insert_subset parts_analz parts_mono + parts_synth synth_mono synth_subset_iff) lemma Fake_parts_insert_in_Un: - "\Z \ parts (insert X H); X \ synth (analz H)\ + "\Z \ parts (insert X H); X \ synth (analz H)\ \ Z \ synth (analz H) \ parts H" -by (metis Fake_parts_insert subsetD) + by (metis Fake_parts_insert subsetD) text\\<^term>\H\ is sometimes \<^term>\Key ` KK \ spies evs\, so can't put \<^term>\G=H\.\ lemma Fake_analz_insert: - "X\ synth (analz G) ==> + "X\ synth (analz G) \ analz (insert X H) \ synth (analz G) \ analz (G \ H)" -apply (rule subsetI) -apply (subgoal_tac "x \ analz (synth (analz G) \ H)", force) -apply (blast intro: analz_mono [THEN [2] rev_subsetD] analz_mono [THEN synth_mono, THEN [2] rev_subsetD]) -done + by (metis UnCI Un_commute Un_upper1 analz_analz_Un analz_mono analz_synth_Un insert_subset) lemma analz_conj_parts [simp]: - "(X \ analz H \ X \ parts H) = (X \ analz H)" -by (blast intro: analz_subset_parts [THEN subsetD]) + "(X \ analz H \ X \ parts H) = (X \ analz H)" + by (blast intro: analz_subset_parts [THEN subsetD]) lemma analz_disj_parts [simp]: - "(X \ analz H | X \ parts H) = (X \ parts H)" -by (blast intro: analz_subset_parts [THEN subsetD]) + "(X \ analz H | X \ parts H) = (X \ parts H)" + by (blast intro: analz_subset_parts [THEN subsetD]) text\Without this equation, other rules for synth and analz would yield redundant cases\ lemma MPair_synth_analz [iff]: - "(\X,Y\ \ synth (analz H)) = - (X \ synth (analz H) \ Y \ synth (analz H))" -by blast + "\X,Y\ \ synth (analz H) \ X \ synth (analz H) \ Y \ synth (analz H)" + by blast lemma Crypt_synth_analz: - "[| Key K \ analz H; Key (invKey K) \ analz H |] - ==> (Crypt K X \ synth (analz H)) = (X \ synth (analz H))" -by blast - + "\Key K \ analz H; Key (invKey K) \ analz H\ + \ (Crypt K X \ synth (analz H)) = (X \ synth (analz H))" + by blast lemma Hash_synth_analz [simp]: - "X \ synth (analz H) - ==> (Hash\X,Y\ \ synth (analz H)) = (Hash\X,Y\ \ analz H)" -by blast + "X \ synth (analz H) + \ (Hash\X,Y\ \ synth (analz H)) = (Hash\X,Y\ \ analz H)" + by blast subsection\HPair: a combination of Hash and MPair\ @@ -734,43 +724,43 @@ unfolding HPair_def by simp lemmas HPair_neqs = Agent_neq_HPair Nonce_neq_HPair Number_neq_HPair - Key_neq_HPair Hash_neq_HPair Crypt_neq_HPair + Key_neq_HPair Hash_neq_HPair Crypt_neq_HPair declare HPair_neqs [iff] declare HPair_neqs [symmetric, iff] lemma HPair_eq [iff]: "(Hash[X'] Y' = Hash[X] Y) = (X' = X \ Y'=Y)" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma MPair_eq_HPair [iff]: - "(\X',Y'\ = Hash[X] Y) = (X' = Hash\X,Y\ \ Y'=Y)" -by (simp add: HPair_def) + "(\X',Y'\ = Hash[X] Y) = (X' = Hash\X,Y\ \ Y'=Y)" + by (simp add: HPair_def) lemma HPair_eq_MPair [iff]: - "(Hash[X] Y = \X',Y'\) = (X' = Hash\X,Y\ \ Y'=Y)" -by (auto simp add: HPair_def) + "(Hash[X] Y = \X',Y'\) = (X' = Hash\X,Y\ \ Y'=Y)" + by (auto simp add: HPair_def) subsubsection\Specialized laws, proved in terms of those for Hash and MPair\ lemma keysFor_insert_HPair [simp]: "keysFor (insert (Hash[X] Y) H) = keysFor H" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma parts_insert_HPair [simp]: - "parts (insert (Hash[X] Y) H) = + "parts (insert (Hash[X] Y) H) = insert (Hash[X] Y) (insert (Hash\X,Y\) (parts (insert Y H)))" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma analz_insert_HPair [simp]: - "analz (insert (Hash[X] Y) H) = + "analz (insert (Hash[X] Y) H) = insert (Hash[X] Y) (insert (Hash\X,Y\) (analz (insert Y H)))" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma HPair_synth_analz [simp]: - "X \ synth (analz H) - ==> (Hash[X] Y \ synth (analz H)) = + "X \ synth (analz H) + \ (Hash[X] Y \ synth (analz H)) = (Hash \X, Y\ \ analz H \ Y \ synth (analz H))" -by (auto simp add: HPair_def) + by (auto simp add: HPair_def) text\We do NOT want Crypt... messages broken up in protocols!!\ @@ -815,8 +805,8 @@ | Number: "Number N \ keyfree" | Nonce: "Nonce N \ keyfree" | Hash: "Hash X \ keyfree" - | MPair: "[|X \ keyfree; Y \ keyfree|] ==> \X,Y\ \ keyfree" - | Crypt: "[|X \ keyfree|] ==> Crypt K X \ keyfree" + | MPair: "\X \ keyfree; Y \ keyfree\ \ \X,Y\ \ keyfree" + | Crypt: "\X \ keyfree\ \ Crypt K X \ keyfree" declare keyfree.intros [intro] @@ -830,13 +820,15 @@ (*The key-free part of a set of messages can be removed from the scope of the analz operator.*) lemma analz_keyfree_into_Un: "\X \ analz (G \ H); G \ keyfree\ \ X \ parts G \ analz H" -apply (erule analz.induct, auto dest: parts.Body) -apply (metis Un_absorb2 keyfree_KeyE parts_Un parts_keyfree UnI2) -done +proof (induction rule: analz.induct) + case (Decrypt K X) + then show ?case + by (metis Un_iff analz.Decrypt in_mono keyfree_KeyE parts.Body parts_keyfree parts_mono) +qed (auto dest: parts.Body) subsection\Tactics useful for many protocol proofs\ ML -\ + \ (*Analysis of Fake cases. Also works for messages that forward unknown parts, but this application is no longer necessary if analz_insert_eq is used. DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) @@ -882,47 +874,49 @@ lemma Crypt_notin_image_Key [simp]: "Crypt K X \ Key ` A" -by auto + by auto lemma Hash_notin_image_Key [simp] :"Hash X \ Key ` A" -by auto + by auto -lemma synth_analz_mono: "G\H ==> synth (analz(G)) \ synth (analz(H))" -by (iprover intro: synth_mono analz_mono) +lemma synth_analz_mono: "G\H \ synth (analz(G)) \ synth (analz(H))" + by (iprover intro: synth_mono analz_mono) lemma Fake_analz_eq [simp]: - "X \ synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)" -by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute - subset_insertI synth_analz_mono synth_increasing synth_subset_iff) + "X \ synth(analz H) \ synth (analz (insert X H)) = synth (analz H)" + by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute + subset_insertI synth_analz_mono synth_increasing synth_subset_iff) text\Two generalizations of \analz_insert_eq\\ lemma gen_analz_insert_eq [rule_format]: - "X \ analz H \ \G. H \ G \ analz (insert X G) = analz G" -by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) + "X \ analz H \ \G. H \ G \ analz (insert X G) = analz G" + by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) -lemma synth_analz_insert_eq [rule_format]: - "X \ synth (analz H) - \ \G. H \ G \ (Key K \ analz (insert X G)) = (Key K \ analz G)" -apply (erule synth.induct) -apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) -done +lemma synth_analz_insert_eq: + "\X \ synth (analz H); H \ G\ + \ (Key K \ analz (insert X G)) \ (Key K \ analz G)" +proof (induction arbitrary: G rule: synth.induct) + case (Inj X) + then show ?case + using gen_analz_insert_eq by presburger +qed (simp_all add: subset_eq) lemma Fake_parts_sing: - "X \ synth (analz H) ==> parts{X} \ synth (analz H) \ parts H" -by (metis Fake_parts_insert empty_subsetI insert_mono parts_mono subset_trans) + "X \ synth (analz H) \ parts{X} \ synth (analz H) \ parts H" + by (metis Fake_parts_insert empty_subsetI insert_mono parts_mono subset_trans) lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD] method_setup spy_analz = \ Scan.succeed (SIMPLE_METHOD' o spy_analz_tac)\ - "for proving the Fake case when analz is involved" + "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = \ Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac)\ - "for debugging spy_analz" + "for debugging spy_analz" method_setup Fake_insert_simp = \ Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac)\ - "for debugging spy_analz" + "for debugging spy_analz" end diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/NS_Public_Bad.thy --- a/src/HOL/Auth/NS_Public_Bad.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/NS_Public_Bad.thy Thu Oct 13 17:19:56 2022 +0100 @@ -78,9 +78,7 @@ \ Crypt (pubEK C) \NA', Nonce NA\ \ parts (spies evs) \ Crypt (pubEK B) \Nonce NA, Agent A\ \ parts (spies evs) \ Nonce NA \ analz (spies evs)" -apply (erule ns_public.induct, simp_all) -apply (blast intro: analz_insertI)+ -done + by (induct rule: ns_public.induct) (auto intro: analz_insertI) (*Unicity for NS1: nonce NA identifies agents A and B*) @@ -90,9 +88,7 @@ Nonce NA \ analz (spies evs); evs \ ns_public\ \ A=A' \ B=B'" apply (erule rev_mp, erule rev_mp, erule rev_mp) -apply (erule ns_public.induct, simp_all) -(*Fake, NS1*) -apply (blast intro!: analz_insertI)+ +apply (erule ns_public.induct, auto intro: analz_insertI) done @@ -116,9 +112,7 @@ \ Crypt (pubEK A) \Nonce NA, Nonce NB\ \ parts (spies evs) \ Says A B (Crypt(pubEK B) \Nonce NA, Agent A\) \ set evs \ Says B A (Crypt(pubEK A) \Nonce NA, Nonce NB\) \ set evs" -apply (erule ns_public.induct) -apply (auto dest: Spy_not_see_NA unique_NA) -done + by (erule ns_public.induct) (auto dest: Spy_not_see_NA unique_NA) theorem A_trusts_NS2: "\Says A B (Crypt(pubEK B) \Nonce NA, Agent A\) \ set evs; @@ -134,10 +128,7 @@ \ Crypt (pubEK B) \Nonce NA, Agent A\ \ parts (spies evs) \ Nonce NA \ analz (spies evs) \ Says A B (Crypt (pubEK B) \Nonce NA, Agent A\) \ set evs" -apply (erule ns_public.induct, simp_all) -(*Fake*) -apply (blast intro!: analz_insertI) -done + by (induction evs rule: ns_public.induct) (use analz_insertI in auto) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/NS_Shared.thy --- a/src/HOL/Auth/NS_Shared.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/NS_Shared.thy Thu Oct 13 17:19:56 2022 +0100 @@ -89,8 +89,8 @@ text\A "possibility property": there are traces that reach the end\ -lemma "[| A \ Server; Key K \ used []; K \ symKeys |] - ==> \N. \evs \ ns_shared. +lemma "\A \ Server; Key K \ used []; K \ symKeys\ + \ \N. \evs \ ns_shared. Says A B (Crypt K \Nonce N, Nonce N\) \ set evs" apply (intro exI bexI) apply (rule_tac [2] ns_shared.Nil @@ -136,8 +136,8 @@ text\Nobody can have used non-existent keys!\ lemma new_keys_not_used [simp]: - "[|Key K \ used evs; K \ symKeys; evs \ ns_shared|] - ==> K \ keysFor (parts (spies evs))" + "\Key K \ used evs; K \ symKeys; evs \ ns_shared\ + \ K \ keysFor (parts (spies evs))" apply (erule rev_mp) apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all) txt\Fake, NS2, NS4, NS5\ diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/OtwayReesBella.thy --- a/src/HOL/Auth/OtwayReesBella.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/OtwayReesBella.thy Thu Oct 13 17:19:56 2022 +0100 @@ -146,7 +146,7 @@ by auto lemma Spy_see_shrK_D [dest!]: - "[|Key (shrK A) \ parts (knows Spy evs); evs \ orb|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ orb\ \ A \ bad" by (blast dest: Spy_see_shrK) lemma new_keys_not_used [simp]: @@ -185,7 +185,7 @@ A \ bad; evs \ orb\ \ Says A B \Nonce M, Agent A, Agent B, Crypt (shrK A) \Nonce Na, Nonce M, Agent A, Agent B\\ \ set evs" apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all) -apply (blast) +apply blast done @@ -312,7 +312,7 @@ txt\Oops\ prefer 4 apply (blast dest: analz_insert_freshCryptK) txt\OR4 - ii\ -prefer 3 apply (blast) +prefer 3 apply blast txt\OR3\ (*adding Gets_imp_ and Says_imp_ for efficiency*) apply (blast dest: diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/OtwayRees_AN.thy --- a/src/HOL/Auth/OtwayRees_AN.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/OtwayRees_AN.thy Thu Oct 13 17:19:56 2022 +0100 @@ -28,31 +28,31 @@ | Fake: \ \The Spy may say anything he can say. The sender field is correct, but agents don't use that information.\ - "[| evsf \ otway; X \ synth (analz (knows Spy evsf)) |] - ==> Says Spy B X # evsf \ otway" + "\evsf \ otway; X \ synth (analz (knows Spy evsf))\ + \ Says Spy B X # evsf \ otway" | Reception: \ \A message that has been sent can be received by the intended recipient.\ - "[| evsr \ otway; Says A B X \set evsr |] - ==> Gets B X # evsr \ otway" + "\evsr \ otway; Says A B X \set evsr\ + \ Gets B X # evsr \ otway" | OR1: \ \Alice initiates a protocol run\ "evs1 \ otway - ==> Says A B \Agent A, Agent B, Nonce NA\ # evs1 \ otway" + \ Says A B \Agent A, Agent B, Nonce NA\ # evs1 \ otway" | OR2: \ \Bob's response to Alice's message.\ - "[| evs2 \ otway; - Gets B \Agent A, Agent B, Nonce NA\ \set evs2 |] - ==> Says B Server \Agent A, Agent B, Nonce NA, Nonce NB\ + "\evs2 \ otway; + Gets B \Agent A, Agent B, Nonce NA\ \set evs2\ + \ Says B Server \Agent A, Agent B, Nonce NA, Nonce NB\ # evs2 \ otway" | OR3: \ \The Server receives Bob's message. Then he sends a new session key to Bob with a packet for forwarding to Alice.\ - "[| evs3 \ otway; Key KAB \ used evs3; + "\evs3 \ otway; Key KAB \ used evs3; Gets Server \Agent A, Agent B, Nonce NA, Nonce NB\ - \set evs3 |] - ==> Says Server B + \set evs3\ + \ Says Server B \Crypt (shrK A) \Nonce NA, Agent A, Agent B, Key KAB\, Crypt (shrK B) \Nonce NB, Agent A, Agent B, Key KAB\\ # evs3 \ otway" @@ -60,20 +60,20 @@ | OR4: \ \Bob receives the Server's (?) message and compares the Nonces with those in the message he previously sent the Server. Need \<^term>\B \ Server\ because we allow messages to self.\ - "[| evs4 \ otway; B \ Server; + "\evs4 \ otway; B \ Server; Says B Server \Agent A, Agent B, Nonce NA, Nonce NB\ \set evs4; Gets B \X, Crypt(shrK B)\Nonce NB,Agent A,Agent B,Key K\\ - \set evs4 |] - ==> Says B A X # evs4 \ otway" + \set evs4\ + \ Says B A X # evs4 \ otway" | Oops: \ \This message models possible leaks of session keys. The nonces identify the protocol run.\ - "[| evso \ otway; + "\evso \ otway; Says Server B \Crypt (shrK A) \Nonce NA, Agent A, Agent B, Key K\, Crypt (shrK B) \Nonce NB, Agent A, Agent B, Key K\\ - \set evso |] - ==> Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ otway" + \set evso\ + \ Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ otway" declare Says_imp_knows_Spy [THEN analz.Inj, dest] @@ -83,8 +83,8 @@ text\A "possibility property": there are traces that reach the end\ -lemma "[| B \ Server; Key K \ used [] |] - ==> \evs \ otway. +lemma "\B \ Server; Key K \ used []\ + \ \evs \ otway. Says B A (Crypt (shrK A) \Nonce NA, Agent A, Agent B, Key K\) \ set evs" apply (intro exI bexI) @@ -96,7 +96,7 @@ done lemma Gets_imp_Says [dest!]: - "[| Gets B X \ set evs; evs \ otway |] ==> \A. Says A B X \ set evs" + "\Gets B X \ set evs; evs \ otway\ \ \A. Says A B X \ set evs" by (erule rev_mp, erule otway.induct, auto) @@ -104,8 +104,8 @@ text\For reasoning about the encrypted portion of messages\ lemma OR4_analz_knows_Spy: - "[| Gets B \X, Crypt(shrK B) X'\ \ set evs; evs \ otway |] - ==> X \ analz (knows Spy evs)" + "\Gets B \X, Crypt(shrK B) X'\ \ set evs; evs \ otway\ + \ X \ analz (knows Spy evs)" by blast @@ -114,15 +114,15 @@ text\Spy never sees a good agent's shared key!\ lemma Spy_see_shrK [simp]: - "evs \ otway ==> (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" + "evs \ otway \ (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" by (erule otway.induct, simp_all, blast+) lemma Spy_analz_shrK [simp]: - "evs \ otway ==> (Key (shrK A) \ analz (knows Spy evs)) = (A \ bad)" + "evs \ otway \ (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 \ otway|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ otway\ \ A \ bad" by (blast dest: Spy_see_shrK) @@ -130,12 +130,12 @@ text\Describes the form of K and NA when the Server sends this message.\ lemma Says_Server_message_form: - "[| Says Server B + "\Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs; - evs \ otway |] - ==> K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" + evs \ otway\ + \ K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" apply (erule rev_mp) apply (erule otway.induct, auto) done @@ -145,7 +145,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. @@ -156,7 +156,7 @@ text\The equality makes the induction hypothesis easier to apply\ lemma analz_image_freshK [rule_format]: - "evs \ otway ==> + "evs \ otway \ \K KK. KK \ -(range shrK) \ (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" @@ -166,7 +166,7 @@ done lemma analz_insert_freshK: - "[| evs \ otway; KAB \ range shrK |] ==> + "\evs \ otway; 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) @@ -174,7 +174,7 @@ text\The Key K uniquely identifies the Server's message.\ lemma unique_session_keys: - "[| Says Server B + "\Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, K\, Crypt (shrK B) \NB, Agent A, Agent B, K\\ \ set evs; @@ -182,8 +182,8 @@ \Crypt (shrK A') \NA', Agent A', Agent B', K\, Crypt (shrK B') \NB', Agent A', Agent B', K\\ \ set evs; - evs \ otway |] - ==> A=A' \ B=B' \ NA=NA' \ NB=NB'" + evs \ otway\ + \ A=A' \ B=B' \ NA=NA' \ NB=NB'" apply (erule rev_mp, erule rev_mp, erule otway.induct, simp_all) apply blast+ \ \OR3 and OR4\ done @@ -193,8 +193,8 @@ text\If the encrypted message appears then it originated with the Server!\ lemma NA_Crypt_imp_Server_msg [rule_format]: - "[| A \ bad; A \ B; evs \ otway |] - ==> Crypt (shrK A) \NA, Agent A, Agent B, Key K\ \ parts (knows Spy evs) + "\A \ bad; A \ B; evs \ otway\ + \ Crypt (shrK A) \NA, Agent A, Agent B, Key K\ \ parts (knows Spy evs) \ (\NB. Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ @@ -208,9 +208,9 @@ text\Corollary: if A receives B's OR4 message then it originated with the Server. Freshness may be inferred from nonce NA.\ lemma A_trusts_OR4: - "[| Says B' A (Crypt (shrK A) \NA, Agent A, Agent B, Key K\) \ set evs; - A \ bad; A \ B; evs \ otway |] - ==> \NB. Says Server B + "\Says B' A (Crypt (shrK A) \NA, Agent A, Agent B, Key K\) \ set evs; + A \ bad; A \ B; evs \ otway\ + \ \NB. Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs" @@ -221,8 +221,8 @@ Does not in itself guarantee security: an attack could violate the premises, e.g. by having \<^term>\A=Spy\\ lemma secrecy_lemma: - "[| A \ bad; B \ bad; evs \ otway |] - ==> Says Server B + "\A \ bad; B \ bad; evs \ otway\ + \ Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs \ @@ -238,23 +238,23 @@ lemma Spy_not_see_encrypted_key: - "[| Says Server B + "\Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs; Notes Spy \NA, NB, Key K\ \ set evs; - A \ bad; B \ bad; evs \ otway |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ otway\ + \ Key K \ analz (knows Spy evs)" by (metis secrecy_lemma) text\A's guarantee. The Oops premise quantifies over NB because A cannot know what it is.\ lemma A_gets_good_key: - "[| Says B' A (Crypt (shrK A) \NA, Agent A, Agent B, Key K\) \ set evs; + "\Says B' A (Crypt (shrK A) \NA, Agent A, Agent B, Key K\) \ set evs; \NB. Notes Spy \NA, NB, Key K\ \ set evs; - A \ bad; B \ bad; A \ B; evs \ otway |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; A \ B; evs \ otway\ + \ Key K \ analz (knows Spy evs)" by (metis A_trusts_OR4 secrecy_lemma) @@ -263,8 +263,8 @@ text\If the encrypted message appears then it originated with the Server!\ lemma NB_Crypt_imp_Server_msg [rule_format]: - "[| B \ bad; A \ B; evs \ otway |] - ==> Crypt (shrK B) \NB, Agent A, Agent B, Key K\ \ parts (knows Spy evs) + "\B \ bad; A \ B; evs \ otway\ + \ Crypt (shrK B) \NB, Agent A, Agent B, Key K\ \ parts (knows Spy evs) \ (\NA. Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ @@ -278,10 +278,10 @@ text\Guarantee for B: if it gets a well-formed certificate then the Server has sent the correct message in round 3.\ lemma B_trusts_OR3: - "[| Says S B \X, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ + "\Says S B \X, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs; - B \ bad; A \ B; evs \ otway |] - ==> \NA. Says Server B + B \ bad; A \ B; evs \ otway\ + \ \NA. Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs" @@ -291,11 +291,11 @@ text\The obvious combination of \B_trusts_OR3\ with \Spy_not_see_encrypted_key\\ lemma B_gets_good_key: - "[| Gets B \X, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ + "\Gets B \X, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs; \NA. Notes Spy \NA, NB, Key K\ \ set evs; - A \ bad; B \ bad; A \ B; evs \ otway |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; A \ B; evs \ otway\ + \ Key K \ analz (knows Spy evs)" by (blast dest: B_trusts_OR3 Spy_not_see_encrypted_key) end diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/OtwayRees_Bad.thy --- a/src/HOL/Auth/OtwayRees_Bad.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/OtwayRees_Bad.thy Thu Oct 13 17:19:56 2022 +0100 @@ -25,26 +25,26 @@ | Fake: \ \The Spy may say anything he can say. The sender field is correct, but agents don't use that information.\ - "[| evsf \ otway; X \ synth (analz (knows Spy evsf)) |] - ==> Says Spy B X # evsf \ otway" + "\evsf \ otway; X \ synth (analz (knows Spy evsf))\ + \ Says Spy B X # evsf \ otway" | Reception: \ \A message that has been sent can be received by the intended recipient.\ - "[| evsr \ otway; Says A B X \set evsr |] - ==> Gets B X # evsr \ otway" + "\evsr \ otway; Says A B X \set evsr\ + \ Gets B X # evsr \ otway" | OR1: \ \Alice initiates a protocol run\ - "[| evs1 \ otway; Nonce NA \ used evs1 |] - ==> Says A B \Nonce NA, Agent A, Agent B, + "\evs1 \ otway; Nonce NA \ used evs1\ + \ Says A B \Nonce NA, Agent A, Agent B, Crypt (shrK A) \Nonce NA, Agent A, Agent B\\ # evs1 \ otway" | OR2: \ \Bob's response to Alice's message. This variant of the protocol does NOT encrypt NB.\ - "[| evs2 \ otway; Nonce NB \ used evs2; - Gets B \Nonce NA, Agent A, Agent B, X\ \ set evs2 |] - ==> Says B Server + "\evs2 \ otway; Nonce NB \ used evs2; + Gets B \Nonce NA, Agent A, Agent B, X\ \ set evs2\ + \ Says B Server \Nonce NA, Agent A, Agent B, X, Nonce NB, Crypt (shrK B) \Nonce NA, Agent A, Agent B\\ # evs2 \ otway" @@ -52,14 +52,14 @@ | OR3: \ \The Server receives Bob's message and checks that the three NAs match. Then he sends a new session key to Bob with a packet for forwarding to Alice.\ - "[| evs3 \ otway; Key KAB \ used evs3; + "\evs3 \ otway; Key KAB \ used evs3; Gets Server \Nonce NA, Agent A, Agent B, Crypt (shrK A) \Nonce NA, Agent A, Agent B\, Nonce NB, Crypt (shrK B) \Nonce NA, Agent A, Agent B\\ - \ set evs3 |] - ==> Says Server B + \ set evs3\ + \ Says Server B \Nonce NA, Crypt (shrK A) \Nonce NA, Key KAB\, Crypt (shrK B) \Nonce NB, Key KAB\\ @@ -68,20 +68,20 @@ | OR4: \ \Bob receives the Server's (?) message and compares the Nonces with those in the message he previously sent the Server. Need \<^term>\B \ Server\ because we allow messages to self.\ - "[| evs4 \ otway; B \ Server; + "\evs4 \ otway; B \ Server; Says B Server \Nonce NA, Agent A, Agent B, X', Nonce NB, Crypt (shrK B) \Nonce NA, Agent A, Agent B\\ \ set evs4; Gets B \Nonce NA, X, Crypt (shrK B) \Nonce NB, Key K\\ - \ set evs4 |] - ==> Says B A \Nonce NA, X\ # evs4 \ otway" + \ set evs4\ + \ Says B A \Nonce NA, X\ # evs4 \ otway" | Oops: \ \This message models possible leaks of session keys. The nonces identify the protocol run.\ - "[| evso \ otway; + "\evso \ otway; Says Server B \Nonce NA, X, Crypt (shrK B) \Nonce NB, Key K\\ - \ set evso |] - ==> Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ otway" + \ set evso\ + \ Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ otway" declare Says_imp_knows_Spy [THEN analz.Inj, dest] @@ -90,8 +90,8 @@ declare Fake_parts_insert_in_Un [dest] text\A "possibility property": there are traces that reach the end\ -lemma "[| B \ Server; Key K \ used [] |] - ==> \NA. \evs \ otway. +lemma "\B \ Server; Key K \ used []\ + \ \NA. \evs \ otway. Says B A \Nonce NA, Crypt (shrK A) \Nonce NA, Key K\\ \ set evs" apply (intro exI bexI) @@ -103,7 +103,7 @@ done lemma Gets_imp_Says [dest!]: - "[| Gets B X \ set evs; evs \ otway |] ==> \A. Says A B X \ set evs" + "\Gets B X \ set evs; evs \ otway\ \ \A. Says A B X \ set evs" apply (erule rev_mp) apply (erule otway.induct, auto) done @@ -112,18 +112,18 @@ subsection\For reasoning about the encrypted portion of messages\ lemma OR2_analz_knows_Spy: - "[| Gets B \N, Agent A, Agent B, X\ \ set evs; evs \ otway |] - ==> X \ analz (knows Spy evs)" + "\Gets B \N, Agent A, Agent B, X\ \ set evs; evs \ otway\ + \ X \ analz (knows Spy evs)" by blast lemma OR4_analz_knows_Spy: - "[| Gets B \N, X, Crypt (shrK B) X'\ \ set evs; evs \ otway |] - ==> X \ analz (knows Spy evs)" + "\Gets B \N, X, Crypt (shrK B) X'\ \ set evs; evs \ otway\ + \ X \ analz (knows Spy evs)" by blast lemma Oops_parts_knows_Spy: "Says Server B \NA, X, Crypt K' \NB,K\\ \ set evs - ==> K \ parts (knows Spy evs)" + \ K \ parts (knows Spy evs)" by blast text\Forwarding lemma: see comments in OtwayRees.thy\ @@ -136,17 +136,17 @@ text\Spy never sees a good agent's shared key!\ lemma Spy_see_shrK [simp]: - "evs \ otway ==> (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" + "evs \ otway \ (Key (shrK A) \ parts (knows Spy evs)) = (A \ bad)" by (erule otway.induct, force, drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+) lemma Spy_analz_shrK [simp]: - "evs \ otway ==> (Key (shrK A) \ analz (knows Spy evs)) = (A \ bad)" + "evs \ otway \ (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 \ otway|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ otway\ \ A \ bad" by (blast dest: Spy_see_shrK) @@ -155,9 +155,9 @@ text\Describes the form of K and NA when the Server sends this message. Also for Oops case.\ lemma Says_Server_message_form: - "[| Says Server B \NA, X, Crypt (shrK B) \NB, Key K\\ \ set evs; - evs \ otway |] - ==> K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" + "\Says Server B \NA, X, Crypt (shrK B) \NB, Key K\\ \ set evs; + evs \ otway\ + \ K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" apply (erule rev_mp) apply (erule otway.induct, simp_all) done @@ -166,7 +166,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. @@ -177,7 +177,7 @@ text\The equality makes the induction hypothesis easier to apply\ lemma analz_image_freshK [rule_format]: - "evs \ otway ==> + "evs \ otway \ \K KK. KK \ -(range shrK) \ (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" @@ -188,7 +188,7 @@ done lemma analz_insert_freshK: - "[| evs \ otway; KAB \ range shrK |] ==> + "\evs \ otway; 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) @@ -196,9 +196,9 @@ text\The Key K uniquely identifies the Server's message.\ lemma unique_session_keys: - "[| Says Server B \NA, X, Crypt (shrK B) \NB, K\\ \ set evs; + "\Says Server B \NA, X, Crypt (shrK B) \NB, K\\ \ set evs; Says Server B' \NA',X',Crypt (shrK B') \NB',K\\ \ set evs; - evs \ otway |] ==> X=X' \ B=B' \ NA=NA' \ NB=NB'" + evs \ otway\ \ X=X' \ B=B' \ NA=NA' \ NB=NB'" apply (erule rev_mp) apply (erule rev_mp) apply (erule otway.induct, simp_all) @@ -210,8 +210,8 @@ Does not in itself guarantee security: an attack could violate the premises, e.g. by having \<^term>\A=Spy\\ lemma secrecy_lemma: - "[| A \ bad; B \ bad; evs \ otway |] - ==> Says Server B + "\A \ bad; B \ bad; evs \ otway\ + \ Says Server B \NA, Crypt (shrK A) \NA, Key K\, Crypt (shrK B) \NB, Key K\\ \ set evs \ Notes Spy \NA, NB, Key K\ \ set evs \ @@ -227,12 +227,12 @@ lemma Spy_not_see_encrypted_key: - "[| Says Server B + "\Says Server B \NA, Crypt (shrK A) \NA, Key K\, Crypt (shrK B) \NB, Key K\\ \ set evs; Notes Spy \NA, NB, Key K\ \ set evs; - A \ bad; B \ bad; evs \ otway |] - ==> Key K \ analz (knows Spy evs)" + A \ bad; B \ bad; evs \ otway\ + \ Key K \ analz (knows Spy evs)" by (blast dest: Says_Server_message_form secrecy_lemma) @@ -242,8 +242,8 @@ \<^term>\A \ B\ prevents OR2's similar-looking cryptogram from being picked up. Original Otway-Rees doesn't need it.\ lemma Crypt_imp_OR1 [rule_format]: - "[| A \ bad; A \ B; evs \ otway |] - ==> Crypt (shrK A) \NA, Agent A, Agent B\ \ parts (knows Spy evs) \ + "\A \ bad; A \ B; evs \ otway\ + \ Crypt (shrK A) \NA, Agent A, Agent B\ \ parts (knows Spy evs) \ Says A B \NA, Agent A, Agent B, Crypt (shrK A) \NA, Agent A, Agent B\\ \ set evs" by (erule otway.induct, force, @@ -255,8 +255,8 @@ The premise \<^term>\A \ B\ allows use of \Crypt_imp_OR1\\ text\Only it is FALSE. Somebody could make a fake message to Server substituting some other nonce NA' for NB.\ -lemma "[| A \ bad; A \ B; evs \ otway |] - ==> Crypt (shrK A) \NA, Key K\ \ parts (knows Spy evs) \ +lemma "\A \ bad; A \ B; evs \ otway\ + \ Crypt (shrK A) \NA, Key K\ \ parts (knows Spy evs) \ Says A B \NA, Agent A, Agent B, Crypt (shrK A) \NA, Agent A, Agent B\\ \ set evs \ diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Public.thy Thu Oct 13 17:19:56 2022 +0100 @@ -11,7 +11,7 @@ imports Event begin -lemma invKey_K: "K \ symKeys ==> invKey K = K" +lemma invKey_K: "K \ symKeys \ invKey K = K" by (simp add: symKeys_def) subsection\Asymmetric Keys\ @@ -59,7 +59,7 @@ \<^term>\True\False\, no agent has identical signing and encryption keys\ specification (publicKey) injective_publicKey: - "publicKey b A = publicKey c A' ==> b=c \ A=A'" + "publicKey b A = publicKey c A' \ b=c \ A=A'" apply (rule exI [of _ "\b A. 2 * case_agent 0 (\n. n + 2) 1 A + case_keymode 0 1 b"]) apply (auto simp add: inj_on_def split: agent.split keymode.split) @@ -88,18 +88,18 @@ lemma not_symKeys_priK [iff]: "privateKey b A \ symKeys" by (simp add: symKeys_def) -lemma symKey_neq_priEK: "K \ symKeys ==> K \ priEK A" +lemma symKey_neq_priEK: "K \ symKeys \ K \ priEK A" by auto -lemma symKeys_neq_imp_neq: "(K \ symKeys) \ (K' \ symKeys) ==> K \ K'" +lemma symKeys_neq_imp_neq: "(K \ symKeys) \ (K' \ symKeys) \ K \ K'" by blast lemma symKeys_invKey_iff [iff]: "(invKey K \ symKeys) = (K \ symKeys)" -by (unfold symKeys_def, auto) + unfolding symKeys_def by auto lemma analz_symKeys_Decrypt: - "[| Crypt K X \ analz H; K \ symKeys; Key K \ analz H |] - ==> X \ analz H" + "\Crypt K X \ analz H; K \ symKeys; Key K \ analz H\ + \ X \ analz H" by (auto simp add: symKeys_def) @@ -152,11 +152,11 @@ by (simp add: invKey_K) lemma analz_shrK_Decrypt: - "[| Crypt (shrK A) X \ analz H; Key(shrK A) \ analz H |] ==> X \ analz H" + "\Crypt (shrK A) X \ analz H; Key(shrK A) \ analz H\ \ X \ analz H" by auto lemma analz_Decrypt': - "[| Crypt K X \ analz H; K \ symKeys; Key K \ analz H |] ==> X \ analz H" + "\Crypt K X \ analz H; K \ symKeys; Key K \ analz H\ \ X \ analz H" by (auto simp add: invKey_K) lemma priK_neq_shrK [iff]: "shrK A \ privateKey b C" @@ -239,15 +239,15 @@ apply (auto dest!: parts_cut simp add: used_Nil) done -lemma MPair_used_D: "\X,Y\ \ used H ==> X \ used H \ Y \ used H" +lemma MPair_used_D: "\X,Y\ \ used H \ X \ used H \ Y \ used H" by (drule used_parts_subset_parts, simp, blast) text\There was a similar theorem in Event.thy, so perhaps this one can be moved up if proved directly by induction.\ lemma MPair_used [elim!]: - "[| \X,Y\ \ used H; - [| X \ used H; Y \ used H |] ==> P |] - ==> P" + "\\X,Y\ \ used H; + \X \ used H; Y \ used H\ \ P\ + \ P" by (blast dest: MPair_used_D) @@ -255,7 +255,7 @@ that expression is not in normal form.\ lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}" -apply (unfold keysFor_def) +unfolding keysFor_def apply (induct_tac "C") apply (auto intro: range_eqI) done @@ -283,10 +283,10 @@ (*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 \ used evs ==> K \ range shrK" +lemma Key_not_used [simp]: "Key K \ used evs \ K \ range shrK" by blast -lemma shrK_neq: "Key K \ used evs ==> shrK B \ K" +lemma shrK_neq: "Key K \ used evs \ shrK B \ K" by blast lemmas neq_shrK = shrK_neq [THEN not_sym] @@ -317,14 +317,14 @@ text\Spy sees private keys of bad agents!\ lemma Spy_spies_bad_privateKey [intro!]: - "A \ bad ==> Key (privateKey b A) \ spies evs" + "A \ bad \ Key (privateKey b A) \ spies evs" apply (induct_tac "evs") apply (auto simp add: imageI knows_Cons split: event.split) done text\Spy sees long-term shared keys of bad agents!\ lemma Spy_spies_bad_shrK [intro!]: - "A \ bad ==> Key (shrK A) \ spies evs" + "A \ bad \ Key (shrK A) \ spies evs" apply (induct_tac "evs") apply (simp_all add: imageI knows_Cons split: event.split) done @@ -341,8 +341,8 @@ (*For case analysis on whether or not an agent is compromised*) lemma Crypt_Spy_analz_bad: - "[| Crypt (shrK A) X \ analz (knows Spy evs); A \ bad |] - ==> X \ analz (knows Spy evs)" + "\Crypt (shrK A) X \ analz (knows Spy evs); A \ bad\ + \ X \ analz (knows Spy evs)" by force @@ -383,13 +383,13 @@ by blast -lemma Crypt_imp_keysFor :"[|Crypt K X \ H; K \ symKeys|] ==> K \ keysFor H" +lemma Crypt_imp_keysFor :"\Crypt K X \ H; K \ symKeys\ \ K \ keysFor H" by (drule Crypt_imp_invKey_keysFor, simp) text\Lemma for the trivial direction of the if-and-only-if of the Session Key Compromise Theorem\ lemma analz_image_freshK_lemma: - "(Key K \ analz (Key`nE \ H)) \ (K \ nE | Key K \ analz H) ==> + "(Key K \ analz (Key`nE \ H)) \ (K \ nE | Key K \ analz H) \ (Key K \ analz (Key`nE \ H)) = (K \ nE | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Recur.thy --- a/src/HOL/Auth/Recur.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Recur.thy Thu Oct 13 17:19:56 2022 +0100 @@ -21,15 +21,15 @@ for evs :: "event list" where One: "Key KAB \ used evs - ==> (Hash[Key(shrK A)] \Agent A, Agent B, Nonce NA, END\, + \ (Hash[Key(shrK A)] \Agent A, Agent B, Nonce NA, END\, \Crypt (shrK A) \Key KAB, Agent B, Nonce NA\, END\, KAB) \ respond evs" (*The most recent session key is passed up to the caller*) - | Cons: "[| (PA, RA, KAB) \ respond evs; + | Cons: "\(PA, RA, KAB) \ respond evs; Key KBC \ used evs; Key KBC \ parts {RA}; - PA = Hash[Key(shrK A)] \Agent A, Agent B, Nonce NA, P\ |] - ==> (Hash[Key(shrK B)] \Agent B, Agent C, Nonce NB, PA\, + PA = Hash[Key(shrK A)] \Agent A, Agent B, Nonce NA, P\\ + \ (Hash[Key(shrK B)] \Agent B, Agent C, Nonce NB, PA\, \Crypt (shrK B) \Key KBC, Agent C, Nonce NB\, Crypt (shrK B) \Key KAB, Agent A, Nonce NB\, RA\, @@ -47,8 +47,8 @@ (*Server terminates lists*) Nil: "END \ responses evs" - | Cons: "[| RA \ responses evs; Key KAB \ used evs |] - ==> \Crypt (shrK B) \Key KAB, Agent A, Nonce NB\, + | Cons: "\RA \ responses evs; Key KAB \ used evs\ + \ \Crypt (shrK B) \Key KAB, Agent A, Nonce NB\, RA\ \ responses evs" @@ -59,37 +59,37 @@ (*The spy MAY say anything he CAN say. Common to all similar protocols.*) - | Fake: "[| evsf \ recur; X \ synth (analz (knows Spy evsf)) |] - ==> Says Spy B X # evsf \ recur" + | Fake: "\evsf \ recur; X \ synth (analz (knows Spy evsf))\ + \ Says Spy B X # evsf \ recur" (*Alice initiates a protocol run. END is a placeholder to terminate the nesting.*) - | RA1: "[| evs1 \ recur; Nonce NA \ used evs1 |] - ==> Says A B (Hash[Key(shrK A)] \Agent A, Agent B, Nonce NA, END\) + | RA1: "\evs1 \ recur; Nonce NA \ used evs1\ + \ Says A B (Hash[Key(shrK A)] \Agent A, Agent B, Nonce NA, END\) # evs1 \ recur" (*Bob's response to Alice's message. C might be the Server. We omit PA = \XA, Agent A, Agent B, Nonce NA, P\ because it complicates proofs, so B may respond to any message at all!*) - | RA2: "[| evs2 \ recur; Nonce NB \ used evs2; - Says A' B PA \ set evs2 |] - ==> Says B C (Hash[Key(shrK B)] \Agent B, Agent C, Nonce NB, PA\) + | RA2: "\evs2 \ recur; Nonce NB \ used evs2; + Says A' B PA \ set evs2\ + \ Says B C (Hash[Key(shrK B)] \Agent B, Agent C, Nonce NB, PA\) # evs2 \ recur" (*The Server receives Bob's message and prepares a response.*) - | RA3: "[| evs3 \ recur; Says B' Server PB \ set evs3; - (PB,RB,K) \ respond evs3 |] - ==> Says Server B RB # evs3 \ recur" + | RA3: "\evs3 \ recur; Says B' Server PB \ set evs3; + (PB,RB,K) \ respond evs3\ + \ Says Server B RB # evs3 \ recur" (*Bob receives the returned message and compares the Nonces with those in the message he previously sent the Server.*) - | RA4: "[| evs4 \ recur; + | RA4: "\evs4 \ recur; Says B C \XH, Agent B, Agent C, Nonce NB, XA, Agent A, Agent B, Nonce NA, P\ \ set evs4; Says C' B \Crypt (shrK B) \Key KBC, Agent C, Nonce NB\, Crypt (shrK B) \Key KAB, Agent A, Nonce NB\, - RA\ \ set evs4 |] - ==> Says B A RA # evs4 \ recur" + RA\ \ set evs4\ + \ Says B A RA # evs4 \ recur" (*No "oops" message can easily be expressed. Each session key is associated--in two separate messages--with two nonces. This is @@ -99,9 +99,9 @@ the chain. Oops cases proved using parts_cut, Key_in_keysFor_parts, etc. - Oops: "[| evso \ recur; Says Server B RB \ set evso; - RB \ responses evs'; Key K \ parts {RB} |] - ==> Notes Spy \Key K, RB\ # evso \ recur" + Oops: "\evso \ recur; Says Server B RB \ set evso; + RB \ responses evs'; Key K \ parts {RB}\ + \ Notes Spy \Key K, RB\ # evso \ recur" *) @@ -119,7 +119,7 @@ text\Simplest case: Alice goes directly to the server\ lemma "Key K \ used [] - ==> \NA. \evs \ recur. + \ \NA. \evs \ recur. Says Server A \Crypt (shrK A) \Key K, Agent Server, Nonce NA\, END\ \ set evs" apply (intro exI bexI) @@ -130,9 +130,9 @@ text\Case two: Alice, Bob and the server\ -lemma "[| Key K \ used []; Key K' \ used []; K \ K'; - Nonce NA \ used []; Nonce NB \ used []; NA < NB |] - ==> \NA. \evs \ recur. +lemma "\Key K \ used []; Key K' \ used []; K \ K'; + Nonce NA \ used []; Nonce NB \ used []; NA < NB\ + \ \NA. \evs \ recur. Says B A \Crypt (shrK A) \Key K, Agent B, Nonce NA\, END\ \ set evs" apply (intro exI bexI) @@ -147,11 +147,11 @@ done (*Case three: Alice, Bob, Charlie and the server Rather slow (5 seconds)*) -lemma "[| Key K \ used []; Key K' \ used []; +lemma "\Key K \ used []; Key K' \ used []; Key K'' \ used []; K \ K'; K' \ K''; K \ K''; Nonce NA \ used []; Nonce NB \ used []; Nonce NC \ used []; - NA < NB; NB < NC |] - ==> \K. \NA. \evs \ recur. + NA < NB; NB < NC\ + \ \K. \NA. \evs \ recur. Says B A \Crypt (shrK A) \Key K, Agent B, Nonce NA\, END\ \ set evs" apply (intro exI bexI) @@ -167,18 +167,18 @@ done -lemma respond_imp_not_used: "(PA,RB,KAB) \ respond evs ==> Key KAB \ used evs" +lemma respond_imp_not_used: "(PA,RB,KAB) \ respond evs \ Key KAB \ used evs" by (erule respond.induct, simp_all) lemma Key_in_parts_respond [rule_format]: - "[| Key K \ parts {RB}; (PB,RB,K') \ respond evs |] ==> Key K \ used evs" + "\Key K \ parts {RB}; (PB,RB,K') \ respond evs\ \ Key K \ used evs" apply (erule rev_mp, erule respond.induct) apply (auto dest: Key_not_used respond_imp_not_used) done text\Simple inductive reasoning about responses\ lemma respond_imp_responses: - "(PA,RB,KAB) \ respond evs ==> RB \ responses evs" + "(PA,RB,KAB) \ respond evs \ RB \ responses evs" apply (erule respond.induct) apply (blast intro!: respond_imp_not_used responses.intros)+ done @@ -189,7 +189,7 @@ lemmas RA2_analz_spies = Says_imp_spies [THEN analz.Inj] lemma RA4_analz_spies: - "Says C' B \Crypt K X, X', RA\ \ set evs ==> RA \ analz (spies evs)" + "Says C' B \Crypt K X, X', RA\ \ set evs \ RA \ analz (spies evs)" by blast @@ -208,18 +208,18 @@ (** Spy never sees another agent's shared key! (unless it's bad at start) **) lemma Spy_see_shrK [simp]: - "evs \ recur ==> (Key (shrK A) \ parts (spies evs)) = (A \ bad)" + "evs \ recur \ (Key (shrK A) \ parts (spies evs)) = (A \ bad)" apply (erule recur.induct, auto) txt\RA3. It's ugly to call auto twice, but it seems necessary.\ apply (auto dest: Key_in_parts_respond simp add: parts_insert_spies) done lemma Spy_analz_shrK [simp]: - "evs \ recur ==> (Key (shrK A) \ analz (spies evs)) = (A \ bad)" + "evs \ recur \ (Key (shrK A) \ analz (spies evs)) = (A \ bad)" by auto lemma Spy_see_shrK_D [dest!]: - "[|Key (shrK A) \ parts (knows Spy evs); evs \ recur|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ recur\ \ A \ bad" by (blast dest: Spy_see_shrK) @@ -231,11 +231,11 @@ Note that it holds for *any* set H (not just "spies evs") satisfying the inductive hypothesis.*) lemma resp_analz_image_freshK_lemma: - "[| RB \ responses evs; + "\RB \ responses evs; \K KK. KK \ - (range shrK) \ (Key K \ analz (Key`KK \ H)) = - (K \ KK | Key K \ analz H) |] - ==> \K KK. KK \ - (range shrK) \ + (K \ KK | Key K \ analz H)\ + \ \K KK. KK \ - (range shrK) \ (Key K \ analz (insert RB (Key`KK \ H))) = (K \ KK | Key K \ analz (insert RB H))" apply (erule responses.induct) @@ -246,7 +246,7 @@ text\Version for the protocol. Proof is easy, thanks to the lemma.\ lemma raw_analz_image_freshK: - "evs \ recur ==> + "evs \ recur \ \K KK. KK \ - (range shrK) \ (Key K \ analz (Key`KK \ (spies evs))) = (K \ KK | Key K \ analz (spies evs))" @@ -260,8 +260,8 @@ (*Instance of the lemma with H replaced by (spies evs): - [| RB \ responses evs; evs \ recur; |] - ==> KK \ - (range shrK) \ + \RB \ responses evs; evs \ recur;\ + \ KK \ - (range shrK) \ Key K \ analz (insert RB (Key`KK \ spies evs)) = (K \ KK | Key K \ analz (insert RB (spies evs))) *) @@ -269,8 +269,8 @@ resp_analz_image_freshK_lemma [OF _ raw_analz_image_freshK] lemma analz_insert_freshK: - "[| evs \ recur; KAB \ range shrK |] - ==> (Key K \ analz (insert (Key KAB) (spies evs))) = + "\evs \ recur; KAB \ range shrK\ + \ (Key K \ analz (insert (Key KAB) (spies evs))) = (K = KAB | Key K \ analz (spies evs))" by (simp del: image_insert add: analz_image_freshK_simps raw_analz_image_freshK) @@ -278,8 +278,8 @@ text\Everything that's hashed is already in past traffic.\ lemma Hash_imp_body: - "[| Hash \Key(shrK A), X\ \ parts (spies evs); - evs \ recur; A \ bad |] ==> X \ parts (spies evs)" + "\Hash \Key(shrK A), X\ \ parts (spies evs); + evs \ recur; A \ bad\ \ X \ parts (spies evs)" apply (erule rev_mp) apply (erule recur.induct, drule_tac [6] RA4_parts_spies, @@ -299,10 +299,10 @@ **) lemma unique_NA: - "[| Hash \Key(shrK A), Agent A, B, NA, P\ \ parts (spies evs); + "\Hash \Key(shrK A), Agent A, B, NA, P\ \ parts (spies evs); Hash \Key(shrK A), Agent A, B',NA, P'\ \ parts (spies evs); - evs \ recur; A \ bad |] - ==> B=B' \ P=P'" + evs \ recur; A \ bad\ + \ B=B' \ P=P'" apply (erule rev_mp, erule rev_mp) apply (erule recur.induct, drule_tac [5] respond_imp_responses) @@ -321,8 +321,8 @@ ***) lemma shrK_in_analz_respond [simp]: - "[| RB \ responses evs; evs \ recur |] - ==> (Key (shrK B) \ analz (insert RB (spies evs))) = (B\bad)" + "\RB \ responses evs; evs \ recur\ + \ (Key (shrK B) \ analz (insert RB (spies evs))) = (B\bad)" apply (erule responses.induct) apply (simp_all del: image_insert add: analz_image_freshK_simps resp_analz_image_freshK, auto) @@ -330,12 +330,12 @@ lemma resp_analz_insert_lemma: - "[| Key K \ analz (insert RB H); + "\Key K \ analz (insert RB H); \K KK. KK \ - (range shrK) \ (Key K \ analz (Key`KK \ H)) = (K \ KK | Key K \ analz H); - RB \ responses evs |] - ==> (Key K \ parts{RB} | Key K \ analz H)" + RB \ responses evs\ + \ (Key K \ parts{RB} | Key K \ analz H)" apply (erule rev_mp, erule responses.induct) apply (simp_all del: image_insert parts_image add: analz_image_freshK_simps resp_analz_image_freshK_lemma) @@ -349,7 +349,7 @@ text\The last key returned by respond indeed appears in a certificate\ lemma respond_certificate: "(Hash[Key(shrK A)] \Agent A, B, NA, P\, RA, K) \ respond evs - ==> Crypt (shrK A) \Key K, B, NA\ \ parts {RA}" + \ Crypt (shrK A) \Key K, B, NA\ \ parts {RA}" apply (ind_cases "(Hash[Key (shrK A)] \Agent A, B, NA, P\, RA, K) \ respond evs") apply simp_all done @@ -360,7 +360,7 @@ the inductive step complicates the case analysis. Unusually for such proofs, the quantifiers appear to be necessary.*) lemma unique_lemma [rule_format]: - "(PB,RB,KXY) \ respond evs ==> + "(PB,RB,KXY) \ respond evs \ \A B N. Crypt (shrK A) \Key K, Agent B, N\ \ parts {RB} \ (\A' B' N'. Crypt (shrK A') \Key K, Agent B', N'\ \ parts {RB} \ (A'=A \ B'=B) | (A'=B \ B'=A))" @@ -370,10 +370,10 @@ done lemma unique_session_keys: - "[| Crypt (shrK A) \Key K, Agent B, N\ \ parts {RB}; + "\Crypt (shrK A) \Key K, Agent B, N\ \ parts {RB}; Crypt (shrK A') \Key K, Agent B', N'\ \ parts {RB}; - (PB,RB,KXY) \ respond evs |] - ==> (A'=A \ B'=B) | (A'=B \ B'=A)" + (PB,RB,KXY) \ respond evs\ + \ (A'=A \ B'=B) | (A'=B \ B'=A)" by (rule unique_lemma, auto) @@ -382,8 +382,8 @@ the premises, e.g. by having A=Spy **) lemma respond_Spy_not_see_session_key [rule_format]: - "[| (PB,RB,KAB) \ respond evs; evs \ recur |] - ==> \A A' N. A \ bad \ A' \ bad \ + "\(PB,RB,KAB) \ respond evs; evs \ recur\ + \ \A A' N. A \ bad \ A' \ bad \ Crypt (shrK A) \Key K, Agent A', N\ \ parts{RB} \ Key K \ analz (insert RB (spies evs))" apply (erule respond.induct) @@ -405,9 +405,9 @@ lemma Spy_not_see_session_key: - "[| Crypt (shrK A) \Key K, Agent A', N\ \ parts (spies evs); - A \ bad; A' \ bad; evs \ recur |] - ==> Key K \ analz (spies evs)" + "\Crypt (shrK A) \Key K, Agent A', N\ \ parts (spies evs); + A \ bad; A' \ bad; evs \ recur\ + \ Key K \ analz (spies evs)" apply (erule rev_mp) apply (erule recur.induct) apply (drule_tac [4] RA2_analz_spies, @@ -430,9 +430,9 @@ text\The response never contains Hashes\ lemma Hash_in_parts_respond: - "[| Hash \Key (shrK B), M\ \ parts (insert RB H); - (PB,RB,K) \ respond evs |] - ==> Hash \Key (shrK B), M\ \ parts H" + "\Hash \Key (shrK B), M\ \ parts (insert RB H); + (PB,RB,K) \ respond evs\ + \ Hash \Key (shrK B), M\ \ parts H" apply (erule rev_mp) apply (erule respond_imp_responses [THEN responses.induct], auto) done @@ -442,10 +442,10 @@ it can say nothing about how recent A's message is. It might later be used to prove B's presence to A at the run's conclusion.\ lemma Hash_auth_sender [rule_format]: - "[| Hash \Key(shrK A), Agent A, Agent B, NA, P\ \ parts(spies evs); - A \ bad; evs \ recur |] - ==> Says A B (Hash[Key(shrK A)] \Agent A, Agent B, NA, P\) \ set evs" -apply (unfold HPair_def) + "\Hash \Key(shrK A), Agent A, Agent B, NA, P\ \ parts(spies evs); + A \ bad; evs \ recur\ + \ Says A B (Hash[Key(shrK A)] \Agent A, Agent B, NA, P\) \ set evs" +unfolding HPair_def apply (erule rev_mp) apply (erule recur.induct, drule_tac [6] RA4_parts_spies, @@ -462,9 +462,9 @@ text\Certificates can only originate with the Server.\ lemma Cert_imp_Server_msg: - "[| Crypt (shrK A) Y \ parts (spies evs); - A \ bad; evs \ recur |] - ==> \C RC. Says Server C RC \ set evs \ + "\Crypt (shrK A) Y \ parts (spies evs); + A \ bad; evs \ recur\ + \ \C RC. Says Server C RC \ set evs \ Crypt (shrK A) Y \ parts {RC}" apply (erule rev_mp, erule recur.induct, simp_all) txt\Fake\ diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Shared.thy --- a/src/HOL/Auth/Shared.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Shared.thy Thu Oct 13 17:19:56 2022 +0100 @@ -48,7 +48,7 @@ lemma analz_Decrypt' [dest]: - "[| Crypt K X \ analz H; Key K \ analz H |] ==> X \ analz H" + "\Crypt K X \ analz H; Key K \ analz H\ \ X \ analz H" by auto text\Now cancel the \dest\ attribute given to @@ -59,17 +59,17 @@ that expression is not in normal form.\ lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}" -apply (unfold keysFor_def) +unfolding keysFor_def apply (induct_tac "C", auto) done (*Specialized to shared-key model: no @{term invKey}*) lemma keysFor_parts_insert: - "[| K \ keysFor (parts (insert X G)); X \ synth (analz H) |] - ==> K \ keysFor (parts (G \ H)) | Key K \ parts H" + "\K \ keysFor (parts (insert X G)); X \ synth (analz H)\ + \ K \ keysFor (parts (G \ H)) | Key K \ parts H" by (metis invKey_K keysFor_parts_insert) -lemma Crypt_imp_keysFor: "Crypt K X \ H ==> K \ keysFor H" +lemma Crypt_imp_keysFor: "Crypt K X \ H \ K \ keysFor H" by (metis Crypt_imp_invKey_keysFor invKey_K) @@ -82,8 +82,8 @@ done (*For case analysis on whether or not an agent is compromised*) -lemma Crypt_Spy_analz_bad: "[| Crypt (shrK A) X \ analz (knows Spy evs); A \ bad |] - ==> X \ analz (knows Spy evs)" +lemma Crypt_Spy_analz_bad: "\Crypt (shrK A) X \ analz (knows Spy evs); A \ bad\ + \ X \ analz (knows Spy evs)" by (metis Spy_knows_Spy_bad analz.Inj analz_Decrypt') @@ -98,10 +98,10 @@ (*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 \ used evs ==> K \ range shrK" +lemma Key_not_used [simp]: "Key K \ used evs \ K \ range shrK" by blast -lemma shrK_neq [simp]: "Key K \ used evs ==> shrK B \ K" +lemma shrK_neq [simp]: "Key K \ used evs \ shrK B \ K" by blast lemmas shrK_sym_neq = shrK_neq [THEN not_sym] @@ -153,7 +153,7 @@ done text\Unlike the corresponding property of nonces, we cannot prove - \<^term>\finite KK ==> \K. K \ KK \ Key K \ used evs\. + \<^term>\finite KK \ \K. K \ KK \ Key K \ 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.\ @@ -184,7 +184,7 @@ (*Lemma for the trivial direction of the if-and-only-if*) lemma analz_image_freshK_lemma: - "(Key K \ analz (Key`nE \ H)) \ (K \ nE | Key K \ analz H) ==> + "(Key K \ analz (Key`nE \ H)) \ (K \ nE | Key K \ analz H) \ (Key K \ analz (Key`nE \ H)) = (K \ nE | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Thu Oct 13 17:19:56 2022 +0100 @@ -118,7 +118,7 @@ by auto lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}" -apply (unfold keysFor_def) +unfolding keysFor_def apply (induct_tac "C", auto) done diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/TLS.thy --- a/src/HOL/Auth/TLS.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/TLS.thy Thu Oct 13 17:19:56 2022 +0100 @@ -102,14 +102,14 @@ | Fake: \ \The Spy may say anything he can say. The sender field is correct, but agents don't use that information.\ - "[| evsf \ tls; X \ synth (analz (spies evsf)) |] - ==> Says Spy B X # evsf \ tls" + "\evsf \ tls; X \ synth (analz (spies evsf))\ + \ Says Spy B X # evsf \ tls" | SpyKeys: \ \The spy may apply \<^term>\PRF\ and \<^term>\sessionK\ to available nonces\ - "[| evsSK \ tls; - {Nonce NA, Nonce NB, Nonce M} \ analz (spies evsSK) |] - ==> Notes Spy \ Nonce (PRF(M,NA,NB)), + "\evsSK \ tls; + {Nonce NA, Nonce NB, Nonce M} \ analz (spies evsSK)\ + \ Notes Spy \ Nonce (PRF(M,NA,NB)), Key (sessionK((NA,NB,M),role))\ # evsSK \ tls" | ClientHello: @@ -120,8 +120,8 @@ UNIX TIME is omitted because the protocol doesn't use it. May assume \<^term>\NA \ range PRF\ because CLIENT RANDOM is 28 bytes while MASTER SECRET is 48 bytes\ - "[| evsCH \ tls; Nonce NA \ used evsCH; NA \ range PRF |] - ==> Says A B \Agent A, Nonce NA, Number SID, Number PA\ + "\evsCH \ tls; Nonce NA \ used evsCH; NA \ range PRF\ + \ Says A B \Agent A, Nonce NA, Number SID, Number PA\ # evsCH \ tls" | ServerHello: @@ -129,14 +129,14 @@ PB represents \CLIENT_VERSION\, \CIPHER_SUITE\ and \COMPRESSION_METHOD\. SERVER CERTIFICATE (7.4.2) is always present. \CERTIFICATE_REQUEST\ (7.4.4) is implied.\ - "[| evsSH \ tls; Nonce NB \ used evsSH; NB \ range PRF; + "\evsSH \ tls; Nonce NB \ used evsSH; NB \ range PRF; Says A' B \Agent A, Nonce NA, Number SID, Number PA\ - \ set evsSH |] - ==> Says B A \Nonce NB, Number SID, Number PB\ # evsSH \ tls" + \ set evsSH\ + \ Says B A \Nonce NB, Number SID, Number PB\ # evsSH \ tls" | Certificate: \ \SERVER (7.4.2) or CLIENT (7.4.6) CERTIFICATE.\ - "evsC \ tls ==> Says B A (certificate B (pubK B)) # evsC \ tls" + "evsC \ tls \ Says B A (certificate B (pubK B)) # evsC \ tls" | ClientKeyExch: \ \CLIENT KEY EXCHANGE (7.4.7). @@ -147,9 +147,9 @@ both items have the same length, 48 bytes). The Note event records in the trace that she knows PMS (see REMARK at top).\ - "[| evsCX \ tls; Nonce PMS \ used evsCX; PMS \ range PRF; - Says B' A (certificate B KB) \ set evsCX |] - ==> Says A B (Crypt KB (Nonce PMS)) + "\evsCX \ tls; Nonce PMS \ used evsCX; PMS \ range PRF; + Says B' A (certificate B KB) \ set evsCX\ + \ Says A B (Crypt KB (Nonce PMS)) # Notes A \Agent B, Nonce PMS\ # evsCX \ tls" @@ -159,10 +159,10 @@ It adds the pre-master-secret, which is also essential! Checking the signature, which is the only use of A's certificate, assures B of A's presence\ - "[| evsCV \ tls; + "\evsCV \ tls; Says B' A \Nonce NB, Number SID, Number PB\ \ set evsCV; - Notes A \Agent B, Nonce PMS\ \ set evsCV |] - ==> Says A B (Crypt (priK A) (Hash\Nonce NB, Agent B, Nonce PMS\)) + Notes A \Agent B, Nonce PMS\ \ set evsCV\ + \ Says A B (Crypt (priK A) (Hash\Nonce NB, Agent B, Nonce PMS\)) # evsCV \ tls" \ \Finally come the FINISHED messages (7.4.8), confirming PA and PB @@ -176,13 +176,13 @@ Spy does not know PMS and could not send ClientFinished. One could simply put \<^term>\A\Spy\ into the rule, but one should not expect the spy to be well-behaved.\ - "[| evsCF \ tls; + "\evsCF \ tls; Says A B \Agent A, Nonce NA, Number SID, Number PA\ \ set evsCF; Says B' A \Nonce NB, Number SID, Number PB\ \ set evsCF; Notes A \Agent B, Nonce PMS\ \ set evsCF; - M = PRF(PMS,NA,NB) |] - ==> Says A B (Crypt (clientK(NA,NB,M)) + M = PRF(PMS,NA,NB)\ + \ Says A B (Crypt (clientK(NA,NB,M)) (Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\)) @@ -191,13 +191,13 @@ | ServerFinished: \ \Keeping A' and A'' distinct means B cannot even check that the two messages originate from the same source.\ - "[| evsSF \ tls; + "\evsSF \ tls; Says A' B \Agent A, Nonce NA, Number SID, Number PA\ \ set evsSF; Says B A \Nonce NB, Number SID, Number PB\ \ set evsSF; Says A'' B (Crypt (pubK B) (Nonce PMS)) \ set evsSF; - M = PRF(PMS,NA,NB) |] - ==> Says B A (Crypt (serverK(NA,NB,M)) + M = PRF(PMS,NA,NB)\ + \ Says B A (Crypt (serverK(NA,NB,M)) (Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\)) @@ -208,15 +208,15 @@ message encrypted with serverK, the client stores the parameters needed to resume this session. The "Notes A ..." premise is used to prove \Notes_master_imp_Crypt_PMS\.\ - "[| evsCA \ tls; + "\evsCA \ tls; Notes A \Agent B, Nonce PMS\ \ set evsCA; M = PRF(PMS,NA,NB); X = Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\; Says A B (Crypt (clientK(NA,NB,M)) X) \ set evsCA; - Says B' A (Crypt (serverK(NA,NB,M)) X) \ set evsCA |] - ==> + Says B' A (Crypt (serverK(NA,NB,M)) X) \ set evsCA\ + \ Notes A \Number SID, Agent A, Agent B, Nonce M\ # evsCA \ tls" | ServerAccepts: @@ -224,7 +224,7 @@ message encrypted with clientK, the server stores the parameters needed to resume this session. The "Says A'' B ..." premise is used to prove \Notes_master_imp_Crypt_PMS\.\ - "[| evsSA \ tls; + "\evsSA \ tls; A \ B; Says A'' B (Crypt (pubK B) (Nonce PMS)) \ set evsSA; M = PRF(PMS,NA,NB); @@ -232,18 +232,18 @@ Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\; Says B A (Crypt (serverK(NA,NB,M)) X) \ set evsSA; - Says A' B (Crypt (clientK(NA,NB,M)) X) \ set evsSA |] - ==> + Says A' B (Crypt (clientK(NA,NB,M)) X) \ set evsSA\ + \ Notes B \Number SID, Agent A, Agent B, Nonce M\ # evsSA \ tls" | ClientResume: \ \If A recalls the \SESSION_ID\, then she sends a FINISHED message using the new nonces and stored MASTER SECRET.\ - "[| evsCR \ tls; + "\evsCR \ tls; Says A B \Agent A, Nonce NA, Number SID, Number PA\ \ set evsCR; Says B' A \Nonce NB, Number SID, Number PB\ \ set evsCR; - Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evsCR |] - ==> Says A B (Crypt (clientK(NA,NB,M)) + Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evsCR\ + \ Says A B (Crypt (clientK(NA,NB,M)) (Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\)) @@ -252,11 +252,11 @@ | ServerResume: \ \Resumption (7.3): If B finds the \SESSION_ID\ then he can send a FINISHED message using the recovered MASTER SECRET\ - "[| evsSR \ tls; + "\evsSR \ tls; Says A' B \Agent A, Nonce NA, Number SID, Number PA\ \ set evsSR; Says B A \Nonce NB, Number SID, Number PB\ \ set evsSR; - Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evsSR |] - ==> Says B A (Crypt (serverK(NA,NB,M)) + Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evsSR\ + \ Says B A (Crypt (serverK(NA,NB,M)) (Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\)) # evsSR @@ -268,9 +268,9 @@ rather unlikely. The assumption \<^term>\A\Spy\ is essential: otherwise the Spy could learn session keys merely by replaying messages!\ - "[| evso \ tls; A \ Spy; - Says A B (Crypt (sessionK((NA,NB,M),role)) X) \ set evso |] - ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso \ tls" + "\evso \ tls; A \ Spy; + Says A B (Crypt (sessionK((NA,NB,M),role)) X) \ set evso\ + \ Says A Spy (Key (sessionK((NA,NB,M),role))) # evso \ tls" (* Protocol goals: @@ -331,8 +331,8 @@ text\Possibility property ending with ClientAccepts.\ -lemma "[| \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] - ==> \SID M. \evs \ tls. +lemma "\\evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B\ + \ \SID M. \evs \ tls. Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evs" apply (intro exI bexI) apply (rule_tac [2] tls.Nil @@ -344,8 +344,8 @@ text\And one for ServerAccepts. Either FINISHED message may come first.\ -lemma "[| \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] - ==> \SID NA PA NB PB M. \evs \ tls. +lemma "\\evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B\ + \ \SID NA PA NB PB M. \evs \ tls. Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evs" apply (intro exI bexI) apply (rule_tac [2] tls.Nil @@ -357,8 +357,8 @@ text\Another one, for CertVerify (which is optional)\ -lemma "[| \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] - ==> \NB PMS. \evs \ tls. +lemma "\\evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B\ + \ \NB PMS. \evs \ tls. Says A B (Crypt (priK A) (Hash\Nonce NB, Agent B, Nonce PMS\)) \ set evs" apply (intro exI bexI) @@ -371,12 +371,12 @@ text\Another one, for session resumption (both ServerResume and ClientResume). NO tls.Nil here: we refer to a previous session, not the empty trace.\ -lemma "[| evs0 \ tls; +lemma "\evs0 \ tls; Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evs0; Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evs0; \evs. (SOME N. Nonce N \ used evs) \ range PRF; - A \ B |] - ==> \NA PA NB PB X. \evs \ tls. + A \ B\ + \ \NA PA NB PB X. \evs \ tls. X = Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, Nonce NB, Number PB, Agent B\ \ @@ -397,15 +397,15 @@ text\Spy never sees a good agent's private key!\ lemma Spy_see_priK [simp]: - "evs \ tls ==> (Key (privateKey b A) \ parts (spies evs)) = (A \ bad)" + "evs \ tls \ (Key (privateKey b A) \ parts (spies evs)) = (A \ bad)" by (erule tls.induct, force, simp_all, blast) lemma Spy_analz_priK [simp]: - "evs \ tls ==> (Key (privateKey b A) \ analz (spies evs)) = (A \ bad)" + "evs \ tls \ (Key (privateKey b A) \ analz (spies evs)) = (A \ bad)" by auto lemma Spy_see_priK_D [dest!]: - "[|Key (privateKey b A) \ parts (knows Spy evs); evs \ tls|] ==> A \ bad" + "\Key (privateKey b A) \ parts (knows Spy evs); evs \ tls\ \ A \ bad" by (blast dest: Spy_see_priK) @@ -414,7 +414,7 @@ little point in doing so: the loss of their private keys is a worse breach of security.\ lemma certificate_valid: - "[| certificate B KB \ parts (spies evs); evs \ tls |] ==> KB = pubK B" + "\certificate B KB \ parts (spies evs); evs \ tls\ \ KB = pubK B" apply (erule rev_mp) apply (erule tls.induct, force, simp_all, blast) done @@ -425,8 +425,8 @@ subsubsection\Properties of items found in Notes\ lemma Notes_Crypt_parts_spies: - "[| Notes A \Agent B, X\ \ set evs; evs \ tls |] - ==> Crypt (pubK B) X \ parts (spies evs)" + "\Notes A \Agent B, X\ \ set evs; evs \ tls\ + \ Crypt (pubK B) X \ parts (spies evs)" apply (erule rev_mp) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB, force, simp_all) @@ -435,9 +435,9 @@ text\C may be either A or B\ lemma Notes_master_imp_Crypt_PMS: - "[| Notes C \s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\ \ set evs; - evs \ tls |] - ==> Crypt (pubK B) (Nonce PMS) \ parts (spies evs)" + "\Notes C \s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\ \ set evs; + evs \ tls\ + \ Crypt (pubK B) (Nonce PMS) \ parts (spies evs)" apply (erule rev_mp) apply (erule tls.induct, force, simp_all) txt\Fake\ @@ -448,9 +448,9 @@ text\Compared with the theorem above, both premise and conclusion are stronger\ lemma Notes_master_imp_Notes_PMS: - "[| Notes A \s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\ \ set evs; - evs \ tls |] - ==> Notes A \Agent B, Nonce PMS\ \ set evs" + "\Notes A \s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\ \ set evs; + evs \ tls\ + \ Notes A \Agent B, Nonce PMS\ \ set evs" apply (erule rev_mp) apply (erule tls.induct, force, simp_all) txt\ServerAccepts\ @@ -462,45 +462,45 @@ text\B can check A's signature if he has received A's certificate.\ lemma TrustCertVerify_lemma: - "[| X \ parts (spies evs); + "\X \ parts (spies evs); X = Crypt (priK A) (Hash\nb, Agent B, pms\); - evs \ tls; A \ bad |] - ==> Says A B X \ set evs" + evs \ tls; A \ bad\ + \ Says A B X \ set evs" apply (erule rev_mp, erule ssubst) apply (erule tls.induct, force, simp_all, blast) done text\Final version: B checks X using the distributed KA instead of priK A\ lemma TrustCertVerify: - "[| X \ parts (spies evs); + "\X \ parts (spies evs); X = Crypt (invKey KA) (Hash\nb, Agent B, pms\); certificate A KA \ parts (spies evs); - evs \ tls; A \ bad |] - ==> Says A B X \ set evs" + evs \ tls; A \ bad\ + \ Says A B X \ set evs" by (blast dest!: certificate_valid intro!: TrustCertVerify_lemma) text\If CertVerify is present then A has chosen PMS.\ lemma UseCertVerify_lemma: - "[| Crypt (priK A) (Hash\nb, Agent B, Nonce PMS\) \ parts (spies evs); - evs \ tls; A \ bad |] - ==> Notes A \Agent B, Nonce PMS\ \ set evs" + "\Crypt (priK A) (Hash\nb, Agent B, Nonce PMS\) \ parts (spies evs); + evs \ tls; A \ bad\ + \ Notes A \Agent B, Nonce PMS\ \ set evs" apply (erule rev_mp) apply (erule tls.induct, force, simp_all, blast) done text\Final version using the distributed KA instead of priK A\ lemma UseCertVerify: - "[| Crypt (invKey KA) (Hash\nb, Agent B, Nonce PMS\) + "\Crypt (invKey KA) (Hash\nb, Agent B, Nonce PMS\) \ parts (spies evs); certificate A KA \ parts (spies evs); - evs \ tls; A \ bad |] - ==> Notes A \Agent B, Nonce PMS\ \ set evs" + evs \ tls; A \ bad\ + \ Notes A \Agent B, Nonce PMS\ \ set evs" by (blast dest!: certificate_valid intro!: UseCertVerify_lemma) lemma no_Notes_A_PRF [simp]: - "evs \ tls ==> Notes A \Agent B, Nonce (PRF x)\ \ set evs" + "evs \ tls \ Notes A \Agent B, Nonce (PRF x)\ \ set evs" apply (erule tls.induct, force, simp_all) txt\ClientKeyExch: PMS is assumed to differ from any PRF.\ apply blast @@ -508,8 +508,8 @@ lemma MS_imp_PMS [dest!]: - "[| Nonce (PRF (PMS,NA,NB)) \ parts (spies evs); evs \ tls |] - ==> Nonce PMS \ parts (spies evs)" + "\Nonce (PRF (PMS,NA,NB)) \ parts (spies evs); evs \ tls\ + \ Nonce PMS \ parts (spies evs)" apply (erule rev_mp) apply (erule tls.induct, force, simp_all) txt\Fake\ @@ -525,11 +525,11 @@ text\PMS determines B.\ lemma Crypt_unique_PMS: - "[| Crypt(pubK B) (Nonce PMS) \ parts (spies evs); + "\Crypt(pubK B) (Nonce PMS) \ parts (spies evs); Crypt(pubK B') (Nonce PMS) \ parts (spies evs); Nonce PMS \ analz (spies evs); - evs \ tls |] - ==> B=B'" + evs \ tls\ + \ B=B'" apply (erule rev_mp, erule rev_mp, erule rev_mp) apply (erule tls.induct, analz_mono_contra, force, simp_all (no_asm_simp)) txt\Fake, ClientKeyExch\ @@ -545,10 +545,10 @@ text\In A's internal Note, PMS determines A and B.\ lemma Notes_unique_PMS: - "[| Notes A \Agent B, Nonce PMS\ \ set evs; + "\Notes A \Agent B, Nonce PMS\ \ set evs; Notes A' \Agent B', Nonce PMS\ \ set evs; - evs \ tls |] - ==> A=A' \ B=B'" + evs \ tls\ + \ A=A' \ B=B'" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, force, simp_all) txt\ClientKeyExch\ @@ -562,7 +562,7 @@ No collection of keys can help the spy get new private keys.\ lemma analz_image_priK [rule_format]: "evs \ tls - ==> \KK. (Key(priK B) \ analz (Key`KK \ (spies evs))) = + \ \KK. (Key(priK B) \ analz (Key`KK \ (spies evs))) = (priK B \ KK | B \ bad)" apply (erule tls.induct) apply (simp_all (no_asm_simp) @@ -582,7 +582,7 @@ text\Lemma for the trivial direction of the if-and-only-if\ lemma analz_image_keys_lemma: - "(X \ analz (G \ H)) \ (X \ analz H) ==> + "(X \ analz (G \ H)) \ (X \ analz H) \ (X \ analz (G \ H)) = (X \ analz H)" by (blast intro: analz_mono [THEN subsetD]) @@ -592,7 +592,7 @@ **) lemma analz_image_keys [rule_format]: - "evs \ tls ==> + "evs \ tls \ \KK. KK \ range sessionK \ (Nonce N \ analz (Key`KK \ (spies evs))) = (Nonce N \ analz (spies evs))" @@ -611,7 +611,7 @@ text\Knowing some session keys is no help in getting new nonces\ lemma analz_insert_key [simp]: - "evs \ tls ==> + "evs \ tls \ (Nonce N \ analz (insert (Key (sessionK z)) (spies evs))) = (Nonce N \ analz (spies evs))" by (simp del: image_insert @@ -628,10 +628,10 @@ Converse doesn't hold; revealing PMS doesn't force the keys to be sent. THEY ARE NOT SUITABLE AS SAFE ELIM RULES.\ lemma PMS_lemma: - "[| Nonce PMS \ parts (spies evs); + "\Nonce PMS \ parts (spies evs); K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role); - evs \ tls |] - ==> Key K \ parts (spies evs) \ (\Y. Crypt K Y \ parts (spies evs))" + evs \ tls\ + \ Key K \ parts (spies evs) \ (\Y. Crypt K Y \ parts (spies evs))" apply (erule rev_mp, erule ssubst) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) @@ -644,15 +644,15 @@ done lemma PMS_sessionK_not_spied: - "[| Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) \ parts (spies evs); - evs \ tls |] - ==> Nonce PMS \ parts (spies evs)" + "\Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) \ parts (spies evs); + evs \ tls\ + \ Nonce PMS \ parts (spies evs)" by (blast dest: PMS_lemma) lemma PMS_Crypt_sessionK_not_spied: - "[| Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y - \ parts (spies evs); evs \ tls |] - ==> Nonce PMS \ parts (spies evs)" + "\Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y + \ parts (spies evs); evs \ tls\ + \ Nonce PMS \ parts (spies evs)" by (blast dest: PMS_lemma) text\Write keys are never sent if M (MASTER SECRET) is secure. @@ -661,9 +661,9 @@ with some effort. NO LONGER USED: see \clientK_not_spied\ and \serverK_not_spied\\ lemma sessionK_not_spied: - "[| \A. Says A Spy (Key (sessionK((NA,NB,M),role))) \ set evs; - Nonce M \ analz (spies evs); evs \ tls |] - ==> Key (sessionK((NA,NB,M),role)) \ parts (spies evs)" + "\\A. Says A Spy (Key (sessionK((NA,NB,M),role))) \ set evs; + Nonce M \ analz (spies evs); evs \ tls\ + \ Key (sessionK((NA,NB,M),role)) \ parts (spies evs)" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, analz_mono_contra) apply (force, simp_all (no_asm_simp)) @@ -674,9 +674,9 @@ text\If A sends ClientKeyExch to an honest B, then the PMS will stay secret.\ lemma Spy_not_see_PMS: - "[| Notes A \Agent B, Nonce PMS\ \ set evs; - evs \ tls; A \ bad; B \ bad |] - ==> Nonce PMS \ analz (spies evs)" + "\Notes A \Agent B, Nonce PMS\ \ set evs; + evs \ tls; A \ bad; B \ bad\ + \ Nonce PMS \ analz (spies evs)" apply (erule rev_mp, erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) txt\Fake\ @@ -696,9 +696,9 @@ text\If A sends ClientKeyExch to an honest B, then the MASTER SECRET will stay secret.\ lemma Spy_not_see_MS: - "[| Notes A \Agent B, Nonce PMS\ \ set evs; - evs \ tls; A \ bad; B \ bad |] - ==> Nonce (PRF(PMS,NA,NB)) \ analz (spies evs)" + "\Notes A \Agent B, Nonce PMS\ \ set evs; + evs \ tls; A \ bad; B \ bad\ + \ Nonce (PRF(PMS,NA,NB)) \ analz (spies evs)" apply (erule rev_mp, erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) txt\Fake\ @@ -719,10 +719,10 @@ text\If A created PMS then nobody else (except the Spy in replays) would send a message using a clientK generated from that PMS.\ lemma Says_clientK_unique: - "[| Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) \ set evs; + "\Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) \ set evs; Notes A \Agent B, Nonce PMS\ \ set evs; - evs \ tls; A' \ Spy |] - ==> A = A'" + evs \ tls; A' \ Spy\ + \ A = A'" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all) @@ -737,11 +737,11 @@ text\If A created PMS and has not leaked her clientK to the Spy, then it is completely secure: not even in parts!\ lemma clientK_not_spied: - "[| Notes A \Agent B, Nonce PMS\ \ set evs; + "\Notes A \Agent B, Nonce PMS\ \ set evs; Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) \ set evs; A \ bad; B \ bad; - evs \ tls |] - ==> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) \ parts (spies evs)" + evs \ tls\ + \ Key (clientK(Na,Nb,PRF(PMS,NA,NB))) \ parts (spies evs)" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) @@ -761,10 +761,10 @@ text\If A created PMS for B, then nobody other than B or the Spy would send a message using a serverK generated from that PMS.\ lemma Says_serverK_unique: - "[| Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) \ set evs; + "\Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) \ set evs; Notes A \Agent B, Nonce PMS\ \ set evs; - evs \ tls; A \ bad; B \ bad; B' \ Spy |] - ==> B = B'" + evs \ tls; A \ bad; B \ bad; B' \ Spy\ + \ B = B'" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all) @@ -779,10 +779,10 @@ text\If A created PMS for B, and B has not leaked his serverK to the Spy, then it is completely secure: not even in parts!\ lemma serverK_not_spied: - "[| Notes A \Agent B, Nonce PMS\ \ set evs; + "\Notes A \Agent B, Nonce PMS\ \ set evs; Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) \ set evs; - A \ bad; B \ bad; evs \ tls |] - ==> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) \ parts (spies evs)" + A \ bad; B \ bad; evs \ tls\ + \ Key (serverK(Na,Nb,PRF(PMS,NA,NB))) \ parts (spies evs)" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) @@ -803,13 +803,13 @@ text\The mention of her name (A) in X assures A that B knows who she is.\ lemma TrustServerFinished [rule_format]: - "[| X = Crypt (serverK(Na,Nb,M)) + "\X = Crypt (serverK(Na,Nb,M)) (Hash\Number SID, Nonce M, Nonce Na, Number PA, Agent A, Nonce Nb, Number PB, Agent B\); M = PRF(PMS,NA,NB); - evs \ tls; A \ bad; B \ bad |] - ==> Says B Spy (Key(serverK(Na,Nb,M))) \ set evs \ + evs \ tls; A \ bad; B \ bad\ + \ Says B Spy (Key(serverK(Na,Nb,M))) \ set evs \ Notes A \Agent B, Nonce PMS\ \ set evs \ X \ parts (spies evs) \ Says B A X \ set evs" apply (erule ssubst)+ @@ -827,8 +827,8 @@ that B sends his message to A. If CLIENT KEY EXCHANGE were augmented to bind A's identity with PMS, then we could replace A' by A below.\ lemma TrustServerMsg [rule_format]: - "[| M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad |] - ==> Says B Spy (Key(serverK(Na,Nb,M))) \ set evs \ + "\M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad\ + \ Says B Spy (Key(serverK(Na,Nb,M))) \ set evs \ Notes A \Agent B, Nonce PMS\ \ set evs \ Crypt (serverK(Na,Nb,M)) Y \ parts (spies evs) \ (\A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) \ set evs)" @@ -853,8 +853,8 @@ ClientFinished, then B can then check the quoted values PA, PB, etc.\ lemma TrustClientMsg [rule_format]: - "[| M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad |] - ==> Says A Spy (Key(clientK(Na,Nb,M))) \ set evs \ + "\M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad\ + \ Says A Spy (Key(clientK(Na,Nb,M))) \ set evs \ Notes A \Agent B, Nonce PMS\ \ set evs \ Crypt (clientK(Na,Nb,M)) Y \ parts (spies evs) \ Says A B (Crypt (clientK(Na,Nb,M)) Y) \ set evs" @@ -874,14 +874,14 @@ check a CertVerify from A, then A has used the quoted values PA, PB, etc. Even this one requires A to be uncompromised.\ lemma AuthClientFinished: - "[| M = PRF(PMS,NA,NB); + "\M = PRF(PMS,NA,NB); Says A Spy (Key(clientK(Na,Nb,M))) \ set evs; Says A' B (Crypt (clientK(Na,Nb,M)) Y) \ set evs; certificate A KA \ parts (spies evs); Says A'' B (Crypt (invKey KA) (Hash\nb, Agent B, Nonce PMS\)) \ set evs; - evs \ tls; A \ bad; B \ bad |] - ==> Says A B (Crypt (clientK(Na,Nb,M)) Y) \ set evs" + evs \ tls; A \ bad; B \ bad\ + \ Says A B (Crypt (clientK(Na,Nb,M)) Y) \ set evs" by (blast intro!: TrustClientMsg UseCertVerify) (*22/9/97: loads in 622s, which is 10 minutes 22 seconds*) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/WooLam.thy --- a/src/HOL/Auth/WooLam.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/WooLam.thy Thu Oct 13 17:19:56 2022 +0100 @@ -27,38 +27,38 @@ (*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 \ woolam; X \ synth (analz (spies evsf)) |] - ==> Says Spy B X # evsf \ woolam" + | Fake: "\evsf \ woolam; X \ synth (analz (spies evsf))\ + \ Says Spy B X # evsf \ woolam" (*Alice initiates a protocol run*) - | WL1: "evs1 \ woolam ==> Says A B (Agent A) # evs1 \ woolam" + | WL1: "evs1 \ woolam \ Says A B (Agent A) # evs1 \ woolam" (*Bob responds to Alice's message with a challenge.*) - | WL2: "[| evs2 \ woolam; Says A' B (Agent A) \ set evs2 |] - ==> Says B A (Nonce NB) # evs2 \ woolam" + | WL2: "\evs2 \ woolam; Says A' B (Agent A) \ set evs2\ + \ Says B A (Nonce NB) # evs2 \ woolam" (*Alice responds to Bob's challenge by encrypting NB with her key. B is *not* properly determined -- Alice essentially broadcasts her reply.*) - | WL3: "[| evs3 \ woolam; + | WL3: "\evs3 \ woolam; Says A B (Agent A) \ set evs3; - Says B' A (Nonce NB) \ set evs3 |] - ==> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \ woolam" + Says B' A (Nonce NB) \ set evs3\ + \ Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \ woolam" (*Bob forwards Alice's response to the Server. NOTE: usually the messages are shown in chronological order, for clarity. But here, exchanging the two events would cause the lemma WL4_analz_spies to pick up the wrong assumption!*) - | WL4: "[| evs4 \ woolam; + | WL4: "\evs4 \ woolam; Says A' B X \ set evs4; - Says A'' B (Agent A) \ set evs4 |] - ==> Says B Server \Agent A, Agent B, X\ # evs4 \ woolam" + Says A'' B (Agent A) \ set evs4\ + \ Says B Server \Agent A, Agent B, X\ # evs4 \ woolam" (*Server decrypts Alice's response for Bob.*) - | WL5: "[| evs5 \ woolam; + | WL5: "\evs5 \ woolam; Says B' Server \Agent A, Agent B, Crypt (shrK A) (Nonce NB)\ - \ set evs5 |] - ==> Says Server B (Crypt (shrK B) \Agent A, Nonce NB\) + \ set evs5\ + \ Says Server B (Crypt (shrK B) \Agent A, Nonce NB\) # evs5 \ woolam" @@ -86,15 +86,15 @@ (*Spy never sees a good agent's shared key!*) lemma Spy_see_shrK [simp]: - "evs \ woolam ==> (Key (shrK A) \ parts (spies evs)) = (A \ bad)" + "evs \ woolam \ (Key (shrK A) \ parts (spies evs)) = (A \ bad)" by (erule woolam.induct, force, simp_all, blast+) lemma Spy_analz_shrK [simp]: - "evs \ woolam ==> (Key (shrK A) \ analz (spies evs)) = (A \ bad)" + "evs \ woolam \ (Key (shrK A) \ analz (spies evs)) = (A \ bad)" by auto lemma Spy_see_shrK_D [dest!]: - "[|Key (shrK A) \ parts (knows Spy evs); evs \ woolam|] ==> A \ bad" + "\Key (shrK A) \ parts (knows Spy evs); evs \ woolam\ \ A \ bad" by (blast dest: Spy_see_shrK) @@ -104,19 +104,19 @@ (*If the encrypted message appears then it originated with Alice*) lemma NB_Crypt_imp_Alice_msg: - "[| Crypt (shrK A) (Nonce NB) \ parts (spies evs); - A \ bad; evs \ woolam |] - ==> \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" + "\Crypt (shrK A) (Nonce NB) \ parts (spies evs); + A \ bad; evs \ woolam\ + \ \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) (*Guarantee for Server: if it gets a message containing a certificate from Alice, then she originated that certificate. But we DO NOT know that B ever saw it: the Spy may have rerouted the message to the Server.*) lemma Server_trusts_WL4 [dest]: - "[| Says B' Server \Agent A, Agent B, Crypt (shrK A) (Nonce NB)\ + "\Says B' Server \Agent A, Agent B, Crypt (shrK A) (Nonce NB)\ \ set evs; - A \ bad; evs \ woolam |] - ==> \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" + A \ bad; evs \ woolam\ + \ \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" by (blast intro!: NB_Crypt_imp_Alice_msg) @@ -124,17 +124,17 @@ (*Server sent WL5 only if it received the right sort of message*) lemma Server_sent_WL5 [dest]: - "[| Says Server B (Crypt (shrK B) \Agent A, NB\) \ set evs; - evs \ woolam |] - ==> \B'. Says B' Server \Agent A, Agent B, Crypt (shrK A) NB\ + "\Says Server B (Crypt (shrK B) \Agent A, NB\) \ set evs; + evs \ woolam\ + \ \B'. Says B' Server \Agent A, Agent B, Crypt (shrK A) NB\ \ set evs" by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) (*If the encrypted message appears then it originated with the Server!*) lemma NB_Crypt_imp_Server_msg [rule_format]: - "[| Crypt (shrK B) \Agent A, NB\ \ parts (spies evs); - B \ bad; evs \ woolam |] - ==> Says Server B (Crypt (shrK B) \Agent A, NB\) \ set evs" + "\Crypt (shrK B) \Agent A, NB\ \ parts (spies evs); + B \ bad; evs \ woolam\ + \ Says Server B (Crypt (shrK B) \Agent A, NB\) \ set evs" by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) (*Guarantee for B. If B gets the Server's certificate then A has encrypted @@ -142,22 +142,22 @@ But A may have sent the nonce to some other agent and it could have reached the Server via the Spy.*) lemma B_trusts_WL5: - "[| Says S B (Crypt (shrK B) \Agent A, Nonce NB\) \ set evs; - A \ bad; B \ bad; evs \ woolam |] - ==> \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" + "\Says S B (Crypt (shrK B) \Agent A, Nonce NB\) \ set evs; + A \ bad; B \ bad; evs \ woolam\ + \ \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" by (blast dest!: NB_Crypt_imp_Server_msg) (*B only issues challenges in response to WL1. Not used.*) lemma B_said_WL2: - "[| Says B A (Nonce NB) \ set evs; B \ Spy; evs \ woolam |] - ==> \A'. Says A' B (Agent A) \ set evs" + "\Says B A (Nonce NB) \ set evs; B \ Spy; evs \ woolam\ + \ \A'. Says A' B (Agent A) \ set evs" by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) (**CANNOT be proved because A doesn't know where challenges come from...*) -lemma "[| A \ bad; B \ Spy; evs \ woolam |] - ==> Crypt (shrK A) (Nonce NB) \ parts (spies evs) \ +lemma "\A \ bad; B \ Spy; evs \ woolam\ + \ Crypt (shrK A) (Nonce NB) \ parts (spies evs) \ Says B A (Nonce NB) \ set evs \ Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" apply (erule rev_mp, erule woolam.induct, force, simp_all, blast, auto) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Yahalom.thy --- a/src/HOL/Auth/Yahalom.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Yahalom.thy Thu Oct 13 17:19:56 2022 +0100 @@ -332,7 +332,7 @@ "Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB\, X\ \ set evs \ KeyWithNonce K NB evs" -by (unfold KeyWithNonce_def, blast) + unfolding KeyWithNonce_def by blast lemma KeyWithNonce_Says [simp]: "KeyWithNonce K NB (Says S A X # evs) = @@ -354,7 +354,7 @@ (with respect to a given trace).\ lemma fresh_not_KeyWithNonce: "Key K \ used evs \ \ KeyWithNonce K NB evs" -by (unfold KeyWithNonce_def, blast) + unfolding KeyWithNonce_def by blast text\The Server message associates K with NB' and therefore not with any other nonce NB.\ @@ -363,7 +363,7 @@ \ set evs; NB \ NB'; evs \ yahalom\ \ \ KeyWithNonce K NB evs" -by (unfold KeyWithNonce_def, blast dest: unique_session_keys) + unfolding KeyWithNonce_def by (blast dest: unique_session_keys) text\The only nonces that can be found with the help of session keys are diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/Yahalom_Bad.thy --- a/src/HOL/Auth/Yahalom_Bad.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/Yahalom_Bad.thy Thu Oct 13 17:19:56 2022 +0100 @@ -21,32 +21,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, 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 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, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ - \ 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" @@ -54,11 +54,11 @@ (*Alice receives the Server's (?) message, checks her Nonce, and uses the new session key to send Bob his Nonce. The premise A \ Server is needed to prove Says_Server_not_range.*) - | YM4: "[| evs4 \ yahalom; A \ Server; K \ symKeys; + | YM4: "\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" declare Says_imp_knows_Spy [THEN analz.Inj, dest] @@ -68,8 +68,8 @@ text\A "possibility property": there are traces that reach the end\ -lemma "[| A \ Server; Key K \ used []; K \ symKeys |] - ==> \X NB. \evs \ yahalom. +lemma "\A \ Server; Key K \ used []; K \ symKeys\ + \ \X NB. \evs \ yahalom. Says A B \X, Crypt K (Nonce NB)\ \ set evs" apply (intro exI bexI) apply (rule_tac [2] yahalom.Nil @@ -83,12 +83,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) (*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] @@ -98,8 +98,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 = @@ -111,24 +111,24 @@ 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)" apply (erule yahalom.induct, force, drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+) done 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) @@ -142,7 +142,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. @@ -151,7 +151,7 @@ subsection\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 \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" @@ -159,7 +159,7 @@ drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast) 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) @@ -167,12 +167,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\ @@ -182,8 +182,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 \ @@ -195,12 +195,12 @@ 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; - 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) @@ -208,9 +208,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" @@ -224,9 +224,9 @@ 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); - A \ bad; B \ bad; evs \ yahalom |] - ==> Key K \ analz (knows Spy evs)" + "\Crypt (shrK A) \Agent B, Key K, na, nb\ \ parts (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) subsection\Security Guarantees for B upon receiving YM4\ @@ -234,9 +234,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\\ \ set evs" @@ -259,8 +259,8 @@ Secrecy of K is assumed; the valid Yahalom proof uses (and later proves) the secrecy of NB.\ lemma B_trusts_YM4_newK [rule_format]: - "[|Key K \ analz (knows Spy evs); evs \ yahalom|] - ==> Crypt K (Nonce NB) \ parts (knows Spy evs) \ + "\Key K \ analz (knows Spy evs); evs \ yahalom\ + \ Crypt K (Nonce NB) \ parts (knows Spy evs) \ (\A B NA. Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, @@ -285,13 +285,13 @@ text\B's session key guarantee from YM4. The two certificates contribute to a single conclusion about the Server's message.\ 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, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ \ set evs; - A \ bad; B \ bad; evs \ yahalom |] - ==> \na nb. Says Server A + A \ bad; B \ bad; evs \ yahalom\ + \ \na nb. Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, Crypt (shrK B) \Agent A, Key K\\ \ set evs" @@ -302,13 +302,13 @@ 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, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ \ 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) @@ -323,7 +323,7 @@ NB matters for freshness.\ lemma 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 \ @@ -349,13 +349,13 @@ 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\, + "\Gets B \Crypt (shrK B) \Agent A, Key K\, Crypt K (Nonce NB)\ \ set evs; Says B Server \Agent B, Nonce NB, Crypt (shrK B) \Agent A, Nonce NA\\ \ 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 (blast intro!: A_Said_YM3_lemma dest: Spy_not_see_encrypted_key B_trusts_YM4 Gets_imp_Says) diff -r 71bf371a9784 -r 2317e9a19239 src/HOL/Auth/ZhouGollmann.thy --- a/src/HOL/Auth/ZhouGollmann.thy Thu Oct 13 17:00:43 2022 +0200 +++ b/src/HOL/Auth/ZhouGollmann.thy Thu Oct 13 17:19:56 2022 +0100 @@ -32,35 +32,35 @@ Nil: "[] \ zg" -| Fake: "[| evsf \ zg; X \ synth (analz (spies evsf)) |] - ==> Says Spy B X # evsf \ zg" +| Fake: "\evsf \ zg; X \ synth (analz (spies evsf))\ + \ Says Spy B X # evsf \ zg" -| Reception: "[| evsr \ zg; Says A B X \ set evsr |] ==> Gets B X # evsr \ zg" +| Reception: "\evsr \ zg; Says A B X \ set evsr\ \ Gets B X # evsr \ zg" (*L is fresh for honest agents. We don't require K to be fresh because we don't bother to prove secrecy! We just assume that the protocol's objective is to deliver K fairly, rather than to keep M secret.*) -| ZG1: "[| evs1 \ zg; Nonce L \ used evs1; C = Crypt K (Number m); +| ZG1: "\evs1 \ zg; Nonce L \ used evs1; C = Crypt K (Number m); K \ symKeys; - NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\|] - ==> Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ # evs1 \ zg" + NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\\ + \ Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ # evs1 \ zg" (*B must check that NRO is A's signature to learn the sender's name*) -| ZG2: "[| evs2 \ zg; +| ZG2: "\evs2 \ zg; Gets B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs2; NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\; - NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\|] - ==> Says B A \Number f_nrr, Agent A, Nonce L, NRR\ # evs2 \ zg" + NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\\ + \ Says B A \Number f_nrr, Agent A, Nonce L, NRR\ # evs2 \ zg" (*A must check that NRR is B's signature to learn the sender's name; without spy, the matching label would be enough*) -| ZG3: "[| evs3 \ zg; C = Crypt K M; K \ symKeys; +| ZG3: "\evs3 \ zg; C = Crypt K M; K \ symKeys; Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs3; Gets A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs3; NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\; - sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\|] - ==> Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ + sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\\ + \ Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ # evs3 \ zg" (*TTP checks that sub_K is A's signature to learn who issued K, then @@ -69,13 +69,13 @@ give con_K to the Spy. This makes the threat model more dangerous, while also allowing lemma @{text Crypt_used_imp_spies} to omit the condition @{term "K \ priK TTP"}. *) -| ZG4: "[| evs4 \ zg; K \ symKeys; +| ZG4: "\evs4 \ zg; K \ symKeys; Gets TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs4; sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; con_K = Crypt (priK TTP) \Number f_con, Agent A, Agent B, - Nonce L, Key K\|] - ==> Says TTP Spy con_K + Nonce L, Key K\\ + \ Says TTP Spy con_K # Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ # evs4 \ zg" @@ -90,7 +90,7 @@ text\A "possibility property": there are traces that reach the end\ -lemma "[|A \ B; TTP \ A; TTP \ B; K \ symKeys|] ==> +lemma "\A \ B; TTP \ A; TTP \ B; K \ symKeys\ \ \L. \evs \ zg. Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, Crypt (priK TTP) \Number f_con, Agent A, Agent B, Nonce L, Key K\\ @@ -107,45 +107,45 @@ subsection \Basic Lemmas\ lemma Gets_imp_Says: - "[| Gets B X \ set evs; evs \ zg |] ==> \A. Says A B X \ set evs" + "\Gets B X \ set evs; evs \ zg\ \ \A. Says A B X \ set evs" apply (erule rev_mp) apply (erule zg.induct, auto) done lemma Gets_imp_knows_Spy: - "[| Gets B X \ set evs; evs \ zg |] ==> X \ spies evs" + "\Gets B X \ set evs; evs \ zg\ \ X \ spies evs" by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) text\Lets us replace proofs about \<^term>\used evs\ by simpler proofs about \<^term>\parts (spies evs)\.\ lemma Crypt_used_imp_spies: - "[| Crypt K X \ used evs; evs \ zg |] - ==> Crypt K X \ parts (spies evs)" + "\Crypt K X \ used evs; evs \ zg\ + \ Crypt K X \ parts (spies evs)" apply (erule rev_mp) apply (erule zg.induct) apply (simp_all add: parts_insert_knows_A) done lemma Notes_TTP_imp_Gets: - "[|Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ + "\Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ \ set evs; sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; - evs \ zg|] - ==> Gets TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" + evs \ zg\ + \ Gets TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" apply (erule rev_mp) apply (erule zg.induct, auto) done text\For reasoning about C, which is encrypted in message ZG2\ lemma ZG2_msg_in_parts_spies: - "[|Gets B \F, B', L, C, X\ \ set evs; evs \ zg|] - ==> C \ parts (spies evs)" + "\Gets B \F, B', L, C, X\ \ set evs; evs \ zg\ + \ C \ parts (spies evs)" by (blast dest: Gets_imp_Says) (*classical regularity lemma on priK*) lemma Spy_see_priK [simp]: - "evs \ zg ==> (Key (priK A) \ parts (spies evs)) = (A \ bad)" + "evs \ zg \ (Key (priK A) \ parts (spies evs)) = (A \ bad)" apply (erule zg.induct) apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) done @@ -154,7 +154,7 @@ declare Spy_see_priK [THEN [2] rev_iffD1, dest!] lemma Spy_analz_priK [simp]: - "evs \ zg ==> (Key (priK A) \ analz (spies evs)) = (A \ bad)" + "evs \ zg \ (Key (priK A) \ analz (spies evs)) = (A \ bad)" by auto @@ -165,10 +165,10 @@ text\Strong conclusion for a good agent\ lemma NRO_validity_good: - "[|NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\; + "\NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\; NRO \ parts (spies evs); - A \ bad; evs \ zg |] - ==> Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs" + A \ bad; evs \ zg\ + \ Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs" apply clarify apply (erule rev_mp) apply (erule zg.induct) @@ -176,18 +176,18 @@ done lemma NRO_sender: - "[|Says A' B \n, b, l, C, Crypt (priK A) X\ \ set evs; evs \ zg|] - ==> A' \ {A,Spy}" + "\Says A' B \n, b, l, C, Crypt (priK A) X\ \ set evs; evs \ zg\ + \ A' \ {A,Spy}" apply (erule rev_mp) apply (erule zg.induct, simp_all) done text\Holds also for \<^term>\A = Spy\!\ theorem NRO_validity: - "[|Gets B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs; + "\Gets B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs; NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\; - A \ broken; evs \ zg |] - ==> Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs" + A \ broken; evs \ zg\ + \ Says A B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs" apply (drule Gets_imp_Says, assumption) apply clarify apply (frule NRO_sender, auto) @@ -205,10 +205,10 @@ text\Strong conclusion for a good agent\ lemma NRR_validity_good: - "[|NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\; + "\NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\; NRR \ parts (spies evs); - B \ bad; evs \ zg |] - ==> Says B A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" + B \ bad; evs \ zg\ + \ Says B A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" apply clarify apply (erule rev_mp) apply (erule zg.induct) @@ -216,18 +216,18 @@ done lemma NRR_sender: - "[|Says B' A \n, a, l, Crypt (priK B) X\ \ set evs; evs \ zg|] - ==> B' \ {B,Spy}" + "\Says B' A \n, a, l, Crypt (priK B) X\ \ set evs; evs \ zg\ + \ B' \ {B,Spy}" apply (erule rev_mp) apply (erule zg.induct, simp_all) done text\Holds also for \<^term>\B = Spy\!\ theorem NRR_validity: - "[|Says B' A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs; + "\Says B' A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs; NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\; - B \ broken; evs \ zg|] - ==> Says B A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" + B \ broken; evs \ zg\ + \ Says B A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" apply clarify apply (frule NRR_sender, auto) txt\We are left with the case where \<^term>\B' = Spy\ and \<^term>\B' \ B\, @@ -243,10 +243,10 @@ text\Strong conclusion for a good agent\ lemma sub_K_validity_good: - "[|sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; + "\sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; sub_K \ parts (spies evs); - A \ bad; evs \ zg |] - ==> Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" + A \ bad; evs \ zg\ + \ Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" apply clarify apply (erule rev_mp) apply (erule zg.induct) @@ -256,18 +256,18 @@ done lemma sub_K_sender: - "[|Says A' TTP \n, b, l, k, Crypt (priK A) X\ \ set evs; evs \ zg|] - ==> A' \ {A,Spy}" + "\Says A' TTP \n, b, l, k, Crypt (priK A) X\ \ set evs; evs \ zg\ + \ A' \ {A,Spy}" apply (erule rev_mp) apply (erule zg.induct, simp_all) done text\Holds also for \<^term>\A = Spy\!\ theorem sub_K_validity: - "[|Gets TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs; + "\Gets TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs; sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; - A \ broken; evs \ zg |] - ==> Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" + A \ broken; evs \ zg\ + \ Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" apply (drule Gets_imp_Says, assumption) apply clarify apply (frule sub_K_sender, auto) @@ -286,11 +286,11 @@ that \<^term>\A\ sent \<^term>\sub_K\\ lemma con_K_validity: - "[|con_K \ used evs; + "\con_K \ used evs; con_K = Crypt (priK TTP) \Number f_con, Agent A, Agent B, Nonce L, Key K\; - evs \ zg |] - ==> Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ + evs \ zg\ + \ Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ \ set evs" apply clarify apply (erule rev_mp) @@ -306,11 +306,11 @@ \<^term>\sub_K\. We assume that \<^term>\A\ is not broken. Importantly, nothing needs to be assumed about the form of \<^term>\con_K\!\ lemma Notes_TTP_imp_Says_A: - "[|Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ + "\Notes TTP \Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\ \ set evs; sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; - A \ broken; evs \ zg|] - ==> Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" + A \ broken; evs \ zg\ + \ Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" apply clarify apply (erule rev_mp) apply (erule zg.induct) @@ -323,12 +323,12 @@ text\If \<^term>\con_K\ exists, then \<^term>\A\ sent \<^term>\sub_K\. We again assume that \<^term>\A\ is not broken.\ theorem B_sub_K_validity: - "[|con_K \ used evs; + "\con_K \ used evs; con_K = Crypt (priK TTP) \Number f_con, Agent A, Agent B, Nonce L, Key K\; sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; - A \ broken; evs \ zg|] - ==> Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" + A \ broken; evs \ zg\ + \ Says A TTP \Number f_sub, Agent B, Nonce L, Key K, sub_K\ \ set evs" by (blast dest: con_K_validity Notes_TTP_imp_Says_A) @@ -340,12 +340,12 @@ text\Strange: unicity of the label protects \<^term>\A\?\ lemma A_unicity: - "[|NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, Crypt K M\; + "\NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, Crypt K M\; NRO \ parts (spies evs); Says A B \Number f_nro, Agent B, Nonce L, Crypt K M', NRO'\ \ set evs; - A \ bad; evs \ zg |] - ==> M'=M" + A \ bad; evs \ zg\ + \ M'=M" apply clarify apply (erule rev_mp) apply (erule rev_mp) @@ -359,13 +359,13 @@ text\Fairness lemma: if \<^term>\sub_K\ exists, then \<^term>\A\ holds NRR. Relies on unicity of labels.\ lemma sub_K_implies_NRR: - "[| NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, Crypt K M\; + "\NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, Crypt K M\; NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, Crypt K M\; sub_K \ parts (spies evs); NRO \ parts (spies evs); sub_K = Crypt (priK A) \Number f_sub, Agent B, Nonce L, Key K\; - A \ bad; evs \ zg |] - ==> Gets A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" + A \ bad; evs \ zg\ + \ Gets A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" apply clarify apply hypsubst_thin apply (erule rev_mp) @@ -382,8 +382,8 @@ lemma Crypt_used_imp_L_used: - "[| Crypt (priK TTP) \F, A, B, L, K\ \ used evs; evs \ zg |] - ==> L \ used evs" + "\Crypt (priK TTP) \F, A, B, L, K\ \ used evs; evs \ zg\ + \ L \ used evs" apply (erule rev_mp) apply (erule zg.induct, auto) txt\Fake\ @@ -397,14 +397,14 @@ then \<^term>\A\ holds NRR. \<^term>\A\ must be uncompromised, but there is no assumption about \<^term>\B\.\ theorem A_fairness_NRO: - "[|con_K \ used evs; + "\con_K \ used evs; NRO \ parts (spies evs); con_K = Crypt (priK TTP) \Number f_con, Agent A, Agent B, Nonce L, Key K\; NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, Crypt K M\; NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, Crypt K M\; - A \ bad; evs \ zg |] - ==> Gets A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" + A \ bad; evs \ zg\ + \ Gets A \Number f_nrr, Agent A, Nonce L, NRR\ \ set evs" apply clarify apply (erule rev_mp) apply (erule rev_mp) @@ -425,11 +425,11 @@ text\Fairness for \<^term>\B\: NRR exists at all, then \<^term>\B\ holds NRO. \<^term>\B\ must be uncompromised, but there is no assumption about \<^term>\A\.\ theorem B_fairness_NRR: - "[|NRR \ used evs; + "\NRR \ used evs; NRR = Crypt (priK B) \Number f_nrr, Agent A, Nonce L, C\; NRO = Crypt (priK A) \Number f_nro, Agent B, Nonce L, C\; - B \ bad; evs \ zg |] - ==> Gets B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs" + B \ bad; evs \ zg\ + \ Gets B \Number f_nro, Agent B, Nonce L, C, NRO\ \ set evs" apply clarify apply (erule rev_mp) apply (erule zg.induct)