merged
authorpaulson
Fri, 14 Oct 2022 10:30:37 +0100
changeset 76293 f3e30ad850ba
parent 76286 a00c80314b06 (current diff)
parent 76292 2317e9a19239 (diff)
child 76296 eb30869e7228
child 76297 e7f9e5b3a36a
merged
--- a/src/HOL/Auth/CertifiedEmail.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/CertifiedEmail.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -36,33 +36,33 @@
 
 | Fake: \<comment> \<open>The Spy may say anything he can say.  The sender field is correct,
           but agents don't use that information.\<close>
-      "[| evsf \<in> certified_mail; X \<in> synth(analz(spies evsf))|] 
-       ==> Says Spy B X # evsf \<in> certified_mail"
+      "\<lbrakk>evsf \<in> certified_mail; X \<in> synth(analz(spies evsf))\<rbrakk> 
+       \<Longrightarrow> Says Spy B X # evsf \<in> certified_mail"
 
 | FakeSSL: \<comment> \<open>The Spy may open SSL sessions with TTP, who is the only agent
     equipped with the necessary credentials to serve as an SSL server.\<close>
-         "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
-          ==> Notes TTP \<lbrace>Agent Spy, Agent TTP, X\<rbrace> # evsfssl \<in> certified_mail"
+         "\<lbrakk>evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))\<rbrakk>
+          \<Longrightarrow> Notes TTP \<lbrace>Agent Spy, Agent TTP, X\<rbrace> # evsfssl \<in> certified_mail"
 
 | CM1: \<comment> \<open>The sender approaches the recipient.  The message is a number.\<close>
- "[|evs1 \<in> certified_mail;
+ "\<lbrakk>evs1 \<in> certified_mail;
     Key K \<notin> used evs1;
     K \<in> symKeys;
     Nonce q \<notin> used evs1;
     hs = Hash\<lbrace>Number cleartext, Nonce q, response S R q, Crypt K (Number m)\<rbrace>;
-    S2TTP = Crypt(pubEK TTP) \<lbrace>Agent S, Number BothAuth, Key K, Agent R, hs\<rbrace>|]
-  ==> Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, 
+    S2TTP = Crypt(pubEK TTP) \<lbrace>Agent S, Number BothAuth, Key K, Agent R, hs\<rbrace>\<rbrakk>
+  \<Longrightarrow> Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, 
                  Number cleartext, Nonce q, S2TTP\<rbrace> # evs1 
         \<in> certified_mail"
 
 | CM2: \<comment> \<open>The recipient records \<^term>\<open>S2TTP\<close> while transmitting it and her
      password to \<^term>\<open>TTP\<close> over an SSL channel.\<close>
- "[|evs2 \<in> certified_mail;
+ "\<lbrakk>evs2 \<in> certified_mail;
     Gets R \<lbrace>Agent S, Agent TTP, em, Number BothAuth, Number cleartext, 
              Nonce q, S2TTP\<rbrace> \<in> set evs2;
     TTP \<noteq> R;  
-    hr = Hash \<lbrace>Number cleartext, Nonce q, response S R q, em\<rbrace> |]
-  ==> 
+    hr = Hash \<lbrace>Number cleartext, Nonce q, response S R q, em\<rbrace>\<rbrakk>
+  \<Longrightarrow> 
    Notes TTP \<lbrace>Agent R, Agent TTP, S2TTP, Key(RPwd R), hr\<rbrace> # evs2
       \<in> certified_mail"
 
@@ -71,26 +71,26 @@
          the client (\<^term>\<open>R\<close>), but \<^term>\<open>TTP\<close> accepts the message only 
          if the given password is that of the claimed sender, \<^term>\<open>R\<close>.
          He replies over the established SSL channel.\<close>
- "[|evs3 \<in> certified_mail;
+ "\<lbrakk>evs3 \<in> certified_mail;
     Notes TTP \<lbrace>Agent R, Agent TTP, S2TTP, Key(RPwd R), hr\<rbrace> \<in> set evs3;
     S2TTP = Crypt (pubEK TTP) 
                      \<lbrace>Agent S, Number BothAuth, Key k, Agent R, hs\<rbrace>;
-    TTP \<noteq> R;  hs = hr;  k \<in> symKeys|]
-  ==> 
+    TTP \<noteq> R;  hs = hr;  k \<in> symKeys\<rbrakk>
+  \<Longrightarrow> 
    Notes R \<lbrace>Agent TTP, Agent R, Key k, hr\<rbrace> # 
    Gets S (Crypt (priSK TTP) S2TTP) # 
    Says TTP S (Crypt (priSK TTP) S2TTP) # evs3 \<in> certified_mail"
 
 | Reception:
- "[|evsr \<in> certified_mail; Says A B X \<in> set evsr|]
-  ==> Gets B X#evsr \<in> certified_mail"
+ "\<lbrakk>evsr \<in> certified_mail; Says A B X \<in> set evsr\<rbrakk>
+  \<Longrightarrow> Gets B X#evsr \<in> 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 \<notin> used []; K \<in> symKeys |] ==> 
+lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk> \<Longrightarrow> 
        \<exists>S2TTP. \<exists>evs \<in> certified_mail.
            Says TTP S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (intro exI bexI)
@@ -103,23 +103,23 @@
 
 
 lemma Gets_imp_Says:
- "[| Gets B X \<in> set evs; evs \<in> certified_mail |] ==> \<exists>A. Says A B X \<in> set evs"
+ "\<lbrakk>Gets B X \<in> set evs; evs \<in> certified_mail\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
 apply (erule rev_mp)
 apply (erule certified_mail.induct, auto)
 done
 
 
 lemma Gets_imp_parts_knows_Spy:
-     "[|Gets A X \<in> set evs; evs \<in> certified_mail|] ==> X \<in> parts(spies evs)"
+     "\<lbrakk>Gets A X \<in> set evs; evs \<in> certified_mail\<rbrakk> \<Longrightarrow> X \<in> 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 \<lbrace>Agent A, Agent B, em, Number AO, Number cleartext, 
+ "\<lbrakk>Gets R \<lbrace>Agent A, Agent B, em, Number AO, Number cleartext, 
               Nonce q, S2TTP\<rbrace> \<in> set evs;
-    evs \<in> certified_mail|] 
-  ==> S2TTP \<in> analz(spies evs)"
+    evs \<in> certified_mail\<rbrakk> 
+  \<Longrightarrow> S2TTP \<in> 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>\<open>R\<noteq>Spy\<close>.\<close>
 lemma hr_form:
- "[|Notes TTP \<lbrace>Agent R, Agent TTP, S2TTP, pwd, hr\<rbrace> \<in> set evs;
-    evs \<in> certified_mail|]
-  ==> hr \<in> synth (analz (spies evs)) | 
+ "\<lbrakk>Notes TTP \<lbrace>Agent R, Agent TTP, S2TTP, pwd, hr\<rbrace> \<in> set evs;
+    evs \<in> certified_mail\<rbrakk>
+  \<Longrightarrow> hr \<in> synth (analz (spies evs)) | 
       (\<exists>clt q S em. hr = Hash \<lbrace>Number clt, Nonce q, response S R q, em\<rbrace>)"
 by (blast intro: hr_form_lemma) 
 
 lemma Spy_dont_know_private_keys [dest!]:
-    "[|Key (privateKey b A) \<in> parts (spies evs); evs \<in> certified_mail|]
-     ==> A \<in> bad"
+    "\<lbrakk>Key (privateKey b A) \<in> parts (spies evs); evs \<in> certified_mail\<rbrakk>
+     \<Longrightarrow> A \<in> bad"
 apply (erule rev_mp) 
 apply (erule certified_mail.induct, simp_all)
 txt\<open>Fake\<close>
@@ -166,15 +166,15 @@
 
 lemma Spy_know_private_keys_iff [simp]:
     "evs \<in> certified_mail
-     ==> (Key (privateKey b A) \<in> parts (spies evs)) = (A \<in> bad)"
+     \<Longrightarrow> (Key (privateKey b A) \<in> parts (spies evs)) = (A \<in> bad)"
 by blast 
 
 lemma Spy_dont_know_TTPKey_parts [simp]:
-     "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> parts(spies evs)" 
+     "evs \<in> certified_mail \<Longrightarrow> Key (privateKey b TTP) \<notin> parts(spies evs)" 
 by simp
 
 lemma Spy_dont_know_TTPKey_analz [simp]:
-     "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> analz(spies evs)"
+     "evs \<in> certified_mail \<Longrightarrow> Key (privateKey b TTP) \<notin> analz(spies evs)"
 by auto
 
 text\<open>Thus, prove any goal that assumes that \<^term>\<open>Spy\<close> knows a private key
@@ -183,11 +183,11 @@
 
 
 lemma CM3_k_parts_knows_Spy:
- "[| evs \<in> certified_mail;
+ "\<lbrakk>evs \<in> certified_mail;
      Notes TTP \<lbrace>Agent A, Agent TTP,
                  Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, 
-                 Agent R, hs\<rbrace>, Key (RPwd R), hs\<rbrace> \<in> set evs|]
-  ==> Key K \<in> parts(spies evs)"
+                 Agent R, hs\<rbrace>, Key (RPwd R), hs\<rbrace> \<in> set evs\<rbrakk>
+  \<Longrightarrow> Key K \<in> 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 \<in> certified_mail ==> Key (RPwd A) \<in> parts(spies evs) \<longrightarrow> A \<in> bad"
+    "evs \<in> certified_mail \<Longrightarrow> Key (RPwd A) \<in> parts(spies evs) \<longrightarrow> A \<in> bad"
 apply (erule certified_mail.induct, simp_all) 
 txt\<open>Fake\<close>
 apply (blast dest: Fake_parts_insert_in_Un) 
@@ -218,17 +218,17 @@
 
 
 lemma Spy_know_RPwd_iff [simp]:
-    "evs \<in> certified_mail ==> (Key (RPwd A) \<in> parts(spies evs)) = (A\<in>bad)"
+    "evs \<in> certified_mail \<Longrightarrow> (Key (RPwd A) \<in> parts(spies evs)) = (A\<in>bad)"
 by (auto simp add: Spy_dont_know_RPwd) 
 
 lemma Spy_analz_RPwd_iff [simp]:
-    "evs \<in> certified_mail ==> (Key (RPwd A) \<in> analz(spies evs)) = (A\<in>bad)"
+    "evs \<in> certified_mail \<Longrightarrow> (Key (RPwd A) \<in> analz(spies evs)) = (A\<in>bad)"
 by (metis Spy_know_RPwd_iff Spy_spies_bad_shrK analz.Inj analz_into_parts)
 
 text\<open>Unused, but a guarantee of sorts\<close>
 theorem CertAutenticity:
-     "[|Crypt (priSK TTP) X \<in> parts (spies evs); evs \<in> certified_mail|] 
-      ==> \<exists>A. Says TTP A (Crypt (priSK TTP) X) \<in> set evs"
+     "\<lbrakk>Crypt (priSK TTP) X \<in> parts (spies evs); evs \<in> certified_mail\<rbrakk> 
+      \<Longrightarrow> \<exists>A. Says TTP A (Crypt (priSK TTP) X) \<in> set evs"
 apply (erule rev_mp)
 apply (erule certified_mail.induct, simp_all) 
 txt\<open>Fake\<close>
@@ -246,7 +246,7 @@
 subsection\<open>Proving Confidentiality Results\<close>
 
 lemma analz_image_freshK [rule_format]:
- "evs \<in> certified_mail ==>
+ "evs \<in> certified_mail \<Longrightarrow>
    \<forall>K KK. invKey (pubEK TTP) \<notin> KK \<longrightarrow>
           (Key K \<in> analz (Key`KK \<union> (spies evs))) =
           (K \<in> KK | Key K \<in> analz (spies evs))"
@@ -264,7 +264,7 @@
 
 
 lemma analz_insert_freshK:
-  "[| evs \<in> certified_mail;  KAB \<noteq> invKey (pubEK TTP) |] ==>
+  "\<lbrakk>evs \<in> certified_mail;  KAB \<noteq> invKey (pubEK TTP)\<rbrakk> \<Longrightarrow>
       (Key K \<in> analz (insert (Key KAB) (spies evs))) =
       (K = KAB | Key K \<in> analz (spies evs))"
 by (simp only: analz_image_freshK analz_image_freshK_simps)
@@ -273,14 +273,14 @@
     provided \<^term>\<open>K\<close> is secure.  Proof is surprisingly hard.\<close>
 
 lemma Notes_SSL_imp_used:
-     "[|Notes B \<lbrace>Agent A, Agent B, X\<rbrace> \<in> set evs|] ==> X \<in> used evs"
+     "\<lbrakk>Notes B \<lbrace>Agent A, Agent B, X\<rbrace> \<in> set evs\<rbrakk> \<Longrightarrow> X \<in> 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 \<in> certified_mail ==>
+ "evs \<in> certified_mail \<Longrightarrow>
     Key K \<notin> analz (spies evs) \<longrightarrow>
     (\<forall>AO. Crypt (pubEK TTP)
            \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace> \<in> used evs \<longrightarrow>
@@ -310,10 +310,10 @@
 done 
 
 lemma S2TTP_sender:
- "[|Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace> \<in> used evs;
+ "\<lbrakk>Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace> \<in> used evs;
     Key K \<notin> analz (spies evs);
-    evs \<in> certified_mail|]
-  ==> \<exists>m ctxt q. 
+    evs \<in> certified_mail\<rbrakk>
+  \<Longrightarrow> \<exists>m ctxt q. 
         hs = Hash\<lbrace>Number ctxt, Nonce q, response S R q, Crypt K (Number m)\<rbrace> \<and>
         Says S R
            \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO,
@@ -325,8 +325,8 @@
 
 text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
-    "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> certified_mail|]
-     ==> K \<notin> keysFor (parts (spies evs))"
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> certified_mail\<rbrakk>
+     \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
 apply (erule rev_mp) 
 apply (erule certified_mail.induct, simp_all) 
 txt\<open>Fake\<close>
@@ -344,7 +344,7 @@
 theorem for ciphertexts of the form \<^term>\<open>Crypt K (Number m)\<close>, 
 where \<^term>\<open>K\<close> is secure.\<close>
 lemma Key_unique_lemma [rule_format]:
-     "evs \<in> certified_mail ==>
+     "evs \<in> certified_mail \<Longrightarrow>
        Key K \<notin> analz (spies evs) \<longrightarrow>
        (\<forall>m cleartext q hs.
         Says S R
@@ -368,7 +368,7 @@
 
 text\<open>The key determines the sender, recipient and protocol options.\<close>
 lemma Key_unique:
-      "[|Says S R
+      "\<lbrakk>Says S R
            \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO,
              Number cleartext, Nonce q,
              Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace>\<rbrace>
@@ -379,8 +379,8 @@
              Crypt (pubEK TTP) \<lbrace>Agent S', Number AO', Key K, Agent R', hs'\<rbrace>\<rbrace>
           \<in> set evs;
          Key K \<notin> analz (spies evs);
-         evs \<in> certified_mail|]
-       ==> R' = R \<and> S' = S \<and> AO' = AO \<and> hs' = hs"
+         evs \<in> certified_mail\<rbrakk>
+       \<Longrightarrow> R' = R \<and> S' = S \<and> AO' = AO \<and> hs' = hs"
 by (rule Key_unique_lemma, assumption+)
 
 
@@ -390,13 +390,13 @@
       If Spy gets the key then \<^term>\<open>R\<close> is bad and \<^term>\<open>S\<close> moreover
       gets his return receipt (and therefore has no grounds for complaint).\<close>
 theorem S_fairness_bad_R:
-      "[|Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO, 
+      "\<lbrakk>Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                      Number cleartext, Nonce q, S2TTP\<rbrace> \<in> set evs;
          S2TTP = Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace>;
          Key K \<in> analz (spies evs);
          evs \<in> certified_mail;
-         S\<noteq>Spy|]
-      ==> R \<in> bad \<and> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
+         S\<noteq>Spy\<rbrakk>
+      \<Longrightarrow> R \<in> bad \<and> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (erule rev_mp)
 apply (erule ssubst)
 apply (erule rev_mp)
@@ -418,24 +418,24 @@
 
 text\<open>Confidentially for the symmetric key\<close>
 theorem Spy_not_see_encrypted_key:
-      "[|Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO, 
+      "\<lbrakk>Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                      Number cleartext, Nonce q, S2TTP\<rbrace> \<in> set evs;
          S2TTP = Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace>;
          evs \<in> certified_mail;
-         S\<noteq>Spy; R \<notin> bad|]
-      ==> Key K \<notin> analz(spies evs)"
+         S\<noteq>Spy; R \<notin> bad\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz(spies evs)"
 by (blast dest: S_fairness_bad_R) 
 
 
 text\<open>Agent \<^term>\<open>R\<close>, who may be the Spy, doesn't receive the key
  until \<^term>\<open>S\<close> has access to the return receipt.\<close> 
 theorem S_guarantee:
-     "[|Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO, 
+     "\<lbrakk>Says S R \<lbrace>Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                     Number cleartext, Nonce q, S2TTP\<rbrace> \<in> set evs;
         S2TTP = Crypt (pubEK TTP) \<lbrace>Agent S, Number AO, Key K, Agent R, hs\<rbrace>;
         Notes R \<lbrace>Agent TTP, Agent R, Key K, hs\<rbrace> \<in> set evs;
-        S\<noteq>Spy;  evs \<in> certified_mail|]
-     ==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
+        S\<noteq>Spy;  evs \<in> certified_mail\<rbrakk>
+     \<Longrightarrow> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (erule rev_mp)
 apply (erule ssubst)
 apply (erule rev_mp)
@@ -451,13 +451,13 @@
  then \<^term>\<open>R\<close> receives the necessary key.  This result is also important
  to \<^term>\<open>S\<close>, as it confirms the validity of the return receipt.\<close>
 theorem RR_validity:
-  "[|Crypt (priSK TTP) S2TTP \<in> used evs;
+  "\<lbrakk>Crypt (priSK TTP) S2TTP \<in> used evs;
      S2TTP = Crypt (pubEK TTP)
                \<lbrace>Agent S, Number AO, Key K, Agent R, 
                  Hash \<lbrace>Number cleartext, Nonce q, r, em\<rbrace>\<rbrace>;
      hr = Hash \<lbrace>Number cleartext, Nonce q, r, em\<rbrace>;
-     R\<noteq>Spy;  evs \<in> certified_mail|]
-  ==> Notes R \<lbrace>Agent TTP, Agent R, Key K, hr\<rbrace> \<in> set evs"
+     R\<noteq>Spy;  evs \<in> certified_mail\<rbrakk>
+  \<Longrightarrow> Notes R \<lbrace>Agent TTP, Agent R, Key K, hr\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule ssubst)
 apply (erule ssubst)
--- a/src/HOL/Auth/Event.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Event.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -186,7 +186,7 @@
 text\<open>What agents DIFFERENT FROM Spy know 
   was either said, or noted, or got, or known initially\<close>
 lemma knows_imp_Says_Gets_Notes_initState [rule_format]:
-     "[| X \<in> knows A evs; A \<noteq> Spy |] ==> \<exists> B.
+     "\<lbrakk>X \<in> knows A evs; A \<noteq> Spy\<rbrakk> \<Longrightarrow> \<exists> B.
   Says A B X \<in> set evs \<or> Gets A X \<in> set evs \<or> Notes A X \<in> set evs \<or> X \<in> initState A"
 apply (erule rev_mp)
 apply (induct_tac "evs")
@@ -258,8 +258,8 @@
 
 text\<open>For proving \<open>new_keys_not_used\<close>\<close>
 lemma keysFor_parts_insert:
-     "[| K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) |] 
-      ==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"
+     "\<lbrakk>K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H)\<rbrakk> 
+      \<Longrightarrow> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"
 by (force 
     dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
            analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD]
--- a/src/HOL/Auth/Guard/Analz.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Analz.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -16,9 +16,9 @@
   pparts :: "msg set => msg set"
   for H :: "msg set"
 where
-  Inj [intro]: "[| X \<in> H; is_MPair X |] ==> X \<in> pparts H"
-| Fst [dest]: "[| \<lbrace>X,Y\<rbrace> \<in> pparts H; is_MPair X |] ==> X \<in> pparts H"
-| Snd [dest]: "[| \<lbrace>X,Y\<rbrace> \<in> pparts H; is_MPair Y |] ==> Y \<in> pparts H"
+  Inj [intro]: "\<lbrakk>X \<in> H; is_MPair X\<rbrakk> \<Longrightarrow> X \<in> pparts H"
+| Fst [dest]: "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> pparts H; is_MPair X\<rbrakk> \<Longrightarrow> X \<in> pparts H"
+| Snd [dest]: "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> pparts H; is_MPair Y\<rbrakk> \<Longrightarrow> Y \<in> pparts H"
 
 subsection\<open>basic facts about \<^term>\<open>pparts\<close>\<close>
 
@@ -46,7 +46,7 @@
 lemma pparts_insertI [intro]: "X \<in> pparts H \<Longrightarrow> X \<in> pparts (insert Y H)"
 by (erule pparts.induct, auto)
 
-lemma pparts_sub: "[| X \<in> pparts G; G \<subseteq> H |] ==> X \<in> pparts H"
+lemma pparts_sub: "\<lbrakk>X \<in> pparts G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> pparts H"
 by (erule pparts.induct, auto)
 
 lemma pparts_insert2 [iff]: "pparts (insert X (insert Y H))
@@ -100,8 +100,8 @@
 
 subsection\<open>facts about \<^term>\<open>pparts\<close> and \<^term>\<open>parts\<close>\<close>
 
-lemma pparts_no_Nonce [dest]: "[| X \<in> pparts {Y}; Nonce n \<notin> parts {Y} |]
-==> Nonce n \<notin> parts {X}"
+lemma pparts_no_Nonce [dest]: "\<lbrakk>X \<in> pparts {Y}; Nonce n \<notin> parts {Y}\<rbrakk>
+\<Longrightarrow> Nonce n \<notin> parts {X}"
 by (erule pparts.induct, simp_all)
 
 subsection\<open>facts about \<^term>\<open>pparts\<close> and \<^term>\<open>analz\<close>\<close>
@@ -109,7 +109,7 @@
 lemma pparts_analz: "X \<in> pparts H \<Longrightarrow> X \<in> analz H"
 by (erule pparts.induct, auto)
 
-lemma pparts_analz_sub: "[| X \<in> pparts G; G \<subseteq> H |] ==> X \<in> analz H"
+lemma pparts_analz_sub: "\<lbrakk>X \<in> pparts G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 by (auto dest: pparts_sub pparts_analz)
 
 subsection\<open>messages that contribute to analz\<close>
@@ -118,9 +118,9 @@
   kparts :: "msg set => msg set"
   for H :: "msg set"
 where
-  Inj [intro]: "[| X \<in> H; not_MPair X |] ==> X \<in> kparts H"
-| Fst [intro]: "[| \<lbrace>X,Y\<rbrace> \<in> pparts H; not_MPair X |] ==> X \<in> kparts H"
-| Snd [intro]: "[| \<lbrace>X,Y\<rbrace> \<in> pparts H; not_MPair Y |] ==> Y \<in> kparts H"
+  Inj [intro]: "\<lbrakk>X \<in> H; not_MPair X\<rbrakk> \<Longrightarrow> X \<in> kparts H"
+| Fst [intro]: "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> pparts H; not_MPair X\<rbrakk> \<Longrightarrow> X \<in> kparts H"
+| Snd [intro]: "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> pparts H; not_MPair Y\<rbrakk> \<Longrightarrow> Y \<in> kparts H"
 
 subsection\<open>basic facts about \<^term>\<open>kparts\<close>\<close>
 
@@ -172,7 +172,7 @@
 X \<notin> kparts H \<longrightarrow> X \<in> kparts {Z}"
 by (erule kparts.induct, (blast dest: pparts_insert)+)
 
-lemma kparts_sub: "[| X \<in> kparts G; G \<subseteq> H |] ==> X \<in> kparts H"
+lemma kparts_sub: "\<lbrakk>X \<in> kparts G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> kparts H"
 by (erule kparts.induct, auto dest: pparts_sub)
 
 lemma kparts_Un [iff]: "kparts (G \<union> H) = kparts G \<union> kparts H"
@@ -197,8 +197,8 @@
 
 subsection\<open>facts about \<^term>\<open>kparts\<close> and \<^term>\<open>parts\<close>\<close>
 
-lemma kparts_no_Nonce [dest]: "[| X \<in> kparts {Y}; Nonce n \<notin> parts {Y} |]
-==> Nonce n \<notin> parts {X}"
+lemma kparts_no_Nonce [dest]: "\<lbrakk>X \<in> kparts {Y}; Nonce n \<notin> parts {Y}\<rbrakk>
+\<Longrightarrow> Nonce n \<notin> parts {X}"
 by (erule kparts.induct, auto)
 
 lemma kparts_parts: "X \<in> kparts H \<Longrightarrow> X \<in> 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 \<in> kparts {Z};
-Nonce n \<in> parts {Y} |] ==> Nonce n \<in> parts {Z}"
+lemma Crypt_kparts_Nonce_parts [dest]: "\<lbrakk>Crypt K Y \<in> kparts {Z};
+Nonce n \<in> parts {Y}\<rbrakk> \<Longrightarrow> Nonce n \<in> parts {Z}"
 by auto
 
 subsection\<open>facts about \<^term>\<open>kparts\<close> and \<^term>\<open>analz\<close>\<close>
@@ -217,7 +217,7 @@
 lemma kparts_analz: "X \<in> kparts H \<Longrightarrow> X \<in> analz H"
 by (erule kparts.induct, auto dest: pparts_analz)
 
-lemma kparts_analz_sub: "[| X \<in> kparts G; G \<subseteq> H |] ==> X \<in> analz H"
+lemma kparts_analz_sub: "\<lbrakk>X \<in> kparts G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 by (erule kparts.induct, auto dest: pparts_analz_sub)
 
 lemma analz_kparts [rule_format,dest]: "X \<in> analz H \<Longrightarrow>
@@ -234,16 +234,16 @@
 \<Longrightarrow> Nonce n \<in> kparts {Y} \<longrightarrow> Nonce n \<in> analz G"
 by (erule synth.induct, auto)
 
-lemma kparts_insert_synth: "[| Y \<in> parts (insert X G); X \<in> synth (analz G);
-Nonce n \<in> kparts {Y}; Nonce n \<notin> analz G |] ==> Y \<in> parts G"
+lemma kparts_insert_synth: "\<lbrakk>Y \<in> parts (insert X G); X \<in> synth (analz G);
+Nonce n \<in> kparts {Y}; Nonce n \<notin> analz G\<rbrakk> \<Longrightarrow> Y \<in> 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 \<in> parts (insert X G); X \<in> synth (analz G); Nonce n \<in> kparts {Y}; Nonce n \<notin> analz G |] 
-   ==> Crypt K Y \<in> parts G"
+  "\<lbrakk>Crypt K Y \<in> parts (insert X G); X \<in> synth (analz G); Nonce n \<in> kparts {Y}; Nonce n \<notin> analz G\<rbrakk> 
+   \<Longrightarrow> Crypt K Y \<in> parts G"
 by (metis Fake_parts_insert_in_Un Nonce_kparts_synth UnE analz_conj_parts synth_simps(5))
 
 
--- a/src/HOL/Auth/Guard/Extensions.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Extensions.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -11,7 +11,7 @@
 
 subsection\<open>Extensions to Theory \<open>Set\<close>\<close>
 
-lemma eq: "[| \<And>x. x\<in>A \<Longrightarrow> x\<in>B; \<And>x. x\<in>B \<Longrightarrow> x\<in>A |] ==> A=B"
+lemma eq: "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> x\<in>B; \<And>x. x\<in>B \<Longrightarrow> x\<in>A\<rbrakk> \<Longrightarrow> A=B"
 by auto
 
 lemma insert_Un: "P ({x} \<union> A) \<Longrightarrow> 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: "\<lbrakk>is_MPair X \<Longrightarrow> P; not_MPair X \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by auto
 
 declare is_MPair_def [simp del]
@@ -109,7 +109,7 @@
 definition usekeys :: "msg set => key set" where
 "usekeys G \<equiv> {K. \<exists>Y. Crypt K Y \<in> G}"
 
-lemma finite_keysFor [intro]: "finite G ==> finite (keysFor G)"
+lemma finite_keysFor [intro]: "finite G \<Longrightarrow> finite (keysFor G)"
 apply (simp add: keysFor_def)
 apply (rule finite_imageI)
 apply (induct G rule: finite_induct)
@@ -122,14 +122,14 @@
 
 subsubsection\<open>lemmas on parts\<close>
 
-lemma parts_sub: "[| X \<in> parts G; G \<subseteq> H |] ==> X \<in> parts H"
+lemma parts_sub: "\<lbrakk>X \<in> parts G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> parts H"
 by (auto dest: parts_mono)
 
 lemma parts_Diff [dest]: "X \<in> parts (G - H) \<Longrightarrow> X \<in> parts G"
 by (erule parts_sub, auto)
 
-lemma parts_Diff_notin: "[| Y \<notin> H; Nonce n \<notin> parts (H - {Y}) |]
-==> Nonce n \<notin> parts H"
+lemma parts_Diff_notin: "\<lbrakk>Y \<notin> H; Nonce n \<notin> parts (H - {Y})\<rbrakk>
+\<Longrightarrow> Nonce n \<notin> 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 \<in> parts {Y}; Y \<in> parts G |] ==> X \<in> parts G"
+lemma parts_parts: "\<lbrakk>X \<in> parts {Y}; Y \<in> parts G\<rbrakk> \<Longrightarrow> X \<in> parts G"
 by (frule parts_cut, auto) 
 
 
-lemma parts_parts_parts: "[| X \<in> parts {Y}; Y \<in> parts {Z}; Z \<in> parts G |] ==> X \<in> parts G"
+lemma parts_parts_parts: "\<lbrakk>X \<in> parts {Y}; Y \<in> parts {Z}; Z \<in> parts G\<rbrakk> \<Longrightarrow> X \<in> parts G"
 by (auto dest: parts_parts)
 
-lemma parts_parts_Crypt: "[| Crypt K X \<in> parts G; Nonce n \<in> parts {X} |]
-==> Nonce n \<in> parts G"
+lemma parts_parts_Crypt: "\<lbrakk>Crypt K X \<in> parts G; Nonce n \<in> parts {X}\<rbrakk>
+\<Longrightarrow> Nonce n \<in> parts G"
 by (blast intro: parts.Body dest: parts_parts)
 
 subsubsection\<open>lemmas on synth\<close>
 
-lemma synth_sub: "[| X \<in> synth G; G \<subseteq> H |] ==> X \<in> synth H"
+lemma synth_sub: "\<lbrakk>X \<in> synth G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> synth H"
 by (auto dest: synth_mono)
 
-lemma Crypt_synth [rule_format]: "[| X \<in> synth G; Key K \<notin> G |] ==>
+lemma Crypt_synth [rule_format]: "\<lbrakk>X \<in> synth G; Key K \<notin> G\<rbrakk> \<Longrightarrow>
 Crypt K Y \<in> parts {X} \<longrightarrow> Crypt K Y \<in> parts G"
 by (erule synth.induct, auto dest: parts_sub)
 
 subsubsection\<open>lemmas on analz\<close>
 
-lemma analz_UnI1 [intro]: "X \<in> analz G ==> X \<in> analz (G \<union> H)"
+lemma analz_UnI1 [intro]: "X \<in> analz G \<Longrightarrow> X \<in> analz (G \<union> H)"
   by (subgoal_tac "G <= G Un H") (blast dest: analz_mono)+
 
-lemma analz_sub: "[| X \<in> analz G; G \<subseteq> H |] ==> X \<in> analz H"
+lemma analz_sub: "\<lbrakk>X \<in> analz G; G \<subseteq> H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 by (auto dest: analz_mono)
 
 lemma analz_Diff [dest]: "X \<in> analz (G - H) \<Longrightarrow> X \<in> 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' \<Longrightarrow> analz A = analz A'"
 by auto
 
 lemmas insert_commute_substI = insert_commute [THEN ssubst]
 
 lemma analz_insertD:
-     "[| Crypt K Y \<in> H; Key (invKey K) \<in> H |] ==> analz (insert Y H) = analz H"
+     "\<lbrakk>Crypt K Y \<in> H; Key (invKey K) \<in> H\<rbrakk> \<Longrightarrow> analz (insert Y H) = analz H"
 by (blast intro: analz.Decrypt analz_insert_eq)  
 
-lemma must_decrypt [rule_format,dest]: "[| X \<in> analz H; has_no_pair H |] ==>
+lemma must_decrypt [rule_format,dest]: "\<lbrakk>X \<in> analz H; has_no_pair H\<rbrakk> \<Longrightarrow>
 X \<notin> H \<longrightarrow> (\<exists>K Y. Crypt K Y \<in> H \<and> Key (invKey K) \<in> 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 \<in> parts (insert Y H);
-Y \<in> synth (analz H); Key K \<notin> analz H |] ==> Crypt K X \<in> parts H"
+lemma Crypt_synth_insert: "\<lbrakk>Crypt K X \<in> parts (insert Y H);
+Y \<in> synth (analz H); Key K \<notin> analz H\<rbrakk> \<Longrightarrow> Crypt K X \<in> 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 \<equiv> \<forall>X. X \<in> G \<longrightarrow> (\<exists>K. X = Key K)"
 
-lemma keyset_in [dest]: "[| keyset G; X \<in> G |] ==> \<exists>K. X = Key K"
+lemma keyset_in [dest]: "\<lbrakk>keyset G; X \<in> G\<rbrakk> \<Longrightarrow> \<exists>K. X = Key K"
 by (auto simp: keyset_def)
 
-lemma MPair_notin_keyset [simp]: "keyset G ==> \<lbrace>X,Y\<rbrace> \<notin> G"
+lemma MPair_notin_keyset [simp]: "keyset G \<Longrightarrow> \<lbrace>X,Y\<rbrace> \<notin> G"
 by auto
 
 lemma Crypt_notin_keyset [simp]: "keyset G \<Longrightarrow> Crypt K X \<notin> G"
@@ -242,7 +242,7 @@
 lemma Nonce_notin_keyset [simp]: "keyset G \<Longrightarrow> Nonce n \<notin> G"
 by auto
 
-lemma parts_keyset [simp]: "keyset G ==> parts G = G"
+lemma parts_keyset [simp]: "keyset G \<Longrightarrow> parts G = G"
 by (auto, erule parts.induct, auto)
 
 subsubsection\<open>keys a priori necessary for decrypting the messages of G\<close>
@@ -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 \<Longrightarrow> keyset (H - keysfor G)"
 by (auto simp: keyset_def)
 
 lemma keysfor_Crypt: "Crypt K X \<in> parts G \<Longrightarrow> Key (invKey K) \<in> keysfor G"
@@ -262,12 +262,12 @@
 lemma no_key_no_Crypt: "Key K \<notin> keysfor G \<Longrightarrow> Crypt (invKey K) X \<notin> parts G"
 by (auto dest: keysfor_Crypt)
 
-lemma finite_keysfor [intro]: "finite G ==> finite (keysfor G)"
+lemma finite_keysfor [intro]: "finite G \<Longrightarrow> finite (keysfor G)"
 by (auto simp: keysfor_def intro: finite_UN_I)
 
 subsubsection\<open>only the keys necessary for G are useful in analz\<close>
 
-lemma analz_keyset: "keyset H ==>
+lemma analz_keyset: "keyset H \<Longrightarrow>
 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 == \<forall>evs B X. evs \<in> p \<longrightarrow> Gets B X \<in> set evs
 \<longrightarrow> (\<exists>A. Says A B X \<in> set evs)"
 
-lemma Gets_correct_Says: "[| Gets_correct p; Gets B X # evs \<in> p |]
-==> \<exists>A. Says A B X \<in> set evs"
+lemma Gets_correct_Says: "\<lbrakk>Gets_correct p; Gets B X # evs \<in> p\<rbrakk>
+\<Longrightarrow> \<exists>A. Says A B X \<in> 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 == \<forall>evs ev. ev#evs \<in> p \<longrightarrow> evs \<in> p"
 
-lemma one_step_Cons [dest]: "[| one_step p; ev#evs \<in> p |] ==> evs \<in> p"
-by (unfold one_step_def, blast)
+lemma one_step_Cons [dest]: "\<lbrakk>one_step p; ev#evs \<in> p\<rbrakk> \<Longrightarrow> evs \<in> p"
+  unfolding one_step_def by blast
 
-lemma one_step_app: "[| evs@evs' \<in> p; one_step p; [] \<in> p |] ==> evs' \<in> p"
+lemma one_step_app: "\<lbrakk>evs@evs' \<in> p; one_step p; [] \<in> p\<rbrakk> \<Longrightarrow> evs' \<in> p"
 by (induct evs, auto)
 
-lemma trunc: "[| evs @ evs' \<in> p; one_step p |] ==> evs' \<in> p"
+lemma trunc: "\<lbrakk>evs @ evs' \<in> p; one_step p\<rbrakk> \<Longrightarrow> evs' \<in> p"
 by (induct evs, auto)
 
 definition has_only_Says :: "event list set => bool" where
 "has_only_Says p \<equiv> \<forall>evs ev. evs \<in> p \<longrightarrow> ev \<in> set evs
 \<longrightarrow> (\<exists>A B X. ev = Says A B X)"
 
-lemma has_only_SaysD: "[| ev \<in> set evs; evs \<in> p; has_only_Says p |]
-==> \<exists>A B X. ev = Says A B X"
-by (unfold has_only_Says_def, blast)
+lemma has_only_SaysD: "\<lbrakk>ev \<in> set evs; evs \<in> p; has_only_Says p\<rbrakk>
+\<Longrightarrow> \<exists>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 \<in> p; ev \<in> set evs |]
-==> \<exists>A B X. ev = Says A B X"
+lemma in_has_only_Says [dest]: "\<lbrakk>has_only_Says p; evs \<in> p; ev \<in> set evs\<rbrakk>
+\<Longrightarrow> \<exists>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"
+\<Longrightarrow> Gets_correct p"
 by (auto simp: has_only_Says_def Gets_correct_def)
 
 subsubsection\<open>lemma on knows\<close>
@@ -335,8 +335,8 @@
 lemma Says_imp_spies2: "Says A B \<lbrace>X,Y\<rbrace> \<in> set evs \<Longrightarrow> Y \<in> parts (spies evs)"
 by (drule Says_imp_spies, drule parts.Inj, drule parts.Snd, simp)
 
-lemma Says_not_parts: "[| Says A B X \<in> set evs; Y \<notin> parts (spies evs) |]
-==> Y \<notin> parts {X}"
+lemma Says_not_parts: "\<lbrakk>Says A B X \<in> set evs; Y \<notin> parts (spies evs)\<rbrakk>
+\<Longrightarrow> Y \<notin> parts {X}"
 by (auto dest: Says_imp_spies parts_parts)
 
 subsubsection\<open>knows without initState\<close>
@@ -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 \<in> p; has_only_Says p; one_step p |]
-==> knows' A evs \<subseteq> spies' evs"
+lemma knows'_sub_spies': "\<lbrakk>evs \<in> p; has_only_Says p; one_step p\<rbrakk>
+\<Longrightarrow> knows' A evs \<subseteq> spies' evs"
 by (induct evs, auto split: event.splits)
 
 subsubsection\<open>knows' is finite\<close>
@@ -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 \<in> p; has_only_Says p; one_step p |]
-==> knows_max' A evs \<subseteq> spies' evs"
+lemma knows_max'_sub_spies': "\<lbrakk>evs \<in> p; has_only_Says p; one_step p\<rbrakk>
+\<Longrightarrow> knows_max' A evs \<subseteq> spies' evs"
 by (induct evs, auto split: event.splits)
 
-lemma knows_max'_in_spies' [dest]: "[| evs \<in> p; X \<in> knows_max' A evs;
-has_only_Says p; one_step p |] ==> X \<in> spies' evs"
+lemma knows_max'_in_spies' [dest]: "\<lbrakk>evs \<in> p; X \<in> knows_max' A evs;
+has_only_Says p; one_step p\<rbrakk> \<Longrightarrow> X \<in> 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 \<in> used evs \<Longrightarrow> X \<in> used (evs @ evs')"
 by (erule used_sub_app [THEN subsetD])
 
-lemma used_parts: "[| X \<in> parts {Y}; Y \<in> used evs |] ==> X \<in> used evs"
+lemma used_parts: "\<lbrakk>X \<in> parts {Y}; Y \<in> used evs\<rbrakk> \<Longrightarrow> X \<in> 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 \<in> set evs; Y \<in> parts {X} |] ==> Y \<in> used evs"
+lemma parts_Says_used: "\<lbrakk>Says A B X \<in> set evs; Y \<in> parts {X}\<rbrakk> \<Longrightarrow> Y \<in> used evs"
 by (induct evs, simp_all, safe, auto intro: used_ConsI)
 
 lemma parts_used_app: "X \<in> parts {Y} \<Longrightarrow> X \<in> used (evs @ Says A B Y # evs')"
@@ -561,16 +561,16 @@
 lemma not_used_not_spied: "X \<notin> used evs \<Longrightarrow> X \<notin> parts (spies evs)"
 by (induct evs, auto simp: used_Nil)
 
-lemma not_used_not_parts: "[| Y \<notin> used evs; Says A B X \<in> set evs |]
-==> Y \<notin> parts {X}"
+lemma not_used_not_parts: "\<lbrakk>Y \<notin> used evs; Says A B X \<in> set evs\<rbrakk>
+\<Longrightarrow> Y \<notin> parts {X}"
 by (induct evs, auto intro: used_ConsI)
 
-lemma not_used_parts_false: "[| X \<notin> used evs; Y \<in> parts (spies evs) |]
-==> X \<notin> parts {Y}"
+lemma not_used_parts_false: "\<lbrakk>X \<notin> used evs; Y \<in> parts (spies evs)\<rbrakk>
+\<Longrightarrow> X \<notin> parts {Y}"
 by (auto dest: parts_parts)
 
-lemma known_used [rule_format]: "[| evs \<in> p; Gets_correct p; one_step p |]
-==> X \<in> parts (knows A evs) \<longrightarrow> X \<in> used evs"
+lemma known_used [rule_format]: "\<lbrakk>evs \<in> p; Gets_correct p; one_step p\<rbrakk>
+\<Longrightarrow> X \<in> parts (knows A evs) \<longrightarrow> X \<in> 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 \<in> p; Gets_correct p; one_step p |]
-==> X \<in> parts (knows_max A evs) \<longrightarrow> X \<in> used evs"
+lemma known_max_used [rule_format]: "\<lbrakk>evs \<in> p; Gets_correct p; one_step p\<rbrakk>
+\<Longrightarrow> X \<in> parts (knows_max A evs) \<longrightarrow> X \<in> 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 \<in> p; X \<notin> used evs;
-Gets_correct p; one_step p |] ==> X \<notin> parts (knows A evs)"
+lemma not_used_not_known: "\<lbrakk>evs \<in> p; X \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> X \<notin> parts (knows A evs)"
 by (case_tac "A=Spy", auto dest: not_used_not_spied known_used)
 
-lemma not_used_not_known_max: "[| evs \<in> p; X \<notin> used evs;
-Gets_correct p; one_step p |] ==> X \<notin> parts (knows_max A evs)"
+lemma not_used_not_known_max: "\<lbrakk>evs \<in> p; X \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> X \<notin> parts (knows_max A evs)"
 by (case_tac "A=Spy", auto dest: not_used_not_spied known_max_used)
 
 subsubsection\<open>a nonce or key in a message cannot equal a fresh nonce or key\<close>
 
-lemma Nonce_neq [dest]: "[| Nonce n' \<notin> used evs;
-Says A B X \<in> set evs; Nonce n \<in> parts {X} |] ==> n \<noteq> n'"
+lemma Nonce_neq [dest]: "\<lbrakk>Nonce n' \<notin> used evs;
+Says A B X \<in> set evs; Nonce n \<in> parts {X}\<rbrakk> \<Longrightarrow> n \<noteq> n'"
 by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub)
 
-lemma Key_neq [dest]: "[| Key n' \<notin> used evs;
-Says A B X \<in> set evs; Key n \<in> parts {X} |] ==> n ~= n'"
+lemma Key_neq [dest]: "\<lbrakk>Key n' \<notin> used evs;
+Says A B X \<in> set evs; Key n \<in> parts {X}\<rbrakk> \<Longrightarrow> n ~= n'"
 by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub)
 
 subsubsection\<open>message of an event\<close>
@@ -623,7 +623,7 @@
 | "msg (Gets A X) = X"
 | "msg (Notes A X) = X"
 
-lemma used_sub_parts_used: "X \<in> used (ev # evs) ==> X \<in> parts {msg ev} \<union> used evs"
+lemma used_sub_parts_used: "X \<in> used (ev # evs) \<Longrightarrow> X \<in> parts {msg ev} \<union> used evs"
 by (induct ev, auto)
 
 end
--- a/src/HOL/Auth/Guard/Guard.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Guard.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -19,7 +19,7 @@
   No_Nonce [intro]: "Nonce n \<notin> parts {X} \<Longrightarrow> X \<in> guard n Ks"
 | Guard_Nonce [intro]: "invKey K \<in> Ks \<Longrightarrow> Crypt K X \<in> guard n Ks"
 | Crypt [intro]: "X \<in> guard n Ks \<Longrightarrow> Crypt K X \<in> guard n Ks"
-| Pair [intro]: "[| X \<in> guard n Ks; Y \<in> guard n Ks |] ==> \<lbrace>X,Y\<rbrace> \<in> guard n Ks"
+| Pair [intro]: "\<lbrakk>X \<in> guard n Ks; Y \<in> guard n Ks\<rbrakk> \<Longrightarrow> \<lbrace>X,Y\<rbrace> \<in> guard n Ks"
 
 subsection\<open>basic facts about \<^term>\<open>guard\<close>\<close>
 
@@ -38,7 +38,7 @@
 lemma Nonce_notin_guard_iff [iff]: "Nonce n \<notin> guard n Ks"
 by (auto dest: Nonce_notin_guard)
 
-lemma guard_has_Crypt [rule_format]: "X \<in> guard n Ks ==> Nonce n \<in> parts {X}
+lemma guard_has_Crypt [rule_format]: "X \<in> guard n Ks \<Longrightarrow> Nonce n \<in> parts {X}
 \<longrightarrow> (\<exists>K Y. Crypt K Y \<in> kparts {X} \<and> Nonce n \<in> parts {Y})"
 by (erule guard.induct, auto)
 
@@ -55,7 +55,7 @@
 Y \<in> kparts {X} \<longrightarrow> Y \<in> guard n Ks"
 by (erule guard.induct, auto)
 
-lemma guard_Crypt: "[| Crypt K Y \<in> guard n Ks; K \<notin> invKey`Ks |] ==> Y \<in> guard n Ks"
+lemma guard_Crypt: "\<lbrakk>Crypt K Y \<in> guard n Ks; K \<notin> invKey`Ks\<rbrakk> \<Longrightarrow> Y \<in> guard n Ks"
   by (ind_cases "Crypt K Y \<in> guard n Ks") (auto intro!: image_eqI)
 
 lemma guard_MPair [iff]: "(\<lbrace>X,Y\<rbrace> \<in> guard n Ks) = (X \<in> guard n Ks \<and> Y \<in> guard n Ks)"
@@ -65,7 +65,7 @@
 Crypt K Y \<in> kparts {X} \<longrightarrow> Nonce n \<in> kparts {Y} \<longrightarrow> Y \<notin> guard n Ks"
 by (erule guard.induct, auto dest: guard_kparts)
 
-lemma guard_extand: "[| X \<in> guard n Ks; Ks \<subseteq> Ks' |] ==> X \<in> guard n Ks'"
+lemma guard_extand: "\<lbrakk>X \<in> guard n Ks; Ks \<subseteq> Ks'\<rbrakk> \<Longrightarrow> X \<in> guard n Ks'"
 by (erule guard.induct, auto)
 
 subsection\<open>guarded sets\<close>
@@ -86,15 +86,15 @@
 lemma Nonce_notin_kparts [simplified]: "Guard n Ks H \<Longrightarrow> Nonce n \<notin> kparts H"
 by (auto simp: Guard_def dest: in_kparts Nonce_notin_kparts_msg)
 
-lemma Guard_must_decrypt: "[| Guard n Ks H; Nonce n \<in> analz H |] ==>
+lemma Guard_must_decrypt: "\<lbrakk>Guard n Ks H; Nonce n \<in> analz H\<rbrakk> \<Longrightarrow>
 \<exists>K Y. Crypt K Y \<in> kparts H \<and> Key (invKey K) \<in> kparts H"
 apply (drule_tac P="\<lambda>G. Nonce n \<in> 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 \<Longrightarrow> 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: "\<lbrakk>Guard n Ks H; G <= H\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> Guard n Ks (synth G)"
 by (auto simp: Guard_def, erule synth.induct, auto)
 
-lemma Guard_analz [intro]: "[| Guard n Ks G; \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |]
-==> Guard n Ks (analz G)"
+lemma Guard_analz [intro]: "\<lbrakk>Guard n Ks G; \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G\<rbrakk>
+\<Longrightarrow> Guard n Ks (analz G)"
 apply (auto simp: Guard_def)
 apply (erule analz.induct, auto)
 by (ind_cases "Crypt K Xa \<in> guard n Ks" for K Xa, auto)
 
-lemma in_Guard [dest]: "[| X \<in> G; Guard n Ks G |] ==> X \<in> guard n Ks"
+lemma in_Guard [dest]: "\<lbrakk>X \<in> G; Guard n Ks G\<rbrakk> \<Longrightarrow> X \<in> guard n Ks"
 by (auto simp: Guard_def)
 
-lemma in_synth_Guard: "[| X \<in> synth G; Guard n Ks G |] ==> X \<in> guard n Ks"
+lemma in_synth_Guard: "\<lbrakk>X \<in> synth G; Guard n Ks G\<rbrakk> \<Longrightarrow> X \<in> guard n Ks"
 by (drule Guard_synth, auto)
 
-lemma in_analz_Guard: "[| X \<in> analz G; Guard n Ks G;
-\<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |] ==> X \<in> guard n Ks"
+lemma in_analz_Guard: "\<lbrakk>X \<in> analz G; Guard n Ks G;
+\<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G\<rbrakk> \<Longrightarrow> X \<in> guard n Ks"
 by (drule Guard_analz, auto)
 
-lemma Guard_keyset [simp]: "keyset G ==> Guard n Ks G"
+lemma Guard_keyset [simp]: "keyset G \<Longrightarrow> Guard n Ks G"
 by (auto simp: Guard_def)
 
-lemma Guard_Un_keyset: "[| Guard n Ks G; keyset H |] ==> Guard n Ks (G \<union> H)"
+lemma Guard_Un_keyset: "\<lbrakk>Guard n Ks G; keyset H\<rbrakk> \<Longrightarrow> Guard n Ks (G \<union> H)"
 by auto
 
-lemma in_Guard_kparts: "[| X \<in> G; Guard n Ks G; Y \<in> kparts {X} |] ==> Y \<in> guard n Ks"
+lemma in_Guard_kparts: "\<lbrakk>X \<in> G; Guard n Ks G; Y \<in> kparts {X}\<rbrakk> \<Longrightarrow> Y \<in> guard n Ks"
 by blast
 
-lemma in_Guard_kparts_neq: "[| X \<in> G; Guard n Ks G; Nonce n' \<in> kparts {X} |]
-==> n \<noteq> n'"
+lemma in_Guard_kparts_neq: "\<lbrakk>X \<in> G; Guard n Ks G; Nonce n' \<in> kparts {X}\<rbrakk>
+\<Longrightarrow> n \<noteq> n'"
 by (blast dest: in_Guard_kparts)
 
-lemma in_Guard_kparts_Crypt: "[| X \<in> G; Guard n Ks G; is_MPair X;
-Crypt K Y \<in> kparts {X}; Nonce n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
+lemma in_Guard_kparts_Crypt: "\<lbrakk>X \<in> G; Guard n Ks G; is_MPair X;
+Crypt K Y \<in> kparts {X}; Nonce n \<in> kparts {Y}\<rbrakk> \<Longrightarrow> invKey K \<in> Ks"
 apply (drule in_Guard, simp)
 apply (frule guard_not_guard, simp+)
 apply (drule guard_kparts, simp)
 by (ind_cases "Crypt K Y \<in> guard n Ks", auto)
 
-lemma Guard_extand: "[| Guard n Ks G; Ks \<subseteq> Ks' |] ==> Guard n Ks' G"
+lemma Guard_extand: "\<lbrakk>Guard n Ks G; Ks \<subseteq> Ks'\<rbrakk> \<Longrightarrow> Guard n Ks' G"
 by (auto simp: Guard_def dest: guard_extand)
 
-lemma guard_invKey [rule_format]: "[| X \<in> guard n Ks; Nonce n \<in> kparts {Y} |] ==>
+lemma guard_invKey [rule_format]: "\<lbrakk>X \<in> guard n Ks; Nonce n \<in> kparts {Y}\<rbrakk> \<Longrightarrow>
 Crypt K Y \<in> kparts {X} \<longrightarrow> invKey K \<in> Ks"
 by (erule guard.induct, auto)
 
-lemma Crypt_guard_invKey [rule_format]: "[| Crypt K Y \<in> guard n Ks;
-Nonce n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
+lemma Crypt_guard_invKey [rule_format]: "\<lbrakk>Crypt K Y \<in> guard n Ks;
+Nonce n \<in> kparts {Y}\<rbrakk> \<Longrightarrow> invKey K \<in> Ks"
 by (auto dest: guard_invKey)
 
 subsection\<open>set obtained by decrypting a message\<close>
@@ -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 \<in> H; Key (invKey K) \<in> H; Nonce n \<in> analz H |]
-==> Nonce n \<in> analz (decrypt H K Y)"
+lemma analz_decrypt: "\<lbrakk>Crypt K Y \<in> H; Key (invKey K) \<in> H; Nonce n \<in> analz H\<rbrakk>
+\<Longrightarrow> Nonce n \<in> analz (decrypt H K Y)"
 apply (drule_tac P="\<lambda>H. Nonce n \<in> analz H" in ssubst [OF insert_Diff])
 apply assumption
 apply (simp only: analz_Crypt_if, simp)
 done
 
-lemma parts_decrypt: "[| Crypt K Y \<in> H; X \<in> parts (decrypt H K Y) |] ==> X \<in> parts H"
+lemma parts_decrypt: "\<lbrakk>Crypt K Y \<in> H; X \<in> parts (decrypt H K Y)\<rbrakk> \<Longrightarrow> X \<in> parts H"
 by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body)
 
 subsection\<open>number of Crypt's in a message\<close>
@@ -195,12 +195,12 @@
 lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'"
 by (induct l, auto)
 
-lemma mem_cnb_minus: "x \<in> set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)"
+lemma mem_cnb_minus: "x \<in> set l \<Longrightarrow> 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 \<in> set l ==> cnb (remove l x) = cnb l - crypt_nb x"
+lemma cnb_minus [simp]: "x \<in> set l \<Longrightarrow> 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 \<in> analz G; Guard n Ks G; finite G |]
-==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
+lemma Guard_invKey_finite: "\<lbrakk>Nonce n \<in> analz G; Guard n Ks G; finite G\<rbrakk>
+\<Longrightarrow> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
 apply (drule finite_list, clarify)
 by (rule Guard_invKey_by_list, auto)
 
-lemma Guard_invKey: "[| Nonce n \<in> analz G; Guard n Ks G |]
-==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
+lemma Guard_invKey: "\<lbrakk>Nonce n \<in> analz G; Guard n Ks G\<rbrakk>
+\<Longrightarrow> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
 by (auto dest: analz_needs_only_finite Guard_invKey_finite)
 
 subsection\<open>if the analyse of a finite guarded set and a (possibly infinite) set of keys
 gives n then it must also gives Ks\<close>
 
-lemma Guard_invKey_keyset: "[| Nonce n \<in> analz (G \<union> H); Guard n Ks G; finite G;
-keyset H |] ==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz (G \<union> H)"
+lemma Guard_invKey_keyset: "\<lbrakk>Nonce n \<in> analz (G \<union> H); Guard n Ks G; finite G;
+keyset H\<rbrakk> \<Longrightarrow> \<exists>K. K \<in> Ks \<and> Key K \<in> analz (G \<union> H)"
 apply (frule_tac P="\<lambda>G. Nonce n \<in> 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)
--- a/src/HOL/Auth/Guard/GuardK.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/GuardK.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -24,9 +24,9 @@
   for n :: nat and Ks :: "key set"
 where
   No_Key [intro]: "Key n \<notin> parts {X} \<Longrightarrow> X \<in> guardK n Ks"
-| Guard_Key [intro]: "invKey K \<in> Ks ==> Crypt K X \<in> guardK n Ks"
+| Guard_Key [intro]: "invKey K \<in> Ks \<Longrightarrow> Crypt K X \<in> guardK n Ks"
 | Crypt [intro]: "X \<in> guardK n Ks \<Longrightarrow> Crypt K X \<in> guardK n Ks"
-| Pair [intro]: "[| X \<in> guardK n Ks; Y \<in> guardK n Ks |] ==> \<lbrace>X,Y\<rbrace> \<in> guardK n Ks"
+| Pair [intro]: "\<lbrakk>X \<in> guardK n Ks; Y \<in> guardK n Ks\<rbrakk> \<Longrightarrow> \<lbrace>X,Y\<rbrace> \<in> guardK n Ks"
 
 subsection\<open>basic facts about \<^term>\<open>guardK\<close>\<close>
 
@@ -62,7 +62,7 @@
 Y \<in> kparts {X} \<longrightarrow> Y \<in> guardK n Ks"
 by (erule guardK.induct, auto dest: kparts_parts parts_sub)
 
-lemma guardK_Crypt: "[| Crypt K Y \<in> guardK n Ks; K \<notin> invKey`Ks |] ==> Y \<in> guardK n Ks"
+lemma guardK_Crypt: "\<lbrakk>Crypt K Y \<in> guardK n Ks; K \<notin> invKey`Ks\<rbrakk> \<Longrightarrow> Y \<in> guardK n Ks"
   by (ind_cases "Crypt K Y \<in> guardK n Ks") (auto intro!: image_eqI)
 
 lemma guardK_MPair [iff]: "(\<lbrace>X,Y\<rbrace> \<in> guardK n Ks)
@@ -73,8 +73,8 @@
 Crypt K Y \<in> kparts {X} \<longrightarrow> Key n \<in> kparts {Y} \<longrightarrow> Y \<notin> guardK n Ks"
 by (erule guardK.induct, auto dest: guardK_kparts)
 
-lemma guardK_extand: "[| X \<in> guardK n Ks; Ks \<subseteq> Ks';
-[| K \<in> Ks'; K \<notin> Ks |] ==> Key K \<notin> parts {X} |] ==> X \<in> guardK n Ks'"
+lemma guardK_extand: "\<lbrakk>X \<in> guardK n Ks; Ks \<subseteq> Ks';
+\<lbrakk>K \<in> Ks'; K \<notin> Ks\<rbrakk> \<Longrightarrow> Key K \<notin> parts {X}\<rbrakk> \<Longrightarrow> X \<in> guardK n Ks'"
 by (erule guardK.induct, auto)
 
 subsection\<open>guarded sets\<close>
@@ -90,7 +90,7 @@
 lemma Key_notin_kparts [simplified]: "GuardK n Ks H \<Longrightarrow> Key n \<notin> kparts H"
 by (auto simp: GuardK_def dest: in_kparts Key_notin_kparts_msg)
 
-lemma GuardK_must_decrypt: "[| GuardK n Ks H; Key n \<in> analz H |] ==>
+lemma GuardK_must_decrypt: "\<lbrakk>GuardK n Ks H; Key n \<in> analz H\<rbrakk> \<Longrightarrow>
 \<exists>K Y. Crypt K Y \<in> kparts H \<and> Key (invKey K) \<in> kparts H"
 apply (drule_tac P="\<lambda>G. Key n \<in> 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 \<Longrightarrow> GuardK n Ks (kparts H)"
 by (auto simp: GuardK_def dest: in_kparts guardK_kparts)
 
-lemma GuardK_mono: "[| GuardK n Ks H; G \<subseteq> H |] ==> GuardK n Ks G"
+lemma GuardK_mono: "\<lbrakk>GuardK n Ks H; G \<subseteq> H\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> GuardK n Ks (synth G)"
 by (auto simp: GuardK_def, erule synth.induct, auto)
 
-lemma GuardK_analz [intro]: "[| GuardK n Ks G; \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |]
-==> GuardK n Ks (analz G)"
+lemma GuardK_analz [intro]: "\<lbrakk>GuardK n Ks G; \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G\<rbrakk>
+\<Longrightarrow> GuardK n Ks (analz G)"
 apply (auto simp: GuardK_def)
 apply (erule analz.induct, auto)
 by (ind_cases "Crypt K Xa \<in> guardK n Ks" for K Xa, auto)
 
-lemma in_GuardK [dest]: "[| X \<in> G; GuardK n Ks G |] ==> X \<in> guardK n Ks"
+lemma in_GuardK [dest]: "\<lbrakk>X \<in> G; GuardK n Ks G\<rbrakk> \<Longrightarrow> X \<in> guardK n Ks"
 by (auto simp: GuardK_def)
 
-lemma in_synth_GuardK: "[| X \<in> synth G; GuardK n Ks G |] ==> X \<in> guardK n Ks"
+lemma in_synth_GuardK: "\<lbrakk>X \<in> synth G; GuardK n Ks G\<rbrakk> \<Longrightarrow> X \<in> guardK n Ks"
 by (drule GuardK_synth, auto)
 
-lemma in_analz_GuardK: "[| X \<in> analz G; GuardK n Ks G;
-\<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |] ==> X \<in> guardK n Ks"
+lemma in_analz_GuardK: "\<lbrakk>X \<in> analz G; GuardK n Ks G;
+\<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G\<rbrakk> \<Longrightarrow> X \<in> guardK n Ks"
 by (drule GuardK_analz, auto)
 
-lemma GuardK_keyset [simp]: "[| keyset G; Key n \<notin> G |] ==> GuardK n Ks G"
+lemma GuardK_keyset [simp]: "\<lbrakk>keyset G; Key n \<notin> G\<rbrakk> \<Longrightarrow> 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 \<notin> H |]
-==> GuardK n Ks (G Un H)"
+lemma GuardK_Un_keyset: "\<lbrakk>GuardK n Ks G; keyset H; Key n \<notin> H\<rbrakk>
+\<Longrightarrow> GuardK n Ks (G Un H)"
 by auto
 
-lemma in_GuardK_kparts: "[| X \<in> G; GuardK n Ks G; Y \<in> kparts {X} |] ==> Y \<in> guardK n Ks"
+lemma in_GuardK_kparts: "\<lbrakk>X \<in> G; GuardK n Ks G; Y \<in> kparts {X}\<rbrakk> \<Longrightarrow> Y \<in> guardK n Ks"
 by blast
 
-lemma in_GuardK_kparts_neq: "[| X \<in> G; GuardK n Ks G; Key n' \<in> kparts {X} |]
-==> n \<noteq> n'"
+lemma in_GuardK_kparts_neq: "\<lbrakk>X \<in> G; GuardK n Ks G; Key n' \<in> kparts {X}\<rbrakk>
+\<Longrightarrow> n \<noteq> n'"
 by (blast dest: in_GuardK_kparts)
 
-lemma in_GuardK_kparts_Crypt: "[| X \<in> G; GuardK n Ks G; is_MPair X;
-Crypt K Y \<in> kparts {X}; Key n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
+lemma in_GuardK_kparts_Crypt: "\<lbrakk>X \<in> G; GuardK n Ks G; is_MPair X;
+Crypt K Y \<in> kparts {X}; Key n \<in> kparts {Y}\<rbrakk> \<Longrightarrow> invKey K \<in> Ks"
 apply (drule in_GuardK, simp)
 apply (frule guardK_not_guardK, simp+)
 apply (drule guardK_kparts, simp)
 by (ind_cases "Crypt K Y \<in> guardK n Ks", auto)
 
-lemma GuardK_extand: "[| GuardK n Ks G; Ks \<subseteq> Ks';
-[| K \<in> Ks'; K \<notin> Ks |] ==> Key K \<notin> parts G |] ==> GuardK n Ks' G"
+lemma GuardK_extand: "\<lbrakk>GuardK n Ks G; Ks \<subseteq> Ks';
+\<lbrakk>K \<in> Ks'; K \<notin> Ks\<rbrakk> \<Longrightarrow> Key K \<notin> parts G\<rbrakk> \<Longrightarrow> GuardK n Ks' G"
 by (auto simp: GuardK_def dest: guardK_extand parts_sub)
 
 subsection\<open>set obtained by decrypting a message\<close>
@@ -158,14 +158,14 @@
   decrypt :: "msg set \<Rightarrow> key \<Rightarrow> msg \<Rightarrow> msg set" where
   "decrypt H K Y \<equiv> insert Y (H - {Crypt K Y})"
 
-lemma analz_decrypt: "[| Crypt K Y \<in> H; Key (invKey K) \<in> H; Key n \<in> analz H |]
-==> Key n \<in> analz (decrypt H K Y)"
+lemma analz_decrypt: "\<lbrakk>Crypt K Y \<in> H; Key (invKey K) \<in> H; Key n \<in> analz H\<rbrakk>
+\<Longrightarrow> Key n \<in> analz (decrypt H K Y)"
 apply (drule_tac P="\<lambda>H. Key n \<in> analz H" in ssubst [OF insert_Diff])
 apply assumption 
 apply (simp only: analz_Crypt_if, simp)
 done
 
-lemma parts_decrypt: "[| Crypt K Y \<in> H; X \<in> parts (decrypt H K Y) |] ==> X \<in> parts H"
+lemma parts_decrypt: "\<lbrakk>Crypt K Y \<in> H; X \<in> parts (decrypt H K Y)\<rbrakk> \<Longrightarrow> X \<in> parts H"
 by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body)
 
 subsection\<open>number of Crypt's in a message\<close>
@@ -191,12 +191,12 @@
 lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'"
 by (induct l, auto)
 
-lemma mem_cnb_minus: "x \<in> set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)"
+lemma mem_cnb_minus: "x \<in> set l \<Longrightarrow> 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 \<in> set l ==> cnb (remove l x) = cnb l - crypt_nb x"
+lemma cnb_minus [simp]: "x \<in> set l \<Longrightarrow> 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 \<in> analz G; GuardK n Ks G; finite G |]
-==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
+lemma GuardK_invKey_finite: "\<lbrakk>Key n \<in> analz G; GuardK n Ks G; finite G\<rbrakk>
+\<Longrightarrow> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
 apply (drule finite_list, clarify)
 by (rule GuardK_invKey_by_list, auto)
 
-lemma GuardK_invKey: "[| Key n \<in> analz G; GuardK n Ks G |]
-==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
+lemma GuardK_invKey: "\<lbrakk>Key n \<in> analz G; GuardK n Ks G\<rbrakk>
+\<Longrightarrow> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
 by (auto dest: analz_needs_only_finite GuardK_invKey_finite)
 
 text\<open>if the analyse of a finite guarded set and a (possibly infinite) set of
 keys gives n then it must also gives Ks\<close>
 
-lemma GuardK_invKey_keyset: "[| Key n \<in> analz (G \<union> H); GuardK n Ks G; finite G;
-keyset H; Key n \<notin> H |] ==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz (G \<union> H)"
+lemma GuardK_invKey_keyset: "\<lbrakk>Key n \<in> analz (G \<union> H); GuardK n Ks G; finite G;
+keyset H; Key n \<notin> H\<rbrakk> \<Longrightarrow> \<exists>K. K \<in> Ks \<and> Key K \<in> analz (G \<union> H)"
 apply (frule_tac P="\<lambda>G. Key n \<in> 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)
--- a/src/HOL/Auth/Guard/Guard_NS_Public.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Guard_NS_Public.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -39,14 +39,14 @@
 
   Nil: "[] \<in> nsp"
 
-| Fake: "[| evs \<in> nsp; X \<in> synth (analz (spies evs)) |] ==> Says Spy B X # evs \<in> nsp"
+| Fake: "\<lbrakk>evs \<in> nsp; X \<in> synth (analz (spies evs))\<rbrakk> \<Longrightarrow> Says Spy B X # evs \<in> nsp"
 
-| NS1: "[| evs1 \<in> nsp; Nonce NA \<notin> used evs1 |] ==> ns1 A B NA # evs1 \<in> nsp"
+| NS1: "\<lbrakk>evs1 \<in> nsp; Nonce NA \<notin> used evs1\<rbrakk> \<Longrightarrow> ns1 A B NA # evs1 \<in> nsp"
 
-| NS2: "[| evs2 \<in> nsp; Nonce NB \<notin> used evs2; ns1' A' A B NA \<in> set evs2 |] ==>
+| NS2: "\<lbrakk>evs2 \<in> nsp; Nonce NB \<notin> used evs2; ns1' A' A B NA \<in> set evs2\<rbrakk> \<Longrightarrow>
   ns2 B A NA NB # evs2 \<in> nsp"
 
-| NS3: "\<And>A B B' NA NB evs3. [| evs3 \<in> nsp; ns1 A B NA \<in> set evs3; ns2' B' B A NA NB \<in> set evs3 |] ==>
+| NS3: "\<And>A B B' NA NB evs3. \<lbrakk>evs3 \<in> nsp; ns1 A B NA \<in> set evs3; ns2' B' B A NA NB \<in> set evs3\<rbrakk> \<Longrightarrow>
   ns3 A B NB # evs3 \<in> nsp"
 
 subsection\<open>declarations for tactics\<close>
@@ -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 \<in> nsp" for ev evs, auto)
+  unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> nsp" for ev evs, auto)
 
 lemma nsp_has_only_Says' [rule_format]: "evs \<in> nsp \<Longrightarrow>
 ev \<in> set evs \<longrightarrow> (\<exists>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') \<lbrace>Nonce NA', Nonce NA, Agent A'\<rbrace> \<in> parts (spies evs);
-Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs); evs \<in> nsp |]
-==> Nonce NA \<in> analz (spies evs)"
+"\<lbrakk>Crypt (pubK B') \<lbrace>Nonce NA', Nonce NA, Agent A'\<rbrace> \<in> parts (spies evs);
+Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs); evs \<in> nsp\<rbrakk>
+\<Longrightarrow> Nonce NA \<in> analz (spies evs)"
 by (rule no_Nonce_NS1_NS2, auto)
  
 lemma NB_is_uniq [rule_format]: "evs \<in> nsp \<Longrightarrow>
@@ -108,7 +108,7 @@
 
 subsection\<open>guardedness of NA\<close>
 
-lemma ns1_imp_Guard [rule_format]: "[| evs \<in> nsp; A \<notin> bad; B \<notin> bad |] ==>
+lemma ns1_imp_Guard [rule_format]: "\<lbrakk>evs \<in> nsp; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow>
 ns1 A B NA \<in> set evs \<longrightarrow> Guard NA {priK A,priK B} (spies evs)"
 apply (erule nsp.induct)
 (* Nil *)
@@ -135,7 +135,7 @@
 
 subsection\<open>guardedness of NB\<close>
 
-lemma ns2_imp_Guard [rule_format]: "[| evs \<in> nsp; A \<notin> bad; B \<notin> bad |] ==>
+lemma ns2_imp_Guard [rule_format]: "\<lbrakk>evs \<in> nsp; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow>
 ns2 B A NA NB \<in> set evs \<longrightarrow> Guard NB {priK A,priK B} (spies evs)" 
 apply (erule nsp.induct)
 (* Nil *)
@@ -165,13 +165,13 @@
 
 subsection\<open>Agents' Authentication\<close>
 
-lemma B_trusts_NS1: "[| evs \<in> nsp; A \<notin> bad; B \<notin> bad |] ==>
+lemma B_trusts_NS1: "\<lbrakk>evs \<in> nsp; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow>
 Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs)
 \<longrightarrow> Nonce NA \<notin> analz (spies evs) \<longrightarrow> ns1 A B NA \<in> set evs"
 apply (erule nsp.induct, simp_all)
 by (blast intro: analz_insertI)+
 
-lemma A_trusts_NS2: "[| evs \<in> nsp; A \<notin> bad; B \<notin> bad |] ==> ns1 A B NA \<in> set evs
+lemma A_trusts_NS2: "\<lbrakk>evs \<in> nsp; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow> ns1 A B NA \<in> set evs
 \<longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs)
 \<longrightarrow> ns2 B A NA NB \<in> 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 \<in> nsp; A \<notin> bad; B \<notin> bad |] ==> ns2 B A NA NB \<in> set evs
+lemma B_trusts_NS3: "\<lbrakk>evs \<in> nsp; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow> ns2 B A NA NB \<in> set evs
 \<longrightarrow> Crypt (pubK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow> ns3 A B NB \<in> set evs"
 apply (erule nsp.induct, simp_all, safe)
 apply (frule_tac B=B in ns2_imp_Guard, simp+)
--- a/src/HOL/Auth/Guard/Guard_OtwayRees.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Guard_OtwayRees.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -61,18 +61,18 @@
 
   Nil: "[] \<in> or"
 
-| Fake: "[| evs \<in> or; X \<in> synth (analz (spies evs)) |] ==> Says Spy B X # evs \<in> or"
+| Fake: "\<lbrakk>evs \<in> or; X \<in> synth (analz (spies evs))\<rbrakk> \<Longrightarrow> Says Spy B X # evs \<in> or"
 
-| OR1: "[| evs1 \<in> or; Nonce NA \<notin> used evs1 |] ==> or1 A B NA # evs1 \<in> or"
+| OR1: "\<lbrakk>evs1 \<in> or; Nonce NA \<notin> used evs1\<rbrakk> \<Longrightarrow> or1 A B NA # evs1 \<in> or"
 
-| OR2: "[| evs2 \<in> or; or1' A' A B NA X \<in> set evs2; Nonce NB \<notin> used evs2 |]
-  ==> or2 A B NA NB X # evs2 \<in> or"
+| OR2: "\<lbrakk>evs2 \<in> or; or1' A' A B NA X \<in> set evs2; Nonce NB \<notin> used evs2\<rbrakk>
+  \<Longrightarrow> or2 A B NA NB X # evs2 \<in> or"
 
-| OR3: "[| evs3 \<in> or; or2' B' A B NA NB \<in> set evs3; Key K \<notin> used evs3 |]
-  ==> or3 A B NA NB K # evs3 \<in> or"
+| OR3: "\<lbrakk>evs3 \<in> or; or2' B' A B NA NB \<in> set evs3; Key K \<notin> used evs3\<rbrakk>
+  \<Longrightarrow> or3 A B NA NB K # evs3 \<in> or"
 
-| OR4: "[| evs4 \<in> or; or2 A B NA NB X \<in> set evs4; or3' S Y A B NA NB K \<in> set evs4 |]
-  ==> or4 A B NA X # evs4 \<in> or"
+| OR4: "\<lbrakk>evs4 \<in> or; or2 A B NA NB X \<in> set evs4; or3' S Y A B NA NB K \<in> set evs4\<rbrakk>
+  \<Longrightarrow> or4 A B NA X # evs4 \<in> or"
 
 subsection\<open>declarations for tactics\<close>
 
@@ -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 \<in> or" for ev evs, auto)
+  unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> or" for ev evs, auto)
 
 lemma or_has_only_Says' [rule_format]: "evs \<in> or \<Longrightarrow>
 ev \<in> set evs \<longrightarrow> (\<exists>A B X. ev=Says A B X)"
@@ -119,7 +119,7 @@
 
 subsection\<open>guardedness of KAB\<close>
 
-lemma Guard_KAB [rule_format]: "[| evs \<in> or; A \<notin> bad; B \<notin> bad |] ==>
+lemma Guard_KAB [rule_format]: "\<lbrakk>evs \<in> or; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow>
 or3 A B NA NB K \<in> set evs \<longrightarrow> GuardK K {shrK A,shrK B} (spies evs)" 
 apply (erule or.induct)
 (* Nil *)
@@ -140,7 +140,7 @@
 
 subsection\<open>guardedness of NB\<close>
 
-lemma Guard_NB [rule_format]: "[| evs \<in> or; B \<notin> bad |] ==>
+lemma Guard_NB [rule_format]: "\<lbrakk>evs \<in> or; B \<notin> bad\<rbrakk> \<Longrightarrow>
 or2 A B NA NB X \<in> set evs \<longrightarrow> Guard NB {shrK B} (spies evs)" 
 apply (erule or.induct)
 (* Nil *)
--- a/src/HOL/Auth/Guard/Guard_Public.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Guard_Public.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -54,7 +54,7 @@
 definition priK_set :: "key set => bool" where
 "priK_set Ks \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> (\<exists>A. K = priK A)"
 
-lemma in_priK_set: "[| priK_set Ks; K \<in> Ks |] ==> \<exists>A. K = priK A"
+lemma in_priK_set: "\<lbrakk>priK_set Ks; K \<in> Ks\<rbrakk> \<Longrightarrow> \<exists>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 == \<forall>K. K \<in> Ks \<longrightarrow> agt K \<notin> bad"
 
-lemma in_good: "[| good Ks; K \<in> Ks |] ==> agt K \<notin> bad"
+lemma in_good: "\<lbrakk>good Ks; K \<in> Ks\<rbrakk> \<Longrightarrow> agt K \<notin> bad"
 by (simp add: good_def)
 
 lemma good1 [simp]: "A \<notin> bad \<Longrightarrow> good {priK A}"
 by (simp add: good_def)
 
-lemma good2 [simp]: "[| A \<notin> bad; B \<notin> bad |] ==> good {priK A, priK B}"
+lemma good2 [simp]: "\<lbrakk>A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow> good {priK A, priK B}"
 by (simp add: good_def)
 
 subsubsection\<open>greatest nonce used in a trace, 0 if there is no nonce\<close>
@@ -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)"
+\<Longrightarrow> Guard n Ks (knows_max C evs)"
 by (simp add: knows_max_def)
 
 lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \<notin> used evs
 \<Longrightarrow> Guard n Ks (spies evs)"
 by (auto simp: Guard_def dest: not_used_not_known parts_sub)
 
-lemma Nonce_not_used_Guard [dest]: "[| evs \<in> p; Nonce n \<notin> used evs;
-Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)"
+lemma Nonce_not_used_Guard [dest]: "\<lbrakk>evs \<in> p; Nonce n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<in> p; Nonce n \<notin> used evs;
-Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)"
+lemma Nonce_not_used_Guard_max [dest]: "\<lbrakk>evs \<in> p; Nonce n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<in> p; Nonce n \<notin> used evs;
-Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)"
+lemma Nonce_not_used_Guard_max' [dest]: "\<lbrakk>evs \<in> p; Nonce n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<Rightarrow> bool" where
 "regular p \<equiv> \<forall>evs A. evs \<in> p \<longrightarrow> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)"
 
-lemma priK_parts_iff_bad [simp]: "[| evs \<in> p; regular p |] ==>
+lemma priK_parts_iff_bad [simp]: "\<lbrakk>evs \<in> p; regular p\<rbrakk> \<Longrightarrow>
 (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)"
 by (auto simp: regular_def)
 
-lemma priK_analz_iff_bad [simp]: "[| evs \<in> p; regular p |] ==>
+lemma priK_analz_iff_bad [simp]: "\<lbrakk>evs \<in> p; regular p\<rbrakk> \<Longrightarrow>
 (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
-lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \<in> p;
-priK_set Ks; good Ks; regular p |] ==> Nonce n \<notin> analz (spies evs)"
+lemma Guard_Nonce_analz: "\<lbrakk>Guard n Ks (spies evs); evs \<in> p;
+priK_set Ks; good Ks; regular p\<rbrakk> \<Longrightarrow> Nonce n \<notin> analz (spies evs)"
 apply (clarify, simp only: knows_decomp)
 apply (drule Guard_invKey_keyset, simp+, safe)
 apply (drule in_good, simp)
--- a/src/HOL/Auth/Guard/Guard_Shared.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Guard_Shared.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -49,7 +49,7 @@
 definition shrK_set :: "key set => bool" where
 "shrK_set Ks \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> (\<exists>A. K = shrK A)"
 
-lemma in_shrK_set: "[| shrK_set Ks; K \<in> Ks |] ==> \<exists>A. K = shrK A"
+lemma in_shrK_set: "\<lbrakk>shrK_set Ks; K \<in> Ks\<rbrakk> \<Longrightarrow> \<exists>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 \<Rightarrow> bool" where
 "good Ks \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> agt K \<notin> bad"
 
-lemma in_good: "[| good Ks; K \<in> Ks |] ==> agt K \<notin> bad"
+lemma in_good: "\<lbrakk>good Ks; K \<in> Ks\<rbrakk> \<Longrightarrow> agt K \<notin> bad"
 by (simp add: good_def)
 
 lemma good1 [simp]: "A \<notin> bad \<Longrightarrow> good {shrK A}"
 by (simp add: good_def)
 
-lemma good2 [simp]: "[| A \<notin> bad; B \<notin> bad |] ==> good {shrK A, shrK B}"
+lemma good2 [simp]: "\<lbrakk>A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow> 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)"
+\<Longrightarrow> Guard n Ks (knows_max C evs)"
 by (simp add: knows_max_def)
 
 lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \<notin> used evs
 \<Longrightarrow> Guard n Ks (spies evs)"
 by (auto simp: Guard_def dest: not_used_not_known parts_sub)
 
-lemma Nonce_not_used_Guard [dest]: "[| evs \<in> p; Nonce n \<notin> used evs;
-Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)"
+lemma Nonce_not_used_Guard [dest]: "\<lbrakk>evs \<in> p; Nonce n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<in> p; Nonce n \<notin> used evs;
-Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)"
+lemma Nonce_not_used_Guard_max [dest]: "\<lbrakk>evs \<in> p; Nonce n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<in> p; Nonce n \<notin> used evs;
-Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)"
+lemma Nonce_not_used_Guard_max' [dest]: "\<lbrakk>evs \<in> p; Nonce n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<notin> range shrK \<Longrightarrow> 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 \<notin> range shrK |]
-==> GuardK n A (knows_max C evs)"
+lemma GuardK_knows_max': "\<lbrakk>GuardK n A (knows_max' C evs); n \<notin> range shrK\<rbrakk>
+\<Longrightarrow> GuardK n A (knows_max C evs)"
 by (simp add: knows_max_def)
 
 lemma Key_not_used_GuardK_spies [dest]: "Key n \<notin> used evs
 \<Longrightarrow> GuardK n A (spies evs)"
 by (auto simp: GuardK_def dest: not_used_not_known parts_sub)
 
-lemma Key_not_used_GuardK [dest]: "[| evs \<in> p; Key n \<notin> used evs;
-Gets_correct p; one_step p |] ==> GuardK n A (knows (Friend C) evs)"
+lemma Key_not_used_GuardK [dest]: "\<lbrakk>evs \<in> p; Key n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<in> p; Key n \<notin> used evs;
-Gets_correct p; one_step p |] ==> GuardK n A (knows_max (Friend C) evs)"
+lemma Key_not_used_GuardK_max [dest]: "\<lbrakk>evs \<in> p; Key n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<in> p; Key n \<notin> used evs;
-Gets_correct p; one_step p |] ==> GuardK n A (knows_max' (Friend C) evs)"
+lemma Key_not_used_GuardK_max' [dest]: "\<lbrakk>evs \<in> p; Key n \<notin> used evs;
+Gets_correct p; one_step p\<rbrakk> \<Longrightarrow> 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 \<equiv> \<forall>evs A. evs \<in> p \<longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 
-lemma shrK_parts_iff_bad [simp]: "[| evs \<in> p; regular p |] ==>
+lemma shrK_parts_iff_bad [simp]: "\<lbrakk>evs \<in> p; regular p\<rbrakk> \<Longrightarrow>
 (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 by (auto simp: regular_def)
 
-lemma shrK_analz_iff_bad [simp]: "[| evs \<in> p; regular p |] ==>
+lemma shrK_analz_iff_bad [simp]: "\<lbrakk>evs \<in> p; regular p\<rbrakk> \<Longrightarrow>
 (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
-lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \<in> p;
-shrK_set Ks; good Ks; regular p |] ==> Nonce n \<notin> analz (spies evs)"
+lemma Guard_Nonce_analz: "\<lbrakk>Guard n Ks (spies evs); evs \<in> p;
+shrK_set Ks; good Ks; regular p\<rbrakk> \<Longrightarrow> Nonce n \<notin> analz (spies evs)"
 apply (clarify, simp only: knows_decomp)
 apply (drule Guard_invKey_keyset, simp+, safe)
 apply (drule in_good, simp)
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -52,18 +52,18 @@
 
   Nil: "[] \<in> ya"
 
-| Fake: "[| evs \<in> ya; X \<in> synth (analz (spies evs)) |] ==> Says Spy B X # evs \<in> ya"
+| Fake: "\<lbrakk>evs \<in> ya; X \<in> synth (analz (spies evs))\<rbrakk> \<Longrightarrow> Says Spy B X # evs \<in> ya"
 
-| YA1: "[| evs1 \<in> ya; Nonce NA \<notin> used evs1 |] ==> ya1 A B NA # evs1 \<in> ya"
+| YA1: "\<lbrakk>evs1 \<in> ya; Nonce NA \<notin> used evs1\<rbrakk> \<Longrightarrow> ya1 A B NA # evs1 \<in> ya"
 
-| YA2: "[| evs2 \<in> ya; ya1' A' A B NA \<in> set evs2; Nonce NB \<notin> used evs2 |]
-  ==> ya2 A B NA NB # evs2 \<in> ya"
+| YA2: "\<lbrakk>evs2 \<in> ya; ya1' A' A B NA \<in> set evs2; Nonce NB \<notin> used evs2\<rbrakk>
+  \<Longrightarrow> ya2 A B NA NB # evs2 \<in> ya"
 
-| YA3: "[| evs3 \<in> ya; ya2' B' A B NA NB \<in> set evs3; Key K \<notin> used evs3 |]
-  ==> ya3 A B NA NB K # evs3 \<in> ya"
+| YA3: "\<lbrakk>evs3 \<in> ya; ya2' B' A B NA NB \<in> set evs3; Key K \<notin> used evs3\<rbrakk>
+  \<Longrightarrow> ya3 A B NA NB K # evs3 \<in> ya"
 
-| YA4: "[| evs4 \<in> ya; ya1 A B NA \<in> set evs4; ya3' S Y A B NA NB K \<in> set evs4 |]
-  ==> ya4 A B K NB Y # evs4 \<in> ya"
+| YA4: "\<lbrakk>evs4 \<in> ya; ya1 A B NA \<in> set evs4; ya3' S Y A B NA NB K \<in> set evs4\<rbrakk>
+  \<Longrightarrow> ya4 A B K NB Y # evs4 \<in> ya"
 
 subsection\<open>declarations for tactics\<close>
 
@@ -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 \<in> ya" for ev evs, auto)
+  unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> ya" for ev evs, auto)
 
 lemma ya_has_only_Says' [rule_format]: "evs \<in> ya \<Longrightarrow>
 ev \<in> set evs \<longrightarrow> (\<exists>A B X. ev=Says A B X)"
@@ -96,7 +96,7 @@
 
 subsection\<open>guardedness of KAB\<close>
 
-lemma Guard_KAB [rule_format]: "[| evs \<in> ya; A \<notin> bad; B \<notin> bad |] ==>
+lemma Guard_KAB [rule_format]: "\<lbrakk>evs \<in> ya; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow>
 ya3 A B NA NB K \<in> set evs \<longrightarrow> GuardK K {shrK A,shrK B} (spies evs)" 
 apply (erule ya.induct)
 (* Nil *)
@@ -127,18 +127,18 @@
 subsection\<open>ya2' implies ya1'\<close>
 
 lemma ya2'_parts_imp_ya1'_parts [rule_format]:
-     "[| evs \<in> ya; B \<notin> bad |] ==>
+     "\<lbrakk>evs \<in> ya; B \<notin> bad\<rbrakk> \<Longrightarrow>
       Ciph B \<lbrace>Agent A, Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
       \<lbrace>Agent A, Nonce NA\<rbrace> \<in> 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 \<in> set evs; evs \<in> ya; B \<notin> bad |]
-==> \<lbrace>Agent A, Nonce NA\<rbrace> \<in> spies evs"
+lemma ya2'_imp_ya1'_parts: "\<lbrakk>ya2' B' A B NA NB \<in> set evs; evs \<in> ya; B \<notin> bad\<rbrakk>
+\<Longrightarrow> \<lbrace>Agent A, Nonce NA\<rbrace> \<in> spies evs"
 by (blast dest: Says_imp_spies ya2'_parts_imp_ya1'_parts)
 
 subsection\<open>uniqueness of NB\<close>
 
-lemma NB_is_uniq_in_ya2'_parts [rule_format]: "[| evs \<in> ya; B \<notin> bad; B' \<notin> bad |] ==>
+lemma NB_is_uniq_in_ya2'_parts [rule_format]: "\<lbrakk>evs \<in> ya; B \<notin> bad; B' \<notin> bad\<rbrakk> \<Longrightarrow>
 Ciph B \<lbrace>Agent A, Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
 Ciph B' \<lbrace>Agent A', Nonce NA', Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
 A=A' \<and> B=B' \<and> 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 \<in> set evs;
-ya2' C' A' B' NA' NB \<in> set evs; evs \<in> ya; B \<notin> bad; B' \<notin> bad |]
-==> A=A' \<and> B=B' \<and> NA=NA'"
+lemma NB_is_uniq_in_ya2': "\<lbrakk>ya2' C A B NA NB \<in> set evs;
+ya2' C' A' B' NA' NB \<in> set evs; evs \<in> ya; B \<notin> bad; B' \<notin> bad\<rbrakk>
+\<Longrightarrow> A=A' \<and> B=B' \<and> NA=NA'"
 by (drule NB_is_uniq_in_ya2'_parts, auto dest: Says_imp_spies)
 
 subsection\<open>ya3' implies ya2'\<close>
 
-lemma ya3'_parts_imp_ya2'_parts [rule_format]: "[| evs \<in> ya; A \<notin> bad |] ==>
+lemma ya3'_parts_imp_ya2'_parts [rule_format]: "\<lbrakk>evs \<in> ya; A \<notin> bad\<rbrakk> \<Longrightarrow>
 Ciph A \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs)
 \<longrightarrow> Ciph B \<lbrace>Agent A, Nonce NA, Nonce NB\<rbrace> \<in> 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 \<in> ya; A \<notin> bad |] ==>
+lemma ya3'_parts_imp_ya2' [rule_format]: "\<lbrakk>evs \<in> ya; A \<notin> bad\<rbrakk> \<Longrightarrow>
 Ciph A \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B'. ya2' B' A B NA NB \<in> 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 \<in> set evs; evs \<in> ya; A \<notin> bad |]
-==> (\<exists>B'. ya2' B' A B NA NB \<in> set evs)"
+lemma ya3'_imp_ya2': "\<lbrakk>ya3' S Y A B NA NB K \<in> set evs; evs \<in> ya; A \<notin> bad\<rbrakk>
+\<Longrightarrow> (\<exists>B'. ya2' B' A B NA NB \<in> set evs)"
 by (drule ya3'_parts_imp_ya2', auto dest: Says_imp_spies)
 
 subsection\<open>ya3' implies ya3\<close>
 
-lemma ya3'_parts_imp_ya3 [rule_format]: "[| evs \<in> ya; A \<notin> bad |] ==>
+lemma ya3'_parts_imp_ya3 [rule_format]: "\<lbrakk>evs \<in> ya; A \<notin> bad\<rbrakk> \<Longrightarrow>
 Ciph A \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace> \<in> parts(spies evs)
 \<longrightarrow> ya3 A B NA NB K \<in> 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 \<in> set evs; evs \<in> ya; A \<notin> bad |]
-==> ya3 A B NA NB K \<in> set evs"
+lemma ya3'_imp_ya3: "\<lbrakk>ya3' S Y A B NA NB K \<in> set evs; evs \<in> ya; A \<notin> bad\<rbrakk>
+\<Longrightarrow> ya3 A B NA NB K \<in> set evs"
 by (blast dest: Says_imp_spies ya3'_parts_imp_ya3)
 
 subsection\<open>guardedness of NB\<close>
@@ -195,7 +195,7 @@
 definition ya_keys :: "agent \<Rightarrow> agent \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> event list \<Rightarrow> key set" where
 "ya_keys A B NA NB evs \<equiv> {shrK A,shrK B} \<union> {K. ya3 A B NA NB K \<in> set evs}"
 
-lemma Guard_NB [rule_format]: "[| evs \<in> ya; A \<notin> bad; B \<notin> bad |] ==>
+lemma Guard_NB [rule_format]: "\<lbrakk>evs \<in> ya; A \<notin> bad; B \<notin> bad\<rbrakk> \<Longrightarrow>
 ya2 A B NA NB \<in> set evs \<longrightarrow> Guard NB (ya_keys A B NA NB evs) (spies evs)"
 apply (erule ya.induct)
 (* Nil *)
--- a/src/HOL/Auth/Guard/List_Msg.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/List_Msg.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -17,8 +17,8 @@
 
 subsubsection\<open>induction principle\<close>
 
-lemma lmsg_induct: "[| !!x. not_MPair x ==> P x; !!x l. P l ==> P (cons x l) |]
-==> P l"
+lemma lmsg_induct: "\<lbrakk>!!x. not_MPair x \<Longrightarrow> P x; !!x l. P l \<Longrightarrow> P (cons x l)\<rbrakk>
+\<Longrightarrow> P l"
 by (induct l) auto
 
 subsubsection\<open>head\<close>
@@ -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) \<Longrightarrow> 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 \<Longrightarrow> ith (l,0) = head l"
 by (cases l) auto
 
 subsubsection\<open>insertion\<close>
@@ -114,14 +114,14 @@
 inductive_set agl :: "msg set"
 where
   Nil[intro]: "nil \<in> agl"
-| Cons[intro]: "[| A \<in> agent; I \<in> agl |] ==> cons (Agent A) I \<in> agl"
+| Cons[intro]: "\<lbrakk>A \<in> agent; I \<in> agl\<rbrakk> \<Longrightarrow> cons (Agent A) I \<in> agl"
 
 subsubsection\<open>basic facts about agent lists\<close>
 
 lemma del_in_agl [intro]: "I \<in> agl \<Longrightarrow> del (a,I) \<in> agl"
 by (erule agl.induct, auto)
 
-lemma app_in_agl [intro]: "[| I \<in> agl; J \<in> agl |] ==> app (I,J) \<in> agl"
+lemma app_in_agl [intro]: "\<lbrakk>I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow> app (I,J) \<in> agl"
 by (erule agl.induct, auto)
 
 lemma no_Key_in_agl: "I \<in> agl \<Longrightarrow> Key K \<notin> parts {I}"
@@ -130,18 +130,18 @@
 lemma no_Nonce_in_agl: "I \<in> agl \<Longrightarrow> Nonce n \<notin> parts {I}"
 by (erule agl.induct, auto)
 
-lemma no_Key_in_appdel: "[| I \<in> agl; J \<in> agl |] ==>
+lemma no_Key_in_appdel: "\<lbrakk>I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow>
 Key K \<notin> parts {app (J, del (Agent B, I))}"
 by (rule no_Key_in_agl, auto)
 
-lemma no_Nonce_in_appdel: "[| I \<in> agl; J \<in> agl |] ==>
+lemma no_Nonce_in_appdel: "\<lbrakk>I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow>
 Nonce n \<notin> parts {app (J, del (Agent B, I))}"
 by (rule no_Nonce_in_agl, auto)
 
 lemma no_Crypt_in_agl: "I \<in> agl \<Longrightarrow> Crypt K X \<notin> parts {I}"
 by (erule agl.induct, auto)
 
-lemma no_Crypt_in_appdel: "[| I \<in> agl; J \<in> agl |] ==>
+lemma no_Crypt_in_appdel: "\<lbrakk>I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow>
 Crypt K X \<notin> parts {app (J, del (Agent B,I))}"
 by (rule no_Crypt_in_agl, auto)
 
--- a/src/HOL/Auth/Guard/P1.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/P1.thy	Fri Oct 14 10:30:37 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'"
+\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 by (auto simp: prom_def)
 
 lemma Nonce_in_prom [iff]: "Nonce ofr \<in> 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'"
+\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 by (auto simp: pro_def dest: prom_inj)
 
 subsubsection\<open>protocol\<close>
@@ -143,13 +143,13 @@
 
   Nil: "[] \<in> p1"
 
-| Fake: "[| evsf \<in> p1; X \<in> synth (analz (spies evsf)) |] ==> Says Spy B X # evsf \<in> p1"
+| Fake: "\<lbrakk>evsf \<in> p1; X \<in> synth (analz (spies evsf))\<rbrakk> \<Longrightarrow> Says Spy B X # evsf \<in> p1"
 
-| Request: "[| evsr \<in> p1; Nonce n \<notin> used evsr; I \<in> agl |] ==> req A r n I B # evsr \<in> p1"
+| Request: "\<lbrakk>evsr \<in> p1; Nonce n \<notin> used evsr; I \<in> agl\<rbrakk> \<Longrightarrow> req A r n I B # evsr \<in> p1"
 
-| Propose: "[| evsp \<in> p1; Says A' B \<lbrace>Agent A,Number r,I,cons M L\<rbrace> \<in> set evsp;
+| Propose: "\<lbrakk>evsp \<in> p1; Says A' B \<lbrace>Agent A,Number r,I,cons M L\<rbrace> \<in> set evsp;
   I \<in> agl; J \<in> agl; isin (Agent C, app (J, del (Agent B, I)));
-  Nonce ofr \<notin> used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp \<in> p1"
+  Nonce ofr \<notin> used evsp\<rbrakk> \<Longrightarrow> pro B ofr A r I (cons M L) J C # evsp \<in> p1"
 
 subsubsection\<open>Composition of Traces\<close>
 
@@ -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 \<in> p1" for ev evs, auto)
+  unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> p1" for ev evs, auto)
 
 lemma p1_has_only_Says' [rule_format]: "evs \<in> p1 \<Longrightarrow>
 ev \<in> set evs \<longrightarrow> (\<exists>A B X. ev=Says A B X)"
@@ -373,8 +373,8 @@
 subsubsection\<open>private keys are safe\<close>
 
 lemma priK_parts_Friend_imp_bad [rule_format,dest]:
-     "[| evs \<in> p1; Friend B \<noteq> A |]
-      ==> (Key (priK A) \<in> parts (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)"
+     "\<lbrakk>evs \<in> p1; Friend B \<noteq> A\<rbrakk>
+      \<Longrightarrow> (Key (priK A) \<in> parts (knows (Friend B) evs)) \<longrightarrow> (A \<in> 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 \<in> p1; Friend B \<noteq> A |]
-==> (Key (priK A) \<in> analz (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)"
+     "\<lbrakk>evs \<in> p1; Friend B \<noteq> A\<rbrakk>
+\<Longrightarrow> (Key (priK A) \<in> analz (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)"
 by auto
 
-lemma priK_notin_knows_max_Friend: "[| evs \<in> p1; A \<notin> bad; A \<noteq> Friend C |]
-==> Key (priK A) \<notin> analz (knows_max (Friend C) evs)"
+lemma priK_notin_knows_max_Friend: "\<lbrakk>evs \<in> p1; A \<notin> bad; A \<noteq> Friend C\<rbrakk>
+\<Longrightarrow> Key (priK A) \<notin> 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 \<in> agl \<Longrightarrow> I \<in> guard n Ks"
 by (erule agl.induct, auto)
 
-lemma Says_to_knows_max'_guard: "[| Says A' C \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
-Guard n Ks (knows_max' C evs) |] ==> L \<in> guard n Ks"
+lemma Says_to_knows_max'_guard: "\<lbrakk>Says A' C \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
+Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks"
 by (auto dest: Says_to_knows_max')
 
-lemma Says_from_knows_max'_guard: "[| Says C A' \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
-Guard n Ks (knows_max' C evs) |] ==> L \<in> guard n Ks"
+lemma Says_from_knows_max'_guard: "\<lbrakk>Says C A' \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
+Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks"
 by (auto dest: Says_from_knows_max')
 
-lemma Says_Nonce_not_used_guard: "[| Says A' B \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
-Nonce n \<notin> used evs |] ==> L \<in> guard n Ks"
+lemma Says_Nonce_not_used_guard: "\<lbrakk>Says A' B \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
+Nonce n \<notin> used evs\<rbrakk> \<Longrightarrow> L \<in> guard n Ks"
 by (drule not_used_not_parts, auto)
 
 subsubsection\<open>guardedness of messages\<close>
@@ -432,16 +432,16 @@
 lemma reqm_guard [intro]: "I \<in> agl \<Longrightarrow> reqm A r n' I B \<in> guard n {priK A}"
 by (case_tac "n'=n", auto simp: reqm_def)
 
-lemma reqm_guard_Nonce_neq [intro]: "[| n \<noteq> n'; I \<in> agl |]
-==> reqm A' r n' I B \<in> guard n {priK A}"
+lemma reqm_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> n'; I \<in> agl\<rbrakk>
+\<Longrightarrow> reqm A' r n' I B \<in> guard n {priK A}"
 by (auto simp: reqm_def)
 
-lemma prom_guard [intro]: "[| I \<in> agl; J \<in> agl; L \<in> guard n {priK A} |]
-==> prom B ofr A r I L J C \<in> guard n {priK A}"
+lemma prom_guard [intro]: "\<lbrakk>I \<in> agl; J \<in> agl; L \<in> guard n {priK A}\<rbrakk>
+\<Longrightarrow> prom B ofr A r I L J C \<in> guard n {priK A}"
 by (auto simp: prom_def)
 
-lemma prom_guard_Nonce_neq [intro]: "[| n \<noteq> ofr; I \<in> agl; J \<in> agl;
-L \<in> guard n {priK A} |] ==> prom B ofr A' r I L J C \<in> guard n {priK A}"
+lemma prom_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> ofr; I \<in> agl; J \<in> agl;
+L \<in> guard n {priK A}\<rbrakk> \<Longrightarrow> prom B ofr A' r I L J C \<in> guard n {priK A}"
 by (auto simp: prom_def)
 
 subsubsection\<open>Nonce uniqueness\<close>
@@ -452,25 +452,25 @@
 lemma uniq_Nonce_in_anchor [dest]: "Nonce k \<in> parts {anchor A n B} \<Longrightarrow> k=n"
 by (auto simp: anchor_def chain_def sign_def)
 
-lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k \<in> parts {reqm A r n I B};
-I \<in> agl |] ==> k=n"
+lemma uniq_Nonce_in_reqm [dest]: "\<lbrakk>Nonce k \<in> parts {reqm A r n I B};
+I \<in> agl\<rbrakk> \<Longrightarrow> k=n"
 by (auto simp: reqm_def dest: no_Nonce_in_agl)
 
-lemma uniq_Nonce_in_prom [dest]: "[| Nonce k \<in> parts {prom B ofr A r I L J C};
-I \<in> agl; J \<in> agl; Nonce k \<notin> parts {L} |] ==> k=ofr"
+lemma uniq_Nonce_in_prom [dest]: "\<lbrakk>Nonce k \<in> parts {prom B ofr A r I L J C};
+I \<in> agl; J \<in> agl; Nonce k \<notin> parts {L}\<rbrakk> \<Longrightarrow> k=ofr"
 by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel)
 
 subsubsection\<open>requests are guarded\<close>
 
-lemma req_imp_Guard [rule_format]: "[| evs \<in> p1; A \<notin> bad |] ==>
+lemma req_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow>
 req A r n I B \<in> set evs \<longrightarrow> 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 \<in> p1; A \<notin> bad; req A r n I B \<in> set evs |]
-==> Guard n {priK A} (knows_max (Friend C) evs)"
+lemma req_imp_Guard_Friend: "\<lbrakk>evs \<in> p1; A \<notin> bad; req A r n I B \<in> set evs\<rbrakk>
+\<Longrightarrow> 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\<open>propositions are guarded\<close>
 
-lemma pro_imp_Guard [rule_format]: "[| evs \<in> p1; B \<notin> bad; A \<notin> bad |] ==>
+lemma pro_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad\<rbrakk> \<Longrightarrow>
 pro B ofr A r I (cons M L) J C \<in> set evs \<longrightarrow> 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 \<in> p1; B \<notin> bad; A \<notin> bad;
-pro B ofr A r I (cons M L) J C \<in> set evs |]
-==> Guard ofr {priK A} (knows_max (Friend D) evs)"
+lemma pro_imp_Guard_Friend: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad;
+pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk>
+\<Longrightarrow> 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\<open>data confidentiality:
 no one other than the originator can decrypt the offers\<close>
 
-lemma Nonce_req_notin_spies: "[| evs \<in> p1; req A r n I B \<in> set evs; A \<notin> bad |]
-==> Nonce n \<notin> analz (spies evs)"
+lemma Nonce_req_notin_spies: "\<lbrakk>evs \<in> p1; req A r n I B \<in> set evs; A \<notin> bad\<rbrakk>
+\<Longrightarrow> Nonce n \<notin> analz (spies evs)"
 by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
 
-lemma Nonce_req_notin_knows_max_Friend: "[| evs \<in> p1; req A r n I B \<in> set evs;
-A \<notin> bad; A \<noteq> Friend C |] ==> Nonce n \<notin> analz (knows_max (Friend C) evs)"
+lemma Nonce_req_notin_knows_max_Friend: "\<lbrakk>evs \<in> p1; req A r n I B \<in> set evs;
+A \<notin> bad; A \<noteq> Friend C\<rbrakk> \<Longrightarrow> Nonce n \<notin> 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 \<in> p1; B \<notin> bad; A \<notin> bad;
-pro B ofr A r I (cons M L) J C \<in> set evs |] ==> Nonce ofr \<notin> analz (spies evs)"
+lemma Nonce_pro_notin_spies: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad;
+pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> \<Longrightarrow> Nonce ofr \<notin> analz (spies evs)"
 by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
 
-lemma Nonce_pro_notin_knows_max_Friend: "[| evs \<in> p1; B \<notin> bad; A \<notin> bad;
-A \<noteq> Friend D; pro B ofr A r I (cons M L) J C \<in> set evs |]
-==> Nonce ofr \<notin> analz (knows_max (Friend D) evs)"
+lemma Nonce_pro_notin_knows_max_Friend: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad;
+A \<noteq> Friend D; pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk>
+\<Longrightarrow> Nonce ofr \<notin> 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\<open>non repudiability:
 an offer signed by B has been sent by B\<close>
 
-lemma Crypt_reqm: "[| Crypt (priK A) X \<in> parts {reqm A' r n I B}; I \<in> agl |] ==> A=A'"
+lemma Crypt_reqm: "\<lbrakk>Crypt (priK A) X \<in> parts {reqm A' r n I B}; I \<in> agl\<rbrakk> \<Longrightarrow> 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 \<in> parts {prom B ofr A' r I L J C};
-I \<in> agl; J \<in> agl |] ==> A=B \<or> Crypt (priK A) X \<in> parts {L}"
+lemma Crypt_prom: "\<lbrakk>Crypt (priK A) X \<in> parts {prom B ofr A' r I L J C};
+I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow> A=B \<or> Crypt (priK A) X \<in> 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 \<in> p1; A \<notin> bad |] ==> Crypt (priK A) X \<in> parts (spies evs)
+lemma Crypt_safeness: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow> Crypt (priK A) X \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> Crypt (priK A) X \<in> 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 \<in> p1; A \<notin> bad |] ==>
+lemma Crypt_Hash_imp_sign: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow>
 Crypt (priK A) (Hash X) \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})"
 apply (erule p1.induct)
@@ -630,7 +630,7 @@
 apply (blast del: MPair_parts)+
 done
 
-lemma sign_safeness: "[| evs \<in> p1; A \<notin> bad |] ==> sign A X \<in> parts (spies evs)
+lemma sign_safeness: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow> sign A X \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})"
 apply (clarify, simp add: sign_def, frule parts.Snd)
 apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def])
--- a/src/HOL/Auth/Guard/P2.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/P2.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -102,7 +102,7 @@
 app (J, del (Agent B, I)), cons (chain B ofr A L C) L\<rbrace>"
 
 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'"
+\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 by (auto simp: prom_def)
 
 lemma Nonce_in_prom [iff]: "Nonce ofr \<in> 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'"
+\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 by (auto simp: pro_def dest: prom_inj)
 
 subsubsection\<open>protocol\<close>
@@ -123,13 +123,13 @@
 
   Nil: "[] \<in> p2"
 
-| Fake: "[| evsf \<in> p2; X \<in> synth (analz (spies evsf)) |] ==> Says Spy B X # evsf \<in> p2"
+| Fake: "\<lbrakk>evsf \<in> p2; X \<in> synth (analz (spies evsf))\<rbrakk> \<Longrightarrow> Says Spy B X # evsf \<in> p2"
 
-| Request: "[| evsr \<in> p2; Nonce n \<notin> used evsr; I \<in> agl |] ==> req A r n I B # evsr \<in> p2"
+| Request: "\<lbrakk>evsr \<in> p2; Nonce n \<notin> used evsr; I \<in> agl\<rbrakk> \<Longrightarrow> req A r n I B # evsr \<in> p2"
 
-| Propose: "[| evsp \<in> p2; Says A' B \<lbrace>Agent A,Number r,I,cons M L\<rbrace> \<in> set evsp;
+| Propose: "\<lbrakk>evsp \<in> p2; Says A' B \<lbrace>Agent A,Number r,I,cons M L\<rbrace> \<in> set evsp;
   I \<in> agl; J \<in> agl; isin (Agent C, app (J, del (Agent B, I)));
-  Nonce ofr \<notin> used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp \<in> p2"
+  Nonce ofr \<notin> used evsp\<rbrakk> \<Longrightarrow> pro B ofr A r I (cons M L) J C # evsp \<in> p2"
 
 subsubsection\<open>valid offer lists\<close>
 
@@ -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 \<in> p2" for ev evs, auto)
+  unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> p2" for ev evs, auto)
 
 lemma p2_has_only_Says' [rule_format]: "evs \<in> p2 \<Longrightarrow>
 ev \<in> set evs \<longrightarrow> (\<exists>A B X. ev=Says A B X)"
@@ -284,8 +284,8 @@
 subsection\<open>private keys are safe\<close>
 
 lemma priK_parts_Friend_imp_bad [rule_format,dest]:
-     "[| evs \<in> p2; Friend B \<noteq> A |]
-      ==> (Key (priK A) \<in> parts (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)"
+     "\<lbrakk>evs \<in> p2; Friend B \<noteq> A\<rbrakk>
+      \<Longrightarrow> (Key (priK A) \<in> parts (knows (Friend B) evs)) \<longrightarrow> (A \<in> 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 \<in> p2; Friend B \<noteq> A |]
-==> (Key (priK A) \<in> analz (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)"
+     "\<lbrakk>evs \<in> p2; Friend B \<noteq> A\<rbrakk>
+\<Longrightarrow> (Key (priK A) \<in> analz (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)"
 by auto
 
 lemma priK_notin_knows_max_Friend:
-     "[| evs \<in> p2; A \<notin> bad; A \<noteq> Friend C |]
-      ==> Key (priK A) \<notin> analz (knows_max (Friend C) evs)"
+     "\<lbrakk>evs \<in> p2; A \<notin> bad; A \<noteq> Friend C\<rbrakk>
+      \<Longrightarrow> Key (priK A) \<notin> 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 \<in> agl \<Longrightarrow> I \<in> guard n Ks"
 by (erule agl.induct, auto)
 
-lemma Says_to_knows_max'_guard: "[| Says A' C \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
-Guard n Ks (knows_max' C evs) |] ==> L \<in> guard n Ks"
+lemma Says_to_knows_max'_guard: "\<lbrakk>Says A' C \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
+Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks"
 by (auto dest: Says_to_knows_max')
 
-lemma Says_from_knows_max'_guard: "[| Says C A' \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
-Guard n Ks (knows_max' C evs) |] ==> L \<in> guard n Ks"
+lemma Says_from_knows_max'_guard: "\<lbrakk>Says C A' \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
+Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks"
 by (auto dest: Says_from_knows_max')
 
-lemma Says_Nonce_not_used_guard: "[| Says A' B \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
-Nonce n \<notin> used evs |] ==> L \<in> guard n Ks"
+lemma Says_Nonce_not_used_guard: "\<lbrakk>Says A' B \<lbrace>A'',r,I,L\<rbrace> \<in> set evs;
+Nonce n \<notin> used evs\<rbrakk> \<Longrightarrow> L \<in> guard n Ks"
 by (drule not_used_not_parts, auto)
 
 subsection\<open>guardedness of messages\<close>
@@ -344,16 +344,16 @@
 lemma reqm_guard [intro]: "I \<in> agl \<Longrightarrow> reqm A r n' I B \<in> guard n {priK A}"
 by (case_tac "n'=n", auto simp: reqm_def)
 
-lemma reqm_guard_Nonce_neq [intro]: "[| n \<noteq> n'; I \<in> agl |]
-==> reqm A' r n' I B \<in> guard n {priK A}"
+lemma reqm_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> n'; I \<in> agl\<rbrakk>
+\<Longrightarrow> reqm A' r n' I B \<in> guard n {priK A}"
 by (auto simp: reqm_def)
 
-lemma prom_guard [intro]: "[| I \<in> agl; J \<in> agl; L \<in> guard n {priK A} |]
-==> prom B ofr A r I L J C \<in> guard n {priK A}"
+lemma prom_guard [intro]: "\<lbrakk>I \<in> agl; J \<in> agl; L \<in> guard n {priK A}\<rbrakk>
+\<Longrightarrow> prom B ofr A r I L J C \<in> guard n {priK A}"
 by (auto simp: prom_def)
 
-lemma prom_guard_Nonce_neq [intro]: "[| n \<noteq> ofr; I \<in> agl; J \<in> agl;
-L \<in> guard n {priK A} |] ==> prom B ofr A' r I L J C \<in> guard n {priK A}"
+lemma prom_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> ofr; I \<in> agl; J \<in> agl;
+L \<in> guard n {priK A}\<rbrakk> \<Longrightarrow> prom B ofr A' r I L J C \<in> guard n {priK A}"
 by (auto simp: prom_def)
 
 subsection\<open>Nonce uniqueness\<close>
@@ -364,25 +364,25 @@
 lemma uniq_Nonce_in_anchor [dest]: "Nonce k \<in> parts {anchor A n B} \<Longrightarrow> k=n"
 by (auto simp: anchor_def chain_def sign_def)
 
-lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k \<in> parts {reqm A r n I B};
-I \<in> agl |] ==> k=n"
+lemma uniq_Nonce_in_reqm [dest]: "\<lbrakk>Nonce k \<in> parts {reqm A r n I B};
+I \<in> agl\<rbrakk> \<Longrightarrow> k=n"
 by (auto simp: reqm_def dest: no_Nonce_in_agl)
 
-lemma uniq_Nonce_in_prom [dest]: "[| Nonce k \<in> parts {prom B ofr A r I L J C};
-I \<in> agl; J \<in> agl; Nonce k \<notin> parts {L} |] ==> k=ofr"
+lemma uniq_Nonce_in_prom [dest]: "\<lbrakk>Nonce k \<in> parts {prom B ofr A r I L J C};
+I \<in> agl; J \<in> agl; Nonce k \<notin> parts {L}\<rbrakk> \<Longrightarrow> k=ofr"
 by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel)
 
 subsection\<open>requests are guarded\<close>
 
-lemma req_imp_Guard [rule_format]: "[| evs \<in> p2; A \<notin> bad |] ==>
+lemma req_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow>
 req A r n I B \<in> set evs \<longrightarrow> 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 \<in> p2; A \<notin> bad; req A r n I B \<in> set evs |]
-==> Guard n {priK A} (knows_max (Friend C) evs)"
+lemma req_imp_Guard_Friend: "\<lbrakk>evs \<in> p2; A \<notin> bad; req A r n I B \<in> set evs\<rbrakk>
+\<Longrightarrow> 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\<open>propositions are guarded\<close>
 
-lemma pro_imp_Guard [rule_format]: "[| evs \<in> p2; B \<notin> bad; A \<notin> bad |] ==>
+lemma pro_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad\<rbrakk> \<Longrightarrow>
 pro B ofr A r I (cons M L) J C \<in> set evs \<longrightarrow> 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 \<in> p2; B \<notin> bad; A \<notin> bad;
-pro B ofr A r I (cons M L) J C \<in> set evs |]
-==> Guard ofr {priK A} (knows_max (Friend D) evs)"
+lemma pro_imp_Guard_Friend: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad;
+pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk>
+\<Longrightarrow> 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\<open>data confidentiality:
 no one other than the originator can decrypt the offers\<close>
 
-lemma Nonce_req_notin_spies: "[| evs \<in> p2; req A r n I B \<in> set evs; A \<notin> bad |]
-==> Nonce n \<notin> analz (spies evs)"
+lemma Nonce_req_notin_spies: "\<lbrakk>evs \<in> p2; req A r n I B \<in> set evs; A \<notin> bad\<rbrakk>
+\<Longrightarrow> Nonce n \<notin> analz (spies evs)"
 by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
 
-lemma Nonce_req_notin_knows_max_Friend: "[| evs \<in> p2; req A r n I B \<in> set evs;
-A \<notin> bad; A \<noteq> Friend C |] ==> Nonce n \<notin> analz (knows_max (Friend C) evs)"
+lemma Nonce_req_notin_knows_max_Friend: "\<lbrakk>evs \<in> p2; req A r n I B \<in> set evs;
+A \<notin> bad; A \<noteq> Friend C\<rbrakk> \<Longrightarrow> Nonce n \<notin> 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 \<in> p2; B \<notin> bad; A \<notin> bad;
-pro B ofr A r I (cons M L) J C \<in> set evs |] ==> Nonce ofr \<notin> analz (spies evs)"
+lemma Nonce_pro_notin_spies: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad;
+pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> \<Longrightarrow> Nonce ofr \<notin> analz (spies evs)"
 by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
 
-lemma Nonce_pro_notin_knows_max_Friend: "[| evs \<in> p2; B \<notin> bad; A \<notin> bad;
-A \<noteq> Friend D; pro B ofr A r I (cons M L) J C \<in> set evs |]
-==> Nonce ofr \<notin> analz (knows_max (Friend D) evs)"
+lemma Nonce_pro_notin_knows_max_Friend: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad;
+A \<noteq> Friend D; pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk>
+\<Longrightarrow> Nonce ofr \<notin> 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\<open>forward privacy:
 only the originator can know the identity of the shops\<close>
 
-lemma forward_privacy_Spy: "[| evs \<in> p2; B \<notin> bad; A \<notin> bad;
-pro B ofr A r I (cons M L) J C \<in> set evs |]
-==> sign B (Nonce ofr) \<notin> analz (spies evs)"
+lemma forward_privacy_Spy: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad;
+pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk>
+\<Longrightarrow> sign B (Nonce ofr) \<notin> analz (spies evs)"
 by (auto simp:sign_def dest: Nonce_pro_notin_spies)
 
-lemma forward_privacy_Friend: "[| evs \<in> p2; B \<notin> bad; A \<notin> bad; A \<noteq> Friend D;
-pro B ofr A r I (cons M L) J C \<in> set evs |]
-==> sign B (Nonce ofr) \<notin> analz (knows_max (Friend D) evs)"
+lemma forward_privacy_Friend: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad; A \<noteq> Friend D;
+pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk>
+\<Longrightarrow> sign B (Nonce ofr) \<notin> analz (knows_max (Friend D) evs)"
 by (auto simp:sign_def dest:Nonce_pro_notin_knows_max_Friend )
 
 subsection\<open>non repudiability: an offer signed by B has been sent by B\<close>
 
-lemma Crypt_reqm: "[| Crypt (priK A) X \<in> parts {reqm A' r n I B}; I \<in> agl |] ==> A=A'"
+lemma Crypt_reqm: "\<lbrakk>Crypt (priK A) X \<in> parts {reqm A' r n I B}; I \<in> agl\<rbrakk> \<Longrightarrow> 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 \<in> parts {prom B ofr A' r I L J C};
-I \<in> agl; J \<in> agl |] ==> A=B | Crypt (priK A) X \<in> parts {L}"
+lemma Crypt_prom: "\<lbrakk>Crypt (priK A) X \<in> parts {prom B ofr A' r I L J C};
+I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow> A=B | Crypt (priK A) X \<in> 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 \<in> p2; A \<notin> bad |] ==> Crypt (priK A) X \<in> parts (spies evs)
+lemma Crypt_safeness: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow> Crypt (priK A) X \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs & Crypt (priK A) X \<in> 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 \<in> p2; A \<notin> bad |] ==>
+lemma Crypt_Hash_imp_sign: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow>
 Crypt (priK A) (Hash X) \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})"
 apply (erule p2.induct)
@@ -554,7 +554,7 @@
 apply (blast del: MPair_parts)+
 done
 
-lemma sign_safeness: "[| evs \<in> p2; A \<notin> bad |] ==> sign A X \<in> parts (spies evs)
+lemma sign_safeness: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow> sign A X \<in> parts (spies evs)
 \<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})"
 apply (clarify, simp add: sign_def, frule parts.Snd)
 apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def])
--- a/src/HOL/Auth/Guard/Proto.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Guard/Proto.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -52,8 +52,8 @@
 (\<exists>k. Nonce k \<in> parts {X} \<and> nonce s k = n)"
 by (induct X, simp_all, blast)
 
-lemma wdef_Nonce: "[| Nonce n \<in> parts {apm s X}; R \<in> p; msg' R = X; wdef p;
-Nonce n \<notin> parts (apm s `(msg `(fst R))) |] ==>
+lemma wdef_Nonce: "\<lbrakk>Nonce n \<in> parts {apm s X}; R \<in> p; msg' R = X; wdef p;
+Nonce n \<notin> parts (apm s `(msg `(fst R)))\<rbrakk> \<Longrightarrow>
 (\<exists>k. Nonce k \<in> parts {X} \<and> 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]: "[] \<in> tr p"
 
-| Fake [intro]: "[| evsf \<in> tr p; X \<in> synth (analz (spies evsf)) |]
-  ==> Says Spy B X # evsf \<in> tr p"
+| Fake [intro]: "\<lbrakk>evsf \<in> tr p; X \<in> synth (analz (spies evsf))\<rbrakk>
+  \<Longrightarrow> Says Spy B X # evsf \<in> tr p"
 
-| Proto [intro]: "[| evs \<in> tr p; R \<in> p; ok evs R s |] ==> ap' s R # evs \<in> tr p"
+| Proto [intro]: "\<lbrakk>evs \<in> tr p; R \<in> p; ok evs R s\<rbrakk> \<Longrightarrow> ap' s R # evs \<in> tr p"
 
 subsection\<open>general properties\<close>
 
@@ -117,29 +117,29 @@
 definition has_only_Says' :: "proto => bool" where
 "has_only_Says' p \<equiv> \<forall>R. R \<in> p \<longrightarrow> is_Says (snd R)"
 
-lemma has_only_Says'D: "[| R \<in> p; has_only_Says' p |]
-==> (\<exists>A B X. snd R = Says A B X)"
+lemma has_only_Says'D: "\<lbrakk>R \<in> p; has_only_Says' p\<rbrakk>
+\<Longrightarrow> (\<exists>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 \<Longrightarrow> 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 \<in> tr p |]
-==> (\<exists>A B X. ev = Says A B X)"
+lemma has_only_Says'_in_trD: "\<lbrakk>has_only_Says' p; list @ ev # evs1 \<in> tr p\<rbrakk>
+\<Longrightarrow> (\<exists>A B X. ev = Says A B X)"
 by (drule has_only_Says_tr, auto)
 
-lemma ok_not_used: "[| Nonce n \<notin> used evs; ok evs R s;
-\<forall>x. x \<in> fst R \<longrightarrow> is_Says x |] ==> Nonce n \<notin> parts (apm s `(msg `(fst R)))"
+lemma ok_not_used: "\<lbrakk>Nonce n \<notin> used evs; ok evs R s;
+\<forall>x. x \<in> fst R \<longrightarrow> is_Says x\<rbrakk> \<Longrightarrow> Nonce n \<notin> 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 \<in> tr p; ok evs R s; has_only_Says' p;
-R \<in> p; x \<in> fst R |] ==> is_Says x"
+lemma ok_is_Says: "\<lbrakk>evs' @ ev # evs \<in> tr p; ok evs R s; has_only_Says' p;
+R \<in> p; x \<in> fst R\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> (\<exists>evs1 evs2.
 evs = evs2 @ ap' s R # evs1 \<and> Nonce n \<notin> used evs1 \<and> R \<in> p \<and> ok evs1 R s
 \<and> Nonce n \<in> parts {apm' s R} \<and> apm' s R \<in> guard n Ks)"
-by (unfold fresh_def, blast)
+  unfolding fresh_def by blast
 
-lemma freshI [intro]: "[| Nonce n \<notin> used evs1; R \<in> p; Nonce n \<in> parts {apm' s R};
-ok evs1 R s; apm' s R \<in> guard n Ks |]
-==> fresh p R s n Ks (list @ ap' s R # evs1)"
-by (unfold fresh_def, blast)
+lemma freshI [intro]: "\<lbrakk>Nonce n \<notin> used evs1; R \<in> p; Nonce n \<in> parts {apm' s R};
+ok evs1 R s; apm' s R \<in> guard n Ks\<rbrakk>
+\<Longrightarrow> fresh p R s n Ks (list @ ap' s R # evs1)"
+  unfolding fresh_def by blast
 
-lemma freshI': "[| Nonce n \<notin> used evs1; (l,r) \<in> p;
-Nonce n \<in> parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r) \<in> guard n Ks |]
-==> fresh p (l,r) s n Ks (evs2 @ ap s r # evs1)"
+lemma freshI': "\<lbrakk>Nonce n \<notin> used evs1; (l,r) \<in> p;
+Nonce n \<in> parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r) \<in> guard n Ks\<rbrakk>
+\<Longrightarrow> 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 \<in> used evs"
+lemma fresh_used: "\<lbrakk>fresh p R' s' n Ks evs; has_only_Says' p\<rbrakk>
+\<Longrightarrow> Nonce n \<in> 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 \<in> tr p; wdef p; has_only_Says' p;
-Nonce n \<notin> used evs; R \<in> p; ok evs R s; Nonce n \<in> parts {apm' s R} |]
-==> \<exists>k. k \<in> newn R \<and> nonce s k = n"
+lemma fresh_newn: "\<lbrakk>evs' @ ap' s R # evs \<in> tr p; wdef p; has_only_Says' p;
+Nonce n \<notin> used evs; R \<in> p; ok evs R s; Nonce n \<in> parts {apm' s R}\<rbrakk>
+\<Longrightarrow> \<exists>k. k \<in> newn R \<and> 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 \<in> tr p; wdef p; Nonce n \<notin> used evs;
-Nonce n \<in> parts {msg ev} |] ==> \<exists>R s. R \<in> p \<and> ap' s R = ev"
+lemma fresh_rule: "\<lbrakk>evs' @ ev # evs \<in> tr p; wdef p; Nonce n \<notin> used evs;
+Nonce n \<in> parts {msg ev}\<rbrakk> \<Longrightarrow> \<exists>R s. R \<in> p \<and> ap' s R = ev"
 apply (drule trunc, simp, ind_cases "ev # evs \<in> 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 \<subseteq> Ks; wdef p;
+lemma fresh_ruleD: "\<lbrakk>fresh p R' s' n Ks evs; keys R' s' n evs \<subseteq> Ks; wdef p;
 has_only_Says' p; evs \<in> tr p; \<forall>R k s. nonce s k = n \<longrightarrow> Nonce n \<in> used evs \<longrightarrow>
 R \<in> p \<longrightarrow> k \<in> newn R \<longrightarrow> Nonce n \<in> parts {apm' s R} \<longrightarrow> apm' s R \<in> guard n Ks \<longrightarrow>
-apm' s R \<in> parts (spies evs) \<longrightarrow> keys R s n evs \<subseteq> Ks \<longrightarrow> P |] ==> P"
+apm' s R \<in> parts (spies evs) \<longrightarrow> keys R s n evs \<subseteq> Ks \<longrightarrow> P\<rbrakk> \<Longrightarrow> 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 \<Rightarrow> msg set \<Rightarrow> bool" where
 "safe Ks G \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G"
 
-lemma safeD [dest]: "[| safe Ks G; K \<in> Ks |] ==> Key K \<notin> analz G"
-by (unfold safe_def, blast)
+lemma safeD [dest]: "\<lbrakk>safe Ks G; K \<in> Ks\<rbrakk> \<Longrightarrow> Key K \<notin> 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) \<Longrightarrow> safe Ks G"
+  unfolding safe_def by blast
 
-lemma Guard_safe: "[| Guard n Ks G; safe Ks G |] ==> Nonce n \<notin> analz G"
+lemma Guard_safe: "\<lbrakk>Guard n Ks G; safe Ks G\<rbrakk> \<Longrightarrow> Nonce n \<notin> analz G"
 by (blast dest: Guard_invKey)
 
 subsection\<open>guardedness preservation\<close>
@@ -235,14 +235,14 @@
 Guard n Ks (spies evs) \<longrightarrow> safe Ks (spies evs) \<longrightarrow> fresh p R' s' n Ks evs \<longrightarrow>
 keys R' s' n evs \<subseteq> Ks \<longrightarrow> R \<in> p \<longrightarrow> ok evs R s \<longrightarrow> apm' s R \<in> guard n Ks)"
 
-lemma preservD: "[| preserv p keys n Ks; evs \<in> tr p; Guard n Ks (spies evs);
+lemma preservD: "\<lbrakk>preserv p keys n Ks; evs \<in> tr p; Guard n Ks (spies evs);
 safe Ks (spies evs); fresh p R' s' n Ks evs; R \<in> p; ok evs R s;
-keys R' s' n evs \<subseteq> Ks |] ==> apm' s R \<in> guard n Ks"
-by (unfold preserv_def, blast)
+keys R' s' n evs \<subseteq> Ks\<rbrakk> \<Longrightarrow> apm' s R \<in> guard n Ks"
+  unfolding preserv_def by blast
 
-lemma preservD': "[| preserv p keys n Ks; evs \<in> tr p; Guard n Ks (spies evs);
+lemma preservD': "\<lbrakk>preserv p keys n Ks; evs \<in> tr p; Guard n Ks (spies evs);
 safe Ks (spies evs); fresh p R' s' n Ks evs; (l,Says A B X) \<in> p;
-ok evs (l,Says A B X) s; keys R' s' n evs \<subseteq> Ks |] ==> apm s X \<in> guard n Ks"
+ok evs (l,Says A B X) s; keys R' s' n evs \<subseteq> Ks\<rbrakk> \<Longrightarrow> apm s X \<in> guard n Ks"
 by (drule preservD, simp+)
 
 subsection\<open>monotonic keyfun\<close>
@@ -251,14 +251,14 @@
 "monoton p keys \<equiv> \<forall>R' s' n ev evs. ev # evs \<in> tr p \<longrightarrow>
 keys R' s' n evs \<subseteq> keys R' s' n (ev # evs)"
 
-lemma monotonD [dest]: "[| keys R' s' n (ev # evs) \<subseteq> Ks; monoton p keys;
-ev # evs \<in> tr p |] ==> keys R' s' n evs \<subseteq> Ks"
-by (unfold monoton_def, blast)
+lemma monotonD [dest]: "\<lbrakk>keys R' s' n (ev # evs) \<subseteq> Ks; monoton p keys;
+ev # evs \<in> tr p\<rbrakk> \<Longrightarrow> keys R' s' n evs \<subseteq> Ks"
+  unfolding monoton_def by blast
 
 subsection\<open>guardedness theorem\<close>
 
-lemma Guard_tr [rule_format]: "[| evs \<in> tr p; has_only_Says' p;
-preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy) |] ==>
+lemma Guard_tr [rule_format]: "\<lbrakk>evs \<in> tr p; has_only_Says' p;
+preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy)\<rbrakk> \<Longrightarrow>
 safe Ks (spies evs) \<longrightarrow> fresh p R' s' n Ks evs \<longrightarrow> keys R' s' n evs \<subseteq> Ks \<longrightarrow>
 Guard n Ks (spies evs)"
 apply (erule tr.induct)
@@ -297,18 +297,18 @@
 
 subsection\<open>useful properties for guardedness\<close>
 
-lemma newn_neq_used: "[| Nonce n \<in> used evs; ok evs R s; k \<in> newn R |]
-==> n \<noteq> nonce s k"
+lemma newn_neq_used: "\<lbrakk>Nonce n \<in> used evs; ok evs R s; k \<in> newn R\<rbrakk>
+\<Longrightarrow> n \<noteq> nonce s k"
 by (auto simp: ok_def)
 
-lemma ok_Guard: "[| ok evs R s; Guard n Ks (spies evs); x \<in> fst R; is_Says x |]
-==> apm s (msg x) \<in> parts (spies evs) \<and> apm s (msg x) \<in> guard n Ks"
+lemma ok_Guard: "\<lbrakk>ok evs R s; Guard n Ks (spies evs); x \<in> fst R; is_Says x\<rbrakk>
+\<Longrightarrow> apm s (msg x) \<in> parts (spies evs) \<and> apm s (msg x) \<in> 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 \<in> parts (spies evs); Nonce (nonce s n) \<in> parts {Y};
-ok evs R s |] ==> n \<notin> newn R"
+lemma ok_parts_not_new: "\<lbrakk>Y \<in> parts (spies evs); Nonce (nonce s n) \<in> parts {Y};
+ok evs R s\<rbrakk> \<Longrightarrow> n \<notin> newn R"
 by (auto simp: ok_def dest: not_used_not_spied parts_parts)
 
 subsection\<open>unicity\<close>
@@ -322,19 +322,19 @@
 secret R n s Ks \<in> parts (spies evs) \<longrightarrow> secret R' n' s' Ks \<in> parts (spies evs) \<longrightarrow>
 secret R n s Ks = secret R' n' s' Ks"
 
-lemma uniqD: "[| uniq p secret; evs \<in> tr p; R \<in> p; R' \<in> p; n \<in> newn R; n' \<in> newn R';
+lemma uniqD: "\<lbrakk>uniq p secret; evs \<in> tr p; R \<in> p; R' \<in> p; n \<in> newn R; n' \<in> newn R';
 nonce s n = nonce s' n'; Nonce (nonce s n) \<notin> analz (spies evs);
 Nonce (nonce s n) \<in> parts {apm' s R}; Nonce (nonce s n) \<in> parts {apm' s' R'};
 secret R n s Ks \<in> parts (spies evs); secret R' n' s' Ks \<in> parts (spies evs);
-apm' s R \<in> guard (nonce s n) Ks; apm' s' R' \<in> guard (nonce s n) Ks |] ==>
+apm' s R \<in> guard (nonce s n) Ks; apm' s' R' \<in> guard (nonce s n) Ks\<rbrakk> \<Longrightarrow>
 secret R n s Ks = secret R' n' s' Ks"
-by (unfold uniq_def, blast)
+  unfolding uniq_def by blast
 
 definition ord :: "proto \<Rightarrow> (rule \<Rightarrow> rule \<Rightarrow> bool) \<Rightarrow> bool" where
 "ord p inff \<equiv> \<forall>R R'. R \<in> p \<longrightarrow> R' \<in> p \<longrightarrow> \<not> inff R R' \<longrightarrow> inff R' R"
 
-lemma ordD: "[| ord p inff; \<not> inff R R'; R \<in> p; R' \<in> p |] ==> inff R' R"
-by (unfold ord_def, blast)
+lemma ordD: "\<lbrakk>ord p inff; \<not> inff R R'; R \<in> p; R' \<in> p\<rbrakk> \<Longrightarrow> inff R' R"
+  unfolding ord_def by blast
 
 definition uniq' :: "proto \<Rightarrow> (rule \<Rightarrow> rule \<Rightarrow> bool) \<Rightarrow> secfun \<Rightarrow> bool" where
 "uniq' p inff secret \<equiv> \<forall>evs R R' n n' Ks s s'. R \<in> p \<longrightarrow> R' \<in> p \<longrightarrow>
@@ -345,16 +345,16 @@
 secret R n s Ks \<in> parts (spies evs) \<longrightarrow> secret R' n' s' Ks \<in> parts (spies evs) \<longrightarrow>
 secret R n s Ks = secret R' n' s' Ks"
 
-lemma uniq'D: "[| uniq' p inff secret; evs \<in> tr p; inff R R'; R \<in> p; R' \<in> p; n \<in> newn R;
+lemma uniq'D: "\<lbrakk>uniq' p inff secret; evs \<in> tr p; inff R R'; R \<in> p; R' \<in> p; n \<in> newn R;
 n' \<in> newn R'; nonce s n = nonce s' n'; Nonce (nonce s n) \<notin> analz (spies evs);
 Nonce (nonce s n) \<in> parts {apm' s R}; Nonce (nonce s n) \<in> parts {apm' s' R'};
 secret R n s Ks \<in> parts (spies evs); secret R' n' s' Ks \<in> parts (spies evs);
-apm' s R \<in> guard (nonce s n) Ks; apm' s' R' \<in> guard (nonce s n) Ks |] ==>
+apm' s R \<in> guard (nonce s n) Ks; apm' s' R' \<in> guard (nonce s n) Ks\<rbrakk> \<Longrightarrow>
 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: "\<lbrakk>uniq' p inff secret; ord p inff\<rbrakk> \<Longrightarrow> 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\<open>guardedness for NSL\<close>
 
-lemma "uniq ns secret ==> preserv ns keys n Ks"
-apply (unfold preserv_def)
+lemma "uniq ns secret \<Longrightarrow> 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)
--- a/src/HOL/Auth/KerberosIV.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/KerberosIV.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -295,7 +295,7 @@
 subsection\<open>Lemmas about \<^term>\<open>authKeys\<close>\<close>
 
 lemma authKeys_empty: "authKeys [] = {}"
-apply (unfold authKeys_def)
+unfolding authKeys_def
 apply (simp (no_asm))
 done
 
@@ -304,27 +304,27 @@
    ev \<noteq> Says Kas A (Crypt (shrK A) \<lbrace>akey, Agent Peer, Ta,
               (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, akey, Ta\<rbrace>)\<rbrace>))
        \<Longrightarrow> authKeys (ev # evs) = authKeys evs"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeys_insert:
   "authKeys
      (Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Peer, Number Ta,
       (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K, Number Ta\<rbrace>)\<rbrace>) # evs)
        = insert K (authKeys evs)"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeys_simp:
    "K \<in> authKeys
     (Says Kas A (Crypt (shrK A) \<lbrace>Key K', Agent Peer, Number Ta,
      (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K', Number Ta\<rbrace>)\<rbrace>) # evs)
         \<Longrightarrow> K = K' | K \<in> authKeys evs"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeysI:
    "Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Tgs, Number Ta,
      (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key K, Number Ta\<rbrace>)\<rbrace>) \<in> set evs
         \<Longrightarrow> K \<in> authKeys evs"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeys_used: "K \<in> authKeys evs \<Longrightarrow> Key K \<in> used evs"
 by (simp add: authKeys_def, blast)
@@ -478,7 +478,7 @@
   Ta = CT (before 
            Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
            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 \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
               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:
  "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>) \<in> set evs;
      evs \<in> kerbIV \<rbrakk> \<Longrightarrow> AKcryptSK authK servK evs"
-apply (unfold AKcryptSK_def)
+unfolding AKcryptSK_def
 apply (blast dest: Says_Tgs_message_form)
 done
 
@@ -1033,7 +1033,7 @@
 lemma Auth_fresh_not_AKcryptSK:
      "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbIV \<rbrakk>
       \<Longrightarrow> \<not> AKcryptSK authK servK evs"
-apply (unfold AKcryptSK_def)
+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 \<notin> used evs \<Longrightarrow> \<not> AKcryptSK authK servK evs"
-by (unfold AKcryptSK_def, blast)
+  unfolding AKcryptSK_def by blast
 
 lemma authK_not_AKcryptSK:
      "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, tk\<rbrace>
@@ -1080,7 +1080,7 @@
 text\<open>Long term keys are not issued as servKeys\<close>
 lemma shrK_not_AKcryptSK:
      "evs \<in> kerbIV \<Longrightarrow> \<not> 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 @@
            \<in> set evs;
          authK' \<noteq> authK;  evs \<in> kerbIV \<rbrakk>
       \<Longrightarrow> \<not> AKcryptSK authK' servK evs"
-apply (unfold AKcryptSK_def)
+unfolding AKcryptSK_def
 apply (blast dest: unique_servKeys)
 done
 
--- a/src/HOL/Auth/KerberosIV_Gets.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/KerberosIV_Gets.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -270,27 +270,27 @@
    ev \<noteq> Says Kas A (Crypt (shrK A) \<lbrace>akey, Agent Peer, Ta,
               (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, akey, Ta\<rbrace>)\<rbrace>))
        \<Longrightarrow> authKeys (ev # evs) = authKeys evs"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeys_insert:
   "authKeys
      (Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Peer, Number Ta,
       (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K, Number Ta\<rbrace>)\<rbrace>) # evs)
        = insert K (authKeys evs)"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeys_simp:
    "K \<in> authKeys
     (Says Kas A (Crypt (shrK A) \<lbrace>Key K', Agent Peer, Number Ta,
      (Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K', Number Ta\<rbrace>)\<rbrace>) # evs)
         \<Longrightarrow> K = K' | K \<in> authKeys evs"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeysI:
    "Says Kas A (Crypt (shrK A) \<lbrace>Key K, Agent Tgs, Number Ta,
      (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key K, Number Ta\<rbrace>)\<rbrace>) \<in> set evs
         \<Longrightarrow> K \<in> authKeys evs"
-by (unfold authKeys_def, auto)
+  unfolding authKeys_def by auto
 
 lemma authKeys_used: "K \<in> authKeys evs \<Longrightarrow> Key K \<in> used evs"
 by (simp add: authKeys_def, blast)
@@ -876,7 +876,7 @@
 lemma AKcryptSKI:
  "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>) \<in> set evs;
      evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> AKcryptSK authK servK evs"
-apply (unfold AKcryptSK_def)
+unfolding AKcryptSK_def
 apply (blast dest: Says_Tgs_message_form)
 done
 
@@ -894,7 +894,7 @@
 lemma Auth_fresh_not_AKcryptSK:
      "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbIV_gets \<rbrakk>
       \<Longrightarrow> \<not> AKcryptSK authK servK evs"
-apply (unfold AKcryptSK_def)
+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 \<notin> used evs \<Longrightarrow> \<not> AKcryptSK authK servK evs"
-by (unfold AKcryptSK_def, blast)
+  unfolding AKcryptSK_def by blast
 
 lemma authK_not_AKcryptSK:
      "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, tk\<rbrace>
@@ -948,7 +948,7 @@
 text\<open>Long term keys are not issued as servKeys\<close>
 lemma shrK_not_AKcryptSK:
      "evs \<in> kerbIV_gets \<Longrightarrow> \<not> 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 @@
            \<in> set evs;
          authK' \<noteq> authK;  evs \<in> kerbIV_gets \<rbrakk>
       \<Longrightarrow> \<not> AKcryptSK authK' servK evs"
-apply (unfold AKcryptSK_def)
+unfolding AKcryptSK_def
 by (blast dest: unique_servKeys)
 
 text\<open>Equivalently\<close>
--- a/src/HOL/Auth/KerberosV.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/KerberosV.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -778,7 +778,7 @@
 lemma Auth_fresh_not_AKcryptSK:
      "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbV \<rbrakk>
       \<Longrightarrow> \<not> AKcryptSK authK servK evs"
-apply (unfold AKcryptSK_def)
+unfolding AKcryptSK_def
 apply (erule rev_mp)
 apply (erule kerbV.induct)
 apply (frule_tac [7] Says_ticket_parts)
@@ -821,7 +821,7 @@
 text\<open>Long term keys are not issued as servKeys\<close>
 lemma shrK_not_AKcryptSK:
      "evs \<in> kerbV \<Longrightarrow> \<not> 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)
--- a/src/HOL/Auth/Kerberos_BAN.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -295,7 +295,7 @@
           Tk = CT(before 
                   Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
                   on evs)"
-apply (unfold before_def)
+unfolding before_def
 apply (erule rev_mp)
 apply (erule bankerberos.induct, simp_all add: takeWhile_tail)
 apply auto
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -283,7 +283,7 @@
           Tk = CT(before 
                   Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
                   on evs)"
-apply (unfold before_def)
+unfolding before_def
 apply (erule rev_mp)
 apply (erule bankerb_gets.induct, simp_all)
 txt\<open>We need this simplification only for Message 2\<close>
--- a/src/HOL/Auth/Message.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Message.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -14,7 +14,7 @@
 
 (*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*)
 lemma [simp] : "A \<union> (B \<union> A) = B \<union> A"
-by blast
+  by blast
 
 type_synonym
   key = nat
@@ -71,87 +71,87 @@
   parts :: "msg set \<Rightarrow> msg set"
   for H :: "msg set"
   where
-    Inj [intro]: "X \<in> H ==> X \<in> parts H"
-  | Fst:         "\<lbrace>X,Y\<rbrace> \<in> parts H ==> X \<in> parts H"
-  | Snd:         "\<lbrace>X,Y\<rbrace> \<in> parts H ==> Y \<in> parts H"
-  | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
+    Inj [intro]: "X \<in> H \<Longrightarrow> X \<in> parts H"
+  | Fst:         "\<lbrace>X,Y\<rbrace> \<in> parts H \<Longrightarrow> X \<in> parts H"
+  | Snd:         "\<lbrace>X,Y\<rbrace> \<in> parts H \<Longrightarrow> Y \<in> parts H"
+  | Body:        "Crypt K X \<in> parts H \<Longrightarrow> X \<in> parts H"
 
 
 text\<open>Monotonicity\<close>
-lemma parts_mono: "G \<subseteq> H ==> parts(G) \<subseteq> parts(H)"
-apply auto
-apply (erule parts.induct) 
-apply (blast dest: parts.Fst parts.Snd parts.Body)+
-done
+lemma parts_mono_aux: "\<lbrakk>G \<subseteq> H; X \<in> parts G\<rbrakk> \<Longrightarrow> X \<in> parts H"
+  by (erule parts.induct) (auto dest: parts.Fst parts.Snd parts.Body)
+
+lemma parts_mono: "G \<subseteq> H \<Longrightarrow> parts(G) \<subseteq> parts(H)"
+  using parts_mono_aux by blast
 
 
 text\<open>Equations hold because constructors are injective.\<close>
 lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x\<in>A)"
-by auto
+  by auto
 
 lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
-by auto
+  by auto
 
 lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
-by auto
+  by auto
 
 
 subsubsection\<open>Inverse of keys\<close>
 
 lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
-by (metis invKey)
+  by (metis invKey)
 
 
 subsection\<open>keysFor operator\<close>
 
 lemma keysFor_empty [simp]: "keysFor {} = {}"
-by (unfold keysFor_def, blast)
+    unfolding keysFor_def by blast
 
 lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
-by (unfold keysFor_def, blast)
+    unfolding keysFor_def by blast
 
 lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
-by (unfold keysFor_def, blast)
+    unfolding keysFor_def by blast
 
 text\<open>Monotonicity\<close>
-lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
-by (unfold keysFor_def, blast)
+lemma keysFor_mono: "G \<subseteq> H \<Longrightarrow> keysFor(G) \<subseteq> 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 \<lbrace>X,Y\<rbrace> 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 \<in> H ==> invKey K \<in> keysFor H"
-by (unfold keysFor_def, blast)
+lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H \<Longrightarrow> invKey K \<in> keysFor H"
+  unfolding keysFor_def by blast
 
 
 subsection\<open>Inductive relation "parts"\<close>
 
 lemma MPair_parts:
-     "[| \<lbrace>X,Y\<rbrace> \<in> parts H;        
-         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
-by (blast dest: parts.Fst parts.Snd) 
+  "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> parts H;        
+         \<lbrakk>X \<in> parts H; Y \<in> parts H\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+  by (blast dest: parts.Fst parts.Snd) 
 
 declare MPair_parts [elim!]  parts.Body [dest!]
 text\<open>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.\<close>
 
 lemma parts_increasing: "H \<subseteq> parts(H)"
-by blast
+  by blast
 
 lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD]
 
+lemma parts_empty_aux: "X \<in> parts{} \<Longrightarrow> 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\<in> parts{} ==> P"
-by simp
+lemma parts_emptyE [elim!]: "X\<in> parts{} \<Longrightarrow> P"
+  by simp
 
 text\<open>WARNING: loops if H = {Y}, therefore must not be repeated!\<close>
-lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
-by (erule parts.induct, fast+)
+lemma parts_singleton: "X \<in> parts H \<Longrightarrow> \<exists>Y\<in>H. X \<in> parts {Y}"
+  by (erule parts.induct, fast+)
 
 
 subsubsection\<open>Unions\<close>
 
 lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> 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 \<union> H) \<subseteq> parts(G) \<union> 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 \<union> H) = parts(G) \<union> 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} \<union> parts H"
-by (metis insert_is_Un parts_Un)
+  by (metis insert_is_Un parts_Un)
 
 text\<open>TWO inserts to avoid looping.  This rewrite is better than nothing.
   But its behaviour can be strange.\<close>
 lemma parts_insert2:
-     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> 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} \<union> parts {Y} \<union> parts H"
+  by (metis Un_commute Un_empty_right Un_insert_right insert_is_Un parts_Un)
 
 lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
-by (intro UN_least parts_mono UN_upper)
+  by (intro UN_least parts_mono UN_upper)
 
 lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>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 (\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts (H x))"
@@ -214,7 +215,7 @@
 lemma parts_image [simp]:
   "parts (f ` A) = (\<Union>x\<in>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) \<subseteq> parts(insert X H)"
-by (blast intro: parts_mono [THEN [2] rev_subsetD])
+  by (blast intro: parts_mono [THEN [2] rev_subsetD])
 
 subsubsection\<open>Idempotence and transitivity\<close>
 
-lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
-by (erule parts.induct, blast+)
+lemma parts_partsD [dest!]: "X\<in> parts (parts H) \<Longrightarrow> X\<in> 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 \<subseteq> parts H) = (G \<subseteq> 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\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
-by (metis parts_subset_iff subsetD)
+lemma parts_trans: "\<lbrakk>X\<in> parts G;  G \<subseteq> parts H\<rbrakk> \<Longrightarrow> X\<in> parts H"
+  by (metis parts_subset_iff subsetD)
 
 text\<open>Cut\<close>
 lemma parts_cut:
-     "[| Y\<in> parts (insert X G);  X\<in> parts H |] ==> Y\<in> parts (G \<union> H)" 
-by (blast intro: parts_trans) 
+  "\<lbrakk>Y\<in> parts (insert X G);  X\<in> parts H\<rbrakk> \<Longrightarrow> Y\<in> parts (G \<union> H)" 
+  by (blast intro: parts_trans) 
 
-lemma parts_cut_eq [simp]: "X\<in> parts H ==> parts (insert X H) = parts H"
-by (metis insert_absorb parts_idem parts_insert)
+lemma parts_cut_eq [simp]: "X\<in> parts H \<Longrightarrow> parts (insert X H) = parts H"
+  by (metis insert_absorb parts_idem parts_insert)
 
 
 subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
@@ -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 \<lbrace>X,Y\<rbrace> H) =  
+  "parts (insert \<lbrace>X,Y\<rbrace> H) =  
           insert \<lbrace>X,Y\<rbrace> (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\<open>In any message, there is an upper bound N on its greatest nonce.\<close>
 lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n \<longrightarrow> Nonce n \<notin> 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 \<comment> \<open>metis works out the necessary sum itself!\<close>
-      by (simp add: parts_insert2) (metis le_trans nat_le_linear)
+  then show ?case \<comment> \<open>metis works out the necessary sum itself!\<close>
+    by (simp add: parts_insert2) (metis le_trans nat_le_linear)
 qed auto
 
 subsection\<open>Inductive relation "analz"\<close>
@@ -331,230 +332,224 @@
   analz :: "msg set \<Rightarrow> msg set"
   for H :: "msg set"
   where
-    Inj [intro,simp]: "X \<in> H ==> X \<in> analz H"
-  | Fst:     "\<lbrace>X,Y\<rbrace> \<in> analz H ==> X \<in> analz H"
-  | Snd:     "\<lbrace>X,Y\<rbrace> \<in> analz H ==> Y \<in> analz H"
+    Inj [intro,simp]: "X \<in> H \<Longrightarrow> X \<in> analz H"
+  | Fst:     "\<lbrace>X,Y\<rbrace> \<in> analz H \<Longrightarrow> X \<in> analz H"
+  | Snd:     "\<lbrace>X,Y\<rbrace> \<in> analz H \<Longrightarrow> Y \<in> analz H"
   | Decrypt [dest]: 
-             "\<lbrakk>Crypt K X \<in> analz H; Key(invKey K) \<in> analz H\<rbrakk> \<Longrightarrow> X \<in> analz H"
+    "\<lbrakk>Crypt K X \<in> analz H; Key(invKey K) \<in> analz H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 
 
 text\<open>Monotonicity; Lemma 1 of Lowe's paper\<close>
-lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
-apply auto
-apply (erule analz.induct) 
-apply (auto dest: analz.Fst analz.Snd) 
-done
+lemma analz_mono_aux: "\<lbrakk>G \<subseteq> H; X \<in> analz G\<rbrakk> \<Longrightarrow> X \<in> analz H"
+  by (erule analz.induct) (auto dest: analz.Fst analz.Snd) 
+
+lemma analz_mono: "G\<subseteq>H \<Longrightarrow> analz(G) \<subseteq> analz(H)"
+  using analz_mono_aux by blast
 
 text\<open>Making it safe speeds up proofs\<close>
 lemma MPair_analz [elim!]:
-     "[| \<lbrace>X,Y\<rbrace> \<in> analz H;        
-             [| X \<in> analz H; Y \<in> analz H |] ==> P   
-          |] ==> P"
-by (blast dest: analz.Fst analz.Snd)
+  "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> analz H;        
+    \<lbrakk>X \<in> analz H; Y \<in> analz H\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+  by (blast dest: analz.Fst analz.Snd)
 
 lemma analz_increasing: "H \<subseteq> analz(H)"
-by blast
+  by blast
+
+lemma analz_into_parts: "X \<in> analz H \<Longrightarrow> X \<in> parts H"
+  by (erule analz.induct) auto
 
 lemma analz_subset_parts: "analz H \<subseteq> 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\<open>General equational properties\<close>
 
 lemma analz_empty [simp]: "analz{} = {}"
-apply safe
-apply (erule analz.induct, blast+)
-done
+  using analz_parts by fastforce
 
 text\<open>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\<close>
 lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> 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) \<subseteq> analz(insert X H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
+  by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
 subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
 
 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\<open>Can only pull out Keys if they are not needed to decrypt the rest\<close>
 lemma analz_insert_Key [simp]: 
-    "K \<notin> keysFor (analz H) ==>   
+  "K \<notin> keysFor (analz H) \<Longrightarrow>   
           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 \<lbrace>X,Y\<rbrace> H) =  
+  "analz (insert \<lbrace>X,Y\<rbrace> H) =  
           insert \<lbrace>X,Y\<rbrace> (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\<open>Can pull out enCrypted message if the Key is not known\<close>
 lemma analz_insert_Crypt:
-     "Key (invKey K) \<notin> 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) \<notin> analz H 
+      \<Longrightarrow> 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) \<in> analz H ==>   
+lemma lemma1: "Key (invKey K) \<in> analz H \<Longrightarrow>   
                analz (insert (Crypt K X) H) \<subseteq>  
                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) \<in> analz H ==>   
+lemma lemma2: "Key (invKey K) \<in> analz H \<Longrightarrow>   
                insert (Crypt K X) (analz (insert X H)) \<subseteq>  
                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) \<in> analz H ==>   
+  "Key (invKey K) \<in> analz H \<Longrightarrow>   
                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\<open>Case analysis: either the message is secure, or it is not! Effective,
 but can cause subgoals to blow up! Use with \<open>if_split\<close>; apparently
 \<open>split_tac\<close> does not cope with patterns such as \<^term>\<open>analz (insert
 (Crypt K X) H)\<close>\<close> 
 lemma analz_Crypt_if [simp]:
-     "analz (insert (Crypt K X) H) =                 
+  "analz (insert (Crypt K X) H) =                 
           (if (Key (invKey K) \<in> 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\<open>This rule supposes "for the sake of argument" that we have the key.\<close>
 lemma analz_insert_Crypt_subset:
-     "analz (insert (Crypt K X) H) \<subseteq>   
+  "analz (insert (Crypt K X) H) \<subseteq>   
            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\<open>Idempotence and transitivity\<close>
 
-lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
-by (erule analz.induct, blast+)
+lemma analz_analzD [dest!]: "X\<in> analz (analz H) \<Longrightarrow> X\<in> 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 \<subseteq> analz H) = (G \<subseteq> 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\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
-by (drule analz_mono, blast)
+lemma analz_trans: "\<lbrakk>X\<in> analz G;  G \<subseteq> analz H\<rbrakk> \<Longrightarrow> X\<in> analz H"
+  by (drule analz_mono, blast)
 
 text\<open>Cut; Lemma 2 of Lowe\<close>
-lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
-by (erule analz_trans, blast)
+lemma analz_cut: "\<lbrakk>Y\<in> analz (insert X H);  X\<in> analz H\<rbrakk> \<Longrightarrow> Y\<in> analz H"
+  by (erule analz_trans, blast)
 
 (*Cut can be proved easily by induction on
-   "Y: analz (insert X H) ==> X: analz H \<longrightarrow> Y: analz H"
+   "Y: analz (insert X H) \<Longrightarrow> X: analz H \<longrightarrow> Y: analz H"
 *)
 
 text\<open>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.\<close>
-lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
-by (metis analz_cut analz_insert_eq_I insert_absorb)
+lemma analz_insert_eq: "X\<in> analz H \<Longrightarrow> analz (insert X H) = analz H"
+  by (metis analz_cut analz_insert_eq_I insert_absorb)
 
 
 text\<open>A congruence rule for "analz"\<close>
 
 lemma analz_subset_cong:
-     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |] 
-      ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
-by (metis Un_mono analz_Un analz_subset_iff subset_trans)
+  "\<lbrakk>analz G \<subseteq> analz G'; analz H \<subseteq> analz H'\<rbrakk> 
+      \<Longrightarrow> analz (G \<union> H) \<subseteq> analz (G' \<union> 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 \<union> H) = analz (G' \<union> H')"
-by (intro equalityI analz_subset_cong, simp_all) 
+  "\<lbrakk>analz G = analz G'; analz H = analz H'\<rbrakk> 
+      \<Longrightarrow> analz (G \<union> H) = analz (G' \<union> 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' \<Longrightarrow> analz(insert X H) = analz(insert X H')"
+  by (force simp only: insert_def intro!: analz_cong)
 
 text\<open>If there are no pairs or encryptions then analz does nothing\<close>
 lemma analz_trivial:
-     "[| \<forall>X Y. \<lbrace>X,Y\<rbrace> \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
-apply safe
-apply (erule analz.induct, blast+)
-done
+  "\<lbrakk>\<forall>X Y. \<lbrace>X,Y\<rbrace> \<notin> H;  \<forall>X K. Crypt K X \<notin> H\<rbrakk> \<Longrightarrow> analz H = H"
+  apply safe
+   apply (erule analz.induct, blast+)
+  done
 
 text\<open>These two are obsolete (with a single Spy) but cost little to prove...\<close>
 lemma analz_UN_analz_lemma:
-     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
-apply (erule analz.induct)
-apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
-done
+  "X\<in> analz (\<Union>i\<in>A. analz (H i)) \<Longrightarrow> X\<in> analz (\<Union>i\<in>A. H i)"
+  apply (erule analz.induct)
+     apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
+  done
 
 lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>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\<open>Inductive relation "synth"\<close>
@@ -568,147 +563,142 @@
   synth :: "msg set => msg set"
   for H :: "msg set"
   where
-    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
+    Inj    [intro]:   "X \<in> H \<Longrightarrow> X \<in> synth H"
   | Agent  [intro]:   "Agent agt \<in> synth H"
   | Number [intro]:   "Number n  \<in> synth H"
-  | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
-  | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> \<lbrace>X,Y\<rbrace> \<in> synth H"
-  | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
+  | Hash   [intro]:   "X \<in> synth H \<Longrightarrow> Hash X \<in> synth H"
+  | MPair  [intro]:   "\<lbrakk>X \<in> synth H;  Y \<in> synth H\<rbrakk> \<Longrightarrow> \<lbrace>X,Y\<rbrace> \<in> synth H"
+  | Crypt  [intro]:   "\<lbrakk>X \<in> synth H;  Key(K) \<in> H\<rbrakk> \<Longrightarrow> Crypt K X \<in> synth H"
 
 text\<open>Monotonicity\<close>
-lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
+lemma synth_mono: "G\<subseteq>H \<Longrightarrow> synth(G) \<subseteq> synth(H)"
   by (auto, erule synth.induct, auto)  
 
 text\<open>NO \<open>Agent_synth\<close>, as any Agent name can be synthesized.  
   The same holds for \<^term>\<open>Number\<close>\<close>
 
 inductive_simps synth_simps [iff]:
- "Nonce n \<in> synth H"
- "Key K \<in> synth H"
- "Hash X \<in> synth H"
- "\<lbrace>X,Y\<rbrace> \<in> synth H"
- "Crypt K X \<in> synth H"
+  "Nonce n \<in> synth H"
+  "Key K \<in> synth H"
+  "Hash X \<in> synth H"
+  "\<lbrace>X,Y\<rbrace> \<in> synth H"
+  "Crypt K X \<in> synth H"
 
 lemma synth_increasing: "H \<subseteq> synth(H)"
-by blast
+  by blast
 
 subsubsection\<open>Unions\<close>
 
 text\<open>Converse fails: we can synth more from the union than from the 
   separate parts, building a compound message using elements of each.\<close>
 lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> 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) \<subseteq> synth(insert X H)"
-by (blast intro: synth_mono [THEN [2] rev_subsetD])
+  by (blast intro: synth_mono [THEN [2] rev_subsetD])
 
 subsubsection\<open>Idempotence and transitivity\<close>
 
-lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
-by (erule synth.induct, auto)
+lemma synth_synthD [dest!]: "X\<in> synth (synth H) \<Longrightarrow> X\<in> 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 \<subseteq> synth H) = (G \<subseteq> 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\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
-by (drule synth_mono, blast)
+lemma synth_trans: "\<lbrakk>X\<in> synth G;  G \<subseteq> synth H\<rbrakk> \<Longrightarrow> X\<in> synth H"
+  by (drule synth_mono, blast)
 
 text\<open>Cut; Lemma 2 of Lowe\<close>
-lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
-by (erule synth_trans, blast)
+lemma synth_cut: "\<lbrakk>Y\<in> synth (insert X H);  X\<in> synth H\<rbrakk> \<Longrightarrow> Y\<in> synth H"
+  by (erule synth_trans, blast)
 
 lemma Crypt_synth_eq [simp]:
-     "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
-by blast
+  "Key K \<notin> H \<Longrightarrow> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
+  by blast
 
 
 lemma keysFor_synth [simp]: 
-    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
-by (unfold keysFor_def, blast)
+  "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
+  unfolding keysFor_def by blast
 
 
 subsubsection\<open>Combinations of parts, analz and synth\<close>
 
 lemma parts_synth [simp]: "parts (synth H) = parts H \<union> 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 \<union> H) = analz (G \<union> 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 \<union> H) = analz (G \<union> H) \<union> 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 \<union> synth H"
-by (metis Un_empty_right analz_synth_Un)
+  by (metis Un_empty_right analz_synth_Un)
 
 
 subsubsection\<open>For reasoning about the Fake rule in traces\<close>
 
-lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
-by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono)
+lemma parts_insert_subset_Un: "X\<in> G \<Longrightarrow> parts(insert X H) \<subseteq> parts G \<union> parts H"
+  by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono)
 
 text\<open>More specifically for Fake. See also \<open>Fake_parts_sing\<close> below\<close>
 lemma Fake_parts_insert:
-     "X \<in> synth (analz H) ==>  
+  "X \<in> synth (analz H) \<Longrightarrow>  
       parts (insert X H) \<subseteq> synth (analz H) \<union> 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:
-     "\<lbrakk>Z \<in> parts (insert X H);  X \<in> synth (analz H)\<rbrakk> 
+  "\<lbrakk>Z \<in> parts (insert X H);  X \<in> synth (analz H)\<rbrakk> 
       \<Longrightarrow> Z \<in> synth (analz H) \<union> parts H"
-by (metis Fake_parts_insert subsetD)
+  by (metis Fake_parts_insert subsetD)
 
 text\<open>\<^term>\<open>H\<close> is sometimes \<^term>\<open>Key ` KK \<union> spies evs\<close>, so can't put 
   \<^term>\<open>G=H\<close>.\<close>
 lemma Fake_analz_insert:
-     "X\<in> synth (analz G) ==>  
+  "X\<in> synth (analz G) \<Longrightarrow>  
       analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
-apply (rule subsetI)
-apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> 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 \<in> analz H \<and> X \<in> parts H) = (X \<in> analz H)"
-by (blast intro: analz_subset_parts [THEN subsetD])
+  "(X \<in> analz H \<and> X \<in> parts H) = (X \<in> analz H)"
+  by (blast intro: analz_subset_parts [THEN subsetD])
 
 lemma analz_disj_parts [simp]:
-     "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
-by (blast intro: analz_subset_parts [THEN subsetD])
+  "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
+  by (blast intro: analz_subset_parts [THEN subsetD])
 
 text\<open>Without this equation, other rules for synth and analz would yield
   redundant cases\<close>
 lemma MPair_synth_analz [iff]:
-     "(\<lbrace>X,Y\<rbrace> \<in> synth (analz H)) =  
-      (X \<in> synth (analz H) \<and> Y \<in> synth (analz H))"
-by blast
+  "\<lbrace>X,Y\<rbrace> \<in> synth (analz H) \<longleftrightarrow> X \<in> synth (analz H) \<and> Y \<in> synth (analz H)"
+  by blast
 
 lemma Crypt_synth_analz:
-     "[| Key K \<in> analz H;  Key (invKey K) \<in> analz H |]  
-       ==> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
-by blast
-
+  "\<lbrakk>Key K \<in> analz H;  Key (invKey K) \<in> analz H\<rbrakk>  
+       \<Longrightarrow> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
+  by blast
 
 lemma Hash_synth_analz [simp]:
-     "X \<notin> synth (analz H)  
-      ==> (Hash\<lbrace>X,Y\<rbrace> \<in> synth (analz H)) = (Hash\<lbrace>X,Y\<rbrace> \<in> analz H)"
-by blast
+  "X \<notin> synth (analz H)  
+      \<Longrightarrow> (Hash\<lbrace>X,Y\<rbrace> \<in> synth (analz H)) = (Hash\<lbrace>X,Y\<rbrace> \<in> analz H)"
+  by blast
 
 
 subsection\<open>HPair: a combination of Hash and MPair\<close>
@@ -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 \<and> Y'=Y)"
-by (simp add: HPair_def)
+  by (simp add: HPair_def)
 
 lemma MPair_eq_HPair [iff]:
-     "(\<lbrace>X',Y'\<rbrace> = Hash[X] Y) = (X' = Hash\<lbrace>X,Y\<rbrace> \<and> Y'=Y)"
-by (simp add: HPair_def)
+  "(\<lbrace>X',Y'\<rbrace> = Hash[X] Y) = (X' = Hash\<lbrace>X,Y\<rbrace> \<and> Y'=Y)"
+  by (simp add: HPair_def)
 
 lemma HPair_eq_MPair [iff]:
-     "(Hash[X] Y = \<lbrace>X',Y'\<rbrace>) = (X' = Hash\<lbrace>X,Y\<rbrace> \<and> Y'=Y)"
-by (auto simp add: HPair_def)
+  "(Hash[X] Y = \<lbrace>X',Y'\<rbrace>) = (X' = Hash\<lbrace>X,Y\<rbrace> \<and> Y'=Y)"
+  by (auto simp add: HPair_def)
 
 
 subsubsection\<open>Specialized laws, proved in terms of those for Hash and MPair\<close>
 
 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\<lbrace>X,Y\<rbrace>) (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\<lbrace>X,Y\<rbrace>) (analz (insert Y H)))"
-by (simp add: HPair_def)
+  by (simp add: HPair_def)
 
 lemma HPair_synth_analz [simp]:
-     "X \<notin> synth (analz H)  
-    ==> (Hash[X] Y \<in> synth (analz H)) =  
+  "X \<notin> synth (analz H)  
+    \<Longrightarrow> (Hash[X] Y \<in> synth (analz H)) =  
         (Hash \<lbrace>X, Y\<rbrace> \<in> analz H \<and> Y \<in> synth (analz H))"
-by (auto simp add: HPair_def)
+  by (auto simp add: HPair_def)
 
 
 text\<open>We do NOT want Crypt... messages broken up in protocols!!\<close>
@@ -815,8 +805,8 @@
   | Number: "Number N \<in> keyfree"
   | Nonce:  "Nonce N \<in> keyfree"
   | Hash:   "Hash X \<in> keyfree"
-  | MPair:  "[|X \<in> keyfree;  Y \<in> keyfree|] ==> \<lbrace>X,Y\<rbrace> \<in> keyfree"
-  | Crypt:  "[|X \<in> keyfree|] ==> Crypt K X \<in> keyfree"
+  | MPair:  "\<lbrakk>X \<in> keyfree;  Y \<in> keyfree\<rbrakk> \<Longrightarrow> \<lbrace>X,Y\<rbrace> \<in> keyfree"
+  | Crypt:  "\<lbrakk>X \<in> keyfree\<rbrakk> \<Longrightarrow> Crypt K X \<in> 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: "\<lbrakk>X \<in> analz (G \<union> H); G \<subseteq> keyfree\<rbrakk> \<Longrightarrow> X \<in> parts G \<union> 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\<open>Tactics useful for many protocol proofs\<close>
 ML
-\<open>
+  \<open>
 (*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 \<notin> Key ` A"
-by auto
+  by auto
 
 lemma Hash_notin_image_Key [simp] :"Hash X \<notin> Key ` A"
-by auto
+  by auto
 
-lemma synth_analz_mono: "G\<subseteq>H ==> synth (analz(G)) \<subseteq> synth (analz(H))"
-by (iprover intro: synth_mono analz_mono) 
+lemma synth_analz_mono: "G\<subseteq>H \<Longrightarrow> synth (analz(G)) \<subseteq> synth (analz(H))"
+  by (iprover intro: synth_mono analz_mono) 
 
 lemma Fake_analz_eq [simp]:
-     "X \<in> 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 \<in> synth(analz H) \<Longrightarrow> 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\<open>Two generalizations of \<open>analz_insert_eq\<close>\<close>
 lemma gen_analz_insert_eq [rule_format]:
-     "X \<in> analz H \<Longrightarrow> \<forall>G. H \<subseteq> G \<longrightarrow> analz (insert X G) = analz G"
-by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
+  "X \<in> analz H \<Longrightarrow> \<forall>G. H \<subseteq> G \<longrightarrow> 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 \<in> synth (analz H) 
-      \<Longrightarrow> \<forall>G. H \<subseteq> G \<longrightarrow> (Key K \<in> analz (insert X G)) = (Key K \<in> 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:
+  "\<lbrakk>X \<in> synth (analz H); H \<subseteq> G\<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert X G)) \<longleftrightarrow> (Key K \<in> 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 \<in> synth (analz H) ==> parts{X} \<subseteq> synth (analz H) \<union> parts H"
-by (metis Fake_parts_insert empty_subsetI insert_mono parts_mono subset_trans)
+  "X \<in> synth (analz H) \<Longrightarrow> parts{X} \<subseteq> synth (analz H) \<union> 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 = \<open>
     Scan.succeed (SIMPLE_METHOD' o spy_analz_tac)\<close>
-    "for proving the Fake case when analz is involved"
+  "for proving the Fake case when analz is involved"
 
 method_setup atomic_spy_analz = \<open>
     Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac)\<close>
-    "for debugging spy_analz"
+  "for debugging spy_analz"
 
 method_setup Fake_insert_simp = \<open>
     Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac)\<close>
-    "for debugging spy_analz"
+  "for debugging spy_analz"
 
 end
--- a/src/HOL/Auth/NS_Public_Bad.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -78,9 +78,7 @@
        \<Longrightarrow> Crypt (pubEK C) \<lbrace>NA', Nonce NA\<rbrace> \<in> parts (spies evs) \<longrightarrow>
            Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>  
            Nonce NA \<in> 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 \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
       \<Longrightarrow> A=A' \<and> 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 @@
     \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
         Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
         Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> 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: 
      "\<lbrakk>Says A  B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;   
@@ -134,10 +128,7 @@
       \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
           Nonce NA \<notin> analz (spies evs) \<longrightarrow>
           Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> 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)
 
 
 
--- a/src/HOL/Auth/NS_Shared.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/NS_Shared.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -89,8 +89,8 @@
 
 
 text\<open>A "possibility property": there are traces that reach the end\<close>
-lemma "[| A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys |]
-       ==> \<exists>N. \<exists>evs \<in> ns_shared.
+lemma "\<lbrakk>A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys\<rbrakk>
+       \<Longrightarrow> \<exists>N. \<exists>evs \<in> ns_shared.
                     Says A B (Crypt K \<lbrace>Nonce N, Nonce N\<rbrace>) \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2] ns_shared.Nil
@@ -136,8 +136,8 @@
 
 text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
-    "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> ns_shared|]
-     ==> K \<notin> keysFor (parts (spies evs))"
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> ns_shared\<rbrakk>
+     \<Longrightarrow> K \<notin> 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\<open>Fake, NS2, NS4, NS5\<close>
--- a/src/HOL/Auth/OtwayReesBella.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/OtwayReesBella.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -146,7 +146,7 @@
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> orb|] ==> A \<in> bad"
+     "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> orb\<rbrakk> \<Longrightarrow> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
 lemma new_keys_not_used [simp]:
@@ -185,7 +185,7 @@
     A \<notin> bad; evs \<in> orb\<rbrakk>                   
  \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs"
 apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
-apply (blast)
+apply blast
 done
 
 
@@ -312,7 +312,7 @@
 txt\<open>Oops\<close>
 prefer 4 apply (blast dest: analz_insert_freshCryptK)
 txt\<open>OR4 - ii\<close>
-prefer 3 apply (blast)
+prefer 3 apply blast
 txt\<open>OR3\<close>
 (*adding Gets_imp_ and Says_imp_ for efficiency*)
 apply (blast dest: 
--- a/src/HOL/Auth/OtwayRees_AN.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/OtwayRees_AN.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -28,31 +28,31 @@
 
  | Fake: \<comment> \<open>The Spy may say anything he can say.  The sender field is correct,
             but agents don't use that information.\<close>
-         "[| evsf \<in> otway;  X \<in> synth (analz (knows Spy evsf)) |]
-          ==> Says Spy B X  # evsf \<in> otway"
+         "\<lbrakk>evsf \<in> otway;  X \<in> synth (analz (knows Spy evsf))\<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> otway"
 
         
  | Reception: \<comment> \<open>A message that has been sent can be received by the
                   intended recipient.\<close>
-              "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
-               ==> Gets B X # evsr \<in> otway"
+              "\<lbrakk>evsr \<in> otway;  Says A B X \<in>set evsr\<rbrakk>
+               \<Longrightarrow> Gets B X # evsr \<in> otway"
 
  | OR1:  \<comment> \<open>Alice initiates a protocol run\<close>
          "evs1 \<in> otway
-          ==> Says A B \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> otway"
+          \<Longrightarrow> Says A B \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> otway"
 
  | OR2:  \<comment> \<open>Bob's response to Alice's message.\<close>
-         "[| evs2 \<in> otway;
-             Gets B \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in>set evs2 |]
-          ==> Says B Server \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB\<rbrace>
+         "\<lbrakk>evs2 \<in> otway;
+             Gets B \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in>set evs2\<rbrakk>
+          \<Longrightarrow> Says B Server \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB\<rbrace>
                  # evs2 \<in> otway"
 
  | OR3:  \<comment> \<open>The Server receives Bob's message.  Then he sends a new
            session key to Bob with a packet for forwarding to Alice.\<close>
-         "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
+         "\<lbrakk>evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB\<rbrace>
-               \<in>set evs3 |]
-          ==> Says Server B
+               \<in>set evs3\<rbrakk>
+          \<Longrightarrow> Says Server B
                \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B, Key KAB\<rbrace>,
                  Crypt (shrK B) \<lbrace>Nonce NB, Agent A, Agent B, Key KAB\<rbrace>\<rbrace>
               # evs3 \<in> otway"
@@ -60,20 +60,20 @@
  | OR4:  \<comment> \<open>Bob receives the Server's (?) message and compares the Nonces with
              those in the message he previously sent the Server.
              Need \<^term>\<open>B \<noteq> Server\<close> because we allow messages to self.\<close>
-         "[| evs4 \<in> otway;  B \<noteq> Server;
+         "\<lbrakk>evs4 \<in> otway;  B \<noteq> Server;
              Says B Server \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB\<rbrace> \<in>set evs4;
              Gets B \<lbrace>X, Crypt(shrK B)\<lbrace>Nonce NB,Agent A,Agent B,Key K\<rbrace>\<rbrace>
-               \<in>set evs4 |]
-          ==> Says B A X # evs4 \<in> otway"
+               \<in>set evs4\<rbrakk>
+          \<Longrightarrow> Says B A X # evs4 \<in> otway"
 
  | Oops: \<comment> \<open>This message models possible leaks of session keys.  The nonces
              identify the protocol run.\<close>
-         "[| evso \<in> otway;
+         "\<lbrakk>evso \<in> otway;
              Says Server B
                       \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B, Key K\<rbrace>,
                         Crypt (shrK B) \<lbrace>Nonce NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
-               \<in>set evso |]
-          ==> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> otway"
+               \<in>set evso\<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> otway"
 
 
 declare Says_imp_knows_Spy [THEN analz.Inj, dest]
@@ -83,8 +83,8 @@
 
 
 text\<open>A "possibility property": there are traces that reach the end\<close>
-lemma "[| B \<noteq> Server; Key K \<notin> used [] |]
-      ==> \<exists>evs \<in> otway.
+lemma "\<lbrakk>B \<noteq> Server; Key K \<notin> used []\<rbrakk>
+      \<Longrightarrow> \<exists>evs \<in> otway.
            Says B A (Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B, Key K\<rbrace>)
              \<in> set evs"
 apply (intro exI bexI)
@@ -96,7 +96,7 @@
 done
 
 lemma Gets_imp_Says [dest!]:
-     "[| Gets B X \<in> set evs; evs \<in> otway |] ==> \<exists>A. Says A B X \<in> set evs"
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> otway\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
 by (erule rev_mp, erule otway.induct, auto)
 
 
@@ -104,8 +104,8 @@
 text\<open>For reasoning about the encrypted portion of messages\<close>
 
 lemma OR4_analz_knows_Spy:
-     "[| Gets B \<lbrace>X, Crypt(shrK B) X'\<rbrace> \<in> set evs;  evs \<in> otway |]
-      ==> X \<in> analz (knows Spy evs)"
+     "\<lbrakk>Gets B \<lbrace>X, Crypt(shrK B) X'\<rbrace> \<in> set evs;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> X \<in> analz (knows Spy evs)"
 by blast
 
 
@@ -114,15 +114,15 @@
 
 text\<open>Spy never sees a good agent's shared key!\<close>
 lemma Spy_see_shrK [simp]:
-     "evs \<in> otway ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
+     "evs \<in> otway \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
 by (erule otway.induct, simp_all, blast+)
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> otway ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
+     "evs \<in> otway \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> otway|] ==> A \<in> bad"
+     "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> otway\<rbrakk> \<Longrightarrow> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
 
@@ -130,12 +130,12 @@
 
 text\<open>Describes the form of K and NA when the Server sends this message.\<close>
 lemma Says_Server_message_form:
-     "[| Says Server B
+     "\<lbrakk>Says Server B
             \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
               Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
            \<in> set evs;
-         evs \<in> otway |]
-      ==> K \<notin> range shrK \<and> (\<exists>i. NA = Nonce i) \<and> (\<exists>j. NB = Nonce j)"
+         evs \<in> otway\<rbrakk>
+      \<Longrightarrow> K \<notin> range shrK \<and> (\<exists>i. NA = Nonce i) \<and> (\<exists>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 \<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+  Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow>
   Key K \<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
@@ -156,7 +156,7 @@
 
 text\<open>The equality makes the induction hypothesis easier to apply\<close>
 lemma analz_image_freshK [rule_format]:
- "evs \<in> otway ==>
+ "evs \<in> otway \<Longrightarrow>
    \<forall>K KK. KK \<subseteq> -(range shrK) \<longrightarrow>
           (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) =
           (K \<in> KK | Key K \<in> analz (knows Spy evs))"
@@ -166,7 +166,7 @@
 done
 
 lemma analz_insert_freshK:
-  "[| evs \<in> otway;  KAB \<notin> range shrK |] ==>
+  "\<lbrakk>evs \<in> otway;  KAB \<notin> range shrK\<rbrakk> \<Longrightarrow>
       (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
       (K = KAB | Key K \<in> analz (knows Spy evs))"
 by (simp only: analz_image_freshK analz_image_freshK_simps)
@@ -174,7 +174,7 @@
 
 text\<open>The Key K uniquely identifies the Server's message.\<close>
 lemma unique_session_keys:
-     "[| Says Server B
+     "\<lbrakk>Says Server B
           \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, K\<rbrace>,
             Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, K\<rbrace>\<rbrace>
          \<in> set evs;
@@ -182,8 +182,8 @@
           \<lbrace>Crypt (shrK A') \<lbrace>NA', Agent A', Agent B', K\<rbrace>,
             Crypt (shrK B') \<lbrace>NB', Agent A', Agent B', K\<rbrace>\<rbrace>
          \<in> set evs;
-        evs \<in> otway |]
-     ==> A=A' \<and> B=B' \<and> NA=NA' \<and> NB=NB'"
+        evs \<in> otway\<rbrakk>
+     \<Longrightarrow> A=A' \<and> B=B' \<and> NA=NA' \<and> NB=NB'"
 apply (erule rev_mp, erule rev_mp, erule otway.induct, simp_all)
 apply blast+  \<comment> \<open>OR3 and OR4\<close>
 done
@@ -193,8 +193,8 @@
 
 text\<open>If the encrypted message appears then it originated with the Server!\<close>
 lemma NA_Crypt_imp_Server_msg [rule_format]:
-    "[| A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-     ==> Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace> \<in> parts (knows Spy evs)
+    "\<lbrakk>A \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+     \<Longrightarrow> Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace> \<in> parts (knows Spy evs)
        \<longrightarrow> (\<exists>NB. Says Server B
                     \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
                       Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
@@ -208,9 +208,9 @@
 text\<open>Corollary: if A receives B's OR4 message then it originated with the
       Server. Freshness may be inferred from nonce NA.\<close>
 lemma A_trusts_OR4:
-     "[| Says B' A (Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>) \<in> set evs;
-         A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-      ==> \<exists>NB. Says Server B
+     "\<lbrakk>Says B' A (Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>) \<in> set evs;
+         A \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> \<exists>NB. Says Server B
                   \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
                     Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
                  \<in> set evs"
@@ -221,8 +221,8 @@
     Does not in itself guarantee security: an attack could violate
     the premises, e.g. by having \<^term>\<open>A=Spy\<close>\<close>
 lemma secrecy_lemma:
-     "[| A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
-      ==> Says Server B
+     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> Says Server B
            \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
              Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
           \<in> set evs \<longrightarrow>
@@ -238,23 +238,23 @@
 
 
 lemma Spy_not_see_encrypted_key:
-     "[| Says Server B
+     "\<lbrakk>Says Server B
             \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
               Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
            \<in> set evs;
          Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+         A \<notin> bad;  B \<notin> bad;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
   by (metis secrecy_lemma)
 
 
 text\<open>A's guarantee.  The Oops premise quantifies over NB because A cannot know
   what it is.\<close>
 lemma A_gets_good_key:
-     "[| Says B' A (Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>) \<in> set evs;
+     "\<lbrakk>Says B' A (Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>) \<in> set evs;
          \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
-         A \<notin> bad;  B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+         A \<notin> bad;  B \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
   by (metis A_trusts_OR4 secrecy_lemma)
 
 
@@ -263,8 +263,8 @@
 
 text\<open>If the encrypted message appears then it originated with the Server!\<close>
 lemma NB_Crypt_imp_Server_msg [rule_format]:
- "[| B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-  ==> Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace> \<in> parts (knows Spy evs)
+ "\<lbrakk>B \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+  \<Longrightarrow> Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace> \<in> parts (knows Spy evs)
       \<longrightarrow> (\<exists>NA. Says Server B
                    \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
                      Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
@@ -278,10 +278,10 @@
 text\<open>Guarantee for B: if it gets a well-formed certificate then the Server
   has sent the correct message in round 3.\<close>
 lemma B_trusts_OR3:
-     "[| Says S B \<lbrace>X, Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
+     "\<lbrakk>Says S B \<lbrace>X, Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
            \<in> set evs;
-         B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-      ==> \<exists>NA. Says Server B
+         B \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> \<exists>NA. Says Server B
                    \<lbrace>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B, Key K\<rbrace>,
                      Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
                    \<in> set evs"
@@ -291,11 +291,11 @@
 text\<open>The obvious combination of \<open>B_trusts_OR3\<close> with 
       \<open>Spy_not_see_encrypted_key\<close>\<close>
 lemma B_gets_good_key:
-     "[| Gets B \<lbrace>X, Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
+     "\<lbrakk>Gets B \<lbrace>X, Crypt (shrK B) \<lbrace>NB, Agent A, Agent B, Key K\<rbrace>\<rbrace>
           \<in> set evs;
          \<forall>NA. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
-         A \<notin> bad;  B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+         A \<notin> bad;  B \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest: B_trusts_OR3 Spy_not_see_encrypted_key)
 
 end
--- a/src/HOL/Auth/OtwayRees_Bad.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/OtwayRees_Bad.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -25,26 +25,26 @@
 
  | Fake: \<comment> \<open>The Spy may say anything he can say.  The sender field is correct,
             but agents don't use that information.\<close>
-         "[| evsf \<in> otway;  X \<in> synth (analz (knows Spy evsf)) |]
-          ==> Says Spy B X  # evsf \<in> otway"
+         "\<lbrakk>evsf \<in> otway;  X \<in> synth (analz (knows Spy evsf))\<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> otway"
 
         
  | Reception: \<comment> \<open>A message that has been sent can be received by the
                   intended recipient.\<close>
-              "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
-               ==> Gets B X # evsr \<in> otway"
+              "\<lbrakk>evsr \<in> otway;  Says A B X \<in>set evsr\<rbrakk>
+               \<Longrightarrow> Gets B X # evsr \<in> otway"
 
  | OR1:  \<comment> \<open>Alice initiates a protocol run\<close>
-         "[| evs1 \<in> otway;  Nonce NA \<notin> used evs1 |]
-          ==> Says A B \<lbrace>Nonce NA, Agent A, Agent B,
+         "\<lbrakk>evs1 \<in> otway;  Nonce NA \<notin> used evs1\<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Nonce NA, Agent A, Agent B,
                          Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace>\<rbrace>
                  # evs1 \<in> otway"
 
  | OR2:  \<comment> \<open>Bob's response to Alice's message.
              This variant of the protocol does NOT encrypt NB.\<close>
-         "[| evs2 \<in> otway;  Nonce NB \<notin> used evs2;
-             Gets B \<lbrace>Nonce NA, Agent A, Agent B, X\<rbrace> \<in> set evs2 |]
-          ==> Says B Server
+         "\<lbrakk>evs2 \<in> otway;  Nonce NB \<notin> used evs2;
+             Gets B \<lbrace>Nonce NA, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
+          \<Longrightarrow> Says B Server
                   \<lbrace>Nonce NA, Agent A, Agent B, X, Nonce NB,
                     Crypt (shrK B) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace>\<rbrace>
                  # evs2 \<in> otway"
@@ -52,14 +52,14 @@
  | OR3:  \<comment> \<open>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.\<close>
-         "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
+         "\<lbrakk>evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server
                   \<lbrace>Nonce NA, Agent A, Agent B,
                     Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace>,
                     Nonce NB,
                     Crypt (shrK B) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace>\<rbrace>
-               \<in> set evs3 |]
-          ==> Says Server B
+               \<in> set evs3\<rbrakk>
+          \<Longrightarrow> Says Server B
                   \<lbrace>Nonce NA,
                     Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
                     Crypt (shrK B) \<lbrace>Nonce NB, Key KAB\<rbrace>\<rbrace>
@@ -68,20 +68,20 @@
  | OR4:  \<comment> \<open>Bob receives the Server's (?) message and compares the Nonces with
              those in the message he previously sent the Server.
              Need \<^term>\<open>B \<noteq> Server\<close> because we allow messages to self.\<close>
-         "[| evs4 \<in> otway;  B \<noteq> Server;
+         "\<lbrakk>evs4 \<in> otway;  B \<noteq> Server;
              Says B Server \<lbrace>Nonce NA, Agent A, Agent B, X', Nonce NB,
                              Crypt (shrK B) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace>\<rbrace>
                \<in> set evs4;
              Gets B \<lbrace>Nonce NA, X, Crypt (shrK B) \<lbrace>Nonce NB, Key K\<rbrace>\<rbrace>
-               \<in> set evs4 |]
-          ==> Says B A \<lbrace>Nonce NA, X\<rbrace> # evs4 \<in> otway"
+               \<in> set evs4\<rbrakk>
+          \<Longrightarrow> Says B A \<lbrace>Nonce NA, X\<rbrace> # evs4 \<in> otway"
 
  | Oops: \<comment> \<open>This message models possible leaks of session keys.  The nonces
              identify the protocol run.\<close>
-         "[| evso \<in> otway;
+         "\<lbrakk>evso \<in> otway;
              Says Server B \<lbrace>Nonce NA, X, Crypt (shrK B) \<lbrace>Nonce NB, Key K\<rbrace>\<rbrace>
-               \<in> set evso |]
-          ==> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> otway"
+               \<in> set evso\<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> otway"
 
 
 declare Says_imp_knows_Spy [THEN analz.Inj, dest]
@@ -90,8 +90,8 @@
 declare Fake_parts_insert_in_Un  [dest]
 
 text\<open>A "possibility property": there are traces that reach the end\<close>
-lemma "[| B \<noteq> Server; Key K \<notin> used [] |]
-      ==> \<exists>NA. \<exists>evs \<in> otway.
+lemma "\<lbrakk>B \<noteq> Server; Key K \<notin> used []\<rbrakk>
+      \<Longrightarrow> \<exists>NA. \<exists>evs \<in> otway.
             Says B A \<lbrace>Nonce NA, Crypt (shrK A) \<lbrace>Nonce NA, Key K\<rbrace>\<rbrace>
               \<in> set evs"
 apply (intro exI bexI)
@@ -103,7 +103,7 @@
 done
 
 lemma Gets_imp_Says [dest!]:
-     "[| Gets B X \<in> set evs; evs \<in> otway |] ==> \<exists>A. Says A B X \<in> set evs"
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> otway\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
 apply (erule rev_mp)
 apply (erule otway.induct, auto)
 done
@@ -112,18 +112,18 @@
 subsection\<open>For reasoning about the encrypted portion of messages\<close>
 
 lemma OR2_analz_knows_Spy:
-     "[| Gets B \<lbrace>N, Agent A, Agent B, X\<rbrace> \<in> set evs;  evs \<in> otway |]
-      ==> X \<in> analz (knows Spy evs)"
+     "\<lbrakk>Gets B \<lbrace>N, Agent A, Agent B, X\<rbrace> \<in> set evs;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> X \<in> analz (knows Spy evs)"
 by blast
 
 lemma OR4_analz_knows_Spy:
-     "[| Gets B \<lbrace>N, X, Crypt (shrK B) X'\<rbrace> \<in> set evs;  evs \<in> otway |]
-      ==> X \<in> analz (knows Spy evs)"
+     "\<lbrakk>Gets B \<lbrace>N, X, Crypt (shrK B) X'\<rbrace> \<in> set evs;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> X \<in> analz (knows Spy evs)"
 by blast
 
 lemma Oops_parts_knows_Spy:
      "Says Server B \<lbrace>NA, X, Crypt K' \<lbrace>NB,K\<rbrace>\<rbrace> \<in> set evs
-      ==> K \<in> parts (knows Spy evs)"
+      \<Longrightarrow> K \<in> parts (knows Spy evs)"
 by blast
 
 text\<open>Forwarding lemma: see comments in OtwayRees.thy\<close>
@@ -136,17 +136,17 @@
 
 text\<open>Spy never sees a good agent's shared key!\<close>
 lemma Spy_see_shrK [simp]:
-     "evs \<in> otway ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
+     "evs \<in> otway \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
 by (erule otway.induct, force,
     drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+)
 
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> otway ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
+     "evs \<in> otway \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> otway|] ==> A \<in> bad"
+     "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> otway\<rbrakk> \<Longrightarrow> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
 
@@ -155,9 +155,9 @@
 text\<open>Describes the form of K and NA when the Server sends this message.  Also
   for Oops case.\<close>
 lemma Says_Server_message_form:
-     "[| Says Server B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs;
-         evs \<in> otway |]
-      ==> K \<notin> range shrK \<and> (\<exists>i. NA = Nonce i) \<and> (\<exists>j. NB = Nonce j)"
+     "\<lbrakk>Says Server B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs;
+         evs \<in> otway\<rbrakk>
+      \<Longrightarrow> K \<notin> range shrK \<and> (\<exists>i. NA = Nonce i) \<and> (\<exists>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 \<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+  Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow>
   Key K \<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
@@ -177,7 +177,7 @@
 
 text\<open>The equality makes the induction hypothesis easier to apply\<close>
 lemma analz_image_freshK [rule_format]:
- "evs \<in> otway ==>
+ "evs \<in> otway \<Longrightarrow>
    \<forall>K KK. KK \<subseteq> -(range shrK) \<longrightarrow>
           (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) =
           (K \<in> KK | Key K \<in> analz (knows Spy evs))"
@@ -188,7 +188,7 @@
 done
 
 lemma analz_insert_freshK:
-  "[| evs \<in> otway;  KAB \<notin> range shrK |] ==>
+  "\<lbrakk>evs \<in> otway;  KAB \<notin> range shrK\<rbrakk> \<Longrightarrow>
       (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
       (K = KAB | Key K \<in> analz (knows Spy evs))"
 by (simp only: analz_image_freshK analz_image_freshK_simps)
@@ -196,9 +196,9 @@
 
 text\<open>The Key K uniquely identifies the Server's  message.\<close>
 lemma unique_session_keys:
-     "[| Says Server B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, K\<rbrace>\<rbrace>   \<in> set evs;
+     "\<lbrakk>Says Server B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, K\<rbrace>\<rbrace>   \<in> set evs;
          Says Server B' \<lbrace>NA',X',Crypt (shrK B') \<lbrace>NB',K\<rbrace>\<rbrace> \<in> set evs;
-         evs \<in> otway |] ==> X=X' \<and> B=B' \<and> NA=NA' \<and> NB=NB'"
+         evs \<in> otway\<rbrakk> \<Longrightarrow> X=X' \<and> B=B' \<and> NA=NA' \<and> 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>\<open>A=Spy\<close>\<close>
 lemma secrecy_lemma:
- "[| A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
-  ==> Says Server B
+ "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> otway\<rbrakk>
+  \<Longrightarrow> Says Server B
         \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>,
           Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs \<longrightarrow>
       Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs \<longrightarrow>
@@ -227,12 +227,12 @@
 
 
 lemma Spy_not_see_encrypted_key:
-     "[| Says Server B
+     "\<lbrakk>Says Server B
           \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>,
                 Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs;
          Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+         A \<notin> bad;  B \<notin> bad;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest: Says_Server_message_form secrecy_lemma)
 
 
@@ -242,8 +242,8 @@
   \<^term>\<open>A \<noteq> B\<close> prevents OR2's similar-looking cryptogram from being picked 
   up. Original Otway-Rees doesn't need it.\<close>
 lemma Crypt_imp_OR1 [rule_format]:
-     "[| A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-      ==> Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
+     "\<lbrakk>A \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+      \<Longrightarrow> Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
           Says A B \<lbrace>NA, Agent A, Agent B,
                      Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace>  \<in> set evs"
 by (erule otway.induct, force,
@@ -255,8 +255,8 @@
   The premise \<^term>\<open>A \<noteq> B\<close> allows use of \<open>Crypt_imp_OR1\<close>\<close>
 text\<open>Only it is FALSE.  Somebody could make a fake message to Server
           substituting some other nonce NA' for NB.\<close>
-lemma "[| A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
-       ==> Crypt (shrK A) \<lbrace>NA, Key K\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
+lemma "\<lbrakk>A \<notin> bad;  A \<noteq> B;  evs \<in> otway\<rbrakk>
+       \<Longrightarrow> Crypt (shrK A) \<lbrace>NA, Key K\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
            Says A B \<lbrace>NA, Agent A, Agent B,
                       Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace>
             \<in> set evs \<longrightarrow>
--- a/src/HOL/Auth/Public.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Public.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -11,7 +11,7 @@
 imports Event
 begin
 
-lemma invKey_K: "K \<in> symKeys ==> invKey K = K"
+lemma invKey_K: "K \<in> symKeys \<Longrightarrow> invKey K = K"
 by (simp add: symKeys_def)
 
 subsection\<open>Asymmetric Keys\<close>
@@ -59,7 +59,7 @@
   \<^term>\<open>True\<noteq>False\<close>, no agent has identical signing and encryption keys\<close>
 specification (publicKey)
   injective_publicKey:
-    "publicKey b A = publicKey c A' ==> b=c \<and> A=A'"
+    "publicKey b A = publicKey c A' \<Longrightarrow> b=c \<and> A=A'"
    apply (rule exI [of _ 
        "\<lambda>b A. 2 * case_agent 0 (\<lambda>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 \<notin> symKeys"
 by (simp add: symKeys_def)
 
-lemma symKey_neq_priEK: "K \<in> symKeys ==> K \<noteq> priEK A"
+lemma symKey_neq_priEK: "K \<in> symKeys \<Longrightarrow> K \<noteq> priEK A"
 by auto
 
-lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'"
+lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) \<Longrightarrow> K \<noteq> K'"
 by blast
 
 lemma symKeys_invKey_iff [iff]: "(invKey K \<in> symKeys) = (K \<in> symKeys)"
-by (unfold symKeys_def, auto)
+  unfolding symKeys_def by auto
 
 lemma analz_symKeys_Decrypt:
-     "[| Crypt K X \<in> analz H;  K \<in> symKeys;  Key K \<in> analz H |]  
-      ==> X \<in> analz H"
+     "\<lbrakk>Crypt K X \<in> analz H;  K \<in> symKeys;  Key K \<in> analz H\<rbrakk>  
+      \<Longrightarrow> X \<in> 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 \<in> analz H; Key(shrK A) \<in> analz H |] ==> X \<in> analz H"
+     "\<lbrakk>Crypt (shrK A) X \<in> analz H; Key(shrK A) \<in> analz H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 by auto
 
 lemma analz_Decrypt':
-     "[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] ==> X \<in> analz H"
+     "\<lbrakk>Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 by (auto simp add: invKey_K)
 
 lemma priK_neq_shrK [iff]: "shrK A \<noteq> privateKey b C"
@@ -239,15 +239,15 @@
 apply (auto dest!: parts_cut simp add: used_Nil) 
 done
 
-lemma MPair_used_D: "\<lbrace>X,Y\<rbrace> \<in> used H ==> X \<in> used H \<and> Y \<in> used H"
+lemma MPair_used_D: "\<lbrace>X,Y\<rbrace> \<in> used H \<Longrightarrow> X \<in> used H \<and> Y \<in> used H"
 by (drule used_parts_subset_parts, simp, blast)
 
 text\<open>There was a similar theorem in Event.thy, so perhaps this one can
   be moved up if proved directly by induction.\<close>
 lemma MPair_used [elim!]:
-     "[| \<lbrace>X,Y\<rbrace> \<in> used H;
-         [| X \<in> used H; Y \<in> used H |] ==> P |] 
-      ==> P"
+     "\<lbrakk>\<lbrace>X,Y\<rbrace> \<in> used H;
+         \<lbrakk>X \<in> used H; Y \<in> used H\<rbrakk> \<Longrightarrow> P\<rbrakk> 
+      \<Longrightarrow> P"
 by (blast dest: MPair_used_D) 
 
 
@@ -255,7 +255,7 @@
   that expression is not in normal form.\<close>
 
 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 \<notin> used evs ==> K \<notin> range shrK"
+lemma Key_not_used [simp]: "Key K \<notin> used evs \<Longrightarrow> K \<notin> range shrK"
 by blast
 
-lemma shrK_neq: "Key K \<notin> used evs ==> shrK B \<noteq> K"
+lemma shrK_neq: "Key K \<notin> used evs \<Longrightarrow> shrK B \<noteq> K"
 by blast
 
 lemmas neq_shrK = shrK_neq [THEN not_sym]
@@ -317,14 +317,14 @@
 
 text\<open>Spy sees private keys of bad agents!\<close>
 lemma Spy_spies_bad_privateKey [intro!]:
-     "A \<in> bad ==> Key (privateKey b A) \<in> spies evs"
+     "A \<in> bad \<Longrightarrow> Key (privateKey b A) \<in> spies evs"
 apply (induct_tac "evs")
 apply (auto simp add: imageI knows_Cons split: event.split)
 done
 
 text\<open>Spy sees long-term shared keys of bad agents!\<close>
 lemma Spy_spies_bad_shrK [intro!]:
-     "A \<in> bad ==> Key (shrK A) \<in> spies evs"
+     "A \<in> bad \<Longrightarrow> Key (shrK A) \<in> 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 \<in> analz (knows Spy evs);  A \<in> bad |]  
-      ==> X \<in> analz (knows Spy evs)"
+     "\<lbrakk>Crypt (shrK A) X \<in> analz (knows Spy evs);  A \<in> bad\<rbrakk>  
+      \<Longrightarrow> X \<in> analz (knows Spy evs)"
 by force
 
 
@@ -383,13 +383,13 @@
 by blast
 
 
-lemma Crypt_imp_keysFor :"[|Crypt K X \<in> H; K \<in> symKeys|] ==> K \<in> keysFor H"
+lemma Crypt_imp_keysFor :"\<lbrakk>Crypt K X \<in> H; K \<in> symKeys\<rbrakk> \<Longrightarrow> K \<in> keysFor H"
 by (drule Crypt_imp_invKey_keysFor, simp)
 
 text\<open>Lemma for the trivial direction of the if-and-only-if of the 
 Session Key Compromise Theorem\<close>
 lemma analz_image_freshK_lemma:
-     "(Key K \<in> analz (Key`nE \<union> H)) \<longrightarrow> (K \<in> nE | Key K \<in> analz H)  ==>  
+     "(Key K \<in> analz (Key`nE \<union> H)) \<longrightarrow> (K \<in> nE | Key K \<in> analz H)  \<Longrightarrow>  
          (Key K \<in> analz (Key`nE \<union> H)) = (K \<in> nE | Key K \<in> analz H)"
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
--- a/src/HOL/Auth/Recur.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Recur.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -21,15 +21,15 @@
   for evs :: "event list"
   where
    One:  "Key KAB \<notin> used evs
-          ==> (Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, Nonce NA, END\<rbrace>,
+          \<Longrightarrow> (Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, Nonce NA, END\<rbrace>,
                \<lbrace>Crypt (shrK A) \<lbrace>Key KAB, Agent B, Nonce NA\<rbrace>, END\<rbrace>,
                KAB)   \<in> respond evs"
 
     (*The most recent session key is passed up to the caller*)
- | Cons: "[| (PA, RA, KAB) \<in> respond evs;
+ | Cons: "\<lbrakk>(PA, RA, KAB) \<in> respond evs;
              Key KBC \<notin> used evs;  Key KBC \<notin> parts {RA};
-             PA = Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, Nonce NA, P\<rbrace> |]
-          ==> (Hash[Key(shrK B)] \<lbrace>Agent B, Agent C, Nonce NB, PA\<rbrace>,
+             PA = Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, Nonce NA, P\<rbrace>\<rbrakk>
+          \<Longrightarrow> (Hash[Key(shrK B)] \<lbrace>Agent B, Agent C, Nonce NB, PA\<rbrace>,
                \<lbrace>Crypt (shrK B) \<lbrace>Key KBC, Agent C, Nonce NB\<rbrace>,
                  Crypt (shrK B) \<lbrace>Key KAB, Agent A, Nonce NB\<rbrace>,
                  RA\<rbrace>,
@@ -47,8 +47,8 @@
     (*Server terminates lists*)
    Nil:  "END \<in> responses evs"
 
- | Cons: "[| RA \<in> responses evs;  Key KAB \<notin> used evs |]
-          ==> \<lbrace>Crypt (shrK B) \<lbrace>Key KAB, Agent A, Nonce NB\<rbrace>,
+ | Cons: "\<lbrakk>RA \<in> responses evs;  Key KAB \<notin> used evs\<rbrakk>
+          \<Longrightarrow> \<lbrace>Crypt (shrK B) \<lbrace>Key KAB, Agent A, Nonce NB\<rbrace>,
                 RA\<rbrace>  \<in> responses evs"
 
 
@@ -59,37 +59,37 @@
 
          (*The spy MAY say anything he CAN say.  Common to
            all similar protocols.*)
- | Fake: "[| evsf \<in> recur;  X \<in> synth (analz (knows Spy evsf)) |]
-          ==> Says Spy B X  # evsf \<in> recur"
+ | Fake: "\<lbrakk>evsf \<in> recur;  X \<in> synth (analz (knows Spy evsf))\<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> recur"
 
          (*Alice initiates a protocol run.
            END is a placeholder to terminate the nesting.*)
- | RA1:  "[| evs1 \<in> recur;  Nonce NA \<notin> used evs1 |]
-          ==> Says A B (Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, Nonce NA, END\<rbrace>)
+ | RA1:  "\<lbrakk>evs1 \<in> recur;  Nonce NA \<notin> used evs1\<rbrakk>
+          \<Longrightarrow> Says A B (Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, Nonce NA, END\<rbrace>)
               # evs1 \<in> recur"
 
          (*Bob's response to Alice's message.  C might be the Server.
            We omit PA = \<lbrace>XA, Agent A, Agent B, Nonce NA, P\<rbrace> because
            it complicates proofs, so B may respond to any message at all!*)
- | RA2:  "[| evs2 \<in> recur;  Nonce NB \<notin> used evs2;
-             Says A' B PA \<in> set evs2 |]
-          ==> Says B C (Hash[Key(shrK B)] \<lbrace>Agent B, Agent C, Nonce NB, PA\<rbrace>)
+ | RA2:  "\<lbrakk>evs2 \<in> recur;  Nonce NB \<notin> used evs2;
+             Says A' B PA \<in> set evs2\<rbrakk>
+          \<Longrightarrow> Says B C (Hash[Key(shrK B)] \<lbrace>Agent B, Agent C, Nonce NB, PA\<rbrace>)
               # evs2 \<in> recur"
 
          (*The Server receives Bob's message and prepares a response.*)
- | RA3:  "[| evs3 \<in> recur;  Says B' Server PB \<in> set evs3;
-             (PB,RB,K) \<in> respond evs3 |]
-          ==> Says Server B RB # evs3 \<in> recur"
+ | RA3:  "\<lbrakk>evs3 \<in> recur;  Says B' Server PB \<in> set evs3;
+             (PB,RB,K) \<in> respond evs3\<rbrakk>
+          \<Longrightarrow> Says Server B RB # evs3 \<in> recur"
 
          (*Bob receives the returned message and compares the Nonces with
            those in the message he previously sent the Server.*)
- | RA4:  "[| evs4 \<in> recur;
+ | RA4:  "\<lbrakk>evs4 \<in> recur;
              Says B  C \<lbrace>XH, Agent B, Agent C, Nonce NB,
                          XA, Agent A, Agent B, Nonce NA, P\<rbrace> \<in> set evs4;
              Says C' B \<lbrace>Crypt (shrK B) \<lbrace>Key KBC, Agent C, Nonce NB\<rbrace>,
                          Crypt (shrK B) \<lbrace>Key KAB, Agent A, Nonce NB\<rbrace>,
-                         RA\<rbrace> \<in> set evs4 |]
-          ==> Says B A RA # evs4 \<in> recur"
+                         RA\<rbrace> \<in> set evs4\<rbrakk>
+          \<Longrightarrow> Says B A RA # evs4 \<in> 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 \<in> recur;  Says Server B RB \<in> set evso;
-              RB \<in> responses evs';  Key K \<in> parts {RB} |]
-           ==> Notes Spy \<lbrace>Key K, RB\<rbrace> # evso \<in> recur"
+   Oops:  "\<lbrakk>evso \<in> recur;  Says Server B RB \<in> set evso;
+              RB \<in> responses evs';  Key K \<in> parts {RB}\<rbrakk>
+           \<Longrightarrow> Notes Spy \<lbrace>Key K, RB\<rbrace> # evso \<in> recur"
   *)
 
 
@@ -119,7 +119,7 @@
 
 text\<open>Simplest case: Alice goes directly to the server\<close>
 lemma "Key K \<notin> used [] 
-       ==> \<exists>NA. \<exists>evs \<in> recur.
+       \<Longrightarrow> \<exists>NA. \<exists>evs \<in> recur.
               Says Server A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent Server, Nonce NA\<rbrace>,
                     END\<rbrace>  \<in> set evs"
 apply (intro exI bexI)
@@ -130,9 +130,9 @@
 
 
 text\<open>Case two: Alice, Bob and the server\<close>
-lemma "[| Key K \<notin> used []; Key K' \<notin> used []; K \<noteq> K';
-          Nonce NA \<notin> used []; Nonce NB \<notin> used []; NA < NB |]
-       ==> \<exists>NA. \<exists>evs \<in> recur.
+lemma "\<lbrakk>Key K \<notin> used []; Key K' \<notin> used []; K \<noteq> K';
+          Nonce NA \<notin> used []; Nonce NB \<notin> used []; NA < NB\<rbrakk>
+       \<Longrightarrow> \<exists>NA. \<exists>evs \<in> recur.
         Says B A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent B, Nonce NA\<rbrace>,
                    END\<rbrace>  \<in> 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 \<notin> used []; Key K' \<notin> used [];  
+lemma "\<lbrakk>Key K \<notin> used []; Key K' \<notin> used [];  
           Key K'' \<notin> used []; K \<noteq> K'; K' \<noteq> K''; K \<noteq> K'';
           Nonce NA \<notin> used []; Nonce NB \<notin> used []; Nonce NC \<notin> used []; 
-          NA < NB; NB < NC |]
-       ==> \<exists>K. \<exists>NA. \<exists>evs \<in> recur.
+          NA < NB; NB < NC\<rbrakk>
+       \<Longrightarrow> \<exists>K. \<exists>NA. \<exists>evs \<in> recur.
              Says B A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent B, Nonce NA\<rbrace>,
                         END\<rbrace>  \<in> set evs"
 apply (intro exI bexI)
@@ -167,18 +167,18 @@
 done
 
 
-lemma respond_imp_not_used: "(PA,RB,KAB) \<in> respond evs ==> Key KAB \<notin> used evs"
+lemma respond_imp_not_used: "(PA,RB,KAB) \<in> respond evs \<Longrightarrow> Key KAB \<notin> used evs"
 by (erule respond.induct, simp_all)
 
 lemma Key_in_parts_respond [rule_format]:
-   "[| Key K \<in> parts {RB};  (PB,RB,K') \<in> respond evs |] ==> Key K \<notin> used evs"
+   "\<lbrakk>Key K \<in> parts {RB};  (PB,RB,K') \<in> respond evs\<rbrakk> \<Longrightarrow> Key K \<notin> used evs"
 apply (erule rev_mp, erule respond.induct)
 apply (auto dest: Key_not_used respond_imp_not_used)
 done
 
 text\<open>Simple inductive reasoning about responses\<close>
 lemma respond_imp_responses:
-     "(PA,RB,KAB) \<in> respond evs ==> RB \<in> responses evs"
+     "(PA,RB,KAB) \<in> respond evs \<Longrightarrow> RB \<in> 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 \<lbrace>Crypt K X, X', RA\<rbrace> \<in> set evs ==> RA \<in> analz (spies evs)"
+     "Says C' B \<lbrace>Crypt K X, X', RA\<rbrace> \<in> set evs \<Longrightarrow> RA \<in> 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 \<in> recur ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
+     "evs \<in> recur \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 apply (erule recur.induct, auto)
 txt\<open>RA3.  It's ugly to call auto twice, but it seems necessary.\<close>
 apply (auto dest: Key_in_parts_respond simp add: parts_insert_spies)
 done
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> recur ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
+     "evs \<in> recur \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> recur|] ==> A \<in> bad"
+     "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> recur\<rbrakk> \<Longrightarrow> A \<in> 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 \<in> responses evs;
+     "\<lbrakk>RB \<in> responses evs;
          \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
                    (Key K \<in> analz (Key`KK \<union> H)) =
-                   (K \<in> KK | Key K \<in> analz H) |]
-     ==> \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
+                   (K \<in> KK | Key K \<in> analz H)\<rbrakk>
+     \<Longrightarrow> \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
                    (Key K \<in> analz (insert RB (Key`KK \<union> H))) =
                    (K \<in> KK | Key K \<in> analz (insert RB H))"
 apply (erule responses.induct)
@@ -246,7 +246,7 @@
 
 text\<open>Version for the protocol.  Proof is easy, thanks to the lemma.\<close>
 lemma raw_analz_image_freshK:
- "evs \<in> recur ==>
+ "evs \<in> recur \<Longrightarrow>
    \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
           (Key K \<in> analz (Key`KK \<union> (spies evs))) =
           (K \<in> KK | Key K \<in> analz (spies evs))"
@@ -260,8 +260,8 @@
 
 
 (*Instance of the lemma with H replaced by (spies evs):
-   [| RB \<in> responses evs;  evs \<in> recur; |]
-   ==> KK \<subseteq> - (range shrK) \<longrightarrow>
+   \<lbrakk>RB \<in> responses evs;  evs \<in> recur;\<rbrakk>
+   \<Longrightarrow> KK \<subseteq> - (range shrK) \<longrightarrow>
        Key K \<in> analz (insert RB (Key`KK \<union> spies evs)) =
        (K \<in> KK | Key K \<in> analz (insert RB (spies evs)))
 *)
@@ -269,8 +269,8 @@
        resp_analz_image_freshK_lemma [OF _ raw_analz_image_freshK]
 
 lemma analz_insert_freshK:
-     "[| evs \<in> recur;  KAB \<notin> range shrK |]
-      ==> (Key K \<in> analz (insert (Key KAB) (spies evs))) =
+     "\<lbrakk>evs \<in> recur;  KAB \<notin> range shrK\<rbrakk>
+      \<Longrightarrow> (Key K \<in> analz (insert (Key KAB) (spies evs))) =
           (K = KAB | Key K \<in> analz (spies evs))"
 by (simp del: image_insert
          add: analz_image_freshK_simps raw_analz_image_freshK)
@@ -278,8 +278,8 @@
 
 text\<open>Everything that's hashed is already in past traffic.\<close>
 lemma Hash_imp_body:
-     "[| Hash \<lbrace>Key(shrK A), X\<rbrace> \<in> parts (spies evs);
-         evs \<in> recur;  A \<notin> bad |] ==> X \<in> parts (spies evs)"
+     "\<lbrakk>Hash \<lbrace>Key(shrK A), X\<rbrace> \<in> parts (spies evs);
+         evs \<in> recur;  A \<notin> bad\<rbrakk> \<Longrightarrow> X \<in> 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 \<lbrace>Key(shrK A), Agent A, B, NA, P\<rbrace> \<in> parts (spies evs);
+  "\<lbrakk>Hash \<lbrace>Key(shrK A), Agent A, B, NA, P\<rbrace> \<in> parts (spies evs);
       Hash \<lbrace>Key(shrK A), Agent A, B',NA, P'\<rbrace> \<in> parts (spies evs);
-      evs \<in> recur;  A \<notin> bad |]
-    ==> B=B' \<and> P=P'"
+      evs \<in> recur;  A \<notin> bad\<rbrakk>
+    \<Longrightarrow> B=B' \<and> 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 \<in> responses evs;  evs \<in> recur |]
-  ==> (Key (shrK B) \<in> analz (insert RB (spies evs))) = (B\<in>bad)"
+     "\<lbrakk>RB \<in> responses evs;  evs \<in> recur\<rbrakk>
+  \<Longrightarrow> (Key (shrK B) \<in> analz (insert RB (spies evs))) = (B\<in>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 \<in> analz (insert RB H);
+     "\<lbrakk>Key K \<in> analz (insert RB H);
          \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
                    (Key K \<in> analz (Key`KK \<union> H)) =
                    (K \<in> KK | Key K \<in> analz H);
-         RB \<in> responses evs |]
-     ==> (Key K \<in> parts{RB} | Key K \<in> analz H)"
+         RB \<in> responses evs\<rbrakk>
+     \<Longrightarrow> (Key K \<in> parts{RB} | Key K \<in> 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\<open>The last key returned by respond indeed appears in a certificate\<close>
 lemma respond_certificate:
      "(Hash[Key(shrK A)] \<lbrace>Agent A, B, NA, P\<rbrace>, RA, K) \<in> respond evs
-      ==> Crypt (shrK A) \<lbrace>Key K, B, NA\<rbrace> \<in> parts {RA}"
+      \<Longrightarrow> Crypt (shrK A) \<lbrace>Key K, B, NA\<rbrace> \<in> parts {RA}"
 apply (ind_cases "(Hash[Key (shrK A)] \<lbrace>Agent A, B, NA, P\<rbrace>, RA, K) \<in> 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) \<in> respond evs ==>
+     "(PB,RB,KXY) \<in> respond evs \<Longrightarrow>
       \<forall>A B N. Crypt (shrK A) \<lbrace>Key K, Agent B, N\<rbrace> \<in> parts {RB} \<longrightarrow>
       (\<forall>A' B' N'. Crypt (shrK A') \<lbrace>Key K, Agent B', N'\<rbrace> \<in> parts {RB} \<longrightarrow>
       (A'=A \<and> B'=B) | (A'=B \<and> B'=A))"
@@ -370,10 +370,10 @@
 done
 
 lemma unique_session_keys:
-     "[| Crypt (shrK A) \<lbrace>Key K, Agent B, N\<rbrace> \<in> parts {RB};
+     "\<lbrakk>Crypt (shrK A) \<lbrace>Key K, Agent B, N\<rbrace> \<in> parts {RB};
          Crypt (shrK A') \<lbrace>Key K, Agent B', N'\<rbrace> \<in> parts {RB};
-         (PB,RB,KXY) \<in> respond evs |]
-      ==> (A'=A \<and> B'=B) | (A'=B \<and> B'=A)"
+         (PB,RB,KXY) \<in> respond evs\<rbrakk>
+      \<Longrightarrow> (A'=A \<and> B'=B) | (A'=B \<and> 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) \<in> respond evs;  evs \<in> recur |]
-      ==> \<forall>A A' N. A \<notin> bad \<and> A' \<notin> bad \<longrightarrow>
+     "\<lbrakk>(PB,RB,KAB) \<in> respond evs;  evs \<in> recur\<rbrakk>
+      \<Longrightarrow> \<forall>A A' N. A \<notin> bad \<and> A' \<notin> bad \<longrightarrow>
           Crypt (shrK A) \<lbrace>Key K, Agent A', N\<rbrace> \<in> parts{RB} \<longrightarrow>
           Key K \<notin> analz (insert RB (spies evs))"
 apply (erule respond.induct)
@@ -405,9 +405,9 @@
 
 
 lemma Spy_not_see_session_key:
-     "[| Crypt (shrK A) \<lbrace>Key K, Agent A', N\<rbrace> \<in> parts (spies evs);
-         A \<notin> bad;  A' \<notin> bad;  evs \<in> recur |]
-      ==> Key K \<notin> analz (spies evs)"
+     "\<lbrakk>Crypt (shrK A) \<lbrace>Key K, Agent A', N\<rbrace> \<in> parts (spies evs);
+         A \<notin> bad;  A' \<notin> bad;  evs \<in> recur\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (spies evs)"
 apply (erule rev_mp)
 apply (erule recur.induct)
 apply (drule_tac [4] RA2_analz_spies,
@@ -430,9 +430,9 @@
 
 text\<open>The response never contains Hashes\<close>
 lemma Hash_in_parts_respond:
-     "[| Hash \<lbrace>Key (shrK B), M\<rbrace> \<in> parts (insert RB H);
-         (PB,RB,K) \<in> respond evs |]
-      ==> Hash \<lbrace>Key (shrK B), M\<rbrace> \<in> parts H"
+     "\<lbrakk>Hash \<lbrace>Key (shrK B), M\<rbrace> \<in> parts (insert RB H);
+         (PB,RB,K) \<in> respond evs\<rbrakk>
+      \<Longrightarrow> Hash \<lbrace>Key (shrK B), M\<rbrace> \<in> 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.\<close>
 lemma Hash_auth_sender [rule_format]:
-     "[| Hash \<lbrace>Key(shrK A), Agent A, Agent B, NA, P\<rbrace> \<in> parts(spies evs);
-         A \<notin> bad;  evs \<in> recur |]
-      ==> Says A B (Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, NA, P\<rbrace>) \<in> set evs"
-apply (unfold HPair_def)
+     "\<lbrakk>Hash \<lbrace>Key(shrK A), Agent A, Agent B, NA, P\<rbrace> \<in> parts(spies evs);
+         A \<notin> bad;  evs \<in> recur\<rbrakk>
+      \<Longrightarrow> Says A B (Hash[Key(shrK A)] \<lbrace>Agent A, Agent B, NA, P\<rbrace>) \<in> set evs"
+unfolding HPair_def
 apply (erule rev_mp)
 apply (erule recur.induct,
        drule_tac [6] RA4_parts_spies,
@@ -462,9 +462,9 @@
 
 text\<open>Certificates can only originate with the Server.\<close>
 lemma Cert_imp_Server_msg:
-     "[| Crypt (shrK A) Y \<in> parts (spies evs);
-         A \<notin> bad;  evs \<in> recur |]
-      ==> \<exists>C RC. Says Server C RC \<in> set evs  \<and>
+     "\<lbrakk>Crypt (shrK A) Y \<in> parts (spies evs);
+         A \<notin> bad;  evs \<in> recur\<rbrakk>
+      \<Longrightarrow> \<exists>C RC. Says Server C RC \<in> set evs  \<and>
                    Crypt (shrK A) Y \<in> parts {RC}"
 apply (erule rev_mp, erule recur.induct, simp_all)
 txt\<open>Fake\<close>
--- a/src/HOL/Auth/Shared.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Shared.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -48,7 +48,7 @@
 
 
 lemma analz_Decrypt' [dest]:
-     "[| Crypt K X \<in> analz H;  Key K  \<in> analz H |] ==> X \<in> analz H"
+     "\<lbrakk>Crypt K X \<in> analz H;  Key K  \<in> analz H\<rbrakk> \<Longrightarrow> X \<in> analz H"
 by auto
 
 text\<open>Now cancel the \<open>dest\<close> attribute given to
@@ -59,17 +59,17 @@
   that expression is not in normal form.\<close>
 
 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 \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) |]
-      ==> K \<in> keysFor (parts (G \<union> H)) | Key K \<in> parts H"
+     "\<lbrakk>K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H)\<rbrakk>
+      \<Longrightarrow> K \<in> keysFor (parts (G \<union> H)) | Key K \<in> parts H"
 by (metis invKey_K keysFor_parts_insert)
 
-lemma Crypt_imp_keysFor: "Crypt K X \<in> H ==> K \<in> keysFor H"
+lemma Crypt_imp_keysFor: "Crypt K X \<in> H \<Longrightarrow> K \<in> 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 \<in> analz (knows Spy evs);  A \<in> bad |]  
-      ==> X \<in> analz (knows Spy evs)"
+lemma Crypt_Spy_analz_bad: "\<lbrakk>Crypt (shrK A) X \<in> analz (knows Spy evs);  A \<in> bad\<rbrakk>  
+      \<Longrightarrow> X \<in> 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 \<notin> used evs ==> K \<notin> range shrK"
+lemma Key_not_used [simp]: "Key K \<notin> used evs \<Longrightarrow> K \<notin> range shrK"
 by blast
 
-lemma shrK_neq [simp]: "Key K \<notin> used evs ==> shrK B \<noteq> K"
+lemma shrK_neq [simp]: "Key K \<notin> used evs \<Longrightarrow> shrK B \<noteq> K"
 by blast
 
 lemmas shrK_sym_neq = shrK_neq [THEN not_sym]
@@ -153,7 +153,7 @@
 done
 
 text\<open>Unlike the corresponding property of nonces, we cannot prove
-    \<^term>\<open>finite KK ==> \<exists>K. K \<notin> KK \<and> Key K \<notin> used evs\<close>.
+    \<^term>\<open>finite KK \<Longrightarrow> \<exists>K. K \<notin> KK \<and> Key K \<notin> used evs\<close>.
     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.\<close>
@@ -184,7 +184,7 @@
 
 (*Lemma for the trivial direction of the if-and-only-if*)
 lemma analz_image_freshK_lemma:
-     "(Key K \<in> analz (Key`nE \<union> H)) \<longrightarrow> (K \<in> nE | Key K \<in> analz H)  ==>  
+     "(Key K \<in> analz (Key`nE \<union> H)) \<longrightarrow> (K \<in> nE | Key K \<in> analz H)  \<Longrightarrow>  
          (Key K \<in> analz (Key`nE \<union> H)) = (K \<in> nE | Key K \<in> analz H)"
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
--- a/src/HOL/Auth/Smartcard/Smartcard.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Fri Oct 14 10:30:37 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
 
--- a/src/HOL/Auth/TLS.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/TLS.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -102,14 +102,14 @@
 
  | Fake: \<comment> \<open>The Spy may say anything he can say.  The sender field is correct,
           but agents don't use that information.\<close>
-         "[| evsf \<in> tls;  X \<in> synth (analz (spies evsf)) |]
-          ==> Says Spy B X # evsf \<in> tls"
+         "\<lbrakk>evsf \<in> tls;  X \<in> synth (analz (spies evsf))\<rbrakk>
+          \<Longrightarrow> Says Spy B X # evsf \<in> tls"
 
  | SpyKeys: \<comment> \<open>The spy may apply \<^term>\<open>PRF\<close> and \<^term>\<open>sessionK\<close>
                 to available nonces\<close>
-         "[| evsSK \<in> tls;
-             {Nonce NA, Nonce NB, Nonce M} \<subseteq> analz (spies evsSK) |]
-          ==> Notes Spy \<lbrace> Nonce (PRF(M,NA,NB)),
+         "\<lbrakk>evsSK \<in> tls;
+             {Nonce NA, Nonce NB, Nonce M} \<subseteq> analz (spies evsSK)\<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace> Nonce (PRF(M,NA,NB)),
                            Key (sessionK((NA,NB,M),role))\<rbrace> # evsSK \<in> tls"
 
  | ClientHello:
@@ -120,8 +120,8 @@
            UNIX TIME is omitted because the protocol doesn't use it.
            May assume \<^term>\<open>NA \<notin> range PRF\<close> because CLIENT RANDOM is 
            28 bytes while MASTER SECRET is 48 bytes\<close>
-         "[| evsCH \<in> tls;  Nonce NA \<notin> used evsCH;  NA \<notin> range PRF |]
-          ==> Says A B \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace>
+         "\<lbrakk>evsCH \<in> tls;  Nonce NA \<notin> used evsCH;  NA \<notin> range PRF\<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace>
                 # evsCH  \<in>  tls"
 
  | ServerHello:
@@ -129,14 +129,14 @@
            PB represents \<open>CLIENT_VERSION\<close>, \<open>CIPHER_SUITE\<close> and \<open>COMPRESSION_METHOD\<close>.
            SERVER CERTIFICATE (7.4.2) is always present.
            \<open>CERTIFICATE_REQUEST\<close> (7.4.4) is implied.\<close>
-         "[| evsSH \<in> tls;  Nonce NB \<notin> used evsSH;  NB \<notin> range PRF;
+         "\<lbrakk>evsSH \<in> tls;  Nonce NB \<notin> used evsSH;  NB \<notin> range PRF;
              Says A' B \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace>
-               \<in> set evsSH |]
-          ==> Says B A \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> # evsSH  \<in>  tls"
+               \<in> set evsSH\<rbrakk>
+          \<Longrightarrow> Says B A \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> # evsSH  \<in>  tls"
 
  | Certificate:
          \<comment> \<open>SERVER (7.4.2) or CLIENT (7.4.6) CERTIFICATE.\<close>
-         "evsC \<in> tls ==> Says B A (certificate B (pubK B)) # evsC  \<in>  tls"
+         "evsC \<in> tls \<Longrightarrow> Says B A (certificate B (pubK B)) # evsC  \<in>  tls"
 
  | ClientKeyExch:
          \<comment> \<open>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).\<close>
-         "[| evsCX \<in> tls;  Nonce PMS \<notin> used evsCX;  PMS \<notin> range PRF;
-             Says B' A (certificate B KB) \<in> set evsCX |]
-          ==> Says A B (Crypt KB (Nonce PMS))
+         "\<lbrakk>evsCX \<in> tls;  Nonce PMS \<notin> used evsCX;  PMS \<notin> range PRF;
+             Says B' A (certificate B KB) \<in> set evsCX\<rbrakk>
+          \<Longrightarrow> Says A B (Crypt KB (Nonce PMS))
               # Notes A \<lbrace>Agent B, Nonce PMS\<rbrace>
               # evsCX  \<in>  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\<close>
-         "[| evsCV \<in> tls;
+         "\<lbrakk>evsCV \<in> tls;
              Says B' A \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> \<in> set evsCV;
-             Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evsCV |]
-          ==> Says A B (Crypt (priK A) (Hash\<lbrace>Nonce NB, Agent B, Nonce PMS\<rbrace>))
+             Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evsCV\<rbrakk>
+          \<Longrightarrow> Says A B (Crypt (priK A) (Hash\<lbrace>Nonce NB, Agent B, Nonce PMS\<rbrace>))
               # evsCV  \<in>  tls"
 
         \<comment> \<open>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>\<open>A\<noteq>Spy\<close> into the rule, but one should not
           expect the spy to be well-behaved.\<close>
-         "[| evsCF \<in> tls;
+         "\<lbrakk>evsCF \<in> tls;
              Says A  B \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace>
                \<in> set evsCF;
              Says B' A \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> \<in> set evsCF;
              Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evsCF;
-             M = PRF(PMS,NA,NB) |]
-          ==> Says A B (Crypt (clientK(NA,NB,M))
+             M = PRF(PMS,NA,NB)\<rbrakk>
+          \<Longrightarrow> Says A B (Crypt (clientK(NA,NB,M))
                         (Hash\<lbrace>Number SID, Nonce M,
                                Nonce NA, Number PA, Agent A,
                                Nonce NB, Number PB, Agent B\<rbrace>))
@@ -191,13 +191,13 @@
  | ServerFinished:
         \<comment> \<open>Keeping A' and A'' distinct means B cannot even check that the
           two messages originate from the same source.\<close>
-         "[| evsSF \<in> tls;
+         "\<lbrakk>evsSF \<in> tls;
              Says A' B  \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace>
                \<in> set evsSF;
              Says B  A  \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> \<in> set evsSF;
              Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
-             M = PRF(PMS,NA,NB) |]
-          ==> Says B A (Crypt (serverK(NA,NB,M))
+             M = PRF(PMS,NA,NB)\<rbrakk>
+          \<Longrightarrow> Says B A (Crypt (serverK(NA,NB,M))
                         (Hash\<lbrace>Number SID, Nonce M,
                                Nonce NA, Number PA, Agent A,
                                Nonce NB, Number PB, Agent B\<rbrace>))
@@ -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 \<open>Notes_master_imp_Crypt_PMS\<close>.\<close>
-         "[| evsCA \<in> tls;
+         "\<lbrakk>evsCA \<in> tls;
              Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evsCA;
              M = PRF(PMS,NA,NB);
              X = Hash\<lbrace>Number SID, Nonce M,
                        Nonce NA, Number PA, Agent A,
                        Nonce NB, Number PB, Agent B\<rbrace>;
              Says A  B (Crypt (clientK(NA,NB,M)) X) \<in> set evsCA;
-             Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA |]
-          ==>
+             Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA\<rbrakk>
+          \<Longrightarrow>
              Notes A \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> # evsCA  \<in>  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 \<open>Notes_master_imp_Crypt_PMS\<close>.\<close>
-         "[| evsSA \<in> tls;
+         "\<lbrakk>evsSA \<in> tls;
              A \<noteq> B;
              Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
              M = PRF(PMS,NA,NB);
@@ -232,18 +232,18 @@
                        Nonce NA, Number PA, Agent A,
                        Nonce NB, Number PB, Agent B\<rbrace>;
              Says B  A (Crypt (serverK(NA,NB,M)) X) \<in> set evsSA;
-             Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA |]
-          ==>
+             Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA\<rbrakk>
+          \<Longrightarrow>
              Notes B \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> # evsSA  \<in>  tls"
 
  | ClientResume:
          \<comment> \<open>If A recalls the \<open>SESSION_ID\<close>, then she sends a FINISHED
              message using the new nonces and stored MASTER SECRET.\<close>
-         "[| evsCR \<in> tls;
+         "\<lbrakk>evsCR \<in> tls;
              Says A  B \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace> \<in> set evsCR;
              Says B' A \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> \<in> set evsCR;
-             Notes A \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evsCR |]
-          ==> Says A B (Crypt (clientK(NA,NB,M))
+             Notes A \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evsCR\<rbrakk>
+          \<Longrightarrow> Says A B (Crypt (clientK(NA,NB,M))
                         (Hash\<lbrace>Number SID, Nonce M,
                                Nonce NA, Number PA, Agent A,
                                Nonce NB, Number PB, Agent B\<rbrace>))
@@ -252,11 +252,11 @@
  | ServerResume:
          \<comment> \<open>Resumption (7.3):  If B finds the \<open>SESSION_ID\<close> then he can 
              send a FINISHED message using the recovered MASTER SECRET\<close>
-         "[| evsSR \<in> tls;
+         "\<lbrakk>evsSR \<in> tls;
              Says A' B \<lbrace>Agent A, Nonce NA, Number SID, Number PA\<rbrace> \<in> set evsSR;
              Says B  A \<lbrace>Nonce NB, Number SID, Number PB\<rbrace> \<in> set evsSR;
-             Notes B \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evsSR |]
-          ==> Says B A (Crypt (serverK(NA,NB,M))
+             Notes B \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evsSR\<rbrakk>
+          \<Longrightarrow> Says B A (Crypt (serverK(NA,NB,M))
                         (Hash\<lbrace>Number SID, Nonce M,
                                Nonce NA, Number PA, Agent A,
                                Nonce NB, Number PB, Agent B\<rbrace>)) # evsSR
@@ -268,9 +268,9 @@
            rather unlikely.  The assumption \<^term>\<open>A\<noteq>Spy\<close> is essential: 
            otherwise the Spy could learn session keys merely by 
            replaying messages!\<close>
-         "[| evso \<in> tls;  A \<noteq> Spy;
-             Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
-          ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  \<in>  tls"
+         "\<lbrakk>evso \<in> tls;  A \<noteq> Spy;
+             Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso\<rbrakk>
+          \<Longrightarrow> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  \<in>  tls"
 
 (*
 Protocol goals:
@@ -331,8 +331,8 @@
 
 
 text\<open>Possibility property ending with ClientAccepts.\<close>
-lemma "[| \<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B |]
-      ==> \<exists>SID M. \<exists>evs \<in> tls.
+lemma "\<lbrakk>\<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B\<rbrakk>
+      \<Longrightarrow> \<exists>SID M. \<exists>evs \<in> tls.
             Notes A \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2] tls.Nil
@@ -344,8 +344,8 @@
 
 
 text\<open>And one for ServerAccepts.  Either FINISHED message may come first.\<close>
-lemma "[| \<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF; A \<noteq> B |]
-      ==> \<exists>SID NA PA NB PB M. \<exists>evs \<in> tls.
+lemma "\<lbrakk>\<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF; A \<noteq> B\<rbrakk>
+      \<Longrightarrow> \<exists>SID NA PA NB PB M. \<exists>evs \<in> tls.
            Notes B \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2] tls.Nil
@@ -357,8 +357,8 @@
 
 
 text\<open>Another one, for CertVerify (which is optional)\<close>
-lemma "[| \<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B |]
-       ==> \<exists>NB PMS. \<exists>evs \<in> tls.
+lemma "\<lbrakk>\<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B\<rbrakk>
+       \<Longrightarrow> \<exists>NB PMS. \<exists>evs \<in> tls.
               Says A B (Crypt (priK A) (Hash\<lbrace>Nonce NB, Agent B, Nonce PMS\<rbrace>)) 
                 \<in> set evs"
 apply (intro exI bexI)
@@ -371,12 +371,12 @@
 
 text\<open>Another one, for session resumption (both ServerResume and ClientResume).
   NO tls.Nil here: we refer to a previous session, not the empty trace.\<close>
-lemma "[| evs0 \<in> tls;
+lemma "\<lbrakk>evs0 \<in> tls;
           Notes A \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evs0;
           Notes B \<lbrace>Number SID, Agent A, Agent B, Nonce M\<rbrace> \<in> set evs0;
           \<forall>evs. (SOME N. Nonce N \<notin> used evs) \<notin> range PRF;
-          A \<noteq> B |]
-      ==> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
+          A \<noteq> B\<rbrakk>
+      \<Longrightarrow> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
                 X = Hash\<lbrace>Number SID, Nonce M,
                           Nonce NA, Number PA, Agent A,
                           Nonce NB, Number PB, Agent B\<rbrace>  \<and>
@@ -397,15 +397,15 @@
 
 text\<open>Spy never sees a good agent's private key!\<close>
 lemma Spy_see_priK [simp]:
-     "evs \<in> tls ==> (Key (privateKey b A) \<in> parts (spies evs)) = (A \<in> bad)"
+     "evs \<in> tls \<Longrightarrow> (Key (privateKey b A) \<in> parts (spies evs)) = (A \<in> bad)"
 by (erule tls.induct, force, simp_all, blast)
 
 lemma Spy_analz_priK [simp]:
-     "evs \<in> tls ==> (Key (privateKey b A) \<in> analz (spies evs)) = (A \<in> bad)"
+     "evs \<in> tls \<Longrightarrow> (Key (privateKey b A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_priK_D [dest!]:
-    "[|Key (privateKey b A) \<in> parts (knows Spy evs);  evs \<in> tls|] ==> A \<in> bad"
+    "\<lbrakk>Key (privateKey b A) \<in> parts (knows Spy evs);  evs \<in> tls\<rbrakk> \<Longrightarrow> A \<in> 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.\<close>
 lemma certificate_valid:
-    "[| certificate B KB \<in> parts (spies evs);  evs \<in> tls |] ==> KB = pubK B"
+    "\<lbrakk>certificate B KB \<in> parts (spies evs);  evs \<in> tls\<rbrakk> \<Longrightarrow> KB = pubK B"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all, blast) 
 done
@@ -425,8 +425,8 @@
 subsubsection\<open>Properties of items found in Notes\<close>
 
 lemma Notes_Crypt_parts_spies:
-     "[| Notes A \<lbrace>Agent B, X\<rbrace> \<in> set evs;  evs \<in> tls |]
-      ==> Crypt (pubK B) X \<in> parts (spies evs)"
+     "\<lbrakk>Notes A \<lbrace>Agent B, X\<rbrace> \<in> set evs;  evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Crypt (pubK B) X \<in> 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\<open>C may be either A or B\<close>
 lemma Notes_master_imp_Crypt_PMS:
-     "[| Notes C \<lbrace>s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\<rbrace> \<in> set evs;
-         evs \<in> tls |]
-      ==> Crypt (pubK B) (Nonce PMS) \<in> parts (spies evs)"
+     "\<lbrakk>Notes C \<lbrace>s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\<rbrace> \<in> set evs;
+         evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Crypt (pubK B) (Nonce PMS) \<in> parts (spies evs)"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all)
 txt\<open>Fake\<close>
@@ -448,9 +448,9 @@
 
 text\<open>Compared with the theorem above, both premise and conclusion are stronger\<close>
 lemma Notes_master_imp_Notes_PMS:
-     "[| Notes A \<lbrace>s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\<rbrace> \<in> set evs;
-         evs \<in> tls |]
-      ==> Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs"
+     "\<lbrakk>Notes A \<lbrace>s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))\<rbrace> \<in> set evs;
+         evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all)
 txt\<open>ServerAccepts\<close>
@@ -462,45 +462,45 @@
 
 text\<open>B can check A's signature if he has received A's certificate.\<close>
 lemma TrustCertVerify_lemma:
-     "[| X \<in> parts (spies evs);
+     "\<lbrakk>X \<in> parts (spies evs);
          X = Crypt (priK A) (Hash\<lbrace>nb, Agent B, pms\<rbrace>);
-         evs \<in> tls;  A \<notin> bad |]
-      ==> Says A B X \<in> set evs"
+         evs \<in> tls;  A \<notin> bad\<rbrakk>
+      \<Longrightarrow> Says A B X \<in> set evs"
 apply (erule rev_mp, erule ssubst)
 apply (erule tls.induct, force, simp_all, blast)
 done
 
 text\<open>Final version: B checks X using the distributed KA instead of priK A\<close>
 lemma TrustCertVerify:
-     "[| X \<in> parts (spies evs);
+     "\<lbrakk>X \<in> parts (spies evs);
          X = Crypt (invKey KA) (Hash\<lbrace>nb, Agent B, pms\<rbrace>);
          certificate A KA \<in> parts (spies evs);
-         evs \<in> tls;  A \<notin> bad |]
-      ==> Says A B X \<in> set evs"
+         evs \<in> tls;  A \<notin> bad\<rbrakk>
+      \<Longrightarrow> Says A B X \<in> set evs"
 by (blast dest!: certificate_valid intro!: TrustCertVerify_lemma)
 
 
 text\<open>If CertVerify is present then A has chosen PMS.\<close>
 lemma UseCertVerify_lemma:
-     "[| Crypt (priK A) (Hash\<lbrace>nb, Agent B, Nonce PMS\<rbrace>) \<in> parts (spies evs);
-         evs \<in> tls;  A \<notin> bad |]
-      ==> Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs"
+     "\<lbrakk>Crypt (priK A) (Hash\<lbrace>nb, Agent B, Nonce PMS\<rbrace>) \<in> parts (spies evs);
+         evs \<in> tls;  A \<notin> bad\<rbrakk>
+      \<Longrightarrow> Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all, blast)
 done
 
 text\<open>Final version using the distributed KA instead of priK A\<close>
 lemma UseCertVerify:
-     "[| Crypt (invKey KA) (Hash\<lbrace>nb, Agent B, Nonce PMS\<rbrace>)
+     "\<lbrakk>Crypt (invKey KA) (Hash\<lbrace>nb, Agent B, Nonce PMS\<rbrace>)
            \<in> parts (spies evs);
          certificate A KA \<in> parts (spies evs);
-         evs \<in> tls;  A \<notin> bad |]
-      ==> Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs"
+         evs \<in> tls;  A \<notin> bad\<rbrakk>
+      \<Longrightarrow> Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs"
 by (blast dest!: certificate_valid intro!: UseCertVerify_lemma)
 
 
 lemma no_Notes_A_PRF [simp]:
-     "evs \<in> tls ==> Notes A \<lbrace>Agent B, Nonce (PRF x)\<rbrace> \<notin> set evs"
+     "evs \<in> tls \<Longrightarrow> Notes A \<lbrace>Agent B, Nonce (PRF x)\<rbrace> \<notin> set evs"
 apply (erule tls.induct, force, simp_all)
 txt\<open>ClientKeyExch: PMS is assumed to differ from any PRF.\<close>
 apply blast
@@ -508,8 +508,8 @@
 
 
 lemma MS_imp_PMS [dest!]:
-     "[| Nonce (PRF (PMS,NA,NB)) \<in> parts (spies evs);  evs \<in> tls |]
-      ==> Nonce PMS \<in> parts (spies evs)"
+     "\<lbrakk>Nonce (PRF (PMS,NA,NB)) \<in> parts (spies evs);  evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Nonce PMS \<in> parts (spies evs)"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all)
 txt\<open>Fake\<close>
@@ -525,11 +525,11 @@
 
 text\<open>PMS determines B.\<close>
 lemma Crypt_unique_PMS:
-     "[| Crypt(pubK B)  (Nonce PMS) \<in> parts (spies evs);
+     "\<lbrakk>Crypt(pubK B)  (Nonce PMS) \<in> parts (spies evs);
          Crypt(pubK B') (Nonce PMS) \<in> parts (spies evs);
          Nonce PMS \<notin> analz (spies evs);
-         evs \<in> tls |]
-      ==> B=B'"
+         evs \<in> tls\<rbrakk>
+      \<Longrightarrow> 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\<open>Fake, ClientKeyExch\<close>
@@ -545,10 +545,10 @@
 
 text\<open>In A's internal Note, PMS determines A and B.\<close>
 lemma Notes_unique_PMS:
-     "[| Notes A  \<lbrace>Agent B,  Nonce PMS\<rbrace> \<in> set evs;
+     "\<lbrakk>Notes A  \<lbrace>Agent B,  Nonce PMS\<rbrace> \<in> set evs;
          Notes A' \<lbrace>Agent B', Nonce PMS\<rbrace> \<in> set evs;
-         evs \<in> tls |]
-      ==> A=A' \<and> B=B'"
+         evs \<in> tls\<rbrakk>
+      \<Longrightarrow> A=A' \<and> B=B'"
 apply (erule rev_mp, erule rev_mp)
 apply (erule tls.induct, force, simp_all)
 txt\<open>ClientKeyExch\<close>
@@ -562,7 +562,7 @@
   No collection of keys can help the spy get new private keys.\<close>
 lemma analz_image_priK [rule_format]:
      "evs \<in> tls
-      ==> \<forall>KK. (Key(priK B) \<in> analz (Key`KK \<union> (spies evs))) =
+      \<Longrightarrow> \<forall>KK. (Key(priK B) \<in> analz (Key`KK \<union> (spies evs))) =
           (priK B \<in> KK | B \<in> bad)"
 apply (erule tls.induct)
 apply (simp_all (no_asm_simp)
@@ -582,7 +582,7 @@
 
 text\<open>Lemma for the trivial direction of the if-and-only-if\<close>
 lemma analz_image_keys_lemma:
-     "(X \<in> analz (G \<union> H)) \<longrightarrow> (X \<in> analz H)  ==>
+     "(X \<in> analz (G \<union> H)) \<longrightarrow> (X \<in> analz H)  \<Longrightarrow>
       (X \<in> analz (G \<union> H))  =  (X \<in> analz H)"
 by (blast intro: analz_mono [THEN subsetD])
 
@@ -592,7 +592,7 @@
 **)
 
 lemma analz_image_keys [rule_format]:
-     "evs \<in> tls ==>
+     "evs \<in> tls \<Longrightarrow>
       \<forall>KK. KK \<subseteq> range sessionK \<longrightarrow>
               (Nonce N \<in> analz (Key`KK \<union> (spies evs))) =
               (Nonce N \<in> analz (spies evs))"
@@ -611,7 +611,7 @@
 
 text\<open>Knowing some session keys is no help in getting new nonces\<close>
 lemma analz_insert_key [simp]:
-     "evs \<in> tls ==>
+     "evs \<in> tls \<Longrightarrow>
       (Nonce N \<in> analz (insert (Key (sessionK z)) (spies evs))) =
       (Nonce N \<in> 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.\<close>
 lemma PMS_lemma:
-     "[| Nonce PMS \<notin> parts (spies evs);
+     "\<lbrakk>Nonce PMS \<notin> parts (spies evs);
          K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role);
-         evs \<in> tls |]
-   ==> Key K \<notin> parts (spies evs) \<and> (\<forall>Y. Crypt K Y \<notin> parts (spies evs))"
+         evs \<in> tls\<rbrakk>
+   \<Longrightarrow> Key K \<notin> parts (spies evs) \<and> (\<forall>Y. Crypt K Y \<notin> 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)) \<in> parts (spies evs);
-         evs \<in> tls |]
-      ==> Nonce PMS \<in> parts (spies evs)"
+     "\<lbrakk>Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) \<in> parts (spies evs);
+         evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Nonce PMS \<in> parts (spies evs)"
 by (blast dest: PMS_lemma)
 
 lemma PMS_Crypt_sessionK_not_spied:
-     "[| Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y
-           \<in> parts (spies evs);  evs \<in> tls |]
-      ==> Nonce PMS \<in> parts (spies evs)"
+     "\<lbrakk>Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y
+           \<in> parts (spies evs);  evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Nonce PMS \<in> parts (spies evs)"
 by (blast dest: PMS_lemma)
 
 text\<open>Write keys are never sent if M (MASTER SECRET) is secure.
@@ -661,9 +661,9 @@
   with some effort.
   NO LONGER USED: see \<open>clientK_not_spied\<close> and \<open>serverK_not_spied\<close>\<close>
 lemma sessionK_not_spied:
-     "[| \<forall>A. Says A Spy (Key (sessionK((NA,NB,M),role))) \<notin> set evs;
-         Nonce M \<notin> analz (spies evs);  evs \<in> tls |]
-      ==> Key (sessionK((NA,NB,M),role)) \<notin> parts (spies evs)"
+     "\<lbrakk>\<forall>A. Says A Spy (Key (sessionK((NA,NB,M),role))) \<notin> set evs;
+         Nonce M \<notin> analz (spies evs);  evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Key (sessionK((NA,NB,M),role)) \<notin> 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\<open>If A sends ClientKeyExch to an honest B, then the PMS will stay secret.\<close>
 lemma Spy_not_see_PMS:
-     "[| Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
-         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
-      ==> Nonce PMS \<notin> analz (spies evs)"
+     "\<lbrakk>Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad\<rbrakk>
+      \<Longrightarrow> Nonce PMS \<notin> 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\<open>Fake\<close>
@@ -696,9 +696,9 @@
 text\<open>If A sends ClientKeyExch to an honest B, then the MASTER SECRET
   will stay secret.\<close>
 lemma Spy_not_see_MS:
-     "[| Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
-         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
-      ==> Nonce (PRF(PMS,NA,NB)) \<notin> analz (spies evs)"
+     "\<lbrakk>Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad\<rbrakk>
+      \<Longrightarrow> Nonce (PRF(PMS,NA,NB)) \<notin> 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\<open>Fake\<close>
@@ -719,10 +719,10 @@
 text\<open>If A created PMS then nobody else (except the Spy in replays)
   would send a message using a clientK generated from that PMS.\<close>
 lemma Says_clientK_unique:
-     "[| Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) \<in> set evs;
+     "\<lbrakk>Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) \<in> set evs;
          Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
-         evs \<in> tls;  A' \<noteq> Spy |]
-      ==> A = A'"
+         evs \<in> tls;  A' \<noteq> Spy\<rbrakk>
+      \<Longrightarrow> 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\<open>If A created PMS and has not leaked her clientK to the Spy,
   then it is completely secure: not even in parts!\<close>
 lemma clientK_not_spied:
-     "[| Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
+     "\<lbrakk>Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
          Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) \<notin> set evs;
          A \<notin> bad;  B \<notin> bad;
-         evs \<in> tls |]
-      ==> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) \<notin> parts (spies evs)"
+         evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) \<notin> 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\<open>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.\<close>
 lemma Says_serverK_unique:
-     "[| Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) \<in> set evs;
+     "\<lbrakk>Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) \<in> set evs;
          Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
-         evs \<in> tls;  A \<notin> bad;  B \<notin> bad;  B' \<noteq> Spy |]
-      ==> B = B'"
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad;  B' \<noteq> Spy\<rbrakk>
+      \<Longrightarrow> 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\<open>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!\<close>
 lemma serverK_not_spied:
-     "[| Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
+     "\<lbrakk>Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs;
          Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) \<notin> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> tls |]
-      ==> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) \<notin> parts (spies evs)"
+         A \<notin> bad;  B \<notin> bad;  evs \<in> tls\<rbrakk>
+      \<Longrightarrow> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) \<notin> 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\<open>The mention of her name (A) in X assures A that B knows who she is.\<close>
 lemma TrustServerFinished [rule_format]:
-     "[| X = Crypt (serverK(Na,Nb,M))
+     "\<lbrakk>X = Crypt (serverK(Na,Nb,M))
                (Hash\<lbrace>Number SID, Nonce M,
                       Nonce Na, Number PA, Agent A,
                       Nonce Nb, Number PB, Agent B\<rbrace>);
          M = PRF(PMS,NA,NB);
-         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
-      ==> Says B Spy (Key(serverK(Na,Nb,M))) \<notin> set evs \<longrightarrow>
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad\<rbrakk>
+      \<Longrightarrow> Says B Spy (Key(serverK(Na,Nb,M))) \<notin> set evs \<longrightarrow>
           Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs \<longrightarrow>
           X \<in> parts (spies evs) \<longrightarrow> Says B A X \<in> 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.\<close>
 lemma TrustServerMsg [rule_format]:
-     "[| M = PRF(PMS,NA,NB);  evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
-      ==> Says B Spy (Key(serverK(Na,Nb,M))) \<notin> set evs \<longrightarrow>
+     "\<lbrakk>M = PRF(PMS,NA,NB);  evs \<in> tls;  A \<notin> bad;  B \<notin> bad\<rbrakk>
+      \<Longrightarrow> Says B Spy (Key(serverK(Na,Nb,M))) \<notin> set evs \<longrightarrow>
           Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs \<longrightarrow>
           Crypt (serverK(Na,Nb,M)) Y \<in> parts (spies evs)  \<longrightarrow>
           (\<exists>A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) \<in> set evs)"
@@ -853,8 +853,8 @@
      ClientFinished, then B can then check the quoted values PA, PB, etc.\<close>
 
 lemma TrustClientMsg [rule_format]:
-     "[| M = PRF(PMS,NA,NB);  evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
-      ==> Says A Spy (Key(clientK(Na,Nb,M))) \<notin> set evs \<longrightarrow>
+     "\<lbrakk>M = PRF(PMS,NA,NB);  evs \<in> tls;  A \<notin> bad;  B \<notin> bad\<rbrakk>
+      \<Longrightarrow> Says A Spy (Key(clientK(Na,Nb,M))) \<notin> set evs \<longrightarrow>
           Notes A \<lbrace>Agent B, Nonce PMS\<rbrace> \<in> set evs \<longrightarrow>
           Crypt (clientK(Na,Nb,M)) Y \<in> parts (spies evs) \<longrightarrow>
           Says A B (Crypt (clientK(Na,Nb,M)) Y) \<in> 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.\<close>
 lemma AuthClientFinished:
-     "[| M = PRF(PMS,NA,NB);
+     "\<lbrakk>M = PRF(PMS,NA,NB);
          Says A Spy (Key(clientK(Na,Nb,M))) \<notin> set evs;
          Says A' B (Crypt (clientK(Na,Nb,M)) Y) \<in> set evs;
          certificate A KA \<in> parts (spies evs);
          Says A'' B (Crypt (invKey KA) (Hash\<lbrace>nb, Agent B, Nonce PMS\<rbrace>))
            \<in> set evs;
-         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
-      ==> Says A B (Crypt (clientK(Na,Nb,M)) Y) \<in> set evs"
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad\<rbrakk>
+      \<Longrightarrow> Says A B (Crypt (clientK(Na,Nb,M)) Y) \<in> set evs"
 by (blast intro!: TrustClientMsg UseCertVerify)
 
 (*22/9/97: loads in 622s, which is 10 minutes 22 seconds*)
--- a/src/HOL/Auth/WooLam.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/WooLam.thy	Fri Oct 14 10:30:37 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 \<in> woolam;  X \<in> synth (analz (spies evsf)) |]
-          ==> Says Spy B X  # evsf \<in> woolam"
+ | Fake: "\<lbrakk>evsf \<in> woolam;  X \<in> synth (analz (spies evsf))\<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> woolam"
 
          (*Alice initiates a protocol run*)
- | WL1:  "evs1 \<in> woolam ==> Says A B (Agent A) # evs1 \<in> woolam"
+ | WL1:  "evs1 \<in> woolam \<Longrightarrow> Says A B (Agent A) # evs1 \<in> woolam"
 
          (*Bob responds to Alice's message with a challenge.*)
- | WL2:  "[| evs2 \<in> woolam;  Says A' B (Agent A) \<in> set evs2 |]
-          ==> Says B A (Nonce NB) # evs2 \<in> woolam"
+ | WL2:  "\<lbrakk>evs2 \<in> woolam;  Says A' B (Agent A) \<in> set evs2\<rbrakk>
+          \<Longrightarrow> Says B A (Nonce NB) # evs2 \<in> 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 \<in> woolam;
+ | WL3:  "\<lbrakk>evs3 \<in> woolam;
              Says A  B (Agent A)  \<in> set evs3;
-             Says B' A (Nonce NB) \<in> set evs3 |]
-          ==> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \<in> woolam"
+             Says B' A (Nonce NB) \<in> set evs3\<rbrakk>
+          \<Longrightarrow> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \<in> 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 \<in> woolam;
+ | WL4:  "\<lbrakk>evs4 \<in> woolam;
              Says A'  B X         \<in> set evs4;
-             Says A'' B (Agent A) \<in> set evs4 |]
-          ==> Says B Server \<lbrace>Agent A, Agent B, X\<rbrace> # evs4 \<in> woolam"
+             Says A'' B (Agent A) \<in> set evs4\<rbrakk>
+          \<Longrightarrow> Says B Server \<lbrace>Agent A, Agent B, X\<rbrace> # evs4 \<in> woolam"
 
          (*Server decrypts Alice's response for Bob.*)
- | WL5:  "[| evs5 \<in> woolam;
+ | WL5:  "\<lbrakk>evs5 \<in> woolam;
              Says B' Server \<lbrace>Agent A, Agent B, Crypt (shrK A) (Nonce NB)\<rbrace>
-               \<in> set evs5 |]
-          ==> Says Server B (Crypt (shrK B) \<lbrace>Agent A, Nonce NB\<rbrace>)
+               \<in> set evs5\<rbrakk>
+          \<Longrightarrow> Says Server B (Crypt (shrK B) \<lbrace>Agent A, Nonce NB\<rbrace>)
                  # evs5 \<in> woolam"
 
 
@@ -86,15 +86,15 @@
 
 (*Spy never sees a good agent's shared key!*)
 lemma Spy_see_shrK [simp]:
-     "evs \<in> woolam ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
+     "evs \<in> woolam \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 by (erule woolam.induct, force, simp_all, blast+)
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> woolam ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
+     "evs \<in> woolam \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> woolam|] ==> A \<in> bad"
+     "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> woolam\<rbrakk> \<Longrightarrow> A \<in> 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) \<in> parts (spies evs);
-         A \<notin> bad;  evs \<in> woolam |]
-      ==> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs"
+     "\<lbrakk>Crypt (shrK A) (Nonce NB) \<in> parts (spies evs);
+         A \<notin> bad;  evs \<in> woolam\<rbrakk>
+      \<Longrightarrow> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> 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 \<lbrace>Agent A, Agent B, Crypt (shrK A) (Nonce NB)\<rbrace>
+     "\<lbrakk>Says B' Server \<lbrace>Agent A, Agent B, Crypt (shrK A) (Nonce NB)\<rbrace>
            \<in> set evs;
-         A \<notin> bad;  evs \<in> woolam |]
-      ==> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs"
+         A \<notin> bad;  evs \<in> woolam\<rbrakk>
+      \<Longrightarrow> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> 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) \<lbrace>Agent A, NB\<rbrace>) \<in> set evs;
-         evs \<in> woolam |]
-      ==> \<exists>B'. Says B' Server \<lbrace>Agent A, Agent B, Crypt (shrK A) NB\<rbrace>
+     "\<lbrakk>Says Server B (Crypt (shrK B) \<lbrace>Agent A, NB\<rbrace>) \<in> set evs;
+         evs \<in> woolam\<rbrakk>
+      \<Longrightarrow> \<exists>B'. Says B' Server \<lbrace>Agent A, Agent B, Crypt (shrK A) NB\<rbrace>
              \<in> 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) \<lbrace>Agent A, NB\<rbrace> \<in> parts (spies evs);
-         B \<notin> bad;  evs \<in> woolam |]
-      ==> Says Server B (Crypt (shrK B) \<lbrace>Agent A, NB\<rbrace>) \<in> set evs"
+     "\<lbrakk>Crypt (shrK B) \<lbrace>Agent A, NB\<rbrace> \<in> parts (spies evs);
+         B \<notin> bad;  evs \<in> woolam\<rbrakk>
+      \<Longrightarrow> Says Server B (Crypt (shrK B) \<lbrace>Agent A, NB\<rbrace>) \<in> 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) \<lbrace>Agent A, Nonce NB\<rbrace>) \<in> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> woolam  |]
-      ==> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs"
+     "\<lbrakk>Says S B (Crypt (shrK B) \<lbrace>Agent A, Nonce NB\<rbrace>) \<in> set evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> woolam\<rbrakk>
+      \<Longrightarrow> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> 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) \<in> set evs;  B \<noteq> Spy;  evs \<in> woolam |]
-      ==> \<exists>A'. Says A' B (Agent A) \<in> set evs"
+     "\<lbrakk>Says B A (Nonce NB) \<in> set evs;  B \<noteq> Spy;  evs \<in> woolam\<rbrakk>
+      \<Longrightarrow> \<exists>A'. Says A' B (Agent A) \<in> 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 \<notin> bad;  B \<noteq> Spy;  evs \<in> woolam |]
-  ==> Crypt (shrK A) (Nonce NB) \<in> parts (spies evs) \<and>
+lemma "\<lbrakk>A \<notin> bad;  B \<noteq> Spy;  evs \<in> woolam\<rbrakk>
+  \<Longrightarrow> Crypt (shrK A) (Nonce NB) \<in> parts (spies evs) \<and>
       Says B A (Nonce NB) \<in> set evs
       \<longrightarrow> Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs"
 apply (erule rev_mp, erule woolam.induct, force, simp_all, blast, auto)
--- a/src/HOL/Auth/Yahalom.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Yahalom.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -332,7 +332,7 @@
  "Says Server A
           \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, Nonce NB\<rbrace>, X\<rbrace>
         \<in> set evs \<Longrightarrow> 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).\<close>
 lemma fresh_not_KeyWithNonce:
      "Key K \<notin> used evs \<Longrightarrow> \<not> KeyWithNonce K NB evs"
-by (unfold KeyWithNonce_def, blast)
+  unfolding KeyWithNonce_def by blast
 
 text\<open>The Server message associates K with NB' and therefore not with any
   other nonce NB.\<close>
@@ -363,7 +363,7 @@
        \<in> set evs;
      NB \<noteq> NB';  evs \<in> yahalom\<rbrakk>
   \<Longrightarrow> \<not> KeyWithNonce K NB evs"
-by (unfold KeyWithNonce_def, blast dest: unique_session_keys)
+  unfolding KeyWithNonce_def by (blast dest: unique_session_keys)
 
 
 text\<open>The only nonces that can be found with the help of session keys are
--- a/src/HOL/Auth/Yahalom_Bad.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/Yahalom_Bad.thy	Fri Oct 14 10:30:37 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 \<in> yahalom;  X \<in> synth (analz (knows Spy evsf)) |]
-          ==> Says Spy B X  # evsf \<in> yahalom"
+ | Fake: "\<lbrakk>evsf \<in> yahalom;  X \<in> synth (analz (knows Spy evsf))\<rbrakk>
+          \<Longrightarrow> Says Spy B X  # evsf \<in> yahalom"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
- | Reception: "[| evsr \<in> yahalom;  Says A B X \<in> set evsr |]
-               ==> Gets B X # evsr \<in> yahalom"
+ | Reception: "\<lbrakk>evsr \<in> yahalom;  Says A B X \<in> set evsr\<rbrakk>
+               \<Longrightarrow> Gets B X # evsr \<in> yahalom"
 
          (*Alice initiates a protocol run*)
- | YM1:  "[| evs1 \<in> yahalom;  Nonce NA \<notin> used evs1 |]
-          ==> Says A B \<lbrace>Agent A, Nonce NA\<rbrace> # evs1 \<in> yahalom"
+ | YM1:  "\<lbrakk>evs1 \<in> yahalom;  Nonce NA \<notin> used evs1\<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Agent A, Nonce NA\<rbrace> # evs1 \<in> yahalom"
 
          (*Bob's response to Alice's message.*)
- | YM2:  "[| evs2 \<in> yahalom;  Nonce NB \<notin> used evs2;
-             Gets B \<lbrace>Agent A, Nonce NA\<rbrace> \<in> set evs2 |]
-          ==> Says B Server
+ | YM2:  "\<lbrakk>evs2 \<in> yahalom;  Nonce NB \<notin> used evs2;
+             Gets B \<lbrace>Agent A, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
+          \<Longrightarrow> Says B Server
                   \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace>
                 # evs2 \<in> 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 \<in> yahalom;  Key KAB \<notin> used evs3;  KAB \<in> symKeys;
+ | YM3:  "\<lbrakk>evs3 \<in> yahalom;  Key KAB \<notin> used evs3;  KAB \<in> symKeys;
              Gets Server
                   \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace>
-               \<in> set evs3 |]
-          ==> Says Server A
+               \<in> set evs3\<rbrakk>
+          \<Longrightarrow> Says Server A
                    \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key KAB, Nonce NA, Nonce NB\<rbrace>,
                      Crypt (shrK B) \<lbrace>Agent A, Key KAB\<rbrace>\<rbrace>
                 # evs3 \<in> 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 \<noteq> Server is needed to prove Says_Server_not_range.*)
- | YM4:  "[| evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
+ | YM4:  "\<lbrakk>evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
              Gets A \<lbrace>Crypt(shrK A) \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace>, X\<rbrace>
                 \<in> set evs4;
-             Says A B \<lbrace>Agent A, Nonce NA\<rbrace> \<in> set evs4 |]
-          ==> Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> # evs4 \<in> yahalom"
+             Says A B \<lbrace>Agent A, Nonce NA\<rbrace> \<in> set evs4\<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> # evs4 \<in> yahalom"
 
 
 declare Says_imp_knows_Spy [THEN analz.Inj, dest]
@@ -68,8 +68,8 @@
 
 
 text\<open>A "possibility property": there are traces that reach the end\<close>
-lemma "[| A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys |] 
-       ==> \<exists>X NB. \<exists>evs \<in> yahalom.
+lemma "\<lbrakk>A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys\<rbrakk> 
+       \<Longrightarrow> \<exists>X NB. \<exists>evs \<in> yahalom.
               Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2] yahalom.Nil
@@ -83,12 +83,12 @@
 subsection\<open>Regularity Lemmas for Yahalom\<close>
 
 lemma Gets_imp_Says:
-     "[| Gets B X \<in> set evs; evs \<in> yahalom |] ==> \<exists>A. Says A B X \<in> set evs"
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> yahalom\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> 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 \<in> set evs; evs \<in> yahalom |]  ==> X \<in> knows Spy evs"
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> yahalom\<rbrakk>  \<Longrightarrow> X \<in> 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\<open>Lets us treat YM4 using a similar argument as for the Fake case.\<close>
 lemma YM4_analz_knows_Spy:
-     "[| Gets A \<lbrace>Crypt (shrK A) Y, X\<rbrace> \<in> set evs;  evs \<in> yahalom |]
-      ==> X \<in> analz (knows Spy evs)"
+     "\<lbrakk>Gets A \<lbrace>Crypt (shrK A) Y, X\<rbrace> \<in> set evs;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> X \<in> analz (knows Spy evs)"
 by blast
 
 lemmas YM4_parts_knows_Spy =
@@ -111,24 +111,24 @@
 
 text\<open>Spy never sees a good agent's shared key!\<close>
 lemma Spy_see_shrK [simp]:
-     "evs \<in> yahalom ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
+     "evs \<in> yahalom \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
 apply (erule yahalom.induct, force,
        drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+)
 done
 
 lemma Spy_analz_shrK [simp]:
-     "evs \<in> yahalom ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
+     "evs \<in> yahalom \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
 by auto
 
 lemma Spy_see_shrK_D [dest!]:
-     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> yahalom|] ==> A \<in> bad"
+     "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> yahalom\<rbrakk> \<Longrightarrow> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
 text\<open>Nobody can have used non-existent keys!
     Needed to apply \<open>analz_insert_Key\<close>\<close>
 lemma new_keys_not_used [simp]:
-    "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> yahalom|]
-     ==> K \<notin> keysFor (parts (spies evs))"
+    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> yahalom\<rbrakk>
+     \<Longrightarrow> K \<notin> 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 \<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+  Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow>
   Key K \<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
@@ -151,7 +151,7 @@
 subsection\<open>Session keys are not used to encrypt other session keys\<close>
 
 lemma analz_image_freshK [rule_format]:
- "evs \<in> yahalom ==>
+ "evs \<in> yahalom \<Longrightarrow>
    \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
           (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) =
           (K \<in> KK | Key K \<in> 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 \<in> yahalom;  KAB \<notin> range shrK |] ==>
+     "\<lbrakk>evs \<in> yahalom;  KAB \<notin> range shrK\<rbrakk> \<Longrightarrow>
       (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
       (K = KAB | Key K \<in> analz (knows Spy evs))"
 by (simp only: analz_image_freshK analz_image_freshK_simps)
@@ -167,12 +167,12 @@
 
 text\<open>The Key K uniquely identifies the Server's  message.\<close>
 lemma unique_session_keys:
-     "[| Says Server A
+     "\<lbrakk>Says Server A
           \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>, X\<rbrace> \<in> set evs;
         Says Server A'
           \<lbrace>Crypt (shrK A') \<lbrace>Agent B', Key K, na', nb'\<rbrace>, X'\<rbrace> \<in> set evs;
-        evs \<in> yahalom |]
-     ==> A=A' \<and> B=B' \<and> na=na' \<and> nb=nb'"
+        evs \<in> yahalom\<rbrakk>
+     \<Longrightarrow> A=A' \<and> B=B' \<and> na=na' \<and> nb=nb'"
 apply (erule rev_mp, erule rev_mp)
 apply (erule yahalom.induct, simp_all)
 txt\<open>YM3, by freshness, and YM4\<close>
@@ -182,8 +182,8 @@
 
 text\<open>Crucial secrecy property: Spy does not see the keys sent in msg YM3\<close>
 lemma secrecy_lemma:
-     "[| A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
-      ==> Says Server A
+     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> Says Server A
             \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>,
               Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace>
            \<in> set evs \<longrightarrow>
@@ -195,12 +195,12 @@
 
 text\<open>Final version\<close>
 lemma Spy_not_see_encrypted_key:
-     "[| Says Server A
+     "\<lbrakk>Says Server A
             \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>,
               Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace>
            \<in> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest: secrecy_lemma)
 
 
@@ -208,9 +208,9 @@
 
 text\<open>If the encrypted message appears then it originated with the Server\<close>
 lemma A_trusts_YM3:
-     "[| Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace> \<in> parts (knows Spy evs);
-         A \<notin> bad;  evs \<in> yahalom |]
-       ==> Says Server A
+     "\<lbrakk>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace> \<in> parts (knows Spy evs);
+         A \<notin> bad;  evs \<in> yahalom\<rbrakk>
+       \<Longrightarrow> Says Server A
             \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>,
               Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace>
            \<in> set evs"
@@ -224,9 +224,9 @@
 text\<open>The obvious combination of \<open>A_trusts_YM3\<close> with
   \<open>Spy_not_see_encrypted_key\<close>\<close>
 lemma A_gets_good_key:
-     "[| Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace> \<in> parts (knows Spy evs);
-         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+     "\<lbrakk>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace> \<in> parts (knows Spy evs);
+         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key)
 
 subsection\<open>Security Guarantees for B upon receiving YM4\<close>
@@ -234,9 +234,9 @@
 text\<open>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.\<close>
 lemma B_trusts_YM4_shrK:
-     "[| Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace> \<in> parts (knows Spy evs);
-         B \<notin> bad;  evs \<in> yahalom |]
-      ==> \<exists>NA NB. Says Server A
+     "\<lbrakk>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace> \<in> parts (knows Spy evs);
+         B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> \<exists>NA NB. Says Server A
                       \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace>,
                         Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace>
                      \<in> set evs"
@@ -259,8 +259,8 @@
   Secrecy of K is assumed; the valid Yahalom proof uses (and later proves)
   the secrecy of NB.\<close>
 lemma B_trusts_YM4_newK [rule_format]:
-     "[|Key K \<notin> analz (knows Spy evs);  evs \<in> yahalom|]
-      ==> Crypt K (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
+     "\<lbrakk>Key K \<notin> analz (knows Spy evs);  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> Crypt K (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
           (\<exists>A B NA. Says Server A
                       \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K,
                                 Nonce NA, Nonce NB\<rbrace>,
@@ -285,13 +285,13 @@
 text\<open>B's session key guarantee from YM4.  The two certificates contribute to a
   single conclusion about the Server's message.\<close>
 lemma B_trusts_YM4:
-     "[| Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>,
+     "\<lbrakk>Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>,
                   Crypt K (Nonce NB)\<rbrace> \<in> set evs;
          Says B Server
            \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace>
            \<in> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
-       ==> \<exists>na nb. Says Server A
+         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+       \<Longrightarrow> \<exists>na nb. Says Server A
                    \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>,
                      Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace>
              \<in> set evs"
@@ -302,13 +302,13 @@
 text\<open>The obvious combination of \<open>B_trusts_YM4\<close> with 
   \<open>Spy_not_see_encrypted_key\<close>\<close>
 lemma B_gets_good_key:
-     "[| Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>,
+     "\<lbrakk>Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>,
                   Crypt K (Nonce NB)\<rbrace> \<in> set evs;
          Says B Server
            \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace>
            \<in> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
-      ==> Key K \<notin> analz (knows Spy evs)"
+         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest!: B_trusts_YM4 Spy_not_see_encrypted_key)
 
 
@@ -323,7 +323,7 @@
   NB matters for freshness.\<close>
 lemma A_Said_YM3_lemma [rule_format]:
      "evs \<in> yahalom
-      ==> Key K \<notin> analz (knows Spy evs) \<longrightarrow>
+      \<Longrightarrow> Key K \<notin> analz (knows Spy evs) \<longrightarrow>
           Crypt K (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
           Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
           B \<notin> bad \<longrightarrow>
@@ -349,13 +349,13 @@
   Moreover, A associates K with NB (thus is talking about the same run).
   Other premises guarantee secrecy of K.\<close>
 lemma YM4_imp_A_Said_YM3 [rule_format]:
-     "[| Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>,
+     "\<lbrakk>Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>,
                   Crypt K (Nonce NB)\<rbrace> \<in> set evs;
          Says B Server
            \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace>
            \<in> set evs;
-         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
-      ==> \<exists>X. Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> \<in> set evs"
+         A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom\<rbrakk>
+      \<Longrightarrow> \<exists>X. Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> \<in> set evs"
 by (blast intro!: A_Said_YM3_lemma
           dest: Spy_not_see_encrypted_key B_trusts_YM4 Gets_imp_Says)
 
--- a/src/HOL/Auth/ZhouGollmann.thy	Thu Oct 13 17:31:22 2022 +0200
+++ b/src/HOL/Auth/ZhouGollmann.thy	Fri Oct 14 10:30:37 2022 +0100
@@ -32,35 +32,35 @@
 
   Nil:  "[] \<in> zg"
 
-| Fake: "[| evsf \<in> zg;  X \<in> synth (analz (spies evsf)) |]
-         ==> Says Spy B X  # evsf \<in> zg"
+| Fake: "\<lbrakk>evsf \<in> zg;  X \<in> synth (analz (spies evsf))\<rbrakk>
+         \<Longrightarrow> Says Spy B X  # evsf \<in> zg"
 
-| Reception:  "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
+| Reception:  "\<lbrakk>evsr \<in> zg; Says A B X \<in> set evsr\<rbrakk> \<Longrightarrow> Gets B X # evsr \<in> 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 \<in> zg;  Nonce L \<notin> used evs1; C = Crypt K (Number m);
+| ZG1: "\<lbrakk>evs1 \<in> zg;  Nonce L \<notin> used evs1; C = Crypt K (Number m);
            K \<in> symKeys;
-           NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>|]
-       ==> Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> # evs1 \<in> zg"
+           NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>\<rbrakk>
+       \<Longrightarrow> Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> # evs1 \<in> zg"
 
   (*B must check that NRO is A's signature to learn the sender's name*)
-| ZG2: "[| evs2 \<in> zg;
+| ZG2: "\<lbrakk>evs2 \<in> zg;
            Gets B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs2;
            NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>;
-           NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>|]
-       ==> Says B A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> # evs2  \<in>  zg"
+           NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>\<rbrakk>
+       \<Longrightarrow> Says B A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> # evs2  \<in>  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 \<in> zg; C = Crypt K M; K \<in> symKeys;
+| ZG3: "\<lbrakk>evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
            Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs3;
            Gets A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs3;
            NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>;
-           sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>|]
-       ==> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace>
+           sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>\<rbrakk>
+       \<Longrightarrow> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace>
              # evs3 \<in> 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 \<noteq> priK TTP"}. *)
-| ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
+| ZG4: "\<lbrakk>evs4 \<in> zg; K \<in> symKeys;
            Gets TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace>
              \<in> set evs4;
            sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
            con_K = Crypt (priK TTP) \<lbrace>Number f_con, Agent A, Agent B,
-                                      Nonce L, Key K\<rbrace>|]
-       ==> Says TTP Spy con_K
+                                      Nonce L, Key K\<rbrace>\<rbrakk>
+       \<Longrightarrow> Says TTP Spy con_K
            #
            Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
            # evs4 \<in> zg"
@@ -90,7 +90,7 @@
 
 
 text\<open>A "possibility property": there are traces that reach the end\<close>
-lemma "[|A \<noteq> B; TTP \<noteq> A; TTP \<noteq> B; K \<in> symKeys|] ==>
+lemma "\<lbrakk>A \<noteq> B; TTP \<noteq> A; TTP \<noteq> B; K \<in> symKeys\<rbrakk> \<Longrightarrow>
      \<exists>L. \<exists>evs \<in> zg.
            Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K,
                Crypt (priK TTP) \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K\<rbrace>\<rbrace>
@@ -107,45 +107,45 @@
 subsection \<open>Basic Lemmas\<close>
 
 lemma Gets_imp_Says:
-     "[| Gets B X \<in> set evs; evs \<in> zg |] ==> \<exists>A. Says A B X \<in> set evs"
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> zg\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
 apply (erule rev_mp)
 apply (erule zg.induct, auto)
 done
 
 lemma Gets_imp_knows_Spy:
-     "[| Gets B X \<in> set evs; evs \<in> zg |]  ==> X \<in> spies evs"
+     "\<lbrakk>Gets B X \<in> set evs; evs \<in> zg\<rbrakk>  \<Longrightarrow> X \<in> spies evs"
 by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
 
 
 text\<open>Lets us replace proofs about \<^term>\<open>used evs\<close> by simpler proofs 
 about \<^term>\<open>parts (spies evs)\<close>.\<close>
 lemma Crypt_used_imp_spies:
-     "[| Crypt K X \<in> used evs; evs \<in> zg |]
-      ==> Crypt K X \<in> parts (spies evs)"
+     "\<lbrakk>Crypt K X \<in> used evs; evs \<in> zg\<rbrakk>
+      \<Longrightarrow> Crypt K X \<in> 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 \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
+     "\<lbrakk>Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
            \<in> set evs;
         sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
-        evs \<in> zg|]
-    ==> Gets TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
+        evs \<in> zg\<rbrakk>
+    \<Longrightarrow> Gets TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule zg.induct, auto)
 done
 
 text\<open>For reasoning about C, which is encrypted in message ZG2\<close>
 lemma ZG2_msg_in_parts_spies:
-     "[|Gets B \<lbrace>F, B', L, C, X\<rbrace> \<in> set evs; evs \<in> zg|]
-      ==> C \<in> parts (spies evs)"
+     "\<lbrakk>Gets B \<lbrace>F, B', L, C, X\<rbrace> \<in> set evs; evs \<in> zg\<rbrakk>
+      \<Longrightarrow> C \<in> parts (spies evs)"
 by (blast dest: Gets_imp_Says)
 
 (*classical regularity lemma on priK*)
 lemma Spy_see_priK [simp]:
-     "evs \<in> zg ==> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)"
+     "evs \<in> zg \<Longrightarrow> (Key (priK A) \<in> parts (spies evs)) = (A \<in> 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 \<in> zg ==> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)"
+     "evs \<in> zg \<Longrightarrow> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto 
 
 
@@ -165,10 +165,10 @@
 
 text\<open>Strong conclusion for a good agent\<close>
 lemma NRO_validity_good:
-     "[|NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>;
+     "\<lbrakk>NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>;
         NRO \<in> parts (spies evs);
-        A \<notin> bad;  evs \<in> zg |]
-     ==> Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs"
+        A \<notin> bad;  evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule zg.induct)
@@ -176,18 +176,18 @@
 done
 
 lemma NRO_sender:
-     "[|Says A' B \<lbrace>n, b, l, C, Crypt (priK A) X\<rbrace> \<in> set evs; evs \<in> zg|]
-    ==> A' \<in> {A,Spy}"
+     "\<lbrakk>Says A' B \<lbrace>n, b, l, C, Crypt (priK A) X\<rbrace> \<in> set evs; evs \<in> zg\<rbrakk>
+    \<Longrightarrow> A' \<in> {A,Spy}"
 apply (erule rev_mp)  
 apply (erule zg.induct, simp_all)
 done
 
 text\<open>Holds also for \<^term>\<open>A = Spy\<close>!\<close>
 theorem NRO_validity:
-     "[|Gets B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs;
+     "\<lbrakk>Gets B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs;
         NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>;
-        A \<notin> broken;  evs \<in> zg |]
-     ==> Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs"
+        A \<notin> broken;  evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says A B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs"
 apply (drule Gets_imp_Says, assumption) 
 apply clarify 
 apply (frule NRO_sender, auto)
@@ -205,10 +205,10 @@
 
 text\<open>Strong conclusion for a good agent\<close>
 lemma NRR_validity_good:
-     "[|NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>;
+     "\<lbrakk>NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>;
         NRR \<in> parts (spies evs);
-        B \<notin> bad;  evs \<in> zg |]
-     ==> Says B A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
+        B \<notin> bad;  evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says B A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule zg.induct) 
@@ -216,18 +216,18 @@
 done
 
 lemma NRR_sender:
-     "[|Says B' A \<lbrace>n, a, l, Crypt (priK B) X\<rbrace> \<in> set evs; evs \<in> zg|]
-    ==> B' \<in> {B,Spy}"
+     "\<lbrakk>Says B' A \<lbrace>n, a, l, Crypt (priK B) X\<rbrace> \<in> set evs; evs \<in> zg\<rbrakk>
+    \<Longrightarrow> B' \<in> {B,Spy}"
 apply (erule rev_mp)  
 apply (erule zg.induct, simp_all)
 done
 
 text\<open>Holds also for \<^term>\<open>B = Spy\<close>!\<close>
 theorem NRR_validity:
-     "[|Says B' A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs;
+     "\<lbrakk>Says B' A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs;
         NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>;
-        B \<notin> broken; evs \<in> zg|]
-    ==> Says B A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
+        B \<notin> broken; evs \<in> zg\<rbrakk>
+    \<Longrightarrow> Says B A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
 apply clarify 
 apply (frule NRR_sender, auto)
 txt\<open>We are left with the case where \<^term>\<open>B' = Spy\<close> and  \<^term>\<open>B' \<noteq> B\<close>,
@@ -243,10 +243,10 @@
 
 text\<open>Strong conclusion for a good agent\<close>
 lemma sub_K_validity_good:
-     "[|sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
+     "\<lbrakk>sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
         sub_K \<in> parts (spies evs);
-        A \<notin> bad;  evs \<in> zg |]
-     ==> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
+        A \<notin> bad;  evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule zg.induct)
@@ -256,18 +256,18 @@
 done
 
 lemma sub_K_sender:
-     "[|Says A' TTP \<lbrace>n, b, l, k, Crypt (priK A) X\<rbrace> \<in> set evs;  evs \<in> zg|]
-    ==> A' \<in> {A,Spy}"
+     "\<lbrakk>Says A' TTP \<lbrace>n, b, l, k, Crypt (priK A) X\<rbrace> \<in> set evs;  evs \<in> zg\<rbrakk>
+    \<Longrightarrow> A' \<in> {A,Spy}"
 apply (erule rev_mp)  
 apply (erule zg.induct, simp_all)
 done
 
 text\<open>Holds also for \<^term>\<open>A = Spy\<close>!\<close>
 theorem sub_K_validity:
-     "[|Gets TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs;
+     "\<lbrakk>Gets TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs;
         sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
-        A \<notin> broken;  evs \<in> zg |]
-     ==> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
+        A \<notin> broken;  evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
 apply (drule Gets_imp_Says, assumption) 
 apply clarify 
 apply (frule sub_K_sender, auto)
@@ -286,11 +286,11 @@
 that \<^term>\<open>A\<close> sent \<^term>\<open>sub_K\<close>\<close>
 
 lemma con_K_validity:
-     "[|con_K \<in> used evs;
+     "\<lbrakk>con_K \<in> used evs;
         con_K = Crypt (priK TTP)
                   \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K\<rbrace>;
-        evs \<in> zg |]
-    ==> Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
+        evs \<in> zg\<rbrakk>
+    \<Longrightarrow> Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
           \<in> set evs"
 apply clarify
 apply (erule rev_mp)
@@ -306,11 +306,11 @@
  \<^term>\<open>sub_K\<close>.  We assume that \<^term>\<open>A\<close> is not broken.  Importantly, nothing
   needs to be assumed about the form of \<^term>\<open>con_K\<close>!\<close>
 lemma Notes_TTP_imp_Says_A:
-     "[|Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
+     "\<lbrakk>Notes TTP \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K, con_K\<rbrace>
            \<in> set evs;
         sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
-        A \<notin> broken; evs \<in> zg|]
-     ==> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
+        A \<notin> broken; evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule zg.induct)
@@ -323,12 +323,12 @@
 text\<open>If \<^term>\<open>con_K\<close> exists, then \<^term>\<open>A\<close> sent \<^term>\<open>sub_K\<close>.  We again
    assume that \<^term>\<open>A\<close> is not broken.\<close>
 theorem B_sub_K_validity:
-     "[|con_K \<in> used evs;
+     "\<lbrakk>con_K \<in> used evs;
         con_K = Crypt (priK TTP) \<lbrace>Number f_con, Agent A, Agent B,
                                    Nonce L, Key K\<rbrace>;
         sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
-        A \<notin> broken; evs \<in> zg|]
-     ==> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
+        A \<notin> broken; evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Says A TTP \<lbrace>Number f_sub, Agent B, Nonce L, Key K, sub_K\<rbrace> \<in> set evs"
 by (blast dest: con_K_validity Notes_TTP_imp_Says_A)
 
 
@@ -340,12 +340,12 @@
 
 text\<open>Strange: unicity of the label protects \<^term>\<open>A\<close>?\<close>
 lemma A_unicity: 
-     "[|NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, Crypt K M\<rbrace>;
+     "\<lbrakk>NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, Crypt K M\<rbrace>;
         NRO \<in> parts (spies evs);
         Says A B \<lbrace>Number f_nro, Agent B, Nonce L, Crypt K M', NRO'\<rbrace>
           \<in> set evs;
-        A \<notin> bad; evs \<in> zg |]
-     ==> M'=M"
+        A \<notin> bad; evs \<in> zg\<rbrakk>
+     \<Longrightarrow> M'=M"
 apply clarify
 apply (erule rev_mp)
 apply (erule rev_mp)
@@ -359,13 +359,13 @@
 text\<open>Fairness lemma: if \<^term>\<open>sub_K\<close> exists, then \<^term>\<open>A\<close> holds 
 NRR.  Relies on unicity of labels.\<close>
 lemma sub_K_implies_NRR:
-     "[| NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, Crypt K M\<rbrace>;
+     "\<lbrakk>NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, Crypt K M\<rbrace>;
          NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, Crypt K M\<rbrace>;
          sub_K \<in> parts (spies evs);
          NRO \<in> parts (spies evs);
          sub_K = Crypt (priK A) \<lbrace>Number f_sub, Agent B, Nonce L, Key K\<rbrace>;
-         A \<notin> bad;  evs \<in> zg |]
-     ==> Gets A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
+         A \<notin> bad;  evs \<in> zg\<rbrakk>
+     \<Longrightarrow> Gets A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
 apply clarify
 apply hypsubst_thin
 apply (erule rev_mp)
@@ -382,8 +382,8 @@
 
 
 lemma Crypt_used_imp_L_used:
-     "[| Crypt (priK TTP) \<lbrace>F, A, B, L, K\<rbrace> \<in> used evs; evs \<in> zg |]
-      ==> L \<in> used evs"
+     "\<lbrakk>Crypt (priK TTP) \<lbrace>F, A, B, L, K\<rbrace> \<in> used evs; evs \<in> zg\<rbrakk>
+      \<Longrightarrow> L \<in> used evs"
 apply (erule rev_mp)
 apply (erule zg.induct, auto)
 txt\<open>Fake\<close>
@@ -397,14 +397,14 @@
 then \<^term>\<open>A\<close> holds NRR.  \<^term>\<open>A\<close> must be uncompromised, but there is no
 assumption about \<^term>\<open>B\<close>.\<close>
 theorem A_fairness_NRO:
-     "[|con_K \<in> used evs;
+     "\<lbrakk>con_K \<in> used evs;
         NRO \<in> parts (spies evs);
         con_K = Crypt (priK TTP)
                       \<lbrace>Number f_con, Agent A, Agent B, Nonce L, Key K\<rbrace>;
         NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, Crypt K M\<rbrace>;
         NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, Crypt K M\<rbrace>;
-        A \<notin> bad;  evs \<in> zg |]
-    ==> Gets A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
+        A \<notin> bad;  evs \<in> zg\<rbrakk>
+    \<Longrightarrow> Gets A \<lbrace>Number f_nrr, Agent A, Nonce L, NRR\<rbrace> \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule rev_mp)
@@ -425,11 +425,11 @@
 text\<open>Fairness for \<^term>\<open>B\<close>: NRR exists at all, then \<^term>\<open>B\<close> holds NRO.
 \<^term>\<open>B\<close> must be uncompromised, but there is no assumption about \<^term>\<open>A\<close>.\<close>
 theorem B_fairness_NRR:
-     "[|NRR \<in> used evs;
+     "\<lbrakk>NRR \<in> used evs;
         NRR = Crypt (priK B) \<lbrace>Number f_nrr, Agent A, Nonce L, C\<rbrace>;
         NRO = Crypt (priK A) \<lbrace>Number f_nro, Agent B, Nonce L, C\<rbrace>;
-        B \<notin> bad; evs \<in> zg |]
-    ==> Gets B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs"
+        B \<notin> bad; evs \<in> zg\<rbrakk>
+    \<Longrightarrow> Gets B \<lbrace>Number f_nro, Agent B, Nonce L, C, NRO\<rbrace> \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule zg.induct)