isabelle update_cartouches -c -t;
authorwenzelm
Thu, 10 Dec 2015 21:39:33 +0100
changeset 61830 4f5ab843cf5b
parent 61829 55c85d25e18c
child 61831 c43f87119d80
child 61834 2154e6c8d52d
isabelle update_cartouches -c -t;
src/HOL/Auth/All_Symmetric.thy
src/HOL/Auth/Auth_Public.thy
src/HOL/Auth/Auth_Shared.thy
src/HOL/Auth/CertifiedEmail.thy
src/HOL/Auth/Event.thy
src/HOL/Auth/Guard/Analz.thy
src/HOL/Auth/Guard/Auth_Guard_Public.thy
src/HOL/Auth/Guard/Auth_Guard_Shared.thy
src/HOL/Auth/Guard/Extensions.thy
src/HOL/Auth/Guard/Guard.thy
src/HOL/Auth/Guard/GuardK.thy
src/HOL/Auth/Guard/Guard_NS_Public.thy
src/HOL/Auth/Guard/Guard_OtwayRees.thy
src/HOL/Auth/Guard/Guard_Public.thy
src/HOL/Auth/Guard/Guard_Shared.thy
src/HOL/Auth/Guard/Guard_Yahalom.thy
src/HOL/Auth/Guard/List_Msg.thy
src/HOL/Auth/Guard/P1.thy
src/HOL/Auth/Guard/P2.thy
src/HOL/Auth/Guard/Proto.thy
src/HOL/Auth/KerberosIV.thy
src/HOL/Auth/KerberosIV_Gets.thy
src/HOL/Auth/KerberosV.thy
src/HOL/Auth/Kerberos_BAN.thy
src/HOL/Auth/Kerberos_BAN_Gets.thy
src/HOL/Auth/Message.thy
src/HOL/Auth/NS_Public.thy
src/HOL/Auth/NS_Public_Bad.thy
src/HOL/Auth/NS_Shared.thy
src/HOL/Auth/OtwayRees.thy
src/HOL/Auth/OtwayReesBella.thy
src/HOL/Auth/OtwayRees_AN.thy
src/HOL/Auth/OtwayRees_Bad.thy
src/HOL/Auth/Public.thy
src/HOL/Auth/Recur.thy
src/HOL/Auth/Shared.thy
src/HOL/Auth/Smartcard/Auth_Smartcard.thy
src/HOL/Auth/Smartcard/EventSC.thy
src/HOL/Auth/Smartcard/ShoupRubin.thy
src/HOL/Auth/Smartcard/ShoupRubinBella.thy
src/HOL/Auth/Smartcard/Smartcard.thy
src/HOL/Auth/TLS.thy
src/HOL/Auth/WooLam.thy
src/HOL/Auth/Yahalom.thy
src/HOL/Auth/Yahalom2.thy
src/HOL/Auth/Yahalom_Bad.thy
src/HOL/Auth/ZhouGollmann.thy
--- a/src/HOL/Auth/All_Symmetric.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/All_Symmetric.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
 imports Message
 begin
 
-text {* All keys are symmetric *}
+text \<open>All keys are symmetric\<close>
 
 defs all_symmetric_def: "all_symmetric \<equiv> True"
 
--- a/src/HOL/Auth/Auth_Public.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Auth_Public.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
     Copyright   1996  University of Cambridge
 *)
 
-section {* Conventional protocols: rely on conventional Message, Event and Public -- Public-key protocols *}
+section \<open>Conventional protocols: rely on conventional Message, Event and Public -- Public-key protocols\<close>
 
 theory Auth_Public
 imports
--- a/src/HOL/Auth/Auth_Shared.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Auth_Shared.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
     Copyright   1996  University of Cambridge
 *)
 
-section {* Conventional protocols: rely on conventional Message, Event and Public -- Shared-key protocols *}
+section \<open>Conventional protocols: rely on conventional Message, Event and Public -- Shared-key protocols\<close>
 
 theory Auth_Shared
 imports
--- a/src/HOL/Auth/CertifiedEmail.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/CertifiedEmail.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
     Author:     Giampaolo Bella, Christiano Longo and Lawrence C Paulson
 *)
 
-section{*The Certified Electronic Mail Protocol by Abadi et al.*}
+section\<open>The Certified Electronic Mail Protocol by Abadi et al.\<close>
 
 theory CertifiedEmail imports Public begin
 
@@ -23,7 +23,7 @@
   SAuth    :: nat
   BothAuth :: nat
 
-text{*We formalize a fixed way of computing responses.  Could be better.*}
+text\<open>We formalize a fixed way of computing responses.  Could be better.\<close>
 definition "response" :: "agent => agent => nat => msg" where
    "response S R q == Hash {|Agent S, Key (shrK R), Nonce q|}"
 
@@ -31,20 +31,20 @@
 inductive_set certified_mail :: "event list set"
   where
 
-  Nil: --{*The empty trace*}
+  Nil: \<comment>\<open>The empty trace\<close>
      "[] \<in> certified_mail"
 
-| Fake: --{*The Spy may say anything he can say.  The sender field is correct,
-          but agents don't use that information.*}
+| 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"
 
-| FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
-    equipped with the necessary credentials to serve as an SSL server.*}
+| 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 {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
 
-| CM1: --{*The sender approaches the recipient.  The message is a number.*}
+| CM1: \<comment>\<open>The sender approaches the recipient.  The message is a number.\<close>
  "[|evs1 \<in> certified_mail;
     Key K \<notin> used evs1;
     K \<in> symKeys;
@@ -55,8 +55,8 @@
                  Number cleartext, Nonce q, S2TTP|} # evs1 
         \<in> certified_mail"
 
-| CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
-     password to @{term TTP} over an SSL channel.*}
+| CM2: \<comment>\<open>The recipient records @{term S2TTP} while transmitting it and her
+     password to @{term TTP} over an SSL channel.\<close>
  "[|evs2 \<in> certified_mail;
     Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext, 
              Nonce q, S2TTP|} \<in> set evs2;
@@ -66,11 +66,11 @@
    Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} # evs2
       \<in> certified_mail"
 
-| CM3: --{*@{term TTP} simultaneously reveals the key to the recipient and gives
+| CM3: \<comment>\<open>@{term TTP} simultaneously reveals the key to the recipient and gives
          a receipt to the sender.  The SSL channel does not authenticate 
          the client (@{term R}), but @{term TTP} accepts the message only 
          if the given password is that of the claimed sender, @{term R}.
-         He replies over the established SSL channel.*}
+         He replies over the established SSL channel.\<close>
  "[|evs3 \<in> certified_mail;
     Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} \<in> set evs3;
     S2TTP = Crypt (pubEK TTP) 
@@ -137,9 +137,9 @@
 apply (synth_analz_mono_contra, simp_all, blast+)
 done 
 
-text{*Cannot strengthen the first disjunct to @{term "R\<noteq>Spy"} because
+text\<open>Cannot strengthen the first disjunct to @{term "R\<noteq>Spy"} because
 the fakessl rule allows Spy to spoof the sender's name.  Maybe can
-strengthen the second disjunct with @{term "R\<noteq>Spy"}.*}
+strengthen the second disjunct with @{term "R\<noteq>Spy"}.\<close>
 lemma hr_form:
  "[|Notes TTP {|Agent R, Agent TTP, S2TTP, pwd, hr|} \<in> set evs;
     evs \<in> certified_mail|]
@@ -152,11 +152,11 @@
      ==> A \<in> bad"
 apply (erule rev_mp) 
 apply (erule certified_mail.induct, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest: Fake_parts_insert_in_Un) 
-txt{*Message 1*}
+txt\<open>Message 1\<close>
 apply blast  
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (frule_tac hr_form, assumption)
 apply (elim disjE exE) 
 apply (simp_all add: parts_insert2) 
@@ -177,8 +177,8 @@
      "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> analz(spies evs)"
 by auto
 
-text{*Thus, prove any goal that assumes that @{term Spy} knows a private key
-belonging to @{term TTP}*}
+text\<open>Thus, prove any goal that assumes that @{term Spy} knows a private key
+belonging to @{term TTP}\<close>
 declare Spy_dont_know_TTPKey_parts [THEN [2] rev_notE, elim!]
 
 
@@ -192,22 +192,22 @@
 apply (erule rev_mp)
 apply (erule certified_mail.induct, simp_all)
    apply (blast  intro:parts_insertI)
-txt{*Fake SSL*}
+txt\<open>Fake SSL\<close>
 apply (blast dest: parts.Body) 
-txt{*Message 2*}
+txt\<open>Message 2\<close>
 apply (blast dest!: Gets_imp_Says elim!: knows_Spy_partsEs)
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (metis parts_insertI)
 done
 
 lemma Spy_dont_know_RPwd [rule_format]:
     "evs \<in> certified_mail ==> Key (RPwd A) \<in> parts(spies evs) --> A \<in> bad"
 apply (erule certified_mail.induct, simp_all) 
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest: Fake_parts_insert_in_Un) 
-txt{*Message 1*}
+txt\<open>Message 1\<close>
 apply blast  
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (frule CM3_k_parts_knows_Spy, assumption)
 apply (frule_tac hr_form, assumption)
 apply (elim disjE exE) 
@@ -225,17 +225,17 @@
     "evs \<in> certified_mail ==> (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{*Unused, but a guarantee of sorts*}
+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"
 apply (erule rev_mp)
 apply (erule certified_mail.induct, simp_all) 
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest: Spy_dont_know_private_keys Fake_parts_insert_in_Un)
-txt{*Message 1*}
+txt\<open>Message 1\<close>
 apply blast 
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (frule_tac hr_form, assumption)
 apply (elim disjE exE) 
 apply (simp_all add: parts_insert2 parts_insert_knows_A) 
@@ -243,7 +243,7 @@
 done
 
 
-subsection{*Proving Confidentiality Results*}
+subsection\<open>Proving Confidentiality Results\<close>
 
 lemma analz_image_freshK [rule_format]:
  "evs \<in> certified_mail ==>
@@ -269,8 +269,8 @@
       (K = KAB | Key K \<in> analz (spies evs))"
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
-text{*@{term S2TTP} must have originated from a valid sender
-    provided @{term K} is secure.  Proof is surprisingly hard.*}
+text\<open>@{term S2TTP} must have originated from a valid sender
+    provided @{term K} is secure.  Proof is surprisingly hard.\<close>
 
 lemma Notes_SSL_imp_used:
      "[|Notes B {|Agent A, Agent B, X|} \<in> set evs|] ==> X \<in> used evs"
@@ -294,18 +294,18 @@
 apply (erule certified_mail.induct, analz_mono_contra)
 apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp)
 apply (simp add: used_Nil Crypt_notin_initState, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest: Fake_parts_sing [THEN subsetD]
              dest!: analz_subset_parts [THEN subsetD])  
-txt{*Fake SSL*}
+txt\<open>Fake SSL\<close>
 apply (blast dest: Fake_parts_sing [THEN subsetD]
              dest: analz_subset_parts [THEN subsetD])  
-txt{*Message 1*}
+txt\<open>Message 1\<close>
 apply (clarsimp, blast)
-txt{*Message 2*}
+txt\<open>Message 2\<close>
 apply (simp add: parts_insert2, clarify) 
 apply (metis parts_cut Un_empty_left usedI)
-txt{*Message 3*} 
+txt\<open>Message 3\<close> 
 apply (blast dest: Notes_SSL_imp_used used_parts_subset_parts) 
 done 
 
@@ -323,26 +323,26 @@
 by (blast intro: S2TTP_sender_lemma) 
 
 
-text{*Nobody can have used non-existent keys!*}
+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))"
 apply (erule rev_mp) 
 apply (erule certified_mail.induct, simp_all) 
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert) 
-txt{*Message 1*}
+txt\<open>Message 1\<close>
 apply blast 
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (frule CM3_k_parts_knows_Spy, assumption)
 apply (frule_tac hr_form, assumption) 
 apply (force dest!: keysFor_parts_insert)
 done
 
 
-text{*Less easy to prove @{term "m'=m"}.  Maybe needs a separate unicity
+text\<open>Less easy to prove @{term "m'=m"}.  Maybe needs a separate unicity
 theorem for ciphertexts of the form @{term "Crypt K (Number m)"}, 
-where @{term K} is secure.*}
+where @{term K} is secure.\<close>
 lemma Key_unique_lemma [rule_format]:
      "evs \<in> certified_mail ==>
        Key K \<notin> analz (spies evs) -->
@@ -360,13 +360,13 @@
           \<in> set evs --> R' = R & S' = S & AO' = AO & hs' = hs))" 
 apply (erule certified_mail.induct, analz_mono_contra, simp_all)
  prefer 2
- txt{*Message 1*}
+ txt\<open>Message 1\<close>
  apply (blast dest!: Says_imp_knows_Spy [THEN parts.Inj] new_keys_not_used Crypt_imp_keysFor)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (auto dest!: usedI S2TTP_sender analz_subset_parts [THEN subsetD]) 
 done
 
-text{*The key determines the sender, recipient and protocol options.*}
+text\<open>The key determines the sender, recipient and protocol options.\<close>
 lemma Key_unique:
       "[|Says S R
            {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
@@ -384,11 +384,11 @@
 by (rule Key_unique_lemma, assumption+)
 
 
-subsection{*The Guarantees for Sender and Recipient*}
+subsection\<open>The Guarantees for Sender and Recipient\<close>
 
-text{*A Sender's guarantee:
+text\<open>A Sender's guarantee:
       If Spy gets the key then @{term R} is bad and @{term S} moreover
-      gets his return receipt (and therefore has no grounds for complaint).*}
+      gets his return receipt (and therefore has no grounds for complaint).\<close>
 theorem S_fairness_bad_R:
       "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                      Number cleartext, Nonce q, S2TTP|} \<in> set evs;
@@ -401,11 +401,11 @@
 apply (erule ssubst)
 apply (erule rev_mp)
 apply (erule certified_mail.induct, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*Fake SSL*}
+txt\<open>Fake SSL\<close>
 apply spy_analz
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (frule_tac hr_form, assumption)
 apply (elim disjE exE) 
 apply (simp_all add: synth_analz_insert_eq  
@@ -416,7 +416,7 @@
 apply (blast dest: Notes_SSL_imp_used S2TTP_sender Key_unique)+
 done
 
-text{*Confidentially for the symmetric key*}
+text\<open>Confidentially for the symmetric key\<close>
 theorem Spy_not_see_encrypted_key:
       "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                      Number cleartext, Nonce q, S2TTP|} \<in> set evs;
@@ -427,8 +427,8 @@
 by (blast dest: S_fairness_bad_R) 
 
 
-text{*Agent @{term R}, who may be the Spy, doesn't receive the key
- until @{term S} has access to the return receipt.*} 
+text\<open>Agent @{term R}, who may be the Spy, doesn't receive the key
+ until @{term S} has access to the return receipt.\<close> 
 theorem S_guarantee:
      "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                     Number cleartext, Nonce q, S2TTP|} \<in> set evs;
@@ -440,16 +440,16 @@
 apply (erule ssubst)
 apply (erule rev_mp)
 apply (erule certified_mail.induct, simp_all)
-txt{*Message 1*}
+txt\<open>Message 1\<close>
 apply (blast dest: Notes_imp_used) 
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (blast dest: Notes_SSL_imp_used S2TTP_sender Key_unique S_fairness_bad_R) 
 done
 
 
-text{*If @{term R} sends message 2, and a delivery certificate exists, 
+text\<open>If @{term R} sends message 2, and a delivery certificate exists, 
  then @{term R} receives the necessary key.  This result is also important
- to @{term S}, as it confirms the validity of the return receipt.*}
+ to @{term S}, as it confirms the validity of the return receipt.\<close>
 theorem RR_validity:
   "[|Crypt (priSK TTP) S2TTP \<in> used evs;
      S2TTP = Crypt (pubEK TTP)
@@ -462,16 +462,16 @@
 apply (erule ssubst)
 apply (erule ssubst)
 apply (erule certified_mail.induct, simp_all)
-txt{*Fake*} 
+txt\<open>Fake\<close> 
 apply (blast dest: Fake_parts_sing [THEN subsetD]
              dest!: analz_subset_parts [THEN subsetD])  
-txt{*Fake SSL*}
+txt\<open>Fake SSL\<close>
 apply (blast dest: Fake_parts_sing [THEN subsetD]
             dest!: analz_subset_parts [THEN subsetD])  
-txt{*Message 2*}
+txt\<open>Message 2\<close>
 apply (drule CM2_S2TTP_parts_knows_Spy, assumption)
 apply (force dest: parts_cut)
-txt{*Message 3*}
+txt\<open>Message 3\<close>
 apply (frule_tac hr_form, assumption)
 apply (elim disjE exE, simp_all) 
 apply (blast dest: Fake_parts_sing [THEN subsetD]
--- a/src/HOL/Auth/Event.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Event.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -8,7 +8,7 @@
     stores are visible to him
 *)
 
-section{*Theory of Events for Security Protocols*}
+section\<open>Theory of Events for Security Protocols\<close>
 
 theory Event imports Message begin
 
@@ -21,9 +21,9 @@
         | Notes agent       msg
        
 consts 
-  bad    :: "agent set"                         -- {* compromised agents *}
+  bad    :: "agent set"                         \<comment> \<open>compromised agents\<close>
 
-text{*Spy has access to his own key for spoof messages, but Server is secure*}
+text\<open>Spy has access to his own key for spoof messages, but Server is secure\<close>
 specification (bad)
   Spy_in_bad     [iff]: "Spy \<in> bad"
   Server_not_bad [iff]: "Server \<notin> bad"
@@ -54,7 +54,7 @@
   therefore the oops case must use Notes
 *)
 
-text{*The constant "spies" is retained for compatibility's sake*}
+text\<open>The constant "spies" is retained for compatibility's sake\<close>
 
 abbreviation (input)
   spies  :: "event list => msg set" where
@@ -72,9 +72,9 @@
                         Says A B X => parts {X} \<union> used evs
                       | Gets A X   => used evs
                       | Notes A X  => parts {X} \<union> used evs)"
-    --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
+    \<comment>\<open>The case for @{term Gets} seems anomalous, but @{term Gets} always
         follows @{term Says} in real protocols.  Seems difficult to change.
-        See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
+        See \<open>Gets_correct\<close> in theory \<open>Guard/Extensions.thy\<close>.\<close>
 
 lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs --> X \<in> used evs"
 apply (induct_tac evs)
@@ -87,7 +87,7 @@
 done
 
 
-subsection{*Function @{term knows}*}
+subsection\<open>Function @{term knows}\<close>
 
 (*Simplifying   
  parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
@@ -98,8 +98,8 @@
      "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
 by simp
 
-text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
-      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
+text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits
+      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}\<close>
 lemma knows_Spy_Notes [simp]:
      "knows Spy (Notes A X # evs) =  
           (if A:bad then insert X (knows Spy evs) else knows Spy evs)"
@@ -120,7 +120,7 @@
      "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)"
 by (simp add: subset_insertI)
 
-text{*Spy sees what is sent on the traffic*}
+text\<open>Spy sees what is sent on the traffic\<close>
 lemma Says_imp_knows_Spy [rule_format]:
      "Says A B X \<in> set evs --> X \<in> knows Spy evs"
 apply (induct_tac "evs")
@@ -134,8 +134,8 @@
 done
 
 
-text{*Elimination rules: derive contradictions from old Says events containing
-  items known to be fresh*}
+text\<open>Elimination rules: derive contradictions from old Says events containing
+  items known to be fresh\<close>
 lemmas Says_imp_parts_knows_Spy = 
        Says_imp_knows_Spy [THEN parts.Inj, elim_format] 
 
@@ -144,13 +144,13 @@
 
 lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj]
 
-text{*Compatibility for the old "spies" function*}
+text\<open>Compatibility for the old "spies" function\<close>
 lemmas spies_partsEs = knows_Spy_partsEs
 lemmas Says_imp_spies = Says_imp_knows_Spy
 lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy]
 
 
-subsection{*Knowledge of Agents*}
+subsection\<open>Knowledge of Agents\<close>
 
 lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)"
 by (simp add: subset_insertI)
@@ -161,21 +161,21 @@
 lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)"
 by (simp add: subset_insertI)
 
-text{*Agents know what they say*}
+text\<open>Agents know what they say\<close>
 lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs --> X \<in> knows A evs"
 apply (induct_tac "evs")
 apply (simp_all (no_asm_simp) split add: event.split)
 apply blast
 done
 
-text{*Agents know what they note*}
+text\<open>Agents know what they note\<close>
 lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs --> X \<in> knows A evs"
 apply (induct_tac "evs")
 apply (simp_all (no_asm_simp) split add: event.split)
 apply blast
 done
 
-text{*Agents know what they receive*}
+text\<open>Agents know what they receive\<close>
 lemma Gets_imp_knows_agents [rule_format]:
      "A \<noteq> Spy --> Gets A X \<in> set evs --> X \<in> knows A evs"
 apply (induct_tac "evs")
@@ -183,8 +183,8 @@
 done
 
 
-text{*What agents DIFFERENT FROM Spy know 
-  was either said, or noted, or got, or known initially*}
+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 |] ==> EX B.  
   Says A B X \<in> set evs | Gets A X \<in> set evs | Notes A X \<in> set evs | X \<in> initState A"
@@ -194,8 +194,8 @@
 apply blast
 done
 
-text{*What the Spy knows -- for the time being --
-  was either said or noted, or known initially*}
+text\<open>What the Spy knows -- for the time being --
+  was either said or noted, or known initially\<close>
 lemma knows_Spy_imp_Says_Notes_initState [rule_format]:
      "[| X \<in> knows Spy evs |] ==> EX A B.  
   Says A B X \<in> set evs | Notes A X \<in> set evs | X \<in> initState Spy"
@@ -231,15 +231,15 @@
 apply (blast intro: initState_into_used)
 done
 
-text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
+text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close>
 declare knows_Cons [simp del]
         used_Nil [simp del] used_Cons [simp del]
 
 
-text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
+text\<open>For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
   New events added by induction to "evs" are discarded.  Provided 
   this information isn't needed, the proof will be much shorter, since
-  it will omit complicated reasoning about @{term analz}.*}
+  it will omit complicated reasoning about @{term analz}.\<close>
 
 lemmas analz_mono_contra =
        knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
@@ -256,7 +256,7 @@
 done
 
 
-text{*For proving @{text new_keys_not_used}*}
+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"
@@ -269,23 +269,23 @@
 lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs
 
 ML
-{*
+\<open>
 fun analz_mono_contra_tac ctxt = 
   resolve_tac ctxt @{thms analz_impI} THEN' 
   REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
   THEN' (mp_tac ctxt)
-*}
+\<close>
 
-method_setup analz_mono_contra = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt))) *}
+method_setup analz_mono_contra = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close>
     "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P"
 
-subsubsection{*Useful for case analysis on whether a hash is a spoof or not*}
+subsubsection\<open>Useful for case analysis on whether a hash is a spoof or not\<close>
 
 lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))"] for Y evs
 
 ML
-{*
+\<open>
 fun synth_analz_mono_contra_tac ctxt = 
   resolve_tac ctxt @{thms syan_impI} THEN'
   REPEAT1 o 
@@ -295,10 +295,10 @@
       @{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
   THEN'
   mp_tac ctxt
-*}
+\<close>
 
-method_setup synth_analz_mono_contra = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt))) *}
+method_setup synth_analz_mono_contra = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))\<close>
     "for proving theorems of the form X \<notin> synth (analz (knows Spy evs)) --> P"
 
 end
--- a/src/HOL/Auth/Guard/Analz.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Analz.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,14 +3,14 @@
     Copyright   2001  University of Cambridge
 *)
 
-section{*Decomposition of Analz into two parts*}
+section\<open>Decomposition of Analz into two parts\<close>
 
 theory Analz imports Extensions begin
 
-text{*decomposition of @{term analz} into two parts: 
-      @{term pparts} (for pairs) and analz of @{term kparts}*}
+text\<open>decomposition of @{term analz} into two parts: 
+      @{term pparts} (for pairs) and analz of @{term kparts}\<close>
 
-subsection{*messages that do not contribute to analz*}
+subsection\<open>messages that do not contribute to analz\<close>
 
 inductive_set
   pparts :: "msg set => msg set"
@@ -20,7 +20,7 @@
 | Fst [dest]: "[| {|X,Y|}:pparts H; is_MPair X |] ==> X:pparts H"
 | Snd [dest]: "[| {|X,Y|}:pparts H; is_MPair Y |] ==> Y:pparts H"
 
-subsection{*basic facts about @{term pparts}*}
+subsection\<open>basic facts about @{term pparts}\<close>
 
 lemma pparts_is_MPair [dest]: "X:pparts H ==> is_MPair X"
 by (erule pparts.induct, auto)
@@ -98,13 +98,13 @@
 lemma in_pparts: "Y:pparts H ==> EX X. X:H & Y:pparts {X}"
 by (erule pparts.induct, auto)
 
-subsection{*facts about @{term pparts} and @{term parts}*}
+subsection\<open>facts about @{term pparts} and @{term parts}\<close>
 
 lemma pparts_no_Nonce [dest]: "[| X:pparts {Y}; Nonce n ~:parts {Y} |]
 ==> Nonce n ~:parts {X}"
 by (erule pparts.induct, simp_all)
 
-subsection{*facts about @{term pparts} and @{term analz}*}
+subsection\<open>facts about @{term pparts} and @{term analz}\<close>
 
 lemma pparts_analz: "X:pparts H ==> X:analz H"
 by (erule pparts.induct, auto)
@@ -112,7 +112,7 @@
 lemma pparts_analz_sub: "[| X:pparts G; G<=H |] ==> X:analz H"
 by (auto dest: pparts_sub pparts_analz)
 
-subsection{*messages that contribute to analz*}
+subsection\<open>messages that contribute to analz\<close>
 
 inductive_set
   kparts :: "msg set => msg set"
@@ -122,7 +122,7 @@
 | Fst [intro]: "[| {|X,Y|}:pparts H; not_MPair X |] ==> X:kparts H"
 | Snd [intro]: "[| {|X,Y|}:pparts H; not_MPair Y |] ==> Y:kparts H"
 
-subsection{*basic facts about @{term kparts}*}
+subsection\<open>basic facts about @{term kparts}\<close>
 
 lemma kparts_not_MPair [dest]: "X:kparts H ==> not_MPair X"
 by (erule kparts.induct, auto)
@@ -195,7 +195,7 @@
 lemma kparts_has_no_pair [iff]: "has_no_pair (kparts H)"
 by auto
 
-subsection{*facts about @{term kparts} and @{term parts}*}
+subsection\<open>facts about @{term kparts} and @{term parts}\<close>
 
 lemma kparts_no_Nonce [dest]: "[| X:kparts {Y}; Nonce n ~:parts {Y} |]
 ==> Nonce n ~:parts {X}"
@@ -212,7 +212,7 @@
 Nonce n:parts {Y} |] ==> Nonce n:parts {Z}"
 by auto
 
-subsection{*facts about @{term kparts} and @{term analz}*}
+subsection\<open>facts about @{term kparts} and @{term analz}\<close>
 
 lemma kparts_analz: "X:kparts H ==> X:analz H"
 by (erule kparts.induct, auto dest: pparts_analz)
@@ -247,7 +247,7 @@
 by (metis Fake_parts_insert_in_Un Nonce_kparts_synth UnE analz_conj_parts synth_simps(5))
 
 
-subsection{*analz is pparts + analz of kparts*}
+subsection\<open>analz is pparts + analz of kparts\<close>
 
 lemma analz_pparts_kparts: "X:analz H ==> X:pparts H | X:analz (kparts H)"
 by (erule analz.induct, auto) 
--- a/src/HOL/Auth/Guard/Auth_Guard_Public.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Auth_Guard_Public.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
     Copyright   1996  University of Cambridge
 *)
 
-section {* Blanqui's "guard" concept: protocol-independent secrecy *}
+section \<open>Blanqui's "guard" concept: protocol-independent secrecy\<close>
 
 theory Auth_Guard_Public
 imports
--- a/src/HOL/Auth/Guard/Auth_Guard_Shared.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Auth_Guard_Shared.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
     Copyright   1996  University of Cambridge
 *)
 
-section {* Blanqui's "guard" concept: protocol-independent secrecy *}
+section \<open>Blanqui's "guard" concept: protocol-independent secrecy\<close>
 
 theory Auth_Guard_Shared
 imports
--- a/src/HOL/Auth/Guard/Extensions.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Extensions.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,13 +3,13 @@
     Copyright   2001  University of Cambridge
 *)
 
-section {*Extensions to Standard Theories*}
+section \<open>Extensions to Standard Theories\<close>
 
 theory Extensions
 imports "../Event"
 begin
 
-subsection{*Extensions to Theory @{text Set}*}
+subsection\<open>Extensions to Theory \<open>Set\<close>\<close>
 
 lemma eq: "[| !!x. x:A ==> x:B; !!x. x:B ==> x:A |] ==> A=B"
 by auto
@@ -21,9 +21,9 @@
 by auto
 
 
-subsection{*Extensions to Theory @{text List}*}
+subsection\<open>Extensions to Theory \<open>List\<close>\<close>
 
-subsubsection{*"remove l x" erase the first element of "l" equal to "x"*}
+subsubsection\<open>"remove l x" erase the first element of "l" equal to "x"\<close>
 
 primrec remove :: "'a list => 'a => 'a list" where
 "remove [] y = []" |
@@ -32,9 +32,9 @@
 lemma set_remove: "set (remove l x) <= set l"
 by (induct l, auto)
 
-subsection{*Extensions to Theory @{text Message}*}
+subsection\<open>Extensions to Theory \<open>Message\<close>\<close>
 
-subsubsection{*declarations for tactics*}
+subsubsection\<open>declarations for tactics\<close>
 
 declare analz_subset_parts [THEN subsetD, dest]
 declare parts_insert2 [simp]
@@ -43,12 +43,12 @@
 declare analz_insertI [intro]
 declare Un_Diff [simp]
 
-subsubsection{*extract the agent number of an Agent message*}
+subsubsection\<open>extract the agent number of an Agent message\<close>
 
 primrec agt_nb :: "msg => agent" where
 "agt_nb (Agent A) = A"
 
-subsubsection{*messages that are pairs*}
+subsubsection\<open>messages that are pairs\<close>
 
 definition is_MPair :: "msg => bool" where
 "is_MPair X == EX Y Z. X = {|Y,Z|}"
@@ -90,7 +90,7 @@
 
 declare has_no_pair_def [simp]
 
-subsubsection{*well-foundedness of messages*}
+subsubsection\<open>well-foundedness of messages\<close>
 
 lemma wf_Crypt1 [iff]: "Crypt K X ~= X"
 by (induct X, auto)
@@ -104,7 +104,7 @@
 lemma wf_Crypt_parts [iff]: "Crypt K X ~:parts {X}"
 by (auto dest: parts_size)
 
-subsubsection{*lemmas on keysFor*}
+subsubsection\<open>lemmas on keysFor\<close>
 
 definition usekeys :: "msg set => key set" where
 "usekeys G == {K. EX Y. Crypt K Y:G}"
@@ -120,7 +120,7 @@
 by (subgoal_tac "{K. EX X. Crypt K X = x | Crypt K X:F} = usekeys F",
 auto simp: usekeys_def)
 
-subsubsection{*lemmas on parts*}
+subsubsection\<open>lemmas on parts\<close>
 
 lemma parts_sub: "[| X:parts G; G<=H |] ==> X:parts H"
 by (auto dest: parts_mono)
@@ -153,7 +153,7 @@
 ==> Nonce n:parts G"
 by (blast intro: parts.Body dest: parts_parts)
 
-subsubsection{*lemmas on synth*}
+subsubsection\<open>lemmas on synth\<close>
 
 lemma synth_sub: "[| X:synth G; G<=H |] ==> X:synth H"
 by (auto dest: synth_mono)
@@ -162,7 +162,7 @@
 Crypt K Y:parts {X} --> Crypt K Y:parts G"
 by (erule synth.induct, auto dest: parts_sub)
 
-subsubsection{*lemmas on analz*}
+subsubsection\<open>lemmas on analz\<close>
 
 lemma analz_UnI1 [intro]: "X:analz G ==> X:analz (G Un H)"
   by (subgoal_tac "G <= G Un H") (blast dest: analz_mono)+
@@ -194,7 +194,7 @@
 lemma notin_analz_insert: "X ~:analz (insert Y G) ==> X ~:analz G"
 by auto
 
-subsubsection{*lemmas on parts, synth and analz*}
+subsubsection\<open>lemmas on parts, synth and analz\<close>
 
 lemma parts_invKey [rule_format,dest]:"X:parts {Y} ==>
 X:analz (insert (Crypt K Y) H) --> X ~:analz H --> Key (invKey K):analz H"
@@ -213,7 +213,7 @@
 apply auto
 done
 
-subsubsection{*greatest nonce used in a message*}
+subsubsection\<open>greatest nonce used in a message\<close>
 
 fun greatest_msg :: "msg => nat"
 where
@@ -225,7 +225,7 @@
 lemma greatest_msg_is_greatest: "Nonce n:parts {X} ==> n <= greatest_msg X"
 by (induct X, auto)
 
-subsubsection{*sets of keys*}
+subsubsection\<open>sets of keys\<close>
 
 definition keyset :: "msg set => bool" where
 "keyset G == ALL X. X:G --> (EX K. X = Key K)"
@@ -245,7 +245,7 @@
 lemma parts_keyset [simp]: "keyset G ==> parts G = G"
 by (auto, erule parts.induct, auto)
 
-subsubsection{*keys a priori necessary for decrypting the messages of G*}
+subsubsection\<open>keys a priori necessary for decrypting the messages of G\<close>
 
 definition keysfor :: "msg set => msg set" where
 "keysfor G == Key ` keysFor (parts G)"
@@ -265,7 +265,7 @@
 lemma finite_keysfor [intro]: "finite G ==> finite (keysfor G)"
 by (auto simp: keysfor_def intro: finite_UN_I)
 
-subsubsection{*only the keys necessary for G are useful in analz*}
+subsubsection\<open>only the keys necessary for G are useful in analz\<close>
 
 lemma analz_keyset: "keyset H ==>
 analz (G Un H) = H - keysfor G Un (analz (G Un (H Int keysfor G)))"
@@ -280,10 +280,10 @@
 lemmas analz_keyset_substD = analz_keyset [THEN sym, THEN ssubst]
 
 
-subsection{*Extensions to Theory @{text Event}*}
+subsection\<open>Extensions to Theory \<open>Event\<close>\<close>
 
 
-subsubsection{*general protocol properties*}
+subsubsection\<open>general protocol properties\<close>
 
 definition is_Says :: "event => bool" where
 "is_Says ev == (EX A B X. ev = Says A B X)"
@@ -330,7 +330,7 @@
 ==> Gets_correct p"
 by (auto simp: has_only_Says_def Gets_correct_def)
 
-subsubsection{*lemma on knows*}
+subsubsection\<open>lemma on knows\<close>
 
 lemma Says_imp_spies2: "Says A B {|X,Y|}:set evs ==> Y:parts (spies evs)"
 by (drule Says_imp_spies, drule parts.Inj, drule parts.Snd, simp)
@@ -339,7 +339,7 @@
 ==> Y ~:parts {X}"
 by (auto dest: Says_imp_spies parts_parts)
 
-subsubsection{*knows without initState*}
+subsubsection\<open>knows without initState\<close>
 
 primrec knows' :: "agent => event list => msg set" where
   knows'_Nil: "knows' A [] = {}" |
@@ -361,7 +361,7 @@
   spies' :: "event list => msg set" where
   "spies' == knows' Spy"
 
-subsubsection{*decomposition of knows into knows' and initState*}
+subsubsection\<open>decomposition of knows into knows' and initState\<close>
 
 lemma knows_decomp: "knows A evs = knows' A evs Un (initState A)"
 by (induct evs, auto split: event.split simp: knows.simps)
@@ -394,12 +394,12 @@
 ==> knows' A evs <= spies' evs"
 by (induct evs, auto split: event.splits)
 
-subsubsection{*knows' is finite*}
+subsubsection\<open>knows' is finite\<close>
 
 lemma finite_knows' [iff]: "finite (knows' A evs)"
 by (induct evs, auto split: event.split simp: knows.simps)
 
-subsubsection{*monotonicity of knows*}
+subsubsection\<open>monotonicity of knows\<close>
 
 lemma knows_sub_Cons: "knows A evs <= knows A (ev#evs)"
 by(cases A, induct evs, auto simp: knows.simps split:event.split)
@@ -413,8 +413,8 @@
 apply (rename_tac a b c)
 by (case_tac a, auto simp: knows.simps)
 
-subsubsection{*maximum knowledge an agent can have
-includes messages sent to the agent*}
+subsubsection\<open>maximum knowledge an agent can have
+includes messages sent to the agent\<close>
 
 primrec knows_max' :: "agent => event list => msg set" where
 knows_max'_def_Nil: "knows_max' A [] = {}" |
@@ -442,7 +442,7 @@
   spies_max :: "event list => msg set" where
   "spies_max evs == knows_max Spy evs"
 
-subsubsection{*basic facts about @{term knows_max}*}
+subsubsection\<open>basic facts about @{term knows_max}\<close>
 
 lemma spies_max_spies [iff]: "spies_max evs = spies evs"
 by (induct evs, auto simp: knows_max_def split: event.splits)
@@ -484,7 +484,7 @@
 lemma Says_from_knows_max': "Says A B X:set evs ==> X:knows_max' A evs"
 by (simp add: in_set_conv_decomp, clarify, simp add: knows_max'_app)
 
-subsubsection{*used without initState*}
+subsubsection\<open>used without initState\<close>
 
 primrec used' :: "event list => msg set" where
 "used' [] = {}" |
@@ -511,7 +511,7 @@
 apply (blast dest: parts_trans)+ 
 done
 
-subsubsection{*monotonicity of used*}
+subsubsection\<open>monotonicity of used\<close>
 
 lemma used_sub_Cons: "used evs <= used (ev#evs)"
 by (induct evs, (induct ev, auto)+)
@@ -550,7 +550,7 @@
 apply (drule_tac evs'=evs in used_appIL)
 by simp
 
-subsubsection{*lemmas on used and knows*}
+subsubsection\<open>lemmas on used and knows\<close>
 
 lemma initState_used: "X:parts (initState A) ==> X:used evs"
 by (induct evs, auto simp: used.simps split: event.split)
@@ -605,7 +605,7 @@
 Gets_correct p; one_step p |] ==> X ~:parts (knows_max A evs)"
 by (case_tac "A=Spy", auto dest: not_used_not_spied known_max_used)
 
-subsubsection{*a nonce or key in a message cannot equal a fresh nonce or key*}
+subsubsection\<open>a nonce or key in a message cannot equal a fresh nonce or key\<close>
 
 lemma Nonce_neq [dest]: "[| Nonce n' ~:used evs;
 Says A B X:set evs; Nonce n:parts {X} |] ==> n ~= n'"
@@ -615,7 +615,7 @@
 Says A B X:set evs; Key n:parts {X} |] ==> n ~= n'"
 by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub)
 
-subsubsection{*message of an event*}
+subsubsection\<open>message of an event\<close>
 
 primrec msg :: "event => msg"
 where
--- a/src/HOL/Auth/Guard/Guard.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Guard.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,7 +3,7 @@
     Copyright   2002  University of Cambridge
 *)
 
-section{*Protocol-Independent Confidentiality Theorem on Nonces*}
+section\<open>Protocol-Independent Confidentiality Theorem on Nonces\<close>
 
 theory Guard imports Analz Extensions begin
 
@@ -21,7 +21,7 @@
 | Crypt [intro]: "X:guard n Ks ==> Crypt K X:guard n Ks"
 | Pair [intro]: "[| X:guard n Ks; Y:guard n Ks |] ==> {|X,Y|}:guard n Ks"
 
-subsection{*basic facts about @{term guard}*}
+subsection\<open>basic facts about @{term guard}\<close>
 
 lemma Key_is_guard [iff]: "Key K:guard n Ks"
 by auto
@@ -68,12 +68,12 @@
 lemma guard_extand: "[| X:guard n Ks; Ks <= Ks' |] ==> X:guard n Ks'"
 by (erule guard.induct, auto)
 
-subsection{*guarded sets*}
+subsection\<open>guarded sets\<close>
 
 definition Guard :: "nat => key set => msg set => bool" where
 "Guard n Ks H == ALL X. X:H --> X:guard n Ks"
 
-subsection{*basic facts about @{term Guard}*}
+subsection\<open>basic facts about @{term Guard}\<close>
 
 lemma Guard_empty [iff]: "Guard n Ks {}"
 by (simp add: Guard_def)
@@ -154,7 +154,7 @@
 Nonce n:kparts {Y} |] ==> invKey K:Ks"
 by (auto dest: guard_invKey)
 
-subsection{*set obtained by decrypting a message*}
+subsection\<open>set obtained by decrypting a message\<close>
 
 abbreviation (input)
   decrypt :: "msg set => key => msg => msg set" where
@@ -170,7 +170,7 @@
 lemma parts_decrypt: "[| Crypt K Y:H; X:parts (decrypt H K Y) |] ==> X:parts H"
 by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body)
 
-subsection{*number of Crypt's in a message*}
+subsection\<open>number of Crypt's in a message\<close>
 
 fun crypt_nb :: "msg => nat"
 where
@@ -178,19 +178,19 @@
 | "crypt_nb {|X,Y|} = crypt_nb X + crypt_nb Y"
 | "crypt_nb X = 0" (* otherwise *)
 
-subsection{*basic facts about @{term crypt_nb}*}
+subsection\<open>basic facts about @{term crypt_nb}\<close>
 
 lemma non_empty_crypt_msg: "Crypt K Y:parts {X} ==> crypt_nb X \<noteq> 0"
 by (induct X, simp_all, safe, simp_all)
 
-subsection{*number of Crypt's in a message list*}
+subsection\<open>number of Crypt's in a message list\<close>
 
 primrec cnb :: "msg list => nat"
 where
   "cnb [] = 0"
 | "cnb (X#l) = crypt_nb X + cnb l"
 
-subsection{*basic facts about @{term cnb}*}
+subsection\<open>basic facts about @{term cnb}\<close>
 
 lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'"
 by (induct l, auto)
@@ -213,7 +213,7 @@
 lemma non_empty_crypt: "Crypt K Y:parts (set l) ==> cnb l \<noteq> 0"
 by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD)
 
-subsection{*list of kparts*}
+subsection\<open>list of kparts\<close>
 
 lemma kparts_msg_set: "EX l. kparts {X} = set l & cnb l = crypt_nb X"
 apply (induct X, simp_all)
@@ -234,20 +234,20 @@
 apply (rule kparts_insert_substI, simp)
 by (rule kparts_msg_set)
 
-subsection{*list corresponding to "decrypt"*}
+subsection\<open>list corresponding to "decrypt"\<close>
 
 definition decrypt' :: "msg list => key => msg => msg list" where
 "decrypt' l K Y == Y # remove l (Crypt K Y)"
 
 declare decrypt'_def [simp]
 
-subsection{*basic facts about @{term decrypt'}*}
+subsection\<open>basic facts about @{term decrypt'}\<close>
 
 lemma decrypt_minus: "decrypt (set l) K Y <= set (decrypt' l K Y)"
 by (induct l, auto)
 
-subsection{*if the analyse of a finite guarded set gives n then it must also gives
-one of the keys of Ks*}
+subsection\<open>if the analyse of a finite guarded set gives n then it must also gives
+one of the keys of Ks\<close>
 
 lemma Guard_invKey_by_list [rule_format]: "ALL l. cnb l = p
 --> Guard n Ks (set l) --> Nonce n:analz (set l)
@@ -295,8 +295,8 @@
 ==> EX K. K:Ks & Key K:analz G"
 by (auto dest: analz_needs_only_finite Guard_invKey_finite)
 
-subsection{*if the analyse of a finite guarded set and a (possibly infinite) set of keys
-gives n then it must also gives Ks*}
+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:analz (G Un H); Guard n Ks G; finite G;
 keyset H |] ==> EX K. K:Ks & Key K:analz (G Un H)"
--- a/src/HOL/Auth/Guard/GuardK.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/GuardK.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -8,7 +8,7 @@
 - the hypothesis Key n ~:G (keyset G) is added
 *)
 
-section{*protocol-independent confidentiality theorem on keys*}
+section\<open>protocol-independent confidentiality theorem on keys\<close>
 
 theory GuardK
 imports Analz Extensions
@@ -28,7 +28,7 @@
 | Crypt [intro]: "X:guardK n Ks ==> Crypt K X:guardK n Ks"
 | Pair [intro]: "[| X:guardK n Ks; Y:guardK n Ks |] ==> {|X,Y|}:guardK n Ks"
 
-subsection{*basic facts about @{term guardK}*}
+subsection\<open>basic facts about @{term guardK}\<close>
 
 lemma Nonce_is_guardK [iff]: "Nonce p:guardK n Ks"
 by auto
@@ -77,12 +77,12 @@
 [| K:Ks'; K ~:Ks |] ==> Key K ~:parts {X} |] ==> X:guardK n Ks'"
 by (erule guardK.induct, auto)
 
-subsection{*guarded sets*}
+subsection\<open>guarded sets\<close>
 
 definition GuardK :: "nat => key set => msg set => bool" where
 "GuardK n Ks H == ALL X. X:H --> X:guardK n Ks"
 
-subsection{*basic facts about @{term GuardK}*}
+subsection\<open>basic facts about @{term GuardK}\<close>
 
 lemma GuardK_empty [iff]: "GuardK n Ks {}"
 by (simp add: GuardK_def)
@@ -152,7 +152,7 @@
 [| K:Ks'; K ~:Ks |] ==> Key K ~:parts G |] ==> GuardK n Ks' G"
 by (auto simp: GuardK_def dest: guardK_extand parts_sub)
 
-subsection{*set obtained by decrypting a message*}
+subsection\<open>set obtained by decrypting a message\<close>
 
 abbreviation (input)
   decrypt :: "msg set => key => msg => msg set" where
@@ -168,25 +168,25 @@
 lemma parts_decrypt: "[| Crypt K Y:H; X:parts (decrypt H K Y) |] ==> X:parts H"
 by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body)
 
-subsection{*number of Crypt's in a message*}
+subsection\<open>number of Crypt's in a message\<close>
 
 fun crypt_nb :: "msg => nat" where
 "crypt_nb (Crypt K X) = Suc (crypt_nb X)" |
 "crypt_nb {|X,Y|} = crypt_nb X + crypt_nb Y" |
 "crypt_nb X = 0" (* otherwise *)
 
-subsection{*basic facts about @{term crypt_nb}*}
+subsection\<open>basic facts about @{term crypt_nb}\<close>
 
 lemma non_empty_crypt_msg: "Crypt K Y:parts {X} ==> crypt_nb X \<noteq> 0"
 by (induct X, simp_all, safe, simp_all)
 
-subsection{*number of Crypt's in a message list*}
+subsection\<open>number of Crypt's in a message list\<close>
 
 primrec cnb :: "msg list => nat" where
 "cnb [] = 0" |
 "cnb (X#l) = crypt_nb X + cnb l"
 
-subsection{*basic facts about @{term cnb}*}
+subsection\<open>basic facts about @{term cnb}\<close>
 
 lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'"
 by (induct l, auto)
@@ -207,7 +207,7 @@
 lemma non_empty_crypt: "Crypt K Y:parts (set l) ==> cnb l \<noteq> 0"
 by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD)
 
-subsection{*list of kparts*}
+subsection\<open>list of kparts\<close>
 
 lemma kparts_msg_set: "EX l. kparts {X} = set l & cnb l = crypt_nb X"
 apply (induct X, simp_all)
@@ -228,20 +228,20 @@
 apply (rule kparts_insert_substI, simp)
 by (rule kparts_msg_set)
 
-subsection{*list corresponding to "decrypt"*}
+subsection\<open>list corresponding to "decrypt"\<close>
 
 definition decrypt' :: "msg list => key => msg => msg list" where
 "decrypt' l K Y == Y # remove l (Crypt K Y)"
 
 declare decrypt'_def [simp]
 
-subsection{*basic facts about @{term decrypt'}*}
+subsection\<open>basic facts about @{term decrypt'}\<close>
 
 lemma decrypt_minus: "decrypt (set l) K Y <= set (decrypt' l K Y)"
 by (induct l, auto)
 
-text{*if the analysis of a finite guarded set gives n then it must also give
-one of the keys of Ks*}
+text\<open>if the analysis of a finite guarded set gives n then it must also give
+one of the keys of Ks\<close>
 
 lemma GuardK_invKey_by_list [rule_format]: "ALL l. cnb l = p
 --> GuardK n Ks (set l) --> Key n:analz (set l)
@@ -289,8 +289,8 @@
 ==> EX K. K:Ks & Key K:analz G"
 by (auto dest: analz_needs_only_finite GuardK_invKey_finite)
 
-text{*if the analyse of a finite guarded set and a (possibly infinite) set of
-keys gives n then it must also gives Ks*}
+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:analz (G Un H); GuardK n Ks G; finite G;
 keyset H; Key n ~:H |] ==> EX K. K:Ks & Key K:analz (G Un H)"
--- a/src/HOL/Auth/Guard/Guard_NS_Public.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Guard_NS_Public.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -5,11 +5,11 @@
 Incorporating Lowe's fix (inclusion of B's identity in round 2).
 *)
 
-section{*Needham-Schroeder-Lowe Public-Key Protocol*}
+section\<open>Needham-Schroeder-Lowe Public-Key Protocol\<close>
 
 theory Guard_NS_Public imports Guard_Public begin
 
-subsection{*messages used in the protocol*}
+subsection\<open>messages used in the protocol\<close>
 
 abbreviation (input)
   ns1 :: "agent => agent => nat => event" where
@@ -32,7 +32,7 @@
   "ns3 A B NB == Says A B (Crypt (pubK B) (Nonce NB))"
 
 
-subsection{*definition of the protocol*}
+subsection\<open>definition of the protocol\<close>
 
 inductive_set nsp :: "event list set"
 where
@@ -49,13 +49,13 @@
 | NS3: "!!A B B' NA NB evs3. [| evs3:nsp; ns1 A B NA:set evs3; ns2' B' B A NA NB:set evs3 |] ==>
   ns3 A B NB # evs3:nsp"
 
-subsection{*declarations for tactics*}
+subsection\<open>declarations for tactics\<close>
 
 declare knows_Spy_partsEs [elim]
 declare Fake_parts_insert [THEN subsetD, dest]
 declare initState.simps [simp del]
 
-subsection{*general properties of nsp*}
+subsection\<open>general properties of nsp\<close>
 
 lemma nsp_has_no_Gets: "evs:nsp ==> ALL A X. Gets A X ~:set evs"
 by (erule nsp.induct, auto)
@@ -77,7 +77,7 @@
 apply (simp only: regular_def, clarify)
 by (erule nsp.induct, auto simp: initState.simps knows.simps)
 
-subsection{*nonce are used only once*}
+subsection\<open>nonce are used only once\<close>
 
 lemma NA_is_uniq [rule_format]: "evs:nsp ==>
 Crypt (pubK B) {|Nonce NA, Agent A|}:parts (spies evs)
@@ -106,7 +106,7 @@
 apply (erule nsp.induct, simp_all)
 by (blast intro: analz_insertI)+
 
-subsection{*guardedness of NA*}
+subsection\<open>guardedness of NA\<close>
 
 lemma ns1_imp_Guard [rule_format]: "[| evs:nsp; A ~:bad; B ~:bad |] ==>
 ns1 A B NA:set evs --> Guard NA {priK A,priK B} (spies evs)"
@@ -133,7 +133,7 @@
 apply (drule Says_imp_knows_Spy)+
 by (drule no_Nonce_NS1_NS2, auto)
 
-subsection{*guardedness of NB*}
+subsection\<open>guardedness of NB\<close>
 
 lemma ns2_imp_Guard [rule_format]: "[| evs:nsp; A ~:bad; B ~:bad |] ==>
 ns2 B A NA NB:set evs --> Guard NB {priK A,priK B} (spies evs)" 
@@ -163,7 +163,7 @@
 apply (auto simp add: guard.No_Nonce)
 done
 
-subsection{*Agents' Authentication*}
+subsection\<open>Agents' Authentication\<close>
 
 lemma B_trusts_NS1: "[| evs:nsp; A ~:bad; B ~:bad |] ==>
 Crypt (pubK B) {|Nonce NA, Agent A|}:parts (spies evs)
--- a/src/HOL/Auth/Guard/Guard_OtwayRees.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Guard_OtwayRees.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   2002  University of Cambridge
 *)
 
-section{*Otway-Rees Protocol*}
+section\<open>Otway-Rees Protocol\<close>
 
 theory Guard_OtwayRees imports Guard_Shared begin
 
-subsection{*messages used in the protocol*}
+subsection\<open>messages used in the protocol\<close>
 
 abbreviation
   nil :: "msg" where
@@ -54,7 +54,7 @@
   or4' :: "agent => agent => nat => key => event" where
   "or4' B' A NA K == Says B' A {|Nonce NA, Ciph A {|Nonce NA, Key K|}, nil|}"
 
-subsection{*definition of the protocol*}
+subsection\<open>definition of the protocol\<close>
 
 inductive_set or :: "event list set"
 where
@@ -74,13 +74,13 @@
 | OR4: "[| evs4:or; or2 A B NA NB X:set evs4; or3' S Y A B NA NB K:set evs4 |]
   ==> or4 A B NA X # evs4:or"
 
-subsection{*declarations for tactics*}
+subsection\<open>declarations for tactics\<close>
 
 declare knows_Spy_partsEs [elim]
 declare Fake_parts_insert [THEN subsetD, dest]
 declare initState.simps [simp del]
 
-subsection{*general properties of or*}
+subsection\<open>general properties of or\<close>
 
 lemma or_has_no_Gets: "evs:or ==> ALL A X. Gets A X ~:set evs"
 by (erule or.induct, auto)
@@ -98,7 +98,7 @@
 lemma or_has_only_Says [iff]: "has_only_Says or"
 by (auto simp: has_only_Says_def dest: or_has_only_Says')
 
-subsection{*or is regular*}
+subsection\<open>or is regular\<close>
 
 lemma or1'_parts_spies [dest]: "or1' A' A B NA X:set evs
 ==> X:parts (spies evs)"
@@ -117,7 +117,7 @@
 apply (erule or.induct, simp_all add: initState.simps knows.simps)
 by (auto dest: parts_sub)
 
-subsection{*guardedness of KAB*}
+subsection\<open>guardedness of KAB\<close>
 
 lemma Guard_KAB [rule_format]: "[| evs:or; A ~:bad; B ~:bad |] ==>
 or3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)" 
@@ -138,7 +138,7 @@
 (* OR4 *)
 by (blast dest: Says_imp_spies in_GuardK_kparts)
 
-subsection{*guardedness of NB*}
+subsection\<open>guardedness of NB\<close>
 
 lemma Guard_NB [rule_format]: "[| evs:or; B ~:bad |] ==>
 or2 A B NA NB X:set evs --> Guard NB {shrK B} (spies evs)" 
--- a/src/HOL/Auth/Guard/Guard_Public.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Guard_Public.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -7,11 +7,11 @@
 
 theory Guard_Public imports Guard "../Public" Extensions begin
 
-subsection{*Extensions to Theory @{text Public}*}
+subsection\<open>Extensions to Theory \<open>Public\<close>\<close>
 
 declare initState.simps [simp del]
 
-subsubsection{*signature*}
+subsubsection\<open>signature\<close>
 
 definition sign :: "agent => msg => msg" where
 "sign A X == {|Agent A, X, Crypt (priK A) (Hash X)|}"
@@ -19,7 +19,7 @@
 lemma sign_inj [iff]: "(sign A X = sign A' X') = (A=A' & X=X')"
 by (auto simp: sign_def)
 
-subsubsection{*agent associated to a key*}
+subsubsection\<open>agent associated to a key\<close>
 
 definition agt :: "key => agent" where
 "agt K == @A. K = priK A | K = pubK A"
@@ -30,7 +30,7 @@
 lemma agt_pubK [simp]: "agt (pubK A) = A"
 by (simp add: agt_def)
 
-subsubsection{*basic facts about @{term initState}*}
+subsubsection\<open>basic facts about @{term initState}\<close>
 
 lemma no_Crypt_in_parts_init [simp]: "Crypt K X ~:parts (initState A)"
 by (cases A, auto simp: initState.simps)
@@ -49,7 +49,7 @@
 lemma keyset_init [iff]: "keyset (initState A)"
 by (cases A, auto simp: keyset_def initState.simps)
 
-subsubsection{*sets of private keys*}
+subsubsection\<open>sets of private keys\<close>
 
 definition priK_set :: "key set => bool" where
 "priK_set Ks == ALL K. K:Ks --> (EX A. K = priK A)"
@@ -63,7 +63,7 @@
 lemma priK_set2 [iff]: "priK_set {priK A, priK B}"
 by (simp add: priK_set_def)
 
-subsubsection{*sets of good keys*}
+subsubsection\<open>sets of good keys\<close>
 
 definition good :: "key set => bool" where
 "good Ks == ALL K. K:Ks --> agt K ~:bad"
@@ -77,7 +77,7 @@
 lemma good2 [simp]: "[| A ~:bad; B ~:bad |] ==> good {priK A, priK B}"
 by (simp add: good_def)
 
-subsubsection{*greatest nonce used in a trace, 0 if there is no nonce*}
+subsubsection\<open>greatest nonce used in a trace, 0 if there is no nonce\<close>
 
 primrec greatest :: "event list => nat"
 where
@@ -90,7 +90,7 @@
 apply (drule greatest_msg_is_greatest, arith)
 by simp
 
-subsubsection{*function giving a new nonce*}
+subsubsection\<open>function giving a new nonce\<close>
 
 definition new :: "event list => nat" where
 "new evs == Suc (greatest evs)"
@@ -98,9 +98,9 @@
 lemma new_isnt_used [iff]: "Nonce (new evs) ~:used evs"
 by (clarify, drule greatest_is_greatest, auto simp: new_def)
 
-subsection{*Proofs About Guarded Messages*}
+subsection\<open>Proofs About Guarded Messages\<close>
 
-subsubsection{*small hack necessary because priK is defined as the inverse of pubK*}
+subsubsection\<open>small hack necessary because priK is defined as the inverse of pubK\<close>
 
 lemma pubK_is_invKey_priK: "pubK A = invKey (priK A)"
 by simp
@@ -113,7 +113,7 @@
 apply (rule pubK_is_invKey_priK_substI, rule invKey_invKey_substI)
 by (rule Guard_Nonce, simp+)
 
-subsubsection{*guardedness results*}
+subsubsection\<open>guardedness results\<close>
 
 lemma sign_guard [intro]: "X:guard n Ks ==> sign A X:guard n Ks"
 by (auto simp: sign_def)
@@ -142,7 +142,7 @@
 apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono)
 by (auto simp: knows_max_def)
 
-subsubsection{*regular protocols*}
+subsubsection\<open>regular protocols\<close>
 
 definition regular :: "event list set => bool" where
 "regular p == ALL evs A. evs:p --> (Key (priK A):parts (spies evs)) = (A:bad)"
--- a/src/HOL/Auth/Guard/Guard_Shared.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Guard_Shared.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,21 +3,21 @@
     Copyright   2002  University of Cambridge
 *)
 
-section{*lemmas on guarded messages for protocols with symmetric keys*}
+section\<open>lemmas on guarded messages for protocols with symmetric keys\<close>
 
 theory Guard_Shared imports Guard GuardK "../Shared" begin
 
-subsection{*Extensions to Theory @{text Shared}*}
+subsection\<open>Extensions to Theory \<open>Shared\<close>\<close>
 
 declare initState.simps [simp del]
 
-subsubsection{*a little abbreviation*}
+subsubsection\<open>a little abbreviation\<close>
 
 abbreviation
   Ciph :: "agent => msg => msg" where
   "Ciph A X == Crypt (shrK A) X"
 
-subsubsection{*agent associated to a key*}
+subsubsection\<open>agent associated to a key\<close>
 
 definition agt :: "key => agent" where
 "agt K == @A. K = shrK A"
@@ -25,7 +25,7 @@
 lemma agt_shrK [simp]: "agt (shrK A) = A"
 by (simp add: agt_def)
 
-subsubsection{*basic facts about @{term initState}*}
+subsubsection\<open>basic facts about @{term initState}\<close>
 
 lemma no_Crypt_in_parts_init [simp]: "Crypt K X ~:parts (initState A)"
 by (cases A, auto simp: initState.simps)
@@ -44,7 +44,7 @@
 lemma keyset_init [iff]: "keyset (initState A)"
 by (cases A, auto simp: keyset_def initState.simps)
 
-subsubsection{*sets of symmetric keys*}
+subsubsection\<open>sets of symmetric keys\<close>
 
 definition shrK_set :: "key set => bool" where
 "shrK_set Ks == ALL K. K:Ks --> (EX A. K = shrK A)"
@@ -58,7 +58,7 @@
 lemma shrK_set2 [iff]: "shrK_set {shrK A, shrK B}"
 by (simp add: shrK_set_def)
 
-subsubsection{*sets of good keys*}
+subsubsection\<open>sets of good keys\<close>
 
 definition good :: "key set => bool" where
 "good Ks == ALL K. K:Ks --> agt K ~:bad"
@@ -73,9 +73,9 @@
 by (simp add: good_def)
 
 
-subsection{*Proofs About Guarded Messages*}
+subsection\<open>Proofs About Guarded Messages\<close>
 
-subsubsection{*small hack*}
+subsubsection\<open>small hack\<close>
 
 lemma shrK_is_invKey_shrK: "shrK A = invKey (shrK A)"
 by simp
@@ -88,7 +88,7 @@
 apply (rule shrK_is_invKey_shrK_substI, rule invKey_invKey_substI)
 by (rule Guard_Nonce, simp+)
 
-subsubsection{*guardedness results on nonces*}
+subsubsection\<open>guardedness results on nonces\<close>
 
 lemma guard_ciph [simp]: "shrK A:Ks ==> Ciph A X:guard n Ks"
 by (rule Guard_Nonce, simp)
@@ -120,7 +120,7 @@
 apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono)
 by (auto simp: knows_max_def)
 
-subsubsection{*guardedness results on keys*}
+subsubsection\<open>guardedness results on keys\<close>
 
 lemma GuardK_init [simp]: "n ~:range shrK ==> GuardK n Ks (initState B)"
 by (induct B, auto simp: GuardK_def initState.simps)
@@ -146,7 +146,7 @@
 apply (rule_tac H="knows_max (Friend C) evs" in GuardK_mono)
 by (auto simp: knows_max_def)
 
-subsubsection{*regular protocols*}
+subsubsection\<open>regular protocols\<close>
 
 definition regular :: "event list set => bool" where
 "regular p == ALL evs A. evs:p --> (Key (shrK A):parts (spies evs)) = (A:bad)"
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   2002  University of Cambridge
 *)
 
-section{*Yahalom Protocol*}
+section\<open>Yahalom Protocol\<close>
 
 theory Guard_Yahalom imports "../Shared" Guard_Shared begin
 
-subsection{*messages used in the protocol*}
+subsection\<open>messages used in the protocol\<close>
 
 abbreviation (input)
   ya1 :: "agent => agent => nat => event" where
@@ -45,7 +45,7 @@
   "ya4' A' B K NB Y == Says A' B {|Y, Crypt K (Nonce NB)|}"
 
 
-subsection{*definition of the protocol*}
+subsection\<open>definition of the protocol\<close>
 
 inductive_set ya :: "event list set"
 where
@@ -65,13 +65,13 @@
 | YA4: "[| evs4:ya; ya1 A B NA:set evs4; ya3' S Y A B NA NB K:set evs4 |]
   ==> ya4 A B K NB Y # evs4:ya"
 
-subsection{*declarations for tactics*}
+subsection\<open>declarations for tactics\<close>
 
 declare knows_Spy_partsEs [elim]
 declare Fake_parts_insert [THEN subsetD, dest]
 declare initState.simps [simp del]
 
-subsection{*general properties of ya*}
+subsection\<open>general properties of ya\<close>
 
 lemma ya_has_no_Gets: "evs:ya ==> ALL A X. Gets A X ~:set evs"
 by (erule ya.induct, auto)
@@ -94,7 +94,7 @@
 apply (erule ya.induct, simp_all add: initState.simps knows.simps)
 by (auto dest: parts_sub)
 
-subsection{*guardedness of KAB*}
+subsection\<open>guardedness of KAB\<close>
 
 lemma Guard_KAB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==>
 ya3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)" 
@@ -115,7 +115,7 @@
 apply (blast dest: Says_imp_spies in_GuardK_kparts)
 by blast
 
-subsection{*session keys are not symmetric keys*}
+subsection\<open>session keys are not symmetric keys\<close>
 
 lemma KAB_isnt_shrK [rule_format]: "evs:ya ==>
 ya3 A B NA NB K:set evs --> K ~:range shrK"
@@ -124,7 +124,7 @@
 lemma ya3_shrK: "evs:ya ==> ya3 A B NA NB (shrK C) ~:set evs"
 by (blast dest: KAB_isnt_shrK)
 
-subsection{*ya2' implies ya1'*}
+subsection\<open>ya2' implies ya1'\<close>
 
 lemma ya2'_parts_imp_ya1'_parts [rule_format]:
      "[| evs:ya; B ~:bad |] ==>
@@ -136,7 +136,7 @@
 ==> {|Agent A, Nonce NA|}:spies evs"
 by (blast dest: Says_imp_spies ya2'_parts_imp_ya1'_parts)
 
-subsection{*uniqueness of NB*}
+subsection\<open>uniqueness of NB\<close>
 
 lemma NB_is_uniq_in_ya2'_parts [rule_format]: "[| evs:ya; B ~:bad; B' ~:bad |] ==>
 Ciph B {|Agent A, Nonce NA, Nonce NB|}:parts (spies evs) -->
@@ -153,7 +153,7 @@
 ==> A=A' & B=B' & NA=NA'"
 by (drule NB_is_uniq_in_ya2'_parts, auto dest: Says_imp_spies)
 
-subsection{*ya3' implies ya2'*}
+subsection\<open>ya3' implies ya2'\<close>
 
 lemma ya3'_parts_imp_ya2'_parts [rule_format]: "[| evs:ya; A ~:bad |] ==>
 Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}:parts (spies evs)
@@ -177,7 +177,7 @@
 ==> (EX B'. ya2' B' A B NA NB:set evs)"
 by (drule ya3'_parts_imp_ya2', auto dest: Says_imp_spies)
 
-subsection{*ya3' implies ya3*}
+subsection\<open>ya3' implies ya3\<close>
 
 lemma ya3'_parts_imp_ya3 [rule_format]: "[| evs:ya; A ~:bad |] ==>
 Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}:parts(spies evs)
@@ -190,7 +190,7 @@
 ==> ya3 A B NA NB K:set evs"
 by (blast dest: Says_imp_spies ya3'_parts_imp_ya3)
 
-subsection{*guardedness of NB*}
+subsection\<open>guardedness of NB\<close>
 
 definition ya_keys :: "agent => agent => nat => nat => event list => key set" where
 "ya_keys A B NA NB evs == {shrK A,shrK B} Un {K. ya3 A B NA NB K:set evs}"
--- a/src/HOL/Auth/Guard/List_Msg.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/List_Msg.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,35 +3,35 @@
     Copyright   2001  University of Cambridge
 *)
 
-section{*Lists of Messages and Lists of Agents*}
+section\<open>Lists of Messages and Lists of Agents\<close>
 
 theory List_Msg imports Extensions begin
 
-subsection{*Implementation of Lists by Messages*}
+subsection\<open>Implementation of Lists by Messages\<close>
 
-subsubsection{*nil is represented by any message which is not a pair*}
+subsubsection\<open>nil is represented by any message which is not a pair\<close>
 
 abbreviation (input)
   cons :: "msg => msg => msg" where
   "cons x l == {|x,l|}"
 
-subsubsection{*induction principle*}
+subsubsection\<open>induction principle\<close>
 
 lemma lmsg_induct: "[| !!x. not_MPair x ==> P x; !!x l. P l ==> P (cons x l) |]
 ==> P l"
 by (induct l) auto
 
-subsubsection{*head*}
+subsubsection\<open>head\<close>
 
 primrec head :: "msg => msg" where
 "head (cons x l) = x"
 
-subsubsection{*tail*}
+subsubsection\<open>tail\<close>
 
 primrec tail :: "msg => msg" where
 "tail (cons x l) = l"
 
-subsubsection{*length*}
+subsubsection\<open>length\<close>
 
 fun len :: "msg => nat" where
 "len (cons x l) = Suc (len l)" |
@@ -40,13 +40,13 @@
 lemma len_not_empty: "n < len l ==> EX x l'. l = cons x l'"
 by (cases l) auto
 
-subsubsection{*membership*}
+subsubsection\<open>membership\<close>
 
 fun isin :: "msg * msg => bool" where
 "isin (x, cons y l) = (x=y | isin (x,l))" |
 "isin (x, other) = False"
 
-subsubsection{*delete an element*}
+subsubsection\<open>delete an element\<close>
 
 fun del :: "msg * msg => msg" where
 "del (x, cons y l) = (if x=y then l else cons y (del (x,l)))" |
@@ -58,7 +58,7 @@
 lemma isin_del [rule_format]: "isin (y, del (x,l)) --> isin (y,l)"
 by (induct l) auto
 
-subsubsection{*concatenation*}
+subsubsection\<open>concatenation\<close>
 
 fun app :: "msg * msg => msg" where
 "app (cons x l, l') = cons x (app (l,l'))" |
@@ -67,14 +67,14 @@
 lemma isin_app [iff]: "isin (x, app(l,l')) = (isin (x,l) | isin (x,l'))"
 by (induct l) auto
 
-subsubsection{*replacement*}
+subsubsection\<open>replacement\<close>
 
 fun repl :: "msg * nat * msg => msg" where
 "repl (cons x l, Suc i, x') = cons x (repl (l,i,x'))" |
 "repl (cons x l, 0, x') = cons x' l" |
 "repl (other, i, M') = other"
 
-subsubsection{*ith element*}
+subsubsection\<open>ith element\<close>
 
 fun ith :: "msg * nat => msg" where
 "ith (cons x l, Suc i) = ith (l,i)" |
@@ -84,7 +84,7 @@
 lemma ith_head: "0 < len l ==> ith (l,0) = head l"
 by (cases l) auto
 
-subsubsection{*insertion*}
+subsubsection\<open>insertion\<close>
 
 fun ins :: "msg * nat * msg => msg" where
 "ins (cons x l, Suc i, y) = cons x (ins (l,i,y))" |
@@ -93,7 +93,7 @@
 lemma ins_head [simp]: "ins (l,0,y) = cons y l"
 by (cases l) auto
 
-subsubsection{*truncation*}
+subsubsection\<open>truncation\<close>
 
 fun trunc :: "msg * nat => msg" where
 "trunc (l,0) = l" |
@@ -103,9 +103,9 @@
 by (cases l) auto
 
 
-subsection{*Agent Lists*}
+subsection\<open>Agent Lists\<close>
 
-subsubsection{*set of well-formed agent-list messages*}
+subsubsection\<open>set of well-formed agent-list messages\<close>
 
 abbreviation
   nil :: msg where
@@ -116,7 +116,7 @@
   Nil[intro]: "nil:agl"
 | Cons[intro]: "[| A:agent; I:agl |] ==> cons (Agent A) I :agl"
 
-subsubsection{*basic facts about agent lists*}
+subsubsection\<open>basic facts about agent lists\<close>
 
 lemma del_in_agl [intro]: "I:agl ==> del (a,I):agl"
 by (erule agl.induct, auto)
--- a/src/HOL/Auth/Guard/P1.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/P1.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -7,11 +7,11 @@
 Mobiles Agents 1998, LNCS 1477.
 *)
 
-section{*Protocol P1*}
+section\<open>Protocol P1\<close>
 
 theory P1 imports "../Public" Guard_Public List_Msg begin
 
-subsection{*Protocol Definition*}
+subsection\<open>Protocol Definition\<close>
 
 (******************************************************************************
 
@@ -30,8 +30,8 @@
 (Crypt in injective)
 ******************************************************************************)
 
-subsubsection{*offer chaining:
-B chains his offer for A with the head offer of L for sending it to C*}
+subsubsection\<open>offer chaining:
+B chains his offer for A with the head offer of L for sending it to C\<close>
 
 definition chain :: "agent => nat => agent => msg => agent => msg" where
 "chain B ofr A L C ==
@@ -48,7 +48,7 @@
 lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}"
 by (auto simp: chain_def sign_def)
 
-subsubsection{*agent whose key is used to sign an offer*}
+subsubsection\<open>agent whose key is used to sign an offer\<close>
 
 fun shop :: "msg => msg" where
 "shop {|B,X,Crypt K H|} = Agent (agt K)"
@@ -56,7 +56,7 @@
 lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B"
 by (simp add: chain_def sign_def)
 
-subsubsection{*nonce used in an offer*}
+subsubsection\<open>nonce used in an offer\<close>
 
 fun nonce :: "msg => msg" where
 "nonce {|B,{|Crypt K ofr,m2|},CryptH|} = ofr"
@@ -64,7 +64,7 @@
 lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr"
 by (simp add: chain_def sign_def)
 
-subsubsection{*next shop*}
+subsubsection\<open>next shop\<close>
 
 fun next_shop :: "msg => agent" where
 "next_shop {|B,{|m1,Hash{|headL,Agent C|}|},CryptH|} = C"
@@ -72,7 +72,7 @@
 lemma next_shop_chain [iff]: "next_shop (chain B ofr A L C) = C"
 by (simp add: chain_def sign_def)
 
-subsubsection{*anchor of the offer list*}
+subsubsection\<open>anchor of the offer list\<close>
 
 definition anchor :: "agent => nat => agent => msg" where
 "anchor A n B == chain A n A (cons nil nil) B"
@@ -93,7 +93,7 @@
 lemma next_shop_anchor [iff]: "next_shop (anchor A n B) = B"
 by (simp add: anchor_def)
 
-subsubsection{*request event*}
+subsubsection\<open>request event\<close>
 
 definition reqm :: "agent => nat => nat => msg => agent => msg" where
 "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I),
@@ -113,7 +113,7 @@
 = (A=A' & r=r' & n=n' & I=I' & B=B')"
 by (auto simp: req_def)
 
-subsubsection{*propose event*}
+subsubsection\<open>propose event\<close>
 
 definition prom :: "agent => nat => agent => nat => msg => msg =>
 msg => agent => msg" where
@@ -136,7 +136,7 @@
 ==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 by (auto simp: pro_def dest: prom_inj)
 
-subsubsection{*protocol*}
+subsubsection\<open>protocol\<close>
 
 inductive_set p1 :: "event list set"
 where
@@ -151,7 +151,7 @@
   I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I)));
   Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p1"
 
-subsubsection{*Composition of Traces*}
+subsubsection\<open>Composition of Traces\<close>
 
 lemma "evs':p1 ==> 
        evs:p1 & (ALL n. Nonce n:used evs' --> Nonce n ~:used evs) --> 
@@ -165,7 +165,7 @@
 apply (erule_tac A'=A' in p1.Propose, auto simp: pro_def)
 done
 
-subsubsection{*Valid Offer Lists*}
+subsubsection\<open>Valid Offer Lists\<close>
 
 inductive_set
   valid :: "agent => nat => agent => msg set"
@@ -176,7 +176,7 @@
 | Propose [intro]: "L:valid A n B
 ==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B"
 
-subsubsection{*basic properties of valid*}
+subsubsection\<open>basic properties of valid\<close>
 
 lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'"
 by (erule valid.cases, auto)
@@ -184,24 +184,24 @@
 lemma valid_pos_len: "L:valid A n B ==> 0 < len L"
 by (erule valid.induct, auto)
 
-subsubsection{*offers of an offer list*}
+subsubsection\<open>offers of an offer list\<close>
 
 definition offer_nonces :: "msg => msg set" where
 "offer_nonces L == {X. X:parts {L} & (EX n. X = Nonce n)}"
 
-subsubsection{*the originator can get the offers*}
+subsubsection\<open>the originator can get the offers\<close>
 
 lemma "L:valid A n B ==> offer_nonces L <= analz (insert L (initState A))"
 by (erule valid.induct, auto simp: anchor_def chain_def sign_def
 offer_nonces_def initState.simps)
 
-subsubsection{*list of offers*}
+subsubsection\<open>list of offers\<close>
 
 fun offers :: "msg => msg" where
 "offers (cons M L) = cons {|shop M, nonce M|} (offers L)" |
 "offers other = nil"
 
-subsubsection{*list of agents whose keys are used to sign a list of offers*}
+subsubsection\<open>list of agents whose keys are used to sign a list of offers\<close>
 
 fun shops :: "msg => msg" where
 "shops (cons M L) = cons (shop M) (shops L)" |
@@ -210,7 +210,7 @@
 lemma shops_in_agl: "L:valid A n B ==> shops L:agl"
 by (erule valid.induct, auto simp: anchor_def chain_def sign_def)
 
-subsubsection{*builds a trace from an itinerary*}
+subsubsection\<open>builds a trace from an itinerary\<close>
 
 fun offer_list :: "agent * nat * agent * msg * nat => msg" where
 "offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil" |
@@ -240,20 +240,20 @@
 
 declare trace'_def [simp]
 
-subsubsection{*there is a trace in which the originator receives a valid answer*}
+subsubsection\<open>there is a trace in which the originator receives a valid answer\<close>
 
 lemma p1_not_empty: "evs:p1 ==> req A r n I B:set evs -->
 (EX evs'. evs'@evs:p1 & pro B' ofr A r I' L J A:set evs' & L:valid A n B)"
 oops
 
 
-subsection{*properties of protocol P1*}
+subsection\<open>properties of protocol P1\<close>
 
-text{*publicly verifiable forward integrity:
-anyone can verify the validity of an offer list*}
+text\<open>publicly verifiable forward integrity:
+anyone can verify the validity of an offer list\<close>
 
-subsubsection{*strong forward integrity:
-except the last one, no offer can be modified*}
+subsubsection\<open>strong forward integrity:
+except the last one, no offer can be modified\<close>
 
 lemma strong_forward_integrity: "ALL L. Suc i < len L
 --> L:valid A n B & repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)"
@@ -273,8 +273,8 @@
 apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 by (drule_tac x=l' in spec, simp, blast)
 
-subsubsection{*insertion resilience:
-except at the beginning, no offer can be inserted*}
+subsubsection\<open>insertion resilience:
+except at the beginning, no offer can be inserted\<close>
 
 lemma chain_isnt_head [simp]: "L:valid A n B ==>
 head L ~= chain (next_shop (head L)) ofr A L C"
@@ -298,8 +298,8 @@
 apply (frule len_not_empty, clarsimp)
 by (drule_tac x=l' in spec, clarsimp)
 
-subsubsection{*truncation resilience:
-only shop i can truncate at offer i*}
+subsubsection\<open>truncation resilience:
+only shop i can truncate at offer i\<close>
 
 lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L
 --> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))"
@@ -318,19 +318,19 @@
 apply (frule len_not_empty, clarsimp)
 by (drule_tac x=l' in spec, clarsimp)
 
-subsubsection{*declarations for tactics*}
+subsubsection\<open>declarations for tactics\<close>
 
 declare knows_Spy_partsEs [elim]
 declare Fake_parts_insert [THEN subsetD, dest]
 declare initState.simps [simp del]
 
-subsubsection{*get components of a message*}
+subsubsection\<open>get components of a message\<close>
 
 lemma get_ML [dest]: "Says A' B {|A,r,I,M,L|}:set evs ==>
 M:parts (spies evs) & L:parts (spies evs)"
 by blast
 
-subsubsection{*general properties of p1*}
+subsubsection\<open>general properties of p1\<close>
 
 lemma reqm_neq_prom [iff]:
 "reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C"
@@ -369,7 +369,7 @@
                      req_def reqm_def anchor_def chain_def sign_def)
 by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans)
 
-subsubsection{*private keys are safe*}
+subsubsection\<open>private keys are safe\<close>
 
 lemma priK_parts_Friend_imp_bad [rule_format,dest]:
      "[| evs:p1; Friend B ~= A |]
@@ -395,7 +395,7 @@
 apply (drule_tac H="spies evs" in parts_sub)
 by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend)
 
-subsubsection{*general guardedness properties*}
+subsubsection\<open>general guardedness properties\<close>
 
 lemma agl_guard [intro]: "I:agl ==> I:guard n Ks"
 by (erule agl.induct, auto)
@@ -412,7 +412,7 @@
 Nonce n ~:used evs |] ==> L:guard n Ks"
 by (drule not_used_not_parts, auto)
 
-subsubsection{*guardedness of messages*}
+subsubsection\<open>guardedness of messages\<close>
 
 lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}"
 by (case_tac "ofr=n", auto simp: chain_def sign_def)
@@ -443,7 +443,7 @@
 L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}"
 by (auto simp: prom_def)
 
-subsubsection{*Nonce uniqueness*}
+subsubsection\<open>Nonce uniqueness\<close>
 
 lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr"
 by (auto simp: chain_def sign_def)
@@ -459,7 +459,7 @@
 I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr"
 by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel)
 
-subsubsection{*requests are guarded*}
+subsubsection\<open>requests are guarded\<close>
 
 lemma req_imp_Guard [rule_format]: "[| evs:p1; A ~:bad |] ==>
 req A r n I B:set evs --> Guard n {priK A} (spies evs)"
@@ -477,7 +477,7 @@
 apply (rule_tac p=p1 in knows_max'_sub_spies', simp+)
 by (rule knows'_sub_knows)
 
-subsubsection{*propositions are guarded*}
+subsubsection\<open>propositions are guarded\<close>
 
 lemma pro_imp_Guard [rule_format]: "[| evs:p1; B ~:bad; A ~:bad |] ==>
 pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)"
@@ -526,8 +526,8 @@
 apply (rule_tac p=p1 in knows_max'_sub_spies', simp+)
 by (rule knows'_sub_knows)
 
-subsubsection{*data confidentiality:
-no one other than the originator can decrypt the offers*}
+subsubsection\<open>data confidentiality:
+no one other than the originator can decrypt the offers\<close>
 
 lemma Nonce_req_notin_spies: "[| evs:p1; req A r n I B:set evs; A ~:bad |]
 ==> Nonce n ~:analz (spies evs)"
@@ -550,8 +550,8 @@
 apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+)
 by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def)
 
-subsubsection{*non repudiability:
-an offer signed by B has been sent by B*}
+subsubsection\<open>non repudiability:
+an offer signed by B has been sent by B\<close>
 
 lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'"
 by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl)
--- a/src/HOL/Auth/Guard/P2.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/P2.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -7,18 +7,18 @@
 Mobiles Agents 1998, LNCS 1477.
 *)
 
-section{*Protocol P2*}
+section\<open>Protocol P2\<close>
 
 theory P2 imports Guard_Public List_Msg begin
 
-subsection{*Protocol Definition*}
+subsection\<open>Protocol Definition\<close>
 
 
-text{*Like P1 except the definitions of @{text chain}, @{text shop},
-  @{text next_shop} and @{text nonce}*}
+text\<open>Like P1 except the definitions of \<open>chain\<close>, \<open>shop\<close>,
+  \<open>next_shop\<close> and \<open>nonce\<close>\<close>
 
-subsubsection{*offer chaining:
-B chains his offer for A with the head offer of L for sending it to C*}
+subsubsection\<open>offer chaining:
+B chains his offer for A with the head offer of L for sending it to C\<close>
 
 definition chain :: "agent => nat => agent => msg => agent => msg" where
 "chain B ofr A L C ==
@@ -35,7 +35,7 @@
 lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}"
 by (auto simp: chain_def sign_def)
 
-subsubsection{*agent whose key is used to sign an offer*}
+subsubsection\<open>agent whose key is used to sign an offer\<close>
 
 fun shop :: "msg => msg" where
 "shop {|Crypt K {|B,ofr,Crypt K' H|},m2|} = Agent (agt K')"
@@ -43,7 +43,7 @@
 lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B"
 by (simp add: chain_def sign_def)
 
-subsubsection{*nonce used in an offer*}
+subsubsection\<open>nonce used in an offer\<close>
 
 fun nonce :: "msg => msg" where
 "nonce {|Crypt K {|B,ofr,CryptH|},m2|} = ofr"
@@ -51,7 +51,7 @@
 lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr"
 by (simp add: chain_def sign_def)
 
-subsubsection{*next shop*}
+subsubsection\<open>next shop\<close>
 
 fun next_shop :: "msg => agent" where
 "next_shop {|m1,Hash {|headL,Agent C|}|} = C"
@@ -59,7 +59,7 @@
 lemma "next_shop (chain B ofr A L C) = C"
 by (simp add: chain_def sign_def)
 
-subsubsection{*anchor of the offer list*}
+subsubsection\<open>anchor of the offer list\<close>
 
 definition anchor :: "agent => nat => agent => msg" where
 "anchor A n B == chain A n A (cons nil nil) B"
@@ -74,7 +74,7 @@
 lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A"
 by (simp add: anchor_def)
 
-subsubsection{*request event*}
+subsubsection\<open>request event\<close>
 
 definition reqm :: "agent => nat => nat => msg => agent => msg" where
 "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I),
@@ -94,7 +94,7 @@
 = (A=A' & r=r' & n=n' & I=I' & B=B')"
 by (auto simp: req_def)
 
-subsubsection{*propose event*}
+subsubsection\<open>propose event\<close>
 
 definition prom :: "agent => nat => agent => nat => msg => msg =>
 msg => agent => msg" where
@@ -116,7 +116,7 @@
 ==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 by (auto simp: pro_def dest: prom_inj)
 
-subsubsection{*protocol*}
+subsubsection\<open>protocol\<close>
 
 inductive_set p2 :: "event list set"
 where
@@ -131,7 +131,7 @@
   I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I)));
   Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p2"
 
-subsubsection{*valid offer lists*}
+subsubsection\<open>valid offer lists\<close>
 
 inductive_set
   valid :: "agent => nat => agent => msg set"
@@ -142,7 +142,7 @@
 | Propose [intro]: "L:valid A n B
   ==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B"
 
-subsubsection{*basic properties of valid*}
+subsubsection\<open>basic properties of valid\<close>
 
 lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'"
 by (erule valid.cases, auto)
@@ -150,7 +150,7 @@
 lemma valid_pos_len: "L:valid A n B ==> 0 < len L"
 by (erule valid.induct, auto)
 
-subsubsection{*list of offers*}
+subsubsection\<open>list of offers\<close>
 
 fun offers :: "msg => msg"
 where
@@ -158,13 +158,13 @@
 | "offers other = nil"
 
 
-subsection{*Properties of Protocol P2*}
+subsection\<open>Properties of Protocol P2\<close>
 
-text{*same as @{text P1_Prop} except that publicly verifiable forward
-integrity is replaced by forward privacy*}
+text\<open>same as \<open>P1_Prop\<close> except that publicly verifiable forward
+integrity is replaced by forward privacy\<close>
 
-subsection{*strong forward integrity:
-except the last one, no offer can be modified*}
+subsection\<open>strong forward integrity:
+except the last one, no offer can be modified\<close>
 
 lemma strong_forward_integrity: "ALL L. Suc i < len L
 --> L:valid A n B --> repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)"
@@ -184,8 +184,8 @@
 apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 by (drule_tac x=l' in spec, simp, blast)
 
-subsection{*insertion resilience:
-except at the beginning, no offer can be inserted*}
+subsection\<open>insertion resilience:
+except at the beginning, no offer can be inserted\<close>
 
 lemma chain_isnt_head [simp]: "L:valid A n B ==>
 head L ~= chain (next_shop (head L)) ofr A L C"
@@ -209,8 +209,8 @@
 apply (frule len_not_empty, clarsimp)
 by (drule_tac x=l' in spec, clarsimp)
 
-subsection{*truncation resilience:
-only shop i can truncate at offer i*}
+subsection\<open>truncation resilience:
+only shop i can truncate at offer i\<close>
 
 lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L
 --> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))"
@@ -229,19 +229,19 @@
 apply (frule len_not_empty, clarsimp)
 by (drule_tac x=l' in spec, clarsimp)
 
-subsection{*declarations for tactics*}
+subsection\<open>declarations for tactics\<close>
 
 declare knows_Spy_partsEs [elim]
 declare Fake_parts_insert [THEN subsetD, dest]
 declare initState.simps [simp del]
 
-subsection{*get components of a message*}
+subsection\<open>get components of a message\<close>
 
 lemma get_ML [dest]: "Says A' B {|A,R,I,M,L|}:set evs ==>
 M:parts (spies evs) & L:parts (spies evs)"
 by blast
 
-subsection{*general properties of p2*}
+subsection\<open>general properties of p2\<close>
 
 lemma reqm_neq_prom [iff]:
 "reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C"
@@ -280,7 +280,7 @@
 req_def reqm_def anchor_def chain_def sign_def)
 by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans)
 
-subsection{*private keys are safe*}
+subsection\<open>private keys are safe\<close>
 
 lemma priK_parts_Friend_imp_bad [rule_format,dest]:
      "[| evs:p2; Friend B ~= A |]
@@ -307,7 +307,7 @@
 apply (drule_tac H="spies evs" in parts_sub)
 by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend)
 
-subsection{*general guardedness properties*}
+subsection\<open>general guardedness properties\<close>
 
 lemma agl_guard [intro]: "I:agl ==> I:guard n Ks"
 by (erule agl.induct, auto)
@@ -324,7 +324,7 @@
 Nonce n ~:used evs |] ==> L:guard n Ks"
 by (drule not_used_not_parts, auto)
 
-subsection{*guardedness of messages*}
+subsection\<open>guardedness of messages\<close>
 
 lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}"
 by (case_tac "ofr=n", auto simp: chain_def sign_def)
@@ -355,7 +355,7 @@
 L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}"
 by (auto simp: prom_def)
 
-subsection{*Nonce uniqueness*}
+subsection\<open>Nonce uniqueness\<close>
 
 lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr"
 by (auto simp: chain_def sign_def)
@@ -371,7 +371,7 @@
 I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr"
 by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel)
 
-subsection{*requests are guarded*}
+subsection\<open>requests are guarded\<close>
 
 lemma req_imp_Guard [rule_format]: "[| evs:p2; A ~:bad |] ==>
 req A r n I B:set evs --> Guard n {priK A} (spies evs)"
@@ -389,7 +389,7 @@
 apply (rule_tac p=p2 in knows_max'_sub_spies', simp+)
 by (rule knows'_sub_knows)
 
-subsection{*propositions are guarded*}
+subsection\<open>propositions are guarded\<close>
 
 lemma pro_imp_Guard [rule_format]: "[| evs:p2; B ~:bad; A ~:bad |] ==>
 pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)"
@@ -438,8 +438,8 @@
 apply (rule_tac p=p2 in knows_max'_sub_spies', simp+)
 by (rule knows'_sub_knows)
 
-subsection{*data confidentiality:
-no one other than the originator can decrypt the offers*}
+subsection\<open>data confidentiality:
+no one other than the originator can decrypt the offers\<close>
 
 lemma Nonce_req_notin_spies: "[| evs:p2; req A r n I B:set evs; A ~:bad |]
 ==> Nonce n ~:analz (spies evs)"
@@ -462,8 +462,8 @@
 apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+)
 by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def)
 
-subsection{*forward privacy:
-only the originator can know the identity of the shops*}
+subsection\<open>forward privacy:
+only the originator can know the identity of the shops\<close>
 
 lemma forward_privacy_Spy: "[| evs:p2; B ~:bad; A ~:bad;
 pro B ofr A r I (cons M L) J C:set evs |]
@@ -475,7 +475,7 @@
 ==> sign B (Nonce ofr) ~:analz (knows_max (Friend D) evs)"
 by (auto simp:sign_def dest:Nonce_pro_notin_knows_max_Friend )
 
-subsection{*non repudiability: an offer signed by B has been sent by B*}
+subsection\<open>non repudiability: an offer signed by B has been sent by B\<close>
 
 lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'"
 by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl)
--- a/src/HOL/Auth/Guard/Proto.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Guard/Proto.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   2002  University of Cambridge
 *)
 
-section{*Other Protocol-Independent Results*}
+section\<open>Other Protocol-Independent Results\<close>
 
 theory Proto imports Guard_Public begin
 
-subsection{*protocols*}
+subsection\<open>protocols\<close>
 
 type_synonym rule = "event set * event"
 
@@ -21,7 +21,7 @@
 "wdef p == ALL R k. R:p --> Number k:parts {msg' R}
 --> Number k:parts (msg`(fst R))"
 
-subsection{*substitutions*}
+subsection\<open>substitutions\<close>
 
 record subs =
   agent   :: "agent => agent"
@@ -82,7 +82,7 @@
   pubK' :: "subs => agent => key" where
   "pubK' s A == pubK (agent s A)"
 
-subsection{*nonces generated by a rule*}
+subsection\<open>nonces generated by a rule\<close>
 
 definition newn :: "rule => nat set" where
 "newn R == {n. Nonce n:parts {msg (snd R)} & Nonce n ~:parts (msg`(fst R))}"
@@ -90,7 +90,7 @@
 lemma newn_parts: "n:newn R ==> Nonce (nonce s n):parts {apm' s R}"
 by (auto simp: newn_def dest: apm_parts)
 
-subsection{*traces generated by a protocol*}
+subsection\<open>traces generated by a protocol\<close>
 
 definition ok :: "event list => rule => subs => bool" where
 "ok evs R s == ((ALL x. x:fst R --> ap s x:set evs)
@@ -108,7 +108,7 @@
 
 | Proto [intro]: "[| evs:tr p; R:p; ok evs R s |] ==> ap' s R # evs:tr p"
 
-subsection{*general properties*}
+subsection\<open>general properties\<close>
 
 lemma one_step_tr [iff]: "one_step (tr p)"
 apply (unfold one_step_def, clarify)
@@ -147,13 +147,13 @@
 apply (drule has_only_SaysD, simp+)
 by (clarify, case_tac x, auto)
 
-subsection{*types*}
+subsection\<open>types\<close>
 
 type_synonym keyfun = "rule => subs => nat => event list => key set"
 
 type_synonym secfun = "rule => nat => subs => key set => msg"
 
-subsection{*introduction of a fresh guarded nonce*}
+subsection\<open>introduction of a fresh guarded nonce\<close>
 
 definition fresh :: "proto => rule => subs => nat => key set => event list
 => bool" where
@@ -214,7 +214,7 @@
 apply (rule_tac Y="apm s' X" in parts_parts, blast)
 by (rule parts.Inj, rule Says_imp_spies, simp, blast)
 
-subsection{*safe keys*}
+subsection\<open>safe keys\<close>
 
 definition safe :: "key set => msg set => bool" where
 "safe Ks G == ALL K. K:Ks --> Key K ~:analz G"
@@ -228,7 +228,7 @@
 lemma Guard_safe: "[| Guard n Ks G; safe Ks G |] ==> Nonce n ~:analz G"
 by (blast dest: Guard_invKey)
 
-subsection{*guardedness preservation*}
+subsection\<open>guardedness preservation\<close>
 
 definition preserv :: "proto => keyfun => nat => key set => bool" where
 "preserv p keys n Ks == (ALL evs R' s' R s. evs:tr p -->
@@ -245,7 +245,7 @@
 ok evs (l,Says A B X) s; keys R' s' n evs <= Ks |] ==> apm s X:guard n Ks"
 by (drule preservD, simp+)
 
-subsection{*monotonic keyfun*}
+subsection\<open>monotonic keyfun\<close>
 
 definition monoton :: "proto => keyfun => bool" where
 "monoton p keys == ALL R' s' n ev evs. ev # evs:tr p -->
@@ -255,7 +255,7 @@
 ev # evs:tr p |] ==> keys R' s' n evs <= Ks"
 by (unfold monoton_def, blast)
 
-subsection{*guardedness theorem*}
+subsection\<open>guardedness theorem\<close>
 
 lemma Guard_tr [rule_format]: "[| evs:tr p; has_only_Says' p;
 preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy) |] ==>
@@ -295,7 +295,7 @@
 apply (blast dest: safe_insert)
 by (blast, simp, simp, blast)
 
-subsection{*useful properties for guardedness*}
+subsection\<open>useful properties for guardedness\<close>
 
 lemma newn_neq_used: "[| Nonce n:used evs; ok evs R s; k:newn R |]
 ==> n ~= nonce s k"
@@ -311,7 +311,7 @@
 ok evs R s |] ==> n ~:newn R"
 by (auto simp: ok_def dest: not_used_not_spied parts_parts)
 
-subsection{*unicity*}
+subsection\<open>unicity\<close>
 
 definition uniq :: "proto => secfun => bool" where
 "uniq p secret == ALL evs R R' n n' Ks s s'. R:p --> R':p -->
@@ -360,7 +360,7 @@
 apply (blast dest: uniq'D)
 by (auto dest: ordD uniq'D intro: sym)
 
-subsection{*Needham-Schroeder-Lowe*}
+subsection\<open>Needham-Schroeder-Lowe\<close>
 
 definition a :: agent where "a == Friend 0"
 definition b :: agent where "b == Friend 1"
@@ -419,7 +419,7 @@
 apply (simp add: split_paired_all)
 by (rule impI, erule ns.cases, simp_all)+
 
-subsection{*general properties*}
+subsection\<open>general properties\<close>
 
 lemma ns_has_only_Says' [iff]: "has_only_Says' ns"
 apply (unfold has_only_Says'_def)
@@ -439,7 +439,7 @@
 lemma ns_wdef [iff]: "wdef ns"
 by (auto simp: wdef_def elim: ns.cases)
 
-subsection{*guardedness for NSL*}
+subsection\<open>guardedness for NSL\<close>
 
 lemma "uniq ns secret ==> preserv ns keys n Ks"
 apply (unfold preserv_def)
@@ -511,7 +511,7 @@
 (* fresh with NS3 *)
 by simp
 
-subsection{*unicity for NSL*}
+subsection\<open>unicity for NSL\<close>
 
 lemma "uniq' ns inf secret"
 apply (unfold uniq'_def)
--- a/src/HOL/Auth/KerberosIV.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/KerberosIV.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   1998  University of Cambridge
 *)
 
-section{*The Kerberos Protocol, Version IV*}
+section\<open>The Kerberos Protocol, Version IV\<close>
 
 theory KerberosIV imports Public begin
 
-text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+text\<open>The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.\<close>
 
 abbreviation
   Kas :: agent where "Kas == Server"
@@ -18,7 +18,7 @@
 
 axiomatization where
   Tgs_not_bad [iff]: "Tgs \<notin> bad"
-   --{*Tgs is secure --- we already know that Kas is secure*}
+   \<comment>\<open>Tgs is secure --- we already know that Kas is secure\<close>
 
 definition
  (* authKeys are those contained in an authTicket *)
@@ -252,7 +252,7 @@
 declare Fake_parts_insert_in_Un [dest]
 
 
-subsection{*Lemmas about lists, for reasoning about  Issues*}
+subsection\<open>Lemmas about lists, for reasoning about  Issues\<close>
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
@@ -286,13 +286,13 @@
 apply (induct_tac "evs")
 apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
-txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+txt\<open>Resembles \<open>used_subset_append\<close> in theory Event.\<close>
 done
 
 lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
 
 
-subsection{*Lemmas about @{term authKeys}*}
+subsection\<open>Lemmas about @{term authKeys}\<close>
 
 lemma authKeys_empty: "authKeys [] = {}"
 apply (unfold authKeys_def)
@@ -330,9 +330,9 @@
 by (simp add: authKeys_def, blast)
 
 
-subsection{*Forwarding Lemmas*}
+subsection\<open>Forwarding Lemmas\<close>
 
-text{*--For reasoning about the encrypted portion of message K3--*}
+text\<open>--For reasoning about the encrypted portion of message K3--\<close>
 lemma K3_msg_in_parts_spies:
      "Says Kas' A (Crypt KeyA \<lbrace>authK, Peer, Ta, authTicket\<rbrace>)
                \<in> set evs \<Longrightarrow> authTicket \<in> parts (spies evs)"
@@ -346,7 +346,7 @@
 apply (erule kerbIV.induct, auto)
 done
 
-text{*--For reasoning about the encrypted portion of message K5--*}
+text\<open>--For reasoning about the encrypted portion of message K5--\<close>
 lemma K5_msg_in_parts_spies:
      "Says Tgs' A (Crypt authK \<lbrace>servK, Agent B, Ts, servTicket\<rbrace>)
                \<in> set evs \<Longrightarrow> servTicket \<in> parts (spies evs)"
@@ -384,7 +384,7 @@
 
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
 
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
     "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbIV\<rbrakk>
      \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
@@ -392,9 +392,9 @@
 apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert)
-txt{*Others*}
+txt\<open>Others\<close>
 apply (force dest!: analz_shrK_Decrypt)+
 done
 
@@ -407,7 +407,7 @@
 
 
 
-subsection{*Lemmas for reasoning about predicate "before"*}
+subsection\<open>Lemmas for reasoning about predicate "before"\<close>
 
 lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
@@ -461,10 +461,10 @@
 by auto
 
 
-subsection{*Regularity Lemmas*}
-text{*These concern the form of items passed in messages*}
+subsection\<open>Regularity Lemmas\<close>
+text\<open>These concern the form of items passed in messages\<close>
 
-text{*Describes the form of all components sent by Kas*}
+text\<open>Describes the form of all components sent by Kas\<close>
 lemma Says_Kas_message_form:
      "\<lbrakk> Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
            \<in> set evs;
@@ -482,7 +482,7 @@
 apply (erule rev_mp)
 apply (erule kerbIV.induct)
 apply (simp_all (no_asm) add: authKeys_def authKeys_insert, blast, blast)
-txt{*K2*}
+txt\<open>K2\<close>
 apply (simp (no_asm) add: takeWhile_tail)
 apply (rule conjI)
 apply (metis Key_not_used authKeys_used length_rev set_rev takeWhile_void used_evs_rev)
@@ -521,7 +521,7 @@
 apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake, K4*}
+txt\<open>Fake, K4\<close>
 apply (blast+)
 done
 
@@ -535,7 +535,7 @@
 apply blast
 done
 
-text{*Describes the form of servK, servTicket and authK sent by Tgs*}
+text\<open>Describes the form of servK, servTicket and authK sent by Tgs\<close>
 lemma Says_Tgs_message_form:
      "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
            \<in> set evs;
@@ -554,16 +554,16 @@
 apply (erule rev_mp)
 apply (erule kerbIV.induct)
 apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast)
-txt{*We need this simplification only for Message 4*}
+txt\<open>We need this simplification only for Message 4\<close>
 apply (simp (no_asm) add: takeWhile_tail)
 apply auto
-txt{*Five subcases of Message 4*}
+txt\<open>Five subcases of Message 4\<close>
 apply (blast dest!: SesKey_is_session_key)
 apply (blast dest: authTicket_crypt_authK)
 apply (blast dest!: authKeys_used Says_Kas_message_form)
-txt{*subcase: used before*}
+txt\<open>subcase: used before\<close>
 apply (metis used_evs_rev used_takeWhile_used)
-txt{*subcase: CT before*}
+txt\<open>subcase: CT before\<close>
 apply (metis length_rev set_evs_rev takeWhile_void)
 done
 
@@ -581,7 +581,7 @@
 apply (blast+)
 done
 
-text{* This form holds also over an authTicket, but is not needed below.*}
+text\<open>This form holds also over an authTicket, but is not needed below.\<close>
 lemma servTicket_form:
      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>
               \<in> parts (spies evs);
@@ -596,7 +596,7 @@
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
 done
 
-text{* Essentially the same as @{text authTicket_form} *}
+text\<open>Essentially the same as \<open>authTicket_form\<close>\<close>
 lemma Says_kas_message_form:
      "\<lbrakk> Says Kas' A (Crypt (shrK A)
               \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>) \<in> set evs;
@@ -619,7 +619,7 @@
 by (metis Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Decrypt analz.Snd invKey_K servTicket_form)
 
 
-subsection{*Authenticity theorems: confirm origin of sensitive messages*}
+subsection\<open>Authenticity theorems: confirm origin of sensitive messages\<close>
 
 lemma authK_authentic:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>
@@ -631,13 +631,13 @@
 apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast dest!: authTicket_authentic [THEN Says_Kas_message_form])
 done
 
-text{*If a certain encrypted message appears then it originated with Tgs*}
+text\<open>If a certain encrypted message appears then it originated with Tgs\<close>
 lemma servK_authentic:
      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
@@ -651,11 +651,11 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply auto
 done
 
@@ -672,13 +672,13 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
 done
 
-text{*Authenticity of servK for B*}
+text\<open>Authenticity of servK for B\<close>
 lemma servTicket_authentic_Tgs:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
            \<in> parts (spies evs); B \<noteq> Tgs;  B \<notin> bad;
@@ -695,7 +695,7 @@
 apply blast+
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma K4_imp_K2:
 "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
       \<in> set evs;  evs \<in> kerbIV\<rbrakk>
@@ -711,7 +711,7 @@
 apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma u_K4_imp_K2:
 "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
       \<in> set evs; evs \<in> kerbIV\<rbrakk>
@@ -780,7 +780,7 @@
   by (metis le_less_trans)
 
 
-subsection{* Reliability: friendly agents send something if something else happened*}
+subsection\<open>Reliability: friendly agents send something if something else happened\<close>
 
 lemma K3_imp_K2:
      "\<lbrakk> Says A Tgs
@@ -797,7 +797,7 @@
 apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN authK_authentic])
 done
 
-text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
+text\<open>Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.\<close>
 lemma Key_unique_SesKey:
      "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T, Ticket\<rbrace>
            \<in> parts (spies evs);
@@ -811,7 +811,7 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake, K2, K4*}
+txt\<open>Fake, K2, K4\<close>
 apply (blast+)
 done
 
@@ -829,13 +829,13 @@
 apply (frule_tac [5] Says_ticket_parts)
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K2*}
+txt\<open>K2\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*K3*}
+txt\<open>K3\<close>
 apply (blast dest: Key_unique_SesKey)
-txt{*K5*}
+txt\<open>K5\<close>
 apply (metis K3_imp_K2 Key_unique_SesKey Spy_see_shrK parts.Body parts.Fst 
              Says_imp_knows_Spy [THEN parts.Inj])
 done
@@ -855,15 +855,15 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
 apply blast
-txt{*K3*}
+txt\<open>K3\<close>
 apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*K5*}
+txt\<open>K5\<close>
 apply (blast dest: Key_unique_SesKey)
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma unique_CryptKey:
      "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
            \<in> parts (spies evs);
@@ -877,7 +877,7 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake, K2, K4*}
+txt\<open>Fake, K2, K4\<close>
 apply (blast+)
 done
 
@@ -903,7 +903,7 @@
              unique_CryptKey) 
 done
 
-text{*Needs a unicity theorem, hence moved here*}
+text\<open>Needs a unicity theorem, hence moved here\<close>
 lemma servK_authentic_ter:
  "\<lbrakk> Says Kas A
     (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>) \<in> set evs;
@@ -920,18 +920,18 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
-txt{*K2*}
+txt\<open>K2\<close>
 apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
-txt{*K4 remain*}
+txt\<open>K4 remain\<close>
 apply (blast dest!: unique_CryptKey)
 done
 
 
-subsection{*Unicity Theorems*}
+subsection\<open>Unicity Theorems\<close>
 
-text{* The session key, if secure, uniquely identifies the Ticket
+text\<open>The session key, if secure, uniquely identifies the Ticket
    whether authTicket or servTicket. As a matter of fact, one can read
-   also Tgs in the place of B.                                     *}
+   also Tgs in the place of B.\<close>
 
 
 (*
@@ -963,11 +963,11 @@
 apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
 done
 
-text{* servK uniquely identifies the message from Tgs *}
+text\<open>servK uniquely identifies the message from Tgs\<close>
 lemma unique_servKeys:
      "\<lbrakk> Says Tgs A
               (Crypt K \<lbrace>Key servK, Agent B, Ts, X\<rbrace>) \<in> set evs;
@@ -979,11 +979,11 @@
 apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
 done
 
-text{* Revised unicity theorems *}
+text\<open>Revised unicity theorems\<close>
 
 lemma Kas_Unique:
      "\<lbrakk> Says Kas A
@@ -1006,7 +1006,7 @@
 done
 
 
-subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
+subsection\<open>Lemmas About the Predicate @{term AKcryptSK}\<close>
 
 lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
 by (simp add: AKcryptSK_def)
@@ -1054,15 +1054,15 @@
 apply (erule kerbIV.induct)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K2: by freshness*}
+txt\<open>K2: by freshness\<close>
 apply (simp add: AKcryptSK_def)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast+)
 done
 
-text{*A secure serverkey cannot have been used to encrypt others*}
+text\<open>A secure serverkey cannot have been used to encrypt others\<close>
 lemma servK_not_AKcryptSK:
  "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, Number Ts\<rbrace> \<in> parts (spies evs);
      Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
@@ -1073,11 +1073,11 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (metis Auth_fresh_not_AKcryptSK Crypt_imp_keysFor new_keys_not_used parts.Fst parts.Snd Says_imp_knows_Spy [THEN parts.Inj] unique_CryptKey)
 done
 
-text{*Long term keys are not issued as servKeys*}
+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)
@@ -1086,8 +1086,8 @@
 apply (frule_tac [5] K3_msg_in_parts_spies, auto)
 done
 
-text{*The Tgs message associates servK with authK and therefore not with any
-  other key authK.*}
+text\<open>The Tgs message associates servK with authK and therefore not with any
+  other key authK.\<close>
 lemma Says_Tgs_AKcryptSK:
      "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>)
            \<in> set evs;
@@ -1097,7 +1097,7 @@
 apply (blast dest: unique_servKeys)
 done
 
-text{*Equivalently*}
+text\<open>Equivalently\<close>
 lemma not_different_AKcryptSK:
      "\<lbrakk> AKcryptSK authK servK evs;
         authK' \<noteq> authK;  evs \<in> kerbIV \<rbrakk>
@@ -1117,11 +1117,11 @@
              authKeys_used authTicket_crypt_authK parts.Fst parts.Inj)
 done
 
-text{*The only session keys that can be found with the help of session keys are
-  those sent by Tgs in step K4.  *}
+text\<open>The only session keys that can be found with the help of session keys are
+  those sent by Tgs in step K4.\<close>
 
-text{*We take some pains to express the property
-  as a logical equivalence so that the simplifier can apply it.*}
+text\<open>We take some pains to express the property
+  as a logical equivalence so that the simplifier can apply it.\<close>
 lemma Key_analz_image_Key_lemma:
      "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
       \<Longrightarrow>
@@ -1152,9 +1152,9 @@
 done
 
 
-subsection{*Secrecy Theorems*}
+subsection\<open>Secrecy Theorems\<close>
 
-text{*For the Oops2 case of the next theorem*}
+text\<open>For the Oops2 case of the next theorem\<close>
 lemma Oops2_not_AKcryptSK:
      "\<lbrakk> evs \<in> kerbIV;
          Says Tgs A (Crypt authK
@@ -1163,11 +1163,11 @@
       \<Longrightarrow> \<not> AKcryptSK servK SK evs"
 by (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
    
-text{* Big simplification law for keys SK that are not crypted by keys in KK
+text\<open>Big simplification law for keys SK that are not crypted by keys in KK
  It helps prove three, otherwise hard, facts about keys. These facts are
  exploited as simplification laws for analz, and also "limit the damage"
  in case of loss of a key to the spy. See ESORICS98.
- [simplified by LCP] *}
+ [simplified by LCP]\<close>
 lemma Key_analz_image_Key [rule_format (no_asm)]:
      "evs \<in> kerbIV \<Longrightarrow>
       (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
@@ -1180,38 +1180,38 @@
 apply (frule_tac [7] Says_tgs_message_form)
 apply (frule_tac [5] Says_kas_message_form)
 apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
-txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
- the induction hypothesis*}
+txt\<open>Case-splits for Oops1 and message 5: the negated case simplifies using
+ the induction hypothesis\<close>
 apply (case_tac [11] "AKcryptSK authK SK evsO1")
 apply (case_tac [8] "AKcryptSK servK SK evs5")
 apply (simp_all del: image_insert
         add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
              Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
        Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
-txt{*Fake*} 
+txt\<open>Fake\<close> 
 apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast 
-txt{*K3*}
+txt\<open>K3\<close>
 apply blast 
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast dest!: authK_not_AKcryptSK)
-txt{*K5*}
+txt\<open>K5\<close>
 apply (case_tac "Key servK \<in> analz (spies evs5) ")
-txt{*If servK is compromised then the result follows directly...*}
+txt\<open>If servK is compromised then the result follows directly...\<close>
 apply (simp (no_asm_simp) add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
-txt{*...therefore servK is uncompromised.*}
-txt{*The AKcryptSK servK SK evs5 case leads to a contradiction.*}
+txt\<open>...therefore servK is uncompromised.\<close>
+txt\<open>The AKcryptSK servK SK evs5 case leads to a contradiction.\<close>
 apply (blast elim!: servK_not_AKcryptSK [THEN [2] rev_notE] del: allE ballE)
-txt{*Another K5 case*}
+txt\<open>Another K5 case\<close>
 apply blast 
-txt{*Oops1*}
+txt\<open>Oops1\<close>
 apply simp 
 apply (blast dest!: AKcryptSK_analz_insert)
 done
 
-text{* First simplification law for analz: no session keys encrypt
-authentication keys or shared keys. *}
+text\<open>First simplification law for analz: no session keys encrypt
+authentication keys or shared keys.\<close>
 lemma analz_insert_freshK1:
      "\<lbrakk> evs \<in> kerbIV;  K \<in> authKeys evs Un range shrK;
         SesKey \<notin> range shrK \<rbrakk>
@@ -1223,7 +1223,7 @@
 done
 
 
-text{* Second simplification law for analz: no service keys encrypt any other keys.*}
+text\<open>Second simplification law for analz: no service keys encrypt any other keys.\<close>
 lemma analz_insert_freshK2:
      "\<lbrakk> evs \<in> kerbIV;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
         K \<in> symKeys \<rbrakk>
@@ -1235,7 +1235,7 @@
 done
 
 
-text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
+text\<open>Third simplification law for analz: only one authentication key encrypts a certain service key.\<close>
 
 lemma analz_insert_freshK3:
  "\<lbrakk> AKcryptSK authK servK evs;
@@ -1258,7 +1258,7 @@
 apply (simp add: analz_insert_freshK3)
 done
 
-text{*a weakness of the protocol*}
+text\<open>a weakness of the protocol\<close>
 lemma authK_compromises_servK:
      "\<lbrakk> Says Tgs A
               (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
@@ -1284,8 +1284,8 @@
 done
 
 
-text{*If Spy sees the Authentication Key sent in msg K2, then
-    the Key has expired.*}
+text\<open>If Spy sees the Authentication Key sent in msg K2, then
+    the Key has expired.\<close>
 lemma Confidentiality_Kas_lemma [rule_format]:
      "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbIV \<rbrakk>
       \<Longrightarrow> Says Kas A
@@ -1302,17 +1302,17 @@
 apply (frule_tac [5] Says_kas_message_form)
 apply (safe del: impI conjI impCE)
 apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
-txt{*Level 8: K5*}
+txt\<open>Level 8: K5\<close>
 apply (blast dest: servK_notin_authKeysD Says_Kas_message_form intro: less_SucI)
-txt{*Oops1*}
+txt\<open>Oops1\<close>
 apply (blast dest!: unique_authKeys intro: less_SucI)
-txt{*Oops2*}
+txt\<open>Oops2\<close>
 apply (blast dest: Says_Tgs_message_form Says_Kas_message_form)
 done
 
@@ -1325,8 +1325,8 @@
       \<Longrightarrow> Key authK \<notin> analz (spies evs)"
 by (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
 
-text{*If Spy sees the Service Key sent in msg K4, then
-    the Key has expired.*}
+text\<open>If Spy sees the Service Key sent in msg K4, then
+    the Key has expired.\<close>
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
@@ -1343,11 +1343,11 @@
 apply (erule rev_mp)
 apply (erule kerbIV.induct)
 apply (rule_tac [9] impI)+
-  --{*The Oops1 case is unusual: must simplify
+  \<comment>\<open>The Oops1 case is unusual: must simplify
     @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
-   @{text analz_mono_contra} weaken it to
+   \<open>analz_mono_contra\<close> weaken it to
    @{term "Authkey \<notin> analz (spies evs)"},
-  for we then conclude @{term "authK \<noteq> authKa"}.*}
+  for we then conclude @{term "authK \<noteq> authKa"}.\<close>
 apply analz_mono_contra
 apply (frule_tac [10] Oops_range_spies2)
 apply (frule_tac [9] Oops_range_spies1)
@@ -1355,23 +1355,23 @@
 apply (frule_tac [5] Says_kas_message_form)
 apply (safe del: impI conjI impCE)
 apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
      apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
     apply (blast intro: parts_insertI less_SucI)
-txt{*K4*}
+txt\<open>K4\<close>
    apply (blast dest: authTicket_authentic Confidentiality_Kas)
-txt{*K5*}
+txt\<open>K5\<close>
   apply (metis Says_imp_spies Says_ticket_parts Tgs_not_bad analz_insert_freshK2 
              less_SucI parts.Inj servK_notin_authKeysD unique_CryptKey)
-txt{*Oops1*} 
+txt\<open>Oops1\<close> 
  apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
-txt{*Oops2*}
+txt\<open>Oops2\<close>
 apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
 done
 
 
-text{* In the real world Tgs can't check wheter authK is secure! *}
+text\<open>In the real world Tgs can't check wheter authK is secure!\<close>
 lemma Confidentiality_Tgs:
      "\<lbrakk> Says Tgs A
               (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
@@ -1382,7 +1382,7 @@
       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 by (blast dest: Says_Tgs_message_form Confidentiality_lemma)
 
-text{* In the real world Tgs CAN check what Kas sends! *}
+text\<open>In the real world Tgs CAN check what Kas sends!\<close>
 lemma Confidentiality_Tgs_bis:
      "\<lbrakk> Says Kas A
                (Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
@@ -1395,13 +1395,13 @@
       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 by (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
 
-text{*Most general form*}
+text\<open>Most general form\<close>
 lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
 
 lemmas Confidentiality_Auth_A = authK_authentic [THEN Confidentiality_Kas]
 
-text{*Needs a confidentiality guarantee, hence moved here.
-      Authenticity of servK for A*}
+text\<open>Needs a confidentiality guarantee, hence moved here.
+      Authenticity of servK for A\<close>
 lemma servK_authentic_bis_r:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
            \<in> parts (spies evs);
@@ -1460,15 +1460,15 @@
 
 
 
-subsection{*Parties authentication: each party verifies "the identity of
-       another party who generated some data" (quoted from Neuman and Ts'o).*}
+subsection\<open>Parties authentication: each party verifies "the identity of
+       another party who generated some data" (quoted from Neuman and Ts'o).\<close>
 
-text{*These guarantees don't assess whether two parties agree on
+text\<open>These guarantees don't assess whether two parties agree on
          the same session key: sending a message containing a key
-         doesn't a priori state knowledge of the key.*}
+         doesn't a priori state knowledge of the key.\<close>
 
 
-text{*@{text Tgs_authenticates_A} can be found above*}
+text\<open>\<open>Tgs_authenticates_A\<close> can be found above\<close>
 
 lemma A_authenticates_Tgs:
  "\<lbrakk> Says Kas A
@@ -1486,9 +1486,9 @@
 apply (erule kerbIV.induct, analz_mono_contra)
 apply (frule_tac [7] K5_msg_in_parts_spies)
 apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast)
-txt{*K2*}
+txt\<open>K2\<close>
 apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast dest!: unique_CryptKey)
 done
 
@@ -1503,7 +1503,7 @@
                Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs"
 by (blast dest: servTicket_authentic_Tgs intro: Says_K5)
 
-text{*The second assumption tells B what kind of key servK is.*}
+text\<open>The second assumption tells B what kind of key servK is.\<close>
 lemma B_authenticates_A_r:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
          Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
@@ -1518,7 +1518,7 @@
                   Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
 by (blast intro: Says_K5 dest: Confidentiality_B servTicket_authentic_Tgs)
 
-text{* @{text u_B_authenticates_A} would be the same as @{text B_authenticates_A} because the servK confidentiality assumption is yet unrelaxed*}
+text\<open>\<open>u_B_authenticates_A\<close> would be the same as \<open>B_authenticates_A\<close> because the servK confidentiality assumption is yet unrelaxed\<close>
 
 lemma u_B_authenticates_A_r:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
@@ -1563,10 +1563,10 @@
 done
 
 
-subsection{* Key distribution guarantees
+subsection\<open>Key distribution guarantees
        An agent knows a session key if he used it to issue a cipher.
        These guarantees also convey a stronger form of 
-       authentication - non-injective agreement on the session key*}
+       authentication - non-injective agreement on the session key\<close>
 
 
 lemma Kas_Issues_A:
@@ -1583,7 +1583,7 @@
 apply (frule_tac [5] Says_ticket_parts)
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*K2*}
+txt\<open>K2\<close>
 apply (simp add: takeWhile_tail)
 apply (blast dest: authK_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD])
 done
@@ -1638,9 +1638,9 @@
 apply (frule_tac [5] Says_ticket_parts)
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*fake*}
+txt\<open>fake\<close>
 apply blast
-txt{*K3*}
+txt\<open>K3\<close>
 (*
 apply clarify
 apply (drule Says_imp_knows_Spy [THEN parts.Inj, THEN authK_authentic, THEN Says_Kas_message_form], assumption, assumption, assumption)
@@ -1678,7 +1678,7 @@
 apply (frule_tac [5] Says_ticket_parts)
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (simp add: takeWhile_tail)
 (*Last two thms installed only to derive authK \<notin> range shrK*)
 apply (metis knows_Spy_partsEs(2) parts.Fst usedI used_evs_rev used_takeWhile_used)
@@ -1709,7 +1709,7 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
 apply blast
-txt{*K6 requires numerous lemmas*}
+txt\<open>K6 requires numerous lemmas\<close>
 apply (simp add: takeWhile_tail)
 apply (blast dest: servTicket_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: Says_K6)
 done
@@ -1776,11 +1776,11 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp))
 apply clarify
-txt{*K5*}
+txt\<open>K5\<close>
 apply auto
 apply (simp add: takeWhile_tail)
-txt{*Level 15: case analysis necessary because the assumption doesn't state
-  the form of servTicket. The guarantee becomes stronger.*}
+txt\<open>Level 15: case analysis necessary because the assumption doesn't state
+  the form of servTicket. The guarantee becomes stronger.\<close>
 apply (blast dest: Says_imp_spies [THEN analz.Inj, THEN analz_Decrypt']
                    K3_imp_K2 servK_authentic_ter
                    parts_spies_takeWhile_mono [THEN subsetD]
@@ -1823,8 +1823,8 @@
    \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
 by (blast dest: B_authenticates_A Confidentiality_B A_Issues_B)
 
-text{* @{text u_B_authenticates_and_keydist_to_A} would be the same as @{text B_authenticates_and_keydist_to_A} because the
- servK confidentiality assumption is yet unrelaxed*}
+text\<open>\<open>u_B_authenticates_and_keydist_to_A\<close> would be the same as \<open>B_authenticates_and_keydist_to_A\<close> because the
+ servK confidentiality assumption is yet unrelaxed\<close>
 
 lemma u_B_authenticates_and_keydist_to_A_r:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
--- a/src/HOL/Auth/KerberosIV_Gets.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/KerberosIV_Gets.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   1998  University of Cambridge
 *)
 
-section{*The Kerberos Protocol, Version IV*}
+section\<open>The Kerberos Protocol, Version IV\<close>
 
 theory KerberosIV_Gets imports Public begin
 
-text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+text\<open>The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.\<close>
 
 abbreviation
   Kas :: agent where "Kas == Server"
@@ -18,7 +18,7 @@
 
 axiomatization where
   Tgs_not_bad [iff]: "Tgs \<notin> bad"
-   --{*Tgs is secure --- we already know that Kas is secure*}
+   \<comment>\<open>Tgs is secure --- we already know that Kas is secure\<close>
 
 definition
  (* authKeys are those contained in an authTicket *)
@@ -240,7 +240,7 @@
 declare analz_into_parts [dest]
 declare Fake_parts_insert_in_Un [dest]
 
-subsection{*Lemmas about reception event*}
+subsection\<open>Lemmas about reception event\<close>
 
 lemma Gets_imp_Says :
      "\<lbrakk> Gets B X \<in> set evs; evs \<in> kerbIV_gets \<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
@@ -260,7 +260,7 @@
      "\<lbrakk> Gets B X \<in> set evs; evs \<in> kerbIV_gets \<rbrakk>  \<Longrightarrow> X \<in> knows B evs"
 by (metis Gets_imp_knows_Spy Gets_imp_knows_agents)
 
-subsection{*Lemmas about @{term authKeys}*}
+subsection\<open>Lemmas about @{term authKeys}\<close>
 
 lemma authKeys_empty: "authKeys [] = {}"
 by (simp add: authKeys_def)
@@ -296,7 +296,7 @@
 by (simp add: authKeys_def, blast)
 
 
-subsection{*Forwarding Lemmas*}
+subsection\<open>Forwarding Lemmas\<close>
 
 lemma Says_ticket_parts:
      "Says S A (Crypt K \<lbrace>SesKey, B, TimeStamp, Ticket\<rbrace>) \<in> set evs
@@ -343,7 +343,7 @@
 by (blast dest: Spy_see_shrK)
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
 
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
     "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbIV_gets\<rbrakk>
      \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
@@ -351,9 +351,9 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert)
-txt{*Others*}
+txt\<open>Others\<close>
 apply (force dest!: analz_shrK_Decrypt)+
 done
 
@@ -365,10 +365,10 @@
 by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD])
 
 
-subsection{*Regularity Lemmas*}
-text{*These concern the form of items passed in messages*}
+subsection\<open>Regularity Lemmas\<close>
+text\<open>These concern the form of items passed in messages\<close>
 
-text{*Describes the form of all components sent by Kas*}
+text\<open>Describes the form of all components sent by Kas\<close>
 
 lemma Says_Kas_message_form:
      "\<lbrakk> Says Kas A (Crypt K \<lbrace>Key authK, Agent Peer, Number Ta, authTicket\<rbrace>)
@@ -406,7 +406,7 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake, K4*}
+txt\<open>Fake, K4\<close>
 apply (blast+)
 done
 
@@ -431,7 +431,7 @@
 apply (erule rev_mp)
 apply (erule kerbIV_gets.induct)
 apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast, auto)
-txt{*Three subcases of Message 4*}
+txt\<open>Three subcases of Message 4\<close>
 apply (blast dest!: SesKey_is_session_key)
 apply (blast dest: authTicket_crypt_authK)
 apply (blast dest!: authKeys_used Says_Kas_message_form)
@@ -452,7 +452,7 @@
 apply blast+
 done
 
-text{* This form holds also over an authTicket, but is not needed below.*}
+text\<open>This form holds also over an authTicket, but is not needed below.\<close>
 lemma servTicket_form:
      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts, servTicket\<rbrace>
               \<in> parts (spies evs);
@@ -467,7 +467,7 @@
 apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
 done
 
-text{* Essentially the same as @{text authTicket_form} *}
+text\<open>Essentially the same as \<open>authTicket_form\<close>\<close>
 lemma Says_kas_message_form:
      "\<lbrakk> Gets A (Crypt (shrK A)
               \<lbrace>Key authK, Agent Tgs, Ta, authTicket\<rbrace>) \<in> set evs;
@@ -494,7 +494,7 @@
 done
 
 
-subsection{*Authenticity theorems: confirm origin of sensitive messages*}
+subsection\<open>Authenticity theorems: confirm origin of sensitive messages\<close>
 
 lemma authK_authentic:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta, authTicket\<rbrace>
@@ -506,13 +506,13 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast dest!: authTicket_authentic [THEN Says_Kas_message_form])
 done
 
-text{*If a certain encrypted message appears then it originated with Tgs*}
+text\<open>If a certain encrypted message appears then it originated with Tgs\<close>
 lemma servK_authentic:
      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
            \<in> parts (spies evs);
@@ -526,11 +526,11 @@
 apply (erule kerbIV_gets.induct, analz_mono_contra)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply auto
 done
 
@@ -547,13 +547,13 @@
 apply (erule kerbIV_gets.induct, analz_mono_contra)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
 done
 
-text{*Authenticity of servK for B*}
+text\<open>Authenticity of servK for B\<close>
 lemma servTicket_authentic_Tgs:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
            \<in> parts (spies evs); B \<noteq> Tgs;  B \<notin> bad;
@@ -570,7 +570,7 @@
 apply blast+
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma K4_imp_K2:
 "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
       \<in> set evs;  evs \<in> kerbIV_gets\<rbrakk>
@@ -586,7 +586,7 @@
 apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma u_K4_imp_K2:
 "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
       \<in> set evs; evs \<in> kerbIV_gets\<rbrakk>
@@ -655,7 +655,7 @@
 by (blast dest: leI le_trans dest: leD)
 
 
-subsection{* Reliability: friendly agents send something if something else happened*}
+subsection\<open>Reliability: friendly agents send something if something else happened\<close>
 
 lemma K3_imp_K2:
      "\<lbrakk> Says A Tgs
@@ -672,7 +672,7 @@
 apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN authK_authentic])
 done
 
-text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
+text\<open>Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.\<close>
 lemma Key_unique_SesKey:
      "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T, Ticket\<rbrace>
            \<in> parts (spies evs);
@@ -686,7 +686,7 @@
 apply (erule kerbIV_gets.induct, analz_mono_contra)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake, K2, K4*}
+txt\<open>Fake, K2, K4\<close>
 apply (blast+)
 done
 
@@ -704,22 +704,22 @@
 apply (frule_tac [6] Gets_ticket_parts)
 apply (frule_tac [9] Gets_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*K2*}
+txt\<open>K2\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*K3*}
+txt\<open>K3\<close>
 apply (blast dest: Key_unique_SesKey)
-txt{*K5*}
-txt{*If authKa were compromised, so would be authK*}
+txt\<open>K5\<close>
+txt\<open>If authKa were compromised, so would be authK\<close>
 apply (case_tac "Key authKa \<in> analz (spies evs5)")
 apply (force dest!: Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst])
-txt{*Besides, since authKa originated with Kas anyway...*}
+txt\<open>Besides, since authKa originated with Kas anyway...\<close>
 apply (clarify, drule K3_imp_K2, assumption, assumption)
 apply (clarify, drule Says_Kas_message_form, assumption)
-txt{*...it cannot be a shared key*. Therefore @{term servK_authentic} applies. 
+txt\<open>...it cannot be a shared key*. Therefore @{term servK_authentic} applies. 
      Contradition: Tgs used authK as a servkey, 
-     while Kas used it as an authkey*}
+     while Kas used it as an authkey\<close>
 apply (blast dest: servK_authentic Says_Tgs_message_form)
 done
 
@@ -738,15 +738,15 @@
 apply (frule_tac [9] Gets_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
 apply blast
-txt{*K3*}
+txt\<open>K3\<close>
 apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*K5*}
+txt\<open>K5\<close>
 apply (blast dest: Key_unique_SesKey)
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma unique_CryptKey:
      "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
            \<in> parts (spies evs);
@@ -760,7 +760,7 @@
 apply (erule kerbIV_gets.induct, analz_mono_contra)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake, K2, K4*}
+txt\<open>Fake, K2, K4\<close>
 apply (blast+)
 done
 
@@ -784,7 +784,7 @@
 apply (blast dest: unique_CryptKey)
 done
 
-text{*Needs a unicity theorem, hence moved here*}
+text\<open>Needs a unicity theorem, hence moved here\<close>
 lemma servK_authentic_ter:
  "\<lbrakk> Says Kas A
     (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>) \<in> set evs;
@@ -801,17 +801,17 @@
 apply (erule kerbIV_gets.induct, analz_mono_contra)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
-txt{*K2 and K4 remain*}
+txt\<open>K2 and K4 remain\<close>
 prefer 2 apply (blast dest!: unique_CryptKey)
 apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
 done
 
 
-subsection{*Unicity Theorems*}
+subsection\<open>Unicity Theorems\<close>
 
-text{* The session key, if secure, uniquely identifies the Ticket
+text\<open>The session key, if secure, uniquely identifies the Ticket
    whether authTicket or servTicket. As a matter of fact, one can read
-   also Tgs in the place of B.                                     *}
+   also Tgs in the place of B.\<close>
 
 
 lemma unique_authKeys:
@@ -825,11 +825,11 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
 done
 
-text{* servK uniquely identifies the message from Tgs *}
+text\<open>servK uniquely identifies the message from Tgs\<close>
 lemma unique_servKeys:
      "\<lbrakk> Says Tgs A
               (Crypt K \<lbrace>Key servK, Agent B, Ts, X\<rbrace>) \<in> set evs;
@@ -841,11 +841,11 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
 done
 
-text{* Revised unicity theorems *}
+text\<open>Revised unicity theorems\<close>
 
 lemma Kas_Unique:
      "\<lbrakk> Says Kas A
@@ -868,7 +868,7 @@
 done
 
 
-subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
+subsection\<open>Lemmas About the Predicate @{term AKcryptSK}\<close>
 
 lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
 by (simp add: AKcryptSK_def)
@@ -915,16 +915,16 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*Reception*}
+txt\<open>Reception\<close>
 apply (simp add: AKcryptSK_def)
-txt{*K2: by freshness*}
+txt\<open>K2: by freshness\<close>
 apply (simp add: AKcryptSK_def)
-txt{*K4*}
+txt\<open>K4\<close>
 by (blast+)
 
-text{*A secure serverkey cannot have been used to encrypt others*}
+text\<open>A secure serverkey cannot have been used to encrypt others\<close>
 lemma servK_not_AKcryptSK:
  "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, Number Ts\<rbrace> \<in> parts (spies evs);
      Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
@@ -935,17 +935,17 @@
 apply (erule kerbIV_gets.induct, analz_mono_contra)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts, simp_all, blast)
-txt{*Reception*}
+txt\<open>Reception\<close>
 apply (simp add: AKcryptSK_def)
-txt{*K4 splits into distinct subcases*}
+txt\<open>K4 splits into distinct subcases\<close>
 apply auto
-txt{*servK can't have been enclosed in two certificates*}
+txt\<open>servK can't have been enclosed in two certificates\<close>
  prefer 2 apply (blast dest: unique_CryptKey)
-txt{*servK is fresh and so could not have been used, by
-   @{text new_keys_not_used}*}
+txt\<open>servK is fresh and so could not have been used, by
+   \<open>new_keys_not_used\<close>\<close>
 by (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
 
-text{*Long term keys are not issued as servKeys*}
+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)
@@ -953,8 +953,8 @@
 apply (frule_tac [8] Gets_ticket_parts)
 by (frule_tac [6] Gets_ticket_parts, auto)
 
-text{*The Tgs message associates servK with authK and therefore not with any
-  other key authK.*}
+text\<open>The Tgs message associates servK with authK and therefore not with any
+  other key authK.\<close>
 lemma Says_Tgs_AKcryptSK:
      "\<lbrakk> Says Tgs A (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, X \<rbrace>)
            \<in> set evs;
@@ -963,7 +963,7 @@
 apply (unfold AKcryptSK_def)
 by (blast dest: unique_servKeys)
 
-text{*Equivalently*}
+text\<open>Equivalently\<close>
 lemma not_different_AKcryptSK:
      "\<lbrakk> AKcryptSK authK servK evs;
         authK' \<noteq> authK;  evs \<in> kerbIV_gets \<rbrakk>
@@ -978,23 +978,23 @@
 apply (erule kerbIV_gets.induct)
 apply (frule_tac [8] Gets_ticket_parts)
 apply (frule_tac [6] Gets_ticket_parts)
-txt{*Reception*}
+txt\<open>Reception\<close>
 prefer 3 apply (simp add: AKcryptSK_def)
 apply (simp_all, safe)
-txt{*K4 splits into subcases*}
+txt\<open>K4 splits into subcases\<close>
 prefer 4 apply (blast dest!: authK_not_AKcryptSK)
-txt{*servK is fresh and so could not have been used, by
-   @{text new_keys_not_used}*}
+txt\<open>servK is fresh and so could not have been used, by
+   \<open>new_keys_not_used\<close>\<close>
  prefer 2 
  apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
-txt{*Others by freshness*}
+txt\<open>Others by freshness\<close>
 by (blast+)
 
-text{*The only session keys that can be found with the help of session keys are
-  those sent by Tgs in step K4.  *}
+text\<open>The only session keys that can be found with the help of session keys are
+  those sent by Tgs in step K4.\<close>
 
-text{*We take some pains to express the property
-  as a logical equivalence so that the simplifier can apply it.*}
+text\<open>We take some pains to express the property
+  as a logical equivalence so that the simplifier can apply it.\<close>
 lemma Key_analz_image_Key_lemma:
      "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
       \<Longrightarrow>
@@ -1022,9 +1022,9 @@
 by (blast dest: Says_Tgs_message_form)
 
 
-subsection{*Secrecy Theorems*}
+subsection\<open>Secrecy Theorems\<close>
 
-text{*For the Oops2 case of the next theorem*}
+text\<open>For the Oops2 case of the next theorem\<close>
 lemma Oops2_not_AKcryptSK:
      "\<lbrakk> evs \<in> kerbIV_gets;
          Says Tgs A (Crypt authK
@@ -1033,10 +1033,10 @@
       \<Longrightarrow> \<not> AKcryptSK servK SK evs"
 by (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
    
-text{* Big simplification law for keys SK that are not crypted by keys in KK
+text\<open>Big simplification law for keys SK that are not crypted by keys in KK
  It helps prove three, otherwise hard, facts about keys. These facts are
  exploited as simplification laws for analz, and also "limit the damage"
- in case of loss of a key to the spy. See ESORICS98. *}
+ in case of loss of a key to the spy. See ESORICS98.\<close>
 lemma Key_analz_image_Key [rule_format (no_asm)]:
      "evs \<in> kerbIV_gets \<Longrightarrow>
       (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
@@ -1049,40 +1049,40 @@
 apply (frule_tac [8] Says_tgs_message_form)
 apply (frule_tac [6] Says_kas_message_form)
 apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
-txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
- the induction hypothesis*}
+txt\<open>Case-splits for Oops1 and message 5: the negated case simplifies using
+ the induction hypothesis\<close>
 apply (case_tac [12] "AKcryptSK authK SK evsO1")
 apply (case_tac [9] "AKcryptSK servK SK evs5")
 apply (simp_all del: image_insert
         add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
              Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
        Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
-  --{*18 seconds on a 1.8GHz machine??*}
-txt{*Fake*} 
+  \<comment>\<open>18 seconds on a 1.8GHz machine??\<close>
+txt\<open>Fake\<close> 
 apply spy_analz
-txt{*Reception*}
+txt\<open>Reception\<close>
 apply (simp add: AKcryptSK_def)
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast 
-txt{*K3*}
+txt\<open>K3\<close>
 apply blast 
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast dest!: authK_not_AKcryptSK)
-txt{*K5*}
+txt\<open>K5\<close>
 apply (case_tac "Key servK \<in> analz (spies evs5) ")
-txt{*If servK is compromised then the result follows directly...*}
+txt\<open>If servK is compromised then the result follows directly...\<close>
 apply (simp (no_asm_simp) add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
-txt{*...therefore servK is uncompromised.*}
-txt{*The AKcryptSK servK SK evs5 case leads to a contradiction.*}
+txt\<open>...therefore servK is uncompromised.\<close>
+txt\<open>The AKcryptSK servK SK evs5 case leads to a contradiction.\<close>
 apply (blast elim!: servK_not_AKcryptSK [THEN [2] rev_notE] del: allE ballE)
-txt{*Another K5 case*}
+txt\<open>Another K5 case\<close>
 apply blast 
-txt{*Oops1*}
+txt\<open>Oops1\<close>
 apply simp 
 by (blast dest!: AKcryptSK_analz_insert)
 
-text{* First simplification law for analz: no session keys encrypt
-authentication keys or shared keys. *}
+text\<open>First simplification law for analz: no session keys encrypt
+authentication keys or shared keys.\<close>
 lemma analz_insert_freshK1:
      "\<lbrakk> evs \<in> kerbIV_gets;  K \<in> authKeys evs Un range shrK;
         SesKey \<notin> range shrK \<rbrakk>
@@ -1094,7 +1094,7 @@
 done
 
 
-text{* Second simplification law for analz: no service keys encrypt any other keys.*}
+text\<open>Second simplification law for analz: no service keys encrypt any other keys.\<close>
 lemma analz_insert_freshK2:
      "\<lbrakk> evs \<in> kerbIV_gets;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
         K \<in> symKeys \<rbrakk>
@@ -1106,7 +1106,7 @@
 done
 
 
-text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
+text\<open>Third simplification law for analz: only one authentication key encrypts a certain service key.\<close>
 
 lemma analz_insert_freshK3:
  "\<lbrakk> AKcryptSK authK servK evs;
@@ -1128,7 +1128,7 @@
 apply (frule AKcryptSKI, assumption)
 by (simp add: analz_insert_freshK3)
 
-text{*a weakness of the protocol*}
+text\<open>a weakness of the protocol\<close>
 lemma authK_compromises_servK:
      "\<lbrakk> Says Tgs A
               (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
@@ -1153,8 +1153,8 @@
 by (blast+)
 
 
-text{*If Spy sees the Authentication Key sent in msg K2, then
-    the Key has expired.*}
+text\<open>If Spy sees the Authentication Key sent in msg K2, then
+    the Key has expired.\<close>
 lemma Confidentiality_Kas_lemma [rule_format]:
      "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbIV_gets \<rbrakk>
       \<Longrightarrow> Says Kas A
@@ -1171,17 +1171,17 @@
 apply (frule_tac [6] Says_kas_message_form)
 apply (safe del: impI conjI impCE)
 apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
-txt{*Level 8: K5*}
+txt\<open>Level 8: K5\<close>
 apply (blast dest: servK_notin_authKeysD Says_Kas_message_form intro: less_SucI)
-txt{*Oops1*}
+txt\<open>Oops1\<close>
 apply (blast dest!: unique_authKeys intro: less_SucI)
-txt{*Oops2*}
+txt\<open>Oops2\<close>
 by (blast dest: Says_Tgs_message_form Says_Kas_message_form)
 
 lemma Confidentiality_Kas:
@@ -1193,8 +1193,8 @@
       \<Longrightarrow> Key authK \<notin> analz (spies evs)"
 by (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
 
-text{*If Spy sees the Service Key sent in msg K4, then
-    the Key has expired.*}
+text\<open>If Spy sees the Service Key sent in msg K4, then
+    the Key has expired.\<close>
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
@@ -1211,11 +1211,11 @@
 apply (erule rev_mp)
 apply (erule kerbIV_gets.induct)
 apply (rule_tac [10] impI)+
-  --{*The Oops1 case is unusual: must simplify
+  \<comment>\<open>The Oops1 case is unusual: must simplify
     @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
-   @{text analz_mono_contra} weaken it to
+   \<open>analz_mono_contra\<close> weaken it to
    @{term "Authkey \<notin> analz (spies evs)"},
-  for we then conclude @{term "authK \<noteq> authKa"}.*}
+  for we then conclude @{term "authK \<noteq> authKa"}.\<close>
 apply analz_mono_contra
 apply (frule_tac [11] Oops_range_spies2)
 apply (frule_tac [10] Oops_range_spies1)
@@ -1223,20 +1223,20 @@
 apply (frule_tac [6] Says_kas_message_form)
 apply (safe del: impI conjI impCE)
 apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
 apply (blast intro: parts_insertI less_SucI)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (blast dest: authTicket_authentic Confidentiality_Kas)
-txt{*Oops2*}
+txt\<open>Oops2\<close>
   prefer 3
   apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
-txt{*Oops1*}
+txt\<open>Oops1\<close>
  prefer 2
 apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
-txt{*K5. Not clear how this step could be integrated with the main
-       simplification step. Done in KerberosV.thy*}
+txt\<open>K5. Not clear how this step could be integrated with the main
+       simplification step. Done in KerberosV.thy\<close>
 apply clarify
 apply (erule_tac V = "Says Aa Tgs X \<in> set evs" for X evs in thin_rl)
 apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN servK_notin_authKeysD])
@@ -1246,7 +1246,7 @@
 done
 
 
-text{* In the real world Tgs can't check wheter authK is secure! *}
+text\<open>In the real world Tgs can't check wheter authK is secure!\<close>
 lemma Confidentiality_Tgs:
      "\<lbrakk> Says Tgs A
               (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
@@ -1257,7 +1257,7 @@
       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 by (blast dest: Says_Tgs_message_form Confidentiality_lemma)
 
-text{* In the real world Tgs CAN check what Kas sends! *}
+text\<open>In the real world Tgs CAN check what Kas sends!\<close>
 lemma Confidentiality_Tgs_bis:
      "\<lbrakk> Says Kas A
                (Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>)
@@ -1270,13 +1270,13 @@
       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 by (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
 
-text{*Most general form*}
+text\<open>Most general form\<close>
 lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
 
 lemmas Confidentiality_Auth_A = authK_authentic [THEN Confidentiality_Kas]
 
-text{*Needs a confidentiality guarantee, hence moved here.
-      Authenticity of servK for A*}
+text\<open>Needs a confidentiality guarantee, hence moved here.
+      Authenticity of servK for A\<close>
 lemma servK_authentic_bis_r:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta, authTicket\<rbrace>
            \<in> parts (spies evs);
@@ -1309,11 +1309,11 @@
 
 
 
-subsection{*2. Parties' strong authentication: 
+subsection\<open>2. Parties' strong authentication: 
        non-injective agreement on the session key. The same guarantees also
-       express key distribution, hence their names*}
+       express key distribution, hence their names\<close>
 
-text{*Authentication here still is weak agreement - of B with A*}
+text\<open>Authentication here still is weak agreement - of B with A\<close>
 lemma A_authenticates_B:
      "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
          Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>
--- a/src/HOL/Auth/KerberosV.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/KerberosV.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,11 +2,11 @@
     Author:     Giampaolo Bella, Catania University
 *)
 
-section{*The Kerberos Protocol, Version V*}
+section\<open>The Kerberos Protocol, Version V\<close>
 
 theory KerberosV imports Public begin
 
-text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+text\<open>The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.\<close>
 
 abbreviation
   Kas :: agent where
@@ -19,7 +19,7 @@
 
 axiomatization where
   Tgs_not_bad [iff]: "Tgs \<notin> bad"
-   --{*Tgs is secure --- we already know that Kas is secure*}
+   \<comment>\<open>Tgs is secure --- we already know that Kas is secure\<close>
 
 definition
  (* authKeys are those contained in an authTicket *)
@@ -203,7 +203,7 @@
 
 
 
-subsection{*Lemmas about lists, for reasoning about  Issues*}
+subsection\<open>Lemmas about lists, for reasoning about  Issues\<close>
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
@@ -237,13 +237,13 @@
 apply (induct_tac "evs")
 apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
-txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+txt\<open>Resembles \<open>used_subset_append\<close> in theory Event.\<close>
 done
 
 lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
 
 
-subsection{*Lemmas about @{term authKeys}*}
+subsection\<open>Lemmas about @{term authKeys}\<close>
 
 lemma authKeys_empty: "authKeys [] = {}"
   by (simp add: authKeys_def)
@@ -279,7 +279,7 @@
   by (auto simp add: authKeys_def)
 
 
-subsection{*Forwarding Lemmas*}
+subsection\<open>Forwarding Lemmas\<close>
 
 lemma Says_ticket_parts:
      "Says S A \<lbrace>Crypt K \<lbrace>SesKey, B, TimeStamp\<rbrace>, Ticket\<rbrace>
@@ -327,7 +327,7 @@
 
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
 
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
     "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbV\<rbrakk>
      \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
@@ -335,9 +335,9 @@
 apply (erule kerbV.induct)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert)
-txt{*Others*}
+txt\<open>Others\<close>
 apply (force dest!: analz_shrK_Decrypt)+
 done
 
@@ -350,10 +350,10 @@
 
 
 
-subsection{*Regularity Lemmas*}
-text{*These concern the form of items passed in messages*}
+subsection\<open>Regularity Lemmas\<close>
+text\<open>These concern the form of items passed in messages\<close>
 
-text{*Describes the form of all components sent by Kas*}
+text\<open>Describes the form of all components sent by Kas\<close>
 lemma Says_Kas_message_form:
      "\<lbrakk> Says Kas A \<lbrace>Crypt K \<lbrace>Key authK, Agent Peer, Ta\<rbrace>, authTicket\<rbrace>
            \<in> set evs;
@@ -399,7 +399,7 @@
 apply (erule kerbV.induct)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all)
-txt{*Fake, K4*}
+txt\<open>Fake, K4\<close>
 apply (blast+)
 done
 
@@ -410,7 +410,7 @@
       \<Longrightarrow> authK \<in> authKeys evs"
 by (metis authKeysI authTicket_authentic)
 
-text{*Describes the form of servK, servTicket and authK sent by Tgs*}
+text\<open>Describes the form of servK, servTicket and authK sent by Tgs\<close>
 lemma Says_Tgs_message_form:
      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, servTicket\<rbrace>
            \<in> set evs;
@@ -422,7 +422,7 @@
 apply (erule rev_mp)
 apply (erule kerbV.induct)
 apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast, auto)
-txt{*Three subcases of Message 4*}
+txt\<open>Three subcases of Message 4\<close>
 apply (blast dest!: authKeys_used Says_Kas_message_form)
 apply (blast dest!: SesKey_is_session_key)
 apply (blast dest: authTicket_crypt_authK)
@@ -444,7 +444,7 @@
 servK_notin_authKeysD is no longer needed.
 *)
 
-subsection{*Authenticity theorems: confirm origin of sensitive messages*}
+subsection\<open>Authenticity theorems: confirm origin of sensitive messages\<close>
 
 lemma authK_authentic:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>
@@ -459,7 +459,7 @@
 apply blast+
 done
 
-text{*If a certain encrypted message appears then it originated with Tgs*}
+text\<open>If a certain encrypted message appears then it originated with Tgs\<close>
 lemma servK_authentic:
      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
            \<in> parts (spies evs);
@@ -491,7 +491,7 @@
 apply (frule_tac [5] Says_ticket_parts, simp_all, blast+)
 done
 
-text{*Authenticity of servK for B*}
+text\<open>Authenticity of servK for B\<close>
 lemma servTicket_authentic_Tgs:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>
            \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
@@ -506,7 +506,7 @@
 apply (frule_tac [5] Says_ticket_parts, simp_all, blast+)
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma K4_imp_K2:
 "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
       \<in> set evs;  evs \<in> kerbV\<rbrakk>
@@ -521,7 +521,7 @@
 apply (metis MPair_analz Says_imp_analz_Spy analz_conj_parts authTicket_authentic)
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma u_K4_imp_K2:
 "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>  \<in> set evs; evs \<in> kerbV\<rbrakk>
    \<Longrightarrow> \<exists>Ta. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
@@ -588,7 +588,7 @@
       \<Longrightarrow> \<not> expiredAK Ta evs"
 by (metis order_le_less_trans)
 
-subsection{* Reliability: friendly agents send somthing if something else happened*}
+subsection\<open>Reliability: friendly agents send somthing if something else happened\<close>
 
 lemma K3_imp_K2:
      "\<lbrakk> Says A Tgs
@@ -604,7 +604,7 @@
 apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authK_authentic])
 done
 
-text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
+text\<open>Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.\<close>
 lemma Key_unique_SesKey:
      "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T\<rbrace>
            \<in> parts (spies evs);
@@ -618,11 +618,11 @@
 apply (erule kerbV.induct, analz_mono_contra)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all)
-txt{*Fake, K2, K4*}
+txt\<open>Fake, K2, K4\<close>
 apply (blast+)
 done
 
-text{*This inevitably has an existential form in version V*}
+text\<open>This inevitably has an existential form in version V\<close>
 lemma Says_K5:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
          Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
@@ -638,15 +638,15 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
 apply blast
-txt{*K3*}
+txt\<open>K3\<close>
 apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*K5*}
+txt\<open>K5\<close>
 apply (blast dest: Key_unique_SesKey)
 done
 
-text{*Anticipated here from next subsection*}
+text\<open>Anticipated here from next subsection\<close>
 lemma unique_CryptKey:
      "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
            \<in> parts (spies evs);
@@ -660,7 +660,7 @@
 apply (erule kerbV.induct, analz_mono_contra)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all)
-txt{*Fake, K2, K4*}
+txt\<open>Fake, K2, K4\<close>
 apply (blast+)
 done
 
@@ -680,15 +680,15 @@
 apply (frule_tac [5] Says_ticket_parts)
 apply simp_all
 
-txt{*fake*}
+txt\<open>fake\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*K6*}
+txt\<open>K6\<close>
 apply (metis MPair_parts Says_imp_parts_knows_Spy unique_CryptKey)
 done
 
-text{*Needs a unicity theorem, hence moved here*}
+text\<open>Needs a unicity theorem, hence moved here\<close>
 lemma servK_authentic_ter:
  "\<lbrakk> Says Kas A
        \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, authTicket\<rbrace> \<in> set evs;
@@ -707,17 +707,17 @@
 apply (erule kerbV.induct, analz_mono_contra)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
-txt{*K2 and K4 remain*}
+txt\<open>K2 and K4 remain\<close>
 apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
 apply (blast dest!: unique_CryptKey)
 done
 
 
-subsection{*Unicity Theorems*}
+subsection\<open>Unicity Theorems\<close>
 
-text{* The session key, if secure, uniquely identifies the Ticket
+text\<open>The session key, if secure, uniquely identifies the Ticket
    whether authTicket or servTicket. As a matter of fact, one can read
-   also Tgs in the place of B.                                     *}
+   also Tgs in the place of B.\<close>
 
 
 lemma unique_authKeys:
@@ -734,7 +734,7 @@
 apply blast+
 done
 
-text{* servK uniquely identifies the message from Tgs *}
+text\<open>servK uniquely identifies the message from Tgs\<close>
 lemma unique_servKeys:
      "\<lbrakk> Says Tgs A
               \<lbrace>Crypt K \<lbrace>Key servK, Agent B, Ts\<rbrace>, X\<rbrace> \<in> set evs;
@@ -749,7 +749,7 @@
 apply blast+
 done
 
-subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
+subsection\<open>Lemmas About the Predicate @{term AKcryptSK}\<close>
 
 lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
 apply (simp add: AKcryptSK_def)
@@ -799,11 +799,11 @@
 apply (erule kerbV.induct)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all)
-txt{*Fake,K2,K4*}
+txt\<open>Fake,K2,K4\<close>
 apply (auto simp add: AKcryptSK_def)
 done
 
-text{*A secure serverkey cannot have been used to encrypt others*}
+text\<open>A secure serverkey cannot have been used to encrypt others\<close>
 lemma servK_not_AKcryptSK:
  "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, tt\<rbrace> \<in> parts (spies evs);
      Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
@@ -814,11 +814,11 @@
 apply (erule kerbV.induct, analz_mono_contra)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
-txt{*K4*}
+txt\<open>K4\<close>
 apply (metis Auth_fresh_not_AKcryptSK MPair_parts Says_imp_parts_knows_Spy authKeys_used authTicket_crypt_authK unique_CryptKey)
 done
 
-text{*Long term keys are not issued as servKeys*}
+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)
@@ -827,8 +827,8 @@
 apply (frule_tac [5] Says_ticket_parts, auto)
 done
 
-text{*The Tgs message associates servK with authK and therefore not with any
-  other key authK.*}
+text\<open>The Tgs message associates servK with authK and therefore not with any
+  other key authK.\<close>
 lemma Says_Tgs_AKcryptSK:
      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>, X \<rbrace>
            \<in> set evs;
@@ -844,13 +844,13 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts)
 apply (simp_all, safe)
-txt{*K4 splits into subcases*}
+txt\<open>K4 splits into subcases\<close>
 prefer 4 apply (blast dest!: authK_not_AKcryptSK)
-txt{*servK is fresh and so could not have been used, by
-   @{text new_keys_not_used}*}
+txt\<open>servK is fresh and so could not have been used, by
+   \<open>new_keys_not_used\<close>\<close>
  prefer 2 
  apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
-txt{*Others by freshness*}
+txt\<open>Others by freshness\<close>
 apply (blast+)
 done
 
@@ -862,11 +862,11 @@
 apply (blast dest: unique_servKeys Says_Tgs_message_form)
 done
 
-text{*The only session keys that can be found with the help of session keys are
-  those sent by Tgs in step K4.  *}
+text\<open>The only session keys that can be found with the help of session keys are
+  those sent by Tgs in step K4.\<close>
 
-text{*We take some pains to express the property
-  as a logical equivalence so that the simplifier can apply it.*}
+text\<open>We take some pains to express the property
+  as a logical equivalence so that the simplifier can apply it.\<close>
 lemma Key_analz_image_Key_lemma:
      "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
       \<Longrightarrow>
@@ -897,9 +897,9 @@
 done
 
 
-subsection{*Secrecy Theorems*}
+subsection\<open>Secrecy Theorems\<close>
 
-text{*For the Oops2 case of the next theorem*}
+text\<open>For the Oops2 case of the next theorem\<close>
 lemma Oops2_not_AKcryptSK:
      "\<lbrakk> evs \<in> kerbV;
          Says Tgs A \<lbrace>Crypt authK
@@ -908,10 +908,10 @@
       \<Longrightarrow> \<not> AKcryptSK servK SK evs"
 by (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
    
-text{* Big simplification law for keys SK that are not crypted by keys in KK
+text\<open>Big simplification law for keys SK that are not crypted by keys in KK
  It helps prove three, otherwise hard, facts about keys. These facts are
  exploited as simplification laws for analz, and also "limit the damage"
- in case of loss of a key to the spy. See ESORICS98.*}
+ in case of loss of a key to the spy. See ESORICS98.\<close>
 lemma Key_analz_image_Key [rule_format (no_asm)]:
      "evs \<in> kerbV \<Longrightarrow>
       (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
@@ -928,28 +928,28 @@
   Instead\<dots>*)
 apply (drule_tac [5] Says_ticket_analz)
 apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
-txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
- the induction hypothesis*}
+txt\<open>Case-splits for Oops1 and message 5: the negated case simplifies using
+ the induction hypothesis\<close>
 apply (case_tac [9] "AKcryptSK authK SK evsO1")
 apply (case_tac [7] "AKcryptSK servK SK evs5")
 apply (simp_all del: image_insert
           add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
                Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
                Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
-txt{*Fake*} 
+txt\<open>Fake\<close> 
 apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast 
-txt{*Cases K3 and K5 solved by the simplifier thanks to the ticket being in 
-analz - this strategy is new wrt version IV*} 
-txt{*K4*}
+txt\<open>Cases K3 and K5 solved by the simplifier thanks to the ticket being in 
+analz - this strategy is new wrt version IV\<close> 
+txt\<open>K4\<close>
 apply (blast dest!: authK_not_AKcryptSK)
-txt{*Oops1*}
+txt\<open>Oops1\<close>
 apply (metis AKcryptSK_analz_insert insert_Key_singleton)
 done
 
-text{* First simplification law for analz: no session keys encrypt
-authentication keys or shared keys. *}
+text\<open>First simplification law for analz: no session keys encrypt
+authentication keys or shared keys.\<close>
 lemma analz_insert_freshK1:
      "\<lbrakk> evs \<in> kerbV;  K \<in> authKeys evs Un range shrK;
         SesKey \<notin> range shrK \<rbrakk>
@@ -961,7 +961,7 @@
 done
 
 
-text{* Second simplification law for analz: no service keys encrypt any other keys.*}
+text\<open>Second simplification law for analz: no service keys encrypt any other keys.\<close>
 lemma analz_insert_freshK2:
      "\<lbrakk> evs \<in> kerbV;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
         K \<in> symKeys \<rbrakk>
@@ -973,7 +973,7 @@
 done
 
 
-text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
+text\<open>Third simplification law for analz: only one authentication key encrypts a certain service key.\<close>
 
 lemma analz_insert_freshK3:
  "\<lbrakk> AKcryptSK authK servK evs;
@@ -995,7 +995,7 @@
 apply (simp add: analz_insert_freshK3)
 done
 
-text{*a weakness of the protocol*}
+text\<open>a weakness of the protocol\<close>
 lemma authK_compromises_servK:
      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
         \<in> set evs;  authK \<in> symKeys;
@@ -1004,10 +1004,10 @@
   by (metis Says_imp_analz_Spy analz.Fst analz_Decrypt')
 
 
-text{*lemma @{text servK_notin_authKeysD} not needed in version V*}
+text\<open>lemma \<open>servK_notin_authKeysD\<close> not needed in version V\<close>
 
-text{*If Spy sees the Authentication Key sent in msg K2, then
-    the Key has expired.*}
+text\<open>If Spy sees the Authentication Key sent in msg K2, then
+    the Key has expired.\<close>
 lemma Confidentiality_Kas_lemma [rule_format]:
      "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbV \<rbrakk>
       \<Longrightarrow> Says Kas A
@@ -1023,15 +1023,15 @@
 apply (frule_tac [5] Says_ticket_analz)
 apply (safe del: impI conjI impCE)
 apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*K2*}
+txt\<open>K2\<close>
 apply blast
-txt{*K4*}
+txt\<open>K4\<close>
 apply blast
-txt{*Oops1*}
+txt\<open>Oops1\<close>
 apply (blast dest!: unique_authKeys intro: less_SucI)
-txt{*Oops2*}
+txt\<open>Oops2\<close>
 apply (blast dest: Says_Tgs_message_form Says_Kas_message_form)
 done
 
@@ -1045,8 +1045,8 @@
 apply (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
 done
 
-text{*If Spy sees the Service Key sent in msg K4, then
-    the Key has expired.*}
+text\<open>If Spy sees the Service Key sent in msg K4, then
+    the Key has expired.\<close>
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
@@ -1062,11 +1062,11 @@
 apply (erule rev_mp)
 apply (erule kerbV.induct)
 apply (rule_tac [9] impI)+
-  --{*The Oops1 case is unusual: must simplify
+  \<comment>\<open>The Oops1 case is unusual: must simplify
     @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
-   @{text analz_mono_contra} weaken it to
+   \<open>analz_mono_contra\<close> weaken it to
    @{term "Authkey \<notin> analz (spies evs)"},
-  for we then conclude @{term "authK \<noteq> authKa"}.*}
+  for we then conclude @{term "authK \<noteq> authKa"}.\<close>
 apply analz_mono_contra
 apply (frule_tac [10] Oops_range_spies2)
 apply (frule_tac [9] Oops_range_spies1)
@@ -1074,20 +1074,20 @@
 apply (frule_tac [5] Says_ticket_analz)
 apply (safe del: impI conjI impCE)
 apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
-    txt{*Fake*}
+    txt\<open>Fake\<close>
     apply spy_analz
-   txt{*K2*}
+   txt\<open>K2\<close>
    apply (blast intro: parts_insertI less_SucI)
-  txt{*K4*}
+  txt\<open>K4\<close>
   apply (blast dest: authTicket_authentic Confidentiality_Kas)
- txt{*Oops1*}
+ txt\<open>Oops1\<close>
  apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
-txt{*Oops2*}
+txt\<open>Oops2\<close>
 apply (metis Suc_le_eq linorder_linear linorder_not_le msg.simps(2) unique_servKeys)
 done
 
 
-text{* In the real world Tgs can't check wheter authK is secure! *}
+text\<open>In the real world Tgs can't check wheter authK is secure!\<close>
 lemma Confidentiality_Tgs:
      "\<lbrakk> Says Tgs A
               \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
@@ -1098,7 +1098,7 @@
       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 by (blast dest: Says_Tgs_message_form Confidentiality_lemma)
 
-text{* In the real world Tgs CAN check what Kas sends! *}
+text\<open>In the real world Tgs CAN check what Kas sends!\<close>
 lemma Confidentiality_Tgs_bis:
      "\<lbrakk> Says Kas A
                \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>, authTicket\<rbrace>
@@ -1111,13 +1111,13 @@
       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 by (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
 
-text{*Most general form*}
+text\<open>Most general form\<close>
 lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
 
 lemmas Confidentiality_Auth_A = authK_authentic [THEN exE, THEN Confidentiality_Kas]
 
-text{*Needs a confidentiality guarantee, hence moved here.
-      Authenticity of servK for A*}
+text\<open>Needs a confidentiality guarantee, hence moved here.
+      Authenticity of servK for A\<close>
 lemma servK_authentic_bis_r:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
            \<in> parts (spies evs);
@@ -1168,15 +1168,15 @@
 
 
 
-subsection{*Parties authentication: each party verifies "the identity of
-       another party who generated some data" (quoted from Neuman and Ts'o).*}
+subsection\<open>Parties authentication: each party verifies "the identity of
+       another party who generated some data" (quoted from Neuman and Ts'o).\<close>
 
-text{*These guarantees don't assess whether two parties agree on
+text\<open>These guarantees don't assess whether two parties agree on
       the same session key: sending a message containing a key
-      doesn't a priori state knowledge of the key.*}
+      doesn't a priori state knowledge of the key.\<close>
 
 
-text{*These didn't have existential form in version IV*}
+text\<open>These didn't have existential form in version IV\<close>
 lemma B_authenticates_A:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
@@ -1186,7 +1186,7 @@
   \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
 by (blast dest: servTicket_authentic_Tgs intro: Says_K5)
 
-text{*The second assumption tells B what kind of key servK is.*}
+text\<open>The second assumption tells B what kind of key servK is.\<close>
 lemma B_authenticates_A_r:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
          Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
@@ -1200,8 +1200,8 @@
   \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
 by (blast intro: Says_K5 dest: Confidentiality_B servTicket_authentic_Tgs)
 
-text{* @{text u_B_authenticates_A} would be the same as @{text B_authenticates_A} because the
- servK confidentiality assumption is yet unrelaxed*}
+text\<open>\<open>u_B_authenticates_A\<close> would be the same as \<open>B_authenticates_A\<close> because the
+ servK confidentiality assumption is yet unrelaxed\<close>
 
 lemma u_B_authenticates_A_r:
      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
@@ -1243,10 +1243,10 @@
 
 
 
-subsection{*Parties' knowledge of session keys. 
+subsection\<open>Parties' knowledge of session keys. 
        An agent knows a session key if he used it to issue a cipher. These
        guarantees can be interpreted both in terms of key distribution
-       and of non-injective agreement on the session key.*}
+       and of non-injective agreement on the session key.\<close>
 
 lemma Kas_Issues_A:
    "\<lbrakk> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>, authTicket\<rbrace> \<in> set evs;
@@ -1262,7 +1262,7 @@
 apply (frule_tac [5] Says_ticket_parts)
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*K2*}
+txt\<open>K2\<close>
 apply (simp add: takeWhile_tail)
 apply (metis MPair_parts parts.Body parts_idem parts_spies_takeWhile_mono parts_trans spies_evs_rev usedI)
 done
@@ -1319,7 +1319,7 @@
 apply (erule kerbV.induct, analz_mono_contra)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
 apply blast
-txt{*K6 requires numerous lemmas*}
+txt\<open>K6 requires numerous lemmas\<close>
 apply (simp add: takeWhile_tail)
 apply (blast intro: Says_K6 dest: servTicket_authentic 
         parts_spies_takeWhile_mono [THEN subsetD] 
@@ -1364,7 +1364,7 @@
 *)
 
 
-text{*But can prove a less general fact conerning only authenticators!*}
+text\<open>But can prove a less general fact conerning only authenticators!\<close>
 lemma honest_never_says_newer_timestamp_in_auth:
      "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; A \<notin> bad; evs \<in> kerbV \<rbrakk> 
      \<Longrightarrow> Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
@@ -1394,11 +1394,11 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts)
 apply (simp_all (no_asm_simp))
-txt{*K5*}
+txt\<open>K5\<close>
 apply auto
 apply (simp add: takeWhile_tail)
-txt{*Level 15: case study necessary because the assumption doesn't state
-  the form of servTicket. The guarantee becomes stronger.*}
+txt\<open>Level 15: case study necessary because the assumption doesn't state
+  the form of servTicket. The guarantee becomes stronger.\<close>
 prefer 2 apply (simp add: takeWhile_tail)
 (**This single command of version IV...
 apply (blast dest: Says_imp_spies [THEN analz.Inj, THEN analz_Decrypt']
@@ -1415,11 +1415,11 @@
 apply (frule servK_authentic_ter, blast, assumption+)
 apply (drule parts_spies_takeWhile_mono [THEN subsetD])
 apply (drule parts_spies_evs_revD2 [THEN subsetD])
-txt{* @{term Says_K5} closes the proof in version IV because it is clear which 
-servTicket an authenticator appears with in msg 5. In version V an authenticator can appear with any item that the spy could replace the servTicket with*}
+txt\<open>@{term Says_K5} closes the proof in version IV because it is clear which 
+servTicket an authenticator appears with in msg 5. In version V an authenticator can appear with any item that the spy could replace the servTicket with\<close>
 apply (frule Says_K5, blast)
-txt{*We need to state that an honest agent wouldn't send the wrong timestamp
-within an authenticator, wathever it is paired with*}
+txt\<open>We need to state that an honest agent wouldn't send the wrong timestamp
+within an authenticator, wathever it is paired with\<close>
 apply (auto simp add: honest_never_says_current_timestamp_in_auth)
 done
 
@@ -1434,16 +1434,16 @@
 
 
 
-subsection{*
+subsection\<open>
 Novel guarantees, never studied before. Because honest agents always say
 the right timestamp in authenticators, we can prove unicity guarantees based 
 exactly on timestamps. Classical unicity guarantees are based on nonces.
 Of course assuming the agent to be different from the Spy, rather than not in 
 bad, would suffice below. Similar guarantees must also hold of
-Kerberos IV.*}
+Kerberos IV.\<close>
 
-text{*Notice that an honest agent can send the same timestamp on two
-different traces of the same length, but not on the same trace!*}
+text\<open>Notice that an honest agent can send the same timestamp on two
+different traces of the same length, but not on the same trace!\<close>
 
 lemma unique_timestamp_authenticator1:
      "\<lbrakk> Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs;
@@ -1475,8 +1475,8 @@
 apply (auto simp add: honest_never_says_current_timestamp_in_auth)
 done
 
-text{*The second part of the message is treated as an authenticator by the last
-simplification step, even if it is not an authenticator!*}
+text\<open>The second part of the message is treated as an authenticator by the last
+simplification step, even if it is not an authenticator!\<close>
 lemma unique_timestamp_authticket:
      "\<lbrakk> Says Kas A \<lbrace>X, Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key AK, T\<rbrace>\<rbrace> \<in> set evs;
        Says Kas A' \<lbrace>X', Crypt (shrK Tgs') \<lbrace>Agent A', Agent Tgs', Key AK', T\<rbrace>\<rbrace> \<in> set evs;
@@ -1487,8 +1487,8 @@
 apply (auto simp add: honest_never_says_current_timestamp_in_auth)
 done
 
-text{*The second part of the message is treated as an authenticator by the last
-simplification step, even if it is not an authenticator!*}
+text\<open>The second part of the message is treated as an authenticator by the last
+simplification step, even if it is not an authenticator!\<close>
 lemma unique_timestamp_servticket:
      "\<lbrakk> Says Tgs A \<lbrace>X, Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, T\<rbrace>\<rbrace> \<in> set evs;
        Says Tgs A' \<lbrace>X', Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SK', T\<rbrace>\<rbrace> \<in> set evs;
--- a/src/HOL/Auth/Kerberos_BAN.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,18 +3,18 @@
     Copyright   1998  University of Cambridge
 *)
 
-section{*The Kerberos Protocol, BAN Version*}
+section\<open>The Kerberos Protocol, BAN Version\<close>
 
 theory Kerberos_BAN imports Public begin
 
-text{*From page 251 of
+text\<open>From page 251 of
   Burrows, Abadi and Needham (1989).  A Logic of Authentication.
   Proc. Royal Soc. 426
 
   Confidentiality (secrecy) and authentication properties are also
   given in a termporal version: strong guarantees in a little abstracted 
   - but very realistic - model.
-*}
+\<close>
 
 (* Temporal model of accidents: session keys can be leaked
                                 ONLY when they have expired *)
@@ -27,12 +27,12 @@
     (*Duration of the authenticator*)
     authlife :: nat
 
-text{*The ticket should remain fresh for two journeys on the network at least*}
+text\<open>The ticket should remain fresh for two journeys on the network at least\<close>
 specification (sesKlife)
   sesKlife_LB [iff]: "2 \<le> sesKlife"
     by blast
 
-text{*The authenticator only for one journey*}
+text\<open>The authenticator only for one journey\<close>
 specification (authlife)
   authlife_LB [iff]:    "authlife \<noteq> 0"
     by blast
@@ -122,7 +122,7 @@
 declare analz_into_parts [dest]
 declare Fake_parts_insert_in_Un [dest]
 
-text{*A "possibility property": there are traces that reach the end.*}
+text\<open>A "possibility property": there are traces that reach the end.\<close>
 lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk>
        \<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerberos.
              Says B A (Crypt K (Number Timestamp))
@@ -135,7 +135,7 @@
 apply (possibility, simp_all (no_asm_simp) add: used_Cons)
 done
 
-subsection{*Lemmas for reasoning about predicate "Issues"*}
+subsection\<open>Lemmas for reasoning about predicate "Issues"\<close>
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
@@ -169,13 +169,13 @@
 apply (induct_tac "evs")
 apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
-txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+txt\<open>Resembles \<open>used_subset_append\<close> in theory Event.\<close>
 done
 
 lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
 
 
-text{*Lemmas for reasoning about predicate "before"*}
+text\<open>Lemmas for reasoning about predicate "before"\<close>
 lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
@@ -231,7 +231,7 @@
 
 (**** Inductive proofs about bankerberos ****)
 
-text{*Forwarding Lemma for reasoning about the encrypted portion of message BK3*}
+text\<open>Forwarding Lemma for reasoning about the encrypted portion of message BK3\<close>
 lemma BK3_msg_in_parts_spies:
      "Says S A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs
       \<Longrightarrow> X \<in> parts (spies evs)"
@@ -244,7 +244,7 @@
 apply blast
 done
 
-text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
+text\<open>Spy never sees another agent's shared key! (unless it's bad at start)\<close>
 lemma Spy_see_shrK [simp]:
      "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 apply (erule bankerberos.induct)
@@ -267,7 +267,7 @@
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D,  dest!]
 
 
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
     "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerberos\<rbrakk>
      \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
@@ -275,15 +275,15 @@
 apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
 apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert)
-txt{*BK2, BK3, BK4*}
+txt\<open>BK2, BK3, BK4\<close>
 apply (force dest!: analz_shrK_Decrypt)+
 done
 
-subsection{* Lemmas concerning the form of items passed in messages *}
+subsection\<open>Lemmas concerning the form of items passed in messages\<close>
 
-text{*Describes the form of K, X and K' when the Server sends this message.*}
+text\<open>Describes the form of K, X and K' when the Server sends this message.\<close>
 lemma Says_Server_message_form:
      "\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
          \<in> set evs; evs \<in> bankerberos \<rbrakk>
@@ -303,10 +303,10 @@
 done
 
 
-text{*If the encrypted message appears then it originated with the Server
+text\<open>If the encrypted message appears then it originated with the Server
   PROVIDED that A is NOT compromised!
   This allows A to verify freshness of the session key.
-*}
+\<close>
 lemma Kab_authentic:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
            \<in> parts (spies evs);
@@ -320,8 +320,8 @@
 done
 
 
-text{*If the TICKET appears then it originated with the Server*}
-text{*FRESHNESS OF THE SESSION KEY to B*}
+text\<open>If the TICKET appears then it originated with the Server\<close>
+text\<open>FRESHNESS OF THE SESSION KEY to B\<close>
 lemma ticket_authentic:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs);
          B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
@@ -336,9 +336,9 @@
 done
 
 
-text{*EITHER describes the form of X when the following message is sent,
+text\<open>EITHER describes the form of X when the following message is sent,
   OR     reduces it to the Fake case.
-  Use @{text Says_Server_message_form} if applicable.*}
+  Use \<open>Says_Server_message_form\<close> if applicable.\<close>
 lemma Says_S_message_form:
      "\<lbrakk> Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
             \<in> set evs;
@@ -363,7 +363,7 @@
 
 ****)
 
-text{* Session keys are not used to encrypt other session keys *}
+text\<open>Session keys are not used to encrypt other session keys\<close>
 lemma analz_image_freshK [rule_format (no_asm)]:
      "evs \<in> bankerberos \<Longrightarrow>
    \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
@@ -382,7 +382,7 @@
 apply (simp only: analz_image_freshK analz_image_freshK_simps)
 done
 
-text{* The session key K uniquely identifies the message *}
+text\<open>The session key K uniquely identifies the message\<close>
 lemma unique_session_keys:
      "\<lbrakk> Says Server A
            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;
@@ -394,7 +394,7 @@
 apply (erule bankerberos.induct)
 apply (frule_tac [7] Oops_parts_spies)
 apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all)
-txt{*BK2: it can't be a new key*}
+txt\<open>BK2: it can't be a new key\<close>
 apply blast
 done
 
@@ -409,13 +409,13 @@
 done
 
 
-subsection{*Non-temporal guarantees, explicitly relying on non-occurrence of
-oops events - refined below by temporal guarantees*}
+subsection\<open>Non-temporal guarantees, explicitly relying on non-occurrence of
+oops events - refined below by temporal guarantees\<close>
 
-text{*Non temporal treatment of confidentiality*}
+text\<open>Non temporal treatment of confidentiality\<close>
 
-text{* Lemma: the session key sent in msg BK2 would be lost by oops
-    if the spy could see it! *}
+text\<open>Lemma: the session key sent in msg BK2 would be lost by oops
+    if the spy could see it!\<close>
 lemma lemma_conf [rule_format (no_asm)]:
      "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
   \<Longrightarrow> Says Server A
@@ -427,21 +427,21 @@
 apply (frule_tac [7] Says_Server_message_form)
 apply (frule_tac [5] Says_S_message_form [THEN disjE])
 apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*BK2*}
+txt\<open>BK2\<close>
 apply (blast intro: parts_insertI)
-txt{*BK3*}
+txt\<open>BK3\<close>
 apply (case_tac "Aa \<in> bad")
  prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
 apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)
-txt{*Oops*}
+txt\<open>Oops\<close>
 apply (blast dest: unique_session_keys)
 done
 
 
-text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
-as long as they have not expired.*}
+text\<open>Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.\<close>
 lemma Confidentiality_S:
      "\<lbrakk> Says Server A
           (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
@@ -452,7 +452,7 @@
 apply (blast intro: lemma_conf)
 done
 
-text{*Confidentiality for Alice*}
+text\<open>Confidentiality for Alice\<close>
 lemma Confidentiality_A:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
         Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
@@ -461,7 +461,7 @@
 apply (blast dest!: Kab_authentic Confidentiality_S)
 done
 
-text{*Confidentiality for Bob*}
+text\<open>Confidentiality for Bob\<close>
 lemma Confidentiality_B:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
           \<in> parts (spies evs);
@@ -471,9 +471,9 @@
 apply (blast dest!: ticket_authentic Confidentiality_S)
 done
 
-text{*Non temporal treatment of authentication*}
+text\<open>Non temporal treatment of authentication\<close>
 
-text{*Lemmas @{text lemma_A} and @{text lemma_B} in fact are common to both temporal and non-temporal treatments*}
+text\<open>Lemmas \<open>lemma_A\<close> and \<open>lemma_B\<close> in fact are common to both temporal and non-temporal treatments\<close>
 lemma lemma_A [rule_format]:
      "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk>
       \<Longrightarrow>
@@ -488,11 +488,11 @@
 apply (frule_tac [5] Says_S_message_form)
 apply (frule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*BK2*}
+txt\<open>BK2\<close>
 apply (force dest: Crypt_imp_invKey_keysFor)
-txt{*BK3*}
+txt\<open>BK3\<close>
 apply (blast dest: Kab_authentic unique_session_keys)
 done
 
@@ -508,20 +508,20 @@
 apply (frule_tac [5] Says_S_message_form)
 apply (drule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*BK2*} 
+txt\<open>BK2\<close> 
 apply (force dest: Crypt_imp_invKey_keysFor)
-txt{*BK4*}
+txt\<open>BK4\<close>
 apply (blast dest: ticket_authentic unique_session_keys
                    Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad)
 done
 
 
-text{*The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+text\<open>The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.\<close>
 
 
-text{*Authentication of A to B*}
+text\<open>Authentication of A to B\<close>
 lemma B_authenticates_A_r:
      "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
@@ -535,7 +535,7 @@
 done
 
 
-text{*Authentication of B to A*}
+text\<open>Authentication of B to A\<close>
 lemma A_authenticates_B_r:
      "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
         Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
@@ -565,14 +565,14 @@
 apply (blast dest!: Kab_authentic intro!: lemma_B)
 done
 
-subsection{*Temporal guarantees, relying on a temporal check that insures that
-no oops event occurred. These are available in the sense of goal availability*}
+subsection\<open>Temporal guarantees, relying on a temporal check that insures that
+no oops event occurred. These are available in the sense of goal availability\<close>
 
 
-text{*Temporal treatment of confidentiality*}
+text\<open>Temporal treatment of confidentiality\<close>
 
-text{* Lemma: the session key sent in msg BK2 would be EXPIRED
-    if the spy could see it! *}
+text\<open>Lemma: the session key sent in msg BK2 would be EXPIRED
+    if the spy could see it!\<close>
 lemma lemma_conf_temporal [rule_format (no_asm)]:
      "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
   \<Longrightarrow> Says Server A
@@ -584,20 +584,20 @@
 apply (frule_tac [7] Says_Server_message_form)
 apply (frule_tac [5] Says_S_message_form [THEN disjE])
 apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*BK2*}
+txt\<open>BK2\<close>
 apply (blast intro: parts_insertI less_SucI)
-txt{*BK3*}
+txt\<open>BK3\<close>
 apply (metis Crypt_Spy_analz_bad Kab_authentic Says_imp_analz_Spy 
           Says_imp_parts_knows_Spy analz.Snd less_SucI unique_session_keys)
-txt{*Oops: PROOF FAILS if unsafe intro below*}
+txt\<open>Oops: PROOF FAILS if unsafe intro below\<close>
 apply (blast dest: unique_session_keys intro!: less_SucI)
 done
 
 
-text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
-as long as they have not expired.*}
+text\<open>Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.\<close>
 lemma Confidentiality_S_temporal:
      "\<lbrakk> Says Server A
           (Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs;
@@ -608,7 +608,7 @@
 apply (blast intro: lemma_conf_temporal)
 done
 
-text{*Confidentiality for Alice*}
+text\<open>Confidentiality for Alice\<close>
 lemma Confidentiality_A_temporal:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
          \<not> expiredK T evs;
@@ -617,7 +617,7 @@
 apply (blast dest!: Kab_authentic Confidentiality_S_temporal)
 done
 
-text{*Confidentiality for Bob*}
+text\<open>Confidentiality for Bob\<close>
 lemma Confidentiality_B_temporal:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
           \<in> parts (spies evs);
@@ -627,9 +627,9 @@
 apply (blast dest!: ticket_authentic Confidentiality_S_temporal)
 done
 
-text{*Temporal treatment of authentication*}
+text\<open>Temporal treatment of authentication\<close>
 
-text{*Authentication of A to B*}
+text\<open>Authentication of A to B\<close>
 lemma B_authenticates_A_temporal:
      "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
@@ -643,7 +643,7 @@
           elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
 done
 
-text{*Authentication of B to A*}
+text\<open>Authentication of B to A\<close>
 lemma A_authenticates_B_temporal:
      "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
          Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
@@ -655,10 +655,10 @@
           intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
 done
 
-subsection{*Treatment of the key distribution goal using trace inspection. All
+subsection\<open>Treatment of the key distribution goal using trace inspection. All
 guarantees are in non-temporal form, hence non available, though their temporal
 form is trivial to derive. These guarantees also convey a stronger form of 
-authentication - non-injective agreement on the session key*}
+authentication - non-injective agreement on the session key\<close>
 
 
 lemma B_Issues_A:
@@ -674,9 +674,9 @@
 apply (erule rev_mp)
 apply (erule bankerberos.induct, analz_mono_contra)
 apply (simp_all (no_asm_simp))
-txt{*fake*}
+txt\<open>fake\<close>
 apply blast
-txt{*K4 obviously is the non-trivial case*}
+txt\<open>K4 obviously is the non-trivial case\<close>
 apply (simp add: takeWhile_tail)
 apply (blast dest: ticket_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: A_authenticates_B_temporal)
 done
@@ -705,9 +705,9 @@
 apply (erule rev_mp)
 apply (erule bankerberos.induct, analz_mono_contra)
 apply (simp_all (no_asm_simp))
-txt{*fake*}
+txt\<open>fake\<close>
 apply blast
-txt{*K3 is the non trivial case*}
+txt\<open>K3 is the non trivial case\<close>
 apply (simp add: takeWhile_tail)
 apply auto (*Technically unnecessary, merely clarifies the subgoal as it is presemted in the book*)
 apply (blast dest: Kab_authentic Says_Server_message_form parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] 
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,18 +2,18 @@
     Author:     Giampaolo Bella, Catania University
 *)
 
-section{*The Kerberos Protocol, BAN Version, with Gets event*}
+section\<open>The Kerberos Protocol, BAN Version, with Gets event\<close>
 
 theory Kerberos_BAN_Gets imports Public begin
 
-text{*From page 251 of
+text\<open>From page 251 of
   Burrows, Abadi and Needham (1989).  A Logic of Authentication.
   Proc. Royal Soc. 426
 
   Confidentiality (secrecy) and authentication properties rely on
   temporal checks: strong guarantees in a little abstracted - but
   very realistic - model.
-*}
+\<close>
 
 (* Temporal modelization: session keys can be leaked
                           ONLY when they have expired *)
@@ -26,14 +26,14 @@
     (*Duration of the authenticator*)
     authlife :: nat
 
-text{*The ticket should remain fresh for two journeys on the network at least*}
-text{*The Gets event causes longer traces for the protocol to reach its end*}
+text\<open>The ticket should remain fresh for two journeys on the network at least\<close>
+text\<open>The Gets event causes longer traces for the protocol to reach its end\<close>
 specification (sesKlife)
   sesKlife_LB [iff]: "4 \<le> sesKlife"
     by blast
 
-text{*The authenticator only for one journey*}
-text{*The Gets event causes longer traces for the protocol to reach its end*}
+text\<open>The authenticator only for one journey\<close>
+text\<open>The Gets event causes longer traces for the protocol to reach its end\<close>
 specification (authlife)
   authlife_LB [iff]:    "2 \<le> authlife"
     by blast
@@ -119,7 +119,7 @@
 declare knows_Spy_partsEs [elim]
 
 
-text{*A "possibility property": there are traces that reach the end.*}
+text\<open>A "possibility property": there are traces that reach the end.\<close>
 lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk>
        \<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerb_gets.
              Says B A (Crypt K (Number Timestamp))
@@ -136,8 +136,8 @@
 done
 
 
-text{*Lemmas about reception invariant: if a message is received it certainly
-was sent*}
+text\<open>Lemmas about reception invariant: if a message is received it certainly
+was sent\<close>
 lemma Gets_imp_Says :
      "\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
 apply (erule rev_mp)
@@ -164,7 +164,7 @@
 apply (blast dest: Gets_imp_knows [THEN analz.Inj])
 done
 
-text{*Lemmas for reasoning about predicate "before"*}
+text\<open>Lemmas for reasoning about predicate "before"\<close>
 lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
@@ -220,7 +220,7 @@
 
 (**** Inductive proofs about bankerb_gets ****)
 
-text{*Forwarding Lemma for reasoning about the encrypted portion of message BK3*}
+text\<open>Forwarding Lemma for reasoning about the encrypted portion of message BK3\<close>
 lemma BK3_msg_in_parts_knows_Spy:
      "\<lbrakk>Gets A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs; evs \<in> bankerb_gets \<rbrakk> 
       \<Longrightarrow> X \<in> parts (knows Spy evs)"
@@ -234,7 +234,7 @@
 done
 
 
-text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
+text\<open>Spy never sees another agent's shared key! (unless it's bad at start)\<close>
 lemma Spy_see_shrK [simp]:
      "evs \<in> bankerb_gets \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
 apply (erule bankerb_gets.induct)
@@ -255,7 +255,7 @@
 lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D,  dest!]
 
 
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
 lemma new_keys_not_used [simp]:
     "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerb_gets\<rbrakk>
      \<Longrightarrow> K \<notin> keysFor (parts (knows Spy evs))"
@@ -263,15 +263,15 @@
 apply (erule bankerb_gets.induct)
 apply (frule_tac [8] Oops_parts_knows_Spy)
 apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert)
-txt{*BK2, BK3, BK4*}
+txt\<open>BK2, BK3, BK4\<close>
 apply (force dest!: analz_shrK_Decrypt)+
 done
 
-subsection{* Lemmas concerning the form of items passed in messages *}
+subsection\<open>Lemmas concerning the form of items passed in messages\<close>
 
-text{*Describes the form of K, X and K' when the Server sends this message.*}
+text\<open>Describes the form of K, X and K' when the Server sends this message.\<close>
 lemma Says_Server_message_form:
      "\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
          \<in> set evs; evs \<in> bankerb_gets \<rbrakk>
@@ -286,21 +286,21 @@
 apply (unfold before_def)
 apply (erule rev_mp)
 apply (erule bankerb_gets.induct, simp_all)
-txt{*We need this simplification only for Message 2*}
+txt\<open>We need this simplification only for Message 2\<close>
 apply (simp (no_asm) add: takeWhile_tail)
 apply auto
-txt{*Two subcases of Message 2. Subcase: used before*}
+txt\<open>Two subcases of Message 2. Subcase: used before\<close>
 apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
                    used_takeWhile_used)
-txt{*subcase: CT before*}
+txt\<open>subcase: CT before\<close>
 apply (fastforce dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
 done
 
 
-text{*If the encrypted message appears then it originated with the Server
+text\<open>If the encrypted message appears then it originated with the Server
   PROVIDED that A is NOT compromised!
   This allows A to verify freshness of the session key.
-*}
+\<close>
 lemma Kab_authentic:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
            \<in> parts (knows Spy evs);
@@ -314,8 +314,8 @@
 done
 
 
-text{*If the TICKET appears then it originated with the Server*}
-text{*FRESHNESS OF THE SESSION KEY to B*}
+text\<open>If the TICKET appears then it originated with the Server\<close>
+text\<open>FRESHNESS OF THE SESSION KEY to B\<close>
 lemma ticket_authentic:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (knows Spy evs);
          B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
@@ -330,9 +330,9 @@
 done
 
 
-text{*EITHER describes the form of X when the following message is sent,
+text\<open>EITHER describes the form of X when the following message is sent,
   OR     reduces it to the Fake case.
-  Use @{text Says_Server_message_form} if applicable.*}
+  Use \<open>Says_Server_message_form\<close> if applicable.\<close>
 lemma Gets_Server_message_form:
      "\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
             \<in> set evs;
@@ -345,7 +345,7 @@
 done
 
 
-text{*Reliability guarantees: honest agents act as we expect*}
+text\<open>Reliability guarantees: honest agents act as we expect\<close>
 
 lemma BK3_imp_Gets:
    "\<lbrakk> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;
@@ -394,7 +394,7 @@
 ****)
 
 
-text{* Session keys are not used to encrypt other session keys *}
+text\<open>Session keys are not used to encrypt other session keys\<close>
 lemma analz_image_freshK [rule_format (no_asm)]:
      "evs \<in> bankerb_gets \<Longrightarrow>
    \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
@@ -413,7 +413,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{* The session key K uniquely identifies the message *}
+text\<open>The session key K uniquely identifies the message\<close>
 lemma unique_session_keys:
      "\<lbrakk> Says Server A
            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;
@@ -425,7 +425,7 @@
 apply (erule bankerb_gets.induct)
 apply (frule_tac [8] Oops_parts_knows_Spy)
 apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)
-txt{*BK2: it can't be a new key*}
+txt\<open>BK2: it can't be a new key\<close>
 apply blast
 done
 
@@ -451,13 +451,13 @@
 
 
 
-subsection{*Non-temporal guarantees, explicitly relying on non-occurrence of
-oops events - refined below by temporal guarantees*}
+subsection\<open>Non-temporal guarantees, explicitly relying on non-occurrence of
+oops events - refined below by temporal guarantees\<close>
 
-text{*Non temporal treatment of confidentiality*}
+text\<open>Non temporal treatment of confidentiality\<close>
 
-text{* Lemma: the session key sent in msg BK2 would be lost by oops
-    if the spy could see it! *}
+text\<open>Lemma: the session key sent in msg BK2 would be lost by oops
+    if the spy could see it!\<close>
 lemma lemma_conf [rule_format (no_asm)]:
      "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
   \<Longrightarrow> Says Server A
@@ -469,21 +469,21 @@
 apply (frule_tac [8] Says_Server_message_form)
 apply (frule_tac [6] Gets_Server_message_form [THEN disjE])
 apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*BK2*}
+txt\<open>BK2\<close>
 apply (blast intro: parts_insertI)
-txt{*BK3*}
+txt\<open>BK3\<close>
 apply (case_tac "Aa \<in> bad")
  prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
 apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)
-txt{*Oops*}
+txt\<open>Oops\<close>
 apply (blast dest: unique_session_keys)
 done
 
 
-text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
-as long as they have not expired.*}
+text\<open>Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.\<close>
 lemma Confidentiality_S:
      "\<lbrakk> Says Server A
           (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
@@ -494,7 +494,7 @@
 apply (blast intro: lemma_conf)
 done
 
-text{*Confidentiality for Alice*}
+text\<open>Confidentiality for Alice\<close>
 lemma Confidentiality_A:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);
         Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
@@ -502,7 +502,7 @@
       \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest!: Kab_authentic Confidentiality_S)
 
-text{*Confidentiality for Bob*}
+text\<open>Confidentiality for Bob\<close>
 lemma Confidentiality_B:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
           \<in> parts (knows Spy evs);
@@ -512,9 +512,9 @@
 by (blast dest!: ticket_authentic Confidentiality_S)
 
 
-text{*Non temporal treatment of authentication*}
+text\<open>Non temporal treatment of authentication\<close>
 
-text{*Lemmas @{text lemma_A} and @{text lemma_B} in fact are common to both temporal and non-temporal treatments*}
+text\<open>Lemmas \<open>lemma_A\<close> and \<open>lemma_B\<close> in fact are common to both temporal and non-temporal treatments\<close>
 lemma lemma_A [rule_format]:
      "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
       \<Longrightarrow>
@@ -529,11 +529,11 @@
 apply (frule_tac [6] Gets_Server_message_form)
 apply (frule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*BK2*}
+txt\<open>BK2\<close>
 apply (force dest: Crypt_imp_invKey_keysFor)
-txt{*BK3*}
+txt\<open>BK3\<close>
 apply (blast dest: Kab_authentic unique_session_keys)
 done
 lemma lemma_B [rule_format]:
@@ -548,19 +548,19 @@
 apply (frule_tac [6] Gets_Server_message_form)
 apply (drule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)
 apply (simp_all (no_asm_simp) add: all_conj_distrib)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*BK2*} 
+txt\<open>BK2\<close> 
 apply (force dest: Crypt_imp_invKey_keysFor)
-txt{*BK4*}
+txt\<open>BK4\<close>
 apply (blast dest: ticket_authentic unique_session_keys
                    Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad)
 done
 
 
-text{*The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
+text\<open>The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.\<close>
 
-text{*Authentication of A to B*}
+text\<open>Authentication of A to B\<close>
 lemma B_authenticates_A_r:
      "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs);
          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (knows Spy evs);
@@ -572,7 +572,7 @@
           intro!: lemma_A
           elim!: Confidentiality_S [THEN [2] rev_notE])
 
-text{*Authentication of B to A*}
+text\<open>Authentication of B to A\<close>
 lemma A_authenticates_B_r:
      "\<lbrakk> Crypt K (Number Ta) \<in> parts (knows Spy evs);
         Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);
@@ -602,14 +602,14 @@
 done
 
 
-subsection{*Temporal guarantees, relying on a temporal check that insures that
-no oops event occurred. These are available in the sense of goal availability*}
+subsection\<open>Temporal guarantees, relying on a temporal check that insures that
+no oops event occurred. These are available in the sense of goal availability\<close>
 
 
-text{*Temporal treatment of confidentiality*}
+text\<open>Temporal treatment of confidentiality\<close>
 
-text{* Lemma: the session key sent in msg BK2 would be EXPIRED
-    if the spy could see it! *}
+text\<open>Lemma: the session key sent in msg BK2 would be EXPIRED
+    if the spy could see it!\<close>
 lemma lemma_conf_temporal [rule_format (no_asm)]:
      "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerb_gets \<rbrakk>
   \<Longrightarrow> Says Server A
@@ -621,21 +621,21 @@
 apply (frule_tac [8] Says_Server_message_form)
 apply (frule_tac [6] Gets_Server_message_form [THEN disjE])
 apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*BK2*}
+txt\<open>BK2\<close>
 apply (blast intro: parts_insertI less_SucI)
-txt{*BK3*}
+txt\<open>BK3\<close>
 apply (case_tac "Aa \<in> bad")
  prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
 apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)
-txt{*Oops: PROOF FAILS if unsafe intro below*}
+txt\<open>Oops: PROOF FAILS if unsafe intro below\<close>
 apply (blast dest: unique_session_keys intro!: less_SucI)
 done
 
 
-text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
-as long as they have not expired.*}
+text\<open>Confidentiality for the Server: Spy does not see the keys sent in msg BK2
+as long as they have not expired.\<close>
 lemma Confidentiality_S_temporal:
      "\<lbrakk> Says Server A
           (Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs;
@@ -646,7 +646,7 @@
 apply (blast intro: lemma_conf_temporal)
 done
 
-text{*Confidentiality for Alice*}
+text\<open>Confidentiality for Alice\<close>
 lemma Confidentiality_A_temporal:
      "\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);
          \<not> expiredK T evs;
@@ -654,7 +654,7 @@
       \<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
 by (blast dest!: Kab_authentic Confidentiality_S_temporal)
 
-text{*Confidentiality for Bob*}
+text\<open>Confidentiality for Bob\<close>
 lemma Confidentiality_B_temporal:
      "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
           \<in> parts (knows Spy evs);
@@ -664,9 +664,9 @@
 by (blast dest!: ticket_authentic Confidentiality_S_temporal)
 
 
-text{*Temporal treatment of authentication*}
+text\<open>Temporal treatment of authentication\<close>
 
-text{*Authentication of A to B*}
+text\<open>Authentication of A to B\<close>
 lemma B_authenticates_A_temporal:
      "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs);
          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
@@ -679,7 +679,7 @@
           intro!: lemma_A
           elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
 
-text{*Authentication of B to A*}
+text\<open>Authentication of B to A\<close>
 lemma A_authenticates_B_temporal:
      "\<lbrakk> Crypt K (Number Ta) \<in> parts (knows Spy evs);
          Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
@@ -691,7 +691,7 @@
           intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
 
 
-subsection{*Combined guarantees of key distribution and non-injective agreement on the session keys*}
+subsection\<open>Combined guarantees of key distribution and non-injective agreement on the session keys\<close>
 
 lemma B_authenticates_and_keydist_to_A:
      "\<lbrakk> Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
--- a/src/HOL/Auth/Message.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Message.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -6,7 +6,7 @@
 Inductive relations "parts", "analz" and "synth"
 *)
 
-section{*Theory of Agents and Messages for Security Protocols*}
+section\<open>Theory of Agents and Messages for Security Protocols\<close>
 
 theory Message
 imports Main
@@ -20,8 +20,8 @@
   key = nat
 
 consts
-  all_symmetric :: bool        --{*true if all keys are symmetric*}
-  invKey        :: "key=>key"  --{*inverse of a symmetric key*}
+  all_symmetric :: bool        \<comment>\<open>true if all keys are symmetric\<close>
+  invKey        :: "key=>key"  \<comment>\<open>inverse of a symmetric key\<close>
 
 specification (invKey)
   invKey [simp]: "invKey (invKey K) = K"
@@ -29,26 +29,26 @@
     by (rule exI [of _ id], auto)
 
 
-text{*The inverse of a symmetric key is itself; that of a public key
-      is the private key and vice versa*}
+text\<open>The inverse of a symmetric key is itself; that of a public key
+      is the private key and vice versa\<close>
 
 definition symKeys :: "key set" where
   "symKeys == {K. invKey K = K}"
 
-datatype  --{*We allow any number of friendly agents*}
+datatype  \<comment>\<open>We allow any number of friendly agents\<close>
   agent = Server | Friend nat | Spy
 
 datatype
-     msg = Agent  agent     --{*Agent names*}
-         | Number nat       --{*Ordinary integers, timestamps, ...*}
-         | Nonce  nat       --{*Unguessable nonces*}
-         | Key    key       --{*Crypto keys*}
-         | Hash   msg       --{*Hashing*}
-         | MPair  msg msg   --{*Compound messages*}
-         | Crypt  key msg   --{*Encryption, public- or shared-key*}
+     msg = Agent  agent     \<comment>\<open>Agent names\<close>
+         | Number nat       \<comment>\<open>Ordinary integers, timestamps, ...\<close>
+         | Nonce  nat       \<comment>\<open>Unguessable nonces\<close>
+         | Key    key       \<comment>\<open>Crypto keys\<close>
+         | Hash   msg       \<comment>\<open>Hashing\<close>
+         | MPair  msg msg   \<comment>\<open>Compound messages\<close>
+         | Crypt  key msg   \<comment>\<open>Encryption, public- or shared-key\<close>
 
 
-text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
+text\<open>Concrete syntax: messages appear as {|A,B,NA|}, etc...\<close>
 syntax
   "_MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
 
@@ -61,15 +61,15 @@
 
 
 definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
-    --{*Message Y paired with a MAC computed with the help of X*}
+    \<comment>\<open>Message Y paired with a MAC computed with the help of X\<close>
     "Hash[X] Y == {| Hash{|X,Y|}, Y|}"
 
 definition keysFor :: "msg set => key set" where
-    --{*Keys useful to decrypt elements of a message set*}
+    \<comment>\<open>Keys useful to decrypt elements of a message set\<close>
   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
 
 
-subsubsection{*Inductive Definition of All Parts" of a Message*}
+subsubsection\<open>Inductive Definition of All Parts" of a Message\<close>
 
 inductive_set
   parts :: "msg set => msg set"
@@ -81,7 +81,7 @@
   | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
 
 
-text{*Monotonicity*}
+text\<open>Monotonicity\<close>
 lemma parts_mono: "G \<subseteq> H ==> parts(G) \<subseteq> parts(H)"
 apply auto
 apply (erule parts.induct) 
@@ -89,7 +89,7 @@
 done
 
 
-text{*Equations hold because constructors are injective.*}
+text\<open>Equations hold because constructors are injective.\<close>
 lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
 by auto
 
@@ -100,13 +100,13 @@
 by auto
 
 
-subsubsection{*Inverse of keys *}
+subsubsection\<open>Inverse of keys\<close>
 
 lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
 by (metis invKey)
 
 
-subsection{*keysFor operator*}
+subsection\<open>keysFor operator\<close>
 
 lemma keysFor_empty [simp]: "keysFor {} = {}"
 by (unfold keysFor_def, blast)
@@ -117,7 +117,7 @@
 lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
 by (unfold keysFor_def, blast)
 
-text{*Monotonicity*}
+text\<open>Monotonicity\<close>
 lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
 by (unfold keysFor_def, blast)
 
@@ -150,7 +150,7 @@
 by (unfold keysFor_def, blast)
 
 
-subsection{*Inductive relation "parts"*}
+subsection\<open>Inductive relation "parts"\<close>
 
 lemma MPair_parts:
      "[| {|X,Y|} \<in> parts H;        
@@ -158,10 +158,10 @@
 by (blast dest: parts.Fst parts.Snd) 
 
 declare MPair_parts [elim!]  parts.Body [dest!]
-text{*NB These two rules are UNSAFE in the formal sense, as they discard the
+text\<open>NB These two rules are UNSAFE in the formal sense, as they discard the
      compound message.  They work well on THIS FILE.  
-  @{text MPair_parts} is left as SAFE because it speeds up proofs.
-  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
+  \<open>MPair_parts\<close> is left as SAFE because it speeds up proofs.
+  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\<close>
 
 lemma parts_increasing: "H \<subseteq> parts(H)"
 by blast
@@ -176,12 +176,12 @@
 lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
 by simp
 
-text{*WARNING: loops if H = {Y}, therefore must not be repeated!*}
+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+)
 
 
-subsubsection{*Unions *}
+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)
@@ -197,8 +197,8 @@
 lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
 by (metis insert_is_Un parts_Un)
 
-text{*TWO inserts to avoid looping.  This rewrite is better than nothing.
-  But its behaviour can be strange.*}
+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)
@@ -214,12 +214,12 @@
 lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
 by (intro equalityI parts_UN_subset1 parts_UN_subset2)
 
-text{*Added to simplify arguments to parts, analz and synth.
-  NOTE: the UN versions are no longer used!*}
+text\<open>Added to simplify arguments to parts, analz and synth.
+  NOTE: the UN versions are no longer used!\<close>
 
 
-text{*This allows @{text blast} to simplify occurrences of 
-  @{term "parts(G\<union>H)"} in the assumption.*}
+text\<open>This allows \<open>blast\<close> to simplify occurrences of 
+  @{term "parts(G\<union>H)"} in the assumption.\<close>
 lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] 
 declare in_parts_UnE [elim!]
 
@@ -227,7 +227,7 @@
 lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
 by (blast intro: parts_mono [THEN [2] rev_subsetD])
 
-subsubsection{*Idempotence and transitivity *}
+subsubsection\<open>Idempotence and transitivity\<close>
 
 lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
 by (erule parts.induct, blast+)
@@ -241,7 +241,7 @@
 lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
 by (metis parts_subset_iff set_mp)
 
-text{*Cut*}
+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) 
@@ -250,7 +250,7 @@
 by (metis insert_absorb parts_idem parts_insert)
 
 
-subsubsection{*Rewrite rules for pulling out atomic messages *}
+subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
 
 lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
 
@@ -308,7 +308,7 @@
 done
 
 
-text{*In any message, there is an upper bound N on its greatest nonce.*}
+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 --> Nonce n \<notin> parts {msg}"
 proof (induct msg)
   case (Nonce n)
@@ -316,15 +316,15 @@
       by simp (metis Suc_n_not_le_n)
 next
   case (MPair X Y)
-    then show ?case --{*metis works out the necessary sum itself!*}
+    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{*Inductive relation "analz"*}
+subsection\<open>Inductive relation "analz"\<close>
 
-text{*Inductive definition of "analz" -- what can be broken down from a set of
+text\<open>Inductive definition of "analz" -- what can be broken down from a set of
     messages, including keys.  A form of downward closure.  Pairs can
-    be taken apart; messages decrypted with known keys.  *}
+    be taken apart; messages decrypted with known keys.\<close>
 
 inductive_set
   analz :: "msg set => msg set"
@@ -337,14 +337,14 @@
              "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
 
 
-text{*Monotonicity; Lemma 1 of Lowe's paper*}
+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
 
-text{*Making it safe speeds up proofs*}
+text\<open>Making it safe speeds up proofs\<close>
 lemma MPair_analz [elim!]:
      "[| {|X,Y|} \<in> analz H;        
              [| X \<in> analz H; Y \<in> analz H |] ==> P   
@@ -374,22 +374,22 @@
 
 lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
 
-subsubsection{*General equational properties *}
+subsubsection\<open>General equational properties\<close>
 
 lemma analz_empty [simp]: "analz{} = {}"
 apply safe
 apply (erule analz.induct, blast+)
 done
 
-text{*Converse fails: we can analz more from the union than from the 
-  separate parts, as a key in one might decrypt a message in the other*}
+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)
 
 lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
-subsubsection{*Rewrite rules for pulling out atomic messages *}
+subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
 
 lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
 
@@ -417,7 +417,7 @@
 apply (erule analz.induct, auto) 
 done
 
-text{*Can only pull out Keys if they are not needed to decrypt the rest*}
+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) ==>   
           analz (insert (Key K) H) = insert (Key K) (analz H)"
@@ -436,7 +436,7 @@
 apply (blast intro: analz.Fst analz.Snd)+
 done
 
-text{*Can pull out enCrypted message if the Key is not known*}
+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)"
@@ -466,10 +466,10 @@
                insert (Crypt K X) (analz (insert X H))"
 by (intro equalityI lemma1 lemma2)
 
-text{*Case analysis: either the message is secure, or it is not! Effective,
-but can cause subgoals to blow up! Use with @{text "split_if"}; apparently
-@{text "split_tac"} does not cope with patterns such as @{term"analz (insert
-(Crypt K X) H)"} *} 
+text\<open>Case analysis: either the message is secure, or it is not! Effective,
+but can cause subgoals to blow up! Use with \<open>split_if\<close>; apparently
+\<open>split_tac\<close> does not cope with patterns such as @{term"analz (insert
+(Crypt K X) H)"}\<close> 
 lemma analz_Crypt_if [simp]:
      "analz (insert (Crypt K X) H) =                 
           (if (Key (invKey K) \<in> analz H)                 
@@ -478,7 +478,7 @@
 by (simp add: analz_insert_Crypt analz_insert_Decrypt)
 
 
-text{*This rule supposes "for the sake of argument" that we have the key.*}
+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>   
            insert (Crypt K X) (analz (insert X H))"
@@ -493,7 +493,7 @@
 done
 
 
-subsubsection{*Idempotence and transitivity *}
+subsubsection\<open>Idempotence and transitivity\<close>
 
 lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
 by (erule analz.induct, blast+)
@@ -507,7 +507,7 @@
 lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
 by (drule analz_mono, blast)
 
-text{*Cut; Lemma 2 of Lowe*}
+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)
 
@@ -515,14 +515,14 @@
    "Y: analz (insert X H) ==> X: analz H --> Y: analz H"
 *)
 
-text{*This rewrite rule helps in the simplification of messages that involve
+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. *}
+  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)
 
 
-text{*A congruence rule for "analz" *}
+text\<open>A congruence rule for "analz"\<close>
 
 lemma analz_subset_cong:
      "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |] 
@@ -538,14 +538,14 @@
      "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
 by (force simp only: insert_def intro!: analz_cong)
 
-text{*If there are no pairs or encryptions then analz does nothing*}
+text\<open>If there are no pairs or encryptions then analz does nothing\<close>
 lemma analz_trivial:
      "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
 apply safe
 apply (erule analz.induct, blast+)
 done
 
-text{*These two are obsolete (with a single Spy) but cost little to prove...*}
+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)
@@ -556,12 +556,12 @@
 by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
 
 
-subsection{*Inductive relation "synth"*}
+subsection\<open>Inductive relation "synth"\<close>
 
-text{*Inductive definition of "synth" -- what can be built up from a set of
+text\<open>Inductive definition of "synth" -- what can be built up from a set of
     messages.  A form of upward closure.  Pairs can be built, messages
     encrypted with known keys.  Agent names are public domain.
-    Numbers can be guessed, but Nonces cannot be.  *}
+    Numbers can be guessed, but Nonces cannot be.\<close>
 
 inductive_set
   synth :: "msg set => msg set"
@@ -574,12 +574,12 @@
   | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
   | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
 
-text{*Monotonicity*}
+text\<open>Monotonicity\<close>
 lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
   by (auto, erule synth.induct, auto)  
 
-text{*NO @{text Agent_synth}, as any Agent name can be synthesized.  
-  The same holds for @{term Number}*}
+text\<open>NO \<open>Agent_synth\<close>, as any Agent name can be synthesized.  
+  The same holds for @{term Number}\<close>
 
 inductive_simps synth_simps [iff]:
  "Nonce n \<in> synth H"
@@ -591,17 +591,17 @@
 lemma synth_increasing: "H \<subseteq> synth(H)"
 by blast
 
-subsubsection{*Unions *}
+subsubsection\<open>Unions\<close>
 
-text{*Converse fails: we can synth more from the union than from the 
-  separate parts, building a compound message using elements of each.*}
+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)
 
 lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
 by (blast intro: synth_mono [THEN [2] rev_subsetD])
 
-subsubsection{*Idempotence and transitivity *}
+subsubsection\<open>Idempotence and transitivity\<close>
 
 lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
 by (erule synth.induct, auto)
@@ -615,7 +615,7 @@
 lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
 by (drule synth_mono, blast)
 
-text{*Cut; Lemma 2 of Lowe*}
+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)
 
@@ -629,7 +629,7 @@
 by (unfold keysFor_def, blast)
 
 
-subsubsection{*Combinations of parts, analz and synth *}
+subsubsection\<open>Combinations of parts, analz and synth\<close>
 
 lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
 apply (rule equalityI)
@@ -656,12 +656,12 @@
 by (metis Un_empty_right analz_synth_Un)
 
 
-subsubsection{*For reasoning about the Fake rule in traces *}
+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)
 
-text{*More specifically for Fake. See also @{text Fake_parts_sing} below *}
+text\<open>More specifically for Fake. See also \<open>Fake_parts_sing\<close> below\<close>
 lemma Fake_parts_insert:
      "X \<in> synth (analz H) ==>  
       parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
@@ -673,8 +673,8 @@
       ==> Z \<in>  synth (analz H) \<union> parts H"
 by (metis Fake_parts_insert set_mp)
 
-text{*@{term H} is sometimes @{term"Key ` KK \<union> spies evs"}, so can't put 
-  @{term "G=H"}.*}
+text\<open>@{term H} is sometimes @{term"Key ` KK \<union> spies evs"}, so can't put 
+  @{term "G=H"}.\<close>
 lemma Fake_analz_insert:
      "X\<in> synth (analz G) ==>  
       analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
@@ -691,8 +691,8 @@
      "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
 by (blast intro: analz_subset_parts [THEN subsetD])
 
-text{*Without this equation, other rules for synth and analz would yield
-  redundant cases*}
+text\<open>Without this equation, other rules for synth and analz would yield
+  redundant cases\<close>
 lemma MPair_synth_analz [iff]:
      "({|X,Y|} \<in> synth (analz H)) =  
       (X \<in> synth (analz H) & Y \<in> synth (analz H))"
@@ -710,9 +710,9 @@
 by blast
 
 
-subsection{*HPair: a combination of Hash and MPair*}
+subsection\<open>HPair: a combination of Hash and MPair\<close>
 
-subsubsection{*Freeness *}
+subsubsection\<open>Freeness\<close>
 
 lemma Agent_neq_HPair: "Agent A ~= Hash[X] Y"
   unfolding HPair_def by simp
@@ -750,7 +750,7 @@
 by (auto simp add: HPair_def)
 
 
-subsubsection{*Specialized laws, proved in terms of those for Hash and MPair *}
+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)
@@ -772,12 +772,12 @@
 by (auto simp add: HPair_def)
 
 
-text{*We do NOT want Crypt... messages broken up in protocols!!*}
+text\<open>We do NOT want Crypt... messages broken up in protocols!!\<close>
 declare parts.Body [rule del]
 
 
-text{*Rewrites to push in Key and Crypt messages, so that other messages can
-    be pulled out using the @{text analz_insert} rules*}
+text\<open>Rewrites to push in Key and Crypt messages, so that other messages can
+    be pulled out using the \<open>analz_insert\<close> rules\<close>
 
 lemmas pushKeys =
   insert_commute [of "Key K" "Agent C"]
@@ -797,12 +797,12 @@
   insert_commute [of "Crypt X K" "MPair X' Y"]
   for X K C N X' Y
 
-text{*Cannot be added with @{text "[simp]"} -- messages should not always be
-  re-ordered. *}
+text\<open>Cannot be added with \<open>[simp]\<close> -- messages should not always be
+  re-ordered.\<close>
 lemmas pushes = pushKeys pushCrypts
 
 
-subsection{*The set of key-free messages*}
+subsection\<open>The set of key-free messages\<close>
 
 (*Note that even the encryption of a key-free message remains key-free.
   This concept is valuable because of the theorem analz_keyfree_into_Un, proved below. *)
@@ -833,9 +833,9 @@
 apply (metis Un_absorb2 keyfree_KeyE parts_Un parts_keyfree UnI2)
 done
 
-subsection{*Tactics useful for many protocol proofs*}
+subsection\<open>Tactics useful for many protocol proofs\<close>
 ML
-{*
+\<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 *)
@@ -872,11 +872,11 @@
        simp_tac ctxt 1,
        REPEAT (FIRSTGOAL (resolve_tac ctxt [allI,impI,notI,conjI,iffI])),
        DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i);
-*}
+\<close>
 
-text{*By default only @{text o_apply} is built-in.  But in the presence of
+text\<open>By default only \<open>o_apply\<close> is built-in.  But in the presence of
 eta-expansion this means that some terms displayed as @{term "f o g"} will be
-rewritten, and others will not!*}
+rewritten, and others will not!\<close>
 declare o_def [simp]
 
 
@@ -894,7 +894,7 @@
 by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute 
           subset_insertI synth_analz_mono synth_increasing synth_subset_iff)
 
-text{*Two generalizations of @{text analz_insert_eq}*}
+text\<open>Two generalizations of \<open>analz_insert_eq\<close>\<close>
 lemma gen_analz_insert_eq [rule_format]:
      "X \<in> analz H ==> ALL G. H \<subseteq> G --> analz (insert X G) = analz G"
 by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
@@ -912,16 +912,16 @@
 
 lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
 
-method_setup spy_analz = {*
-    Scan.succeed (SIMPLE_METHOD' o spy_analz_tac) *}
+method_setup spy_analz = \<open>
+    Scan.succeed (SIMPLE_METHOD' o spy_analz_tac)\<close>
     "for proving the Fake case when analz is involved"
 
-method_setup atomic_spy_analz = {*
-    Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac) *}
+method_setup atomic_spy_analz = \<open>
+    Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac)\<close>
     "for debugging spy_analz"
 
-method_setup Fake_insert_simp = {*
-    Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac) *}
+method_setup Fake_insert_simp = \<open>
+    Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac)\<close>
     "for debugging spy_analz"
 
 end
--- a/src/HOL/Auth/NS_Public.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/NS_Public.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -6,7 +6,7 @@
 Version incorporating Lowe's fix (inclusion of B's identity in round 2).
 *)
 
-section{*Verifying the Needham-Schroeder-Lowe Public-Key Protocol*}
+section\<open>Verifying the Needham-Schroeder-Lowe Public-Key Protocol\<close>
 
 theory NS_Public imports Public begin
 
@@ -63,7 +63,7 @@
       "evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> analz (spies evs)) = (A \<in> bad)"
 by auto
 
-subsection{*Authenticity properties obtained from NS2*}
+subsection\<open>Authenticity properties obtained from NS2\<close>
 
 
 (*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
@@ -135,7 +135,7 @@
 done
 
 
-subsection{*Authenticity properties obtained from NS2*}
+subsection\<open>Authenticity properties obtained from NS2\<close>
 
 (*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
   [unicity of B makes Lowe's fix work]
@@ -180,7 +180,7 @@
       \<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"
 by (blast intro: B_trusts_NS3_lemma)
 
-subsection{*Overall guarantee for B*}
+subsection\<open>Overall guarantee for B\<close>
 
 (*If NS3 has been sent and the nonce NB agrees with the nonce B joined with
   NA, then A initiated the run using NA.*)
--- a/src/HOL/Auth/NS_Public_Bad.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/NS_Public_Bad.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -10,7 +10,7 @@
   Proc. Royal Soc. 426 (1989)
 *)
 
-section{*Verifying the Needham-Schroeder Public-Key Protocol*}
+section\<open>Verifying the Needham-Schroeder Public-Key Protocol\<close>
 
 theory NS_Public_Bad imports Public begin
 
@@ -203,9 +203,9 @@
 apply (frule_tac A' = A in 
        Says_imp_knows_Spy [THEN parts.Inj, THEN unique_NB], auto)
 apply (rename_tac evs3 B' C)
-txt{*This is the attack!
+txt\<open>This is the attack!
 @{subgoals[display,indent=0,margin=65]}
-*}
+\<close>
 oops
 
 (*
--- a/src/HOL/Auth/NS_Shared.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/NS_Shared.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,15 +3,15 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*Needham-Schroeder Shared-Key Protocol and the Issues Property*}
+section\<open>Needham-Schroeder Shared-Key Protocol and the Issues Property\<close>
 
 theory NS_Shared imports Public begin
 
-text{*
+text\<open>
 From page 247 of
   Burrows, Abadi and Needham (1989).  A Logic of Authentication.
   Proc. Royal Soc. 426
-*}
+\<close>
 
 definition
  (* A is the true creator of X if she has sent X and X never appeared on
@@ -88,7 +88,7 @@
 declare analz_into_parts [dest]
 
 
-text{*A "possibility property": there are traces that reach the end*}
+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.
                     Says A B (Crypt K \<lbrace>Nonce N, Nonce N\<rbrace>) \<in> set evs"
@@ -105,25 +105,25 @@
 *)
 
 
-subsection{*Inductive proofs about @{term ns_shared}*}
+subsection\<open>Inductive proofs about @{term ns_shared}\<close>
 
-subsubsection{*Forwarding lemmas, to aid simplification*}
+subsubsection\<open>Forwarding lemmas, to aid simplification\<close>
 
-text{*For reasoning about the encrypted portion of message NS3*}
+text\<open>For reasoning about the encrypted portion of message NS3\<close>
 lemma NS3_msg_in_parts_spies:
      "Says S A (Crypt KA \<lbrace>N, B, K, X\<rbrace>) \<in> set evs \<Longrightarrow> X \<in> parts (spies evs)"
 by blast
 
-text{*For reasoning about the Oops message*}
+text\<open>For reasoning about the Oops message\<close>
 lemma Oops_parts_spies:
      "Says Server A (Crypt (shrK A) \<lbrace>NA, B, K, X\<rbrace>) \<in> set evs
             \<Longrightarrow> K \<in> parts (spies evs)"
 by blast
 
-text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that NOBODY
-    sends messages containing @{term X}*}
+text\<open>Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that NOBODY
+    sends messages containing @{term X}\<close>
 
-text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
+text\<open>Spy never sees another agent's shared key! (unless it's bad at start)\<close>
 lemma Spy_see_shrK [simp]:
      "evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all, blast+)
@@ -134,20 +134,20 @@
 by auto
 
 
-text{*Nobody can have used non-existent keys!*}
+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))"
 apply (erule rev_mp)
 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all)
-txt{*Fake, NS2, NS4, NS5*}
+txt\<open>Fake, NS2, NS4, NS5\<close>
 apply (force dest!: keysFor_parts_insert, blast+)
 done
 
 
-subsubsection{*Lemmas concerning the form of items passed in messages*}
+subsubsection\<open>Lemmas concerning the form of items passed in messages\<close>
 
-text{*Describes the form of K, X and K' when the Server sends this message.*}
+text\<open>Describes the form of K, X and K' when the Server sends this message.\<close>
 lemma Says_Server_message_form:
      "\<lbrakk>Says Server A (Crypt K' \<lbrace>N, Agent B, Key K, X\<rbrace>) \<in> set evs;
        evs \<in> ns_shared\<rbrakk>
@@ -157,7 +157,7 @@
 by (erule rev_mp, erule ns_shared.induct, auto)
 
 
-text{*If the encrypted message appears then it originated with the Server*}
+text\<open>If the encrypted message appears then it originated with the Server\<close>
 lemma A_trusts_NS2:
      "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
        A \<notin> bad;  evs \<in> ns_shared\<rbrakk>
@@ -172,9 +172,9 @@
       \<Longrightarrow> K \<notin> range shrK \<and>  X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)"
 by (blast dest!: A_trusts_NS2 Says_Server_message_form)
 
-text{*EITHER describes the form of X when the following message is sent,
+text\<open>EITHER describes the form of X when the following message is sent,
   OR     reduces it to the Fake case.
-  Use @{text Says_Server_message_form} if applicable.*}
+  Use \<open>Says_Server_message_form\<close> if applicable.\<close>
 lemma Says_S_message_form:
      "\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
        evs \<in> ns_shared\<rbrakk>
@@ -204,23 +204,23 @@
  A more general formula must be proved inductively.
 ****)
 
-text{*NOT useful in this form, but it says that session keys are not used
+text\<open>NOT useful in this form, but it says that session keys are not used
   to encrypt messages containing other keys, in the actual protocol.
-  We require that agents should behave like this subsequently also.*}
+  We require that agents should behave like this subsequently also.\<close>
 lemma  "\<lbrakk>evs \<in> ns_shared;  Kab \<notin> range shrK\<rbrakk> \<Longrightarrow>
          (Crypt KAB X) \<in> parts (spies evs) \<and>
          Key K \<in> parts {X} \<longrightarrow> Key K \<in> parts (spies evs)"
 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest: parts_insert_subset_Un)
-txt{*Base, NS4 and NS5*}
+txt\<open>Base, NS4 and NS5\<close>
 apply auto
 done
 
 
-subsubsection{*Session keys are not used to encrypt other session keys*}
+subsubsection\<open>Session keys are not used to encrypt other session keys\<close>
 
-text{*The equality makes the induction hypothesis easier to apply*}
+text\<open>The equality makes the induction hypothesis easier to apply\<close>
 
 lemma analz_image_freshK [rule_format]:
  "evs \<in> ns_shared \<Longrightarrow>
@@ -230,7 +230,7 @@
 apply (erule ns_shared.induct)
 apply (drule_tac [8] Says_Server_message_form)
 apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz)
-txt{*NS2, NS3*}
+txt\<open>NS2, NS3\<close>
 apply blast+ 
 done
 
@@ -242,9 +242,9 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-subsubsection{*The session key K uniquely identifies the message*}
+subsubsection\<open>The session key K uniquely identifies the message\<close>
 
-text{*In messages of this form, the session key uniquely identifies the rest*}
+text\<open>In messages of this form, the session key uniquely identifies the rest\<close>
 lemma unique_session_keys:
      "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
        Says Server A' (Crypt (shrK A') \<lbrace>NA', Agent B', Key K, X'\<rbrace>) \<in> set evs;
@@ -252,9 +252,9 @@
 by (erule rev_mp, erule rev_mp, erule ns_shared.induct, simp_all, blast+)
 
 
-subsubsection{*Crucial secrecy property: Spy doesn't see the keys sent in NS2*}
+subsubsection\<open>Crucial secrecy property: Spy doesn't see the keys sent in NS2\<close>
 
-text{*Beware of @{text "[rule_format]"} and the universal quantifier!*}
+text\<open>Beware of \<open>[rule_format]\<close> and the universal quantifier!\<close>
 lemma secrecy_lemma:
      "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
                                       Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>)
@@ -268,18 +268,18 @@
 apply (frule_tac [4] Says_S_message_form)
 apply (erule_tac [5] disjE)
 apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs, spy_analz)
-txt{*NS2*}
+txt\<open>NS2\<close>
 apply blast
-txt{*NS3*}
+txt\<open>NS3\<close>
 apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
              dest:  Says_imp_knows_Spy analz.Inj unique_session_keys)
-txt{*Oops*}
+txt\<open>Oops\<close>
 apply (blast dest: unique_session_keys)
 done
 
 
 
-text{*Final version: Server's message in the most abstract form*}
+text\<open>Final version: Server's message in the most abstract form\<close>
 lemma Spy_not_see_encrypted_key:
      "\<lbrakk>Says Server A (Crypt K' \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
        \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
@@ -288,9 +288,9 @@
 by (blast dest: Says_Server_message_form secrecy_lemma)
 
 
-subsection{*Guarantees available at various stages of protocol*}
+subsection\<open>Guarantees available at various stages of protocol\<close>
 
-text{*If the encrypted message appears then it originated with the Server*}
+text\<open>If the encrypted message appears then it originated with the Server\<close>
 lemma B_trusts_NS3:
      "\<lbrakk>Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
        B \<notin> bad;  evs \<in> ns_shared\<rbrakk>
@@ -311,14 +311,14 @@
       Says B A (Crypt K (Nonce NB)) \<in> set evs"
 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies)
 apply (analz_mono_contra, simp_all, blast)
-txt{*NS2: contradiction from the assumptions @{term "Key K \<notin> used evs2"} and
-    @{term "Crypt K (Nonce NB) \<in> parts (spies evs2)"} *} 
+txt\<open>NS2: contradiction from the assumptions @{term "Key K \<notin> used evs2"} and
+    @{term "Crypt K (Nonce NB) \<in> parts (spies evs2)"}\<close> 
 apply (force dest!: Crypt_imp_keysFor)
-txt{*NS4*}
+txt\<open>NS4\<close>
 apply (metis B_trusts_NS3 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst unique_session_keys)
 done
 
-text{*This version no longer assumes that K is secure*}
+text\<open>This version no longer assumes that K is secure\<close>
 lemma A_trusts_NS4:
      "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs);
        Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
@@ -328,9 +328,9 @@
 by (blast intro: A_trusts_NS4_lemma
           dest: A_trusts_NS2 Spy_not_see_encrypted_key)
 
-text{*If the session key has been used in NS4 then somebody has forwarded
+text\<open>If the session key has been used in NS4 then somebody has forwarded
   component X in some instance of NS4.  Perhaps an interesting property,
-  but not needed (after all) for the proofs below.*}
+  but not needed (after all) for the proofs below.\<close>
 theorem NS4_implies_NS3 [rule_format]:
   "evs \<in> ns_shared \<Longrightarrow>
      Key K \<notin> analz (spies evs) \<longrightarrow>
@@ -341,9 +341,9 @@
 apply (drule_tac [4] NS3_msg_in_parts_spies)
 apply analz_mono_contra
 apply (simp_all add: ex_disj_distrib, blast)
-txt{*NS2*}
+txt\<open>NS2\<close>
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
-txt{*NS4*}
+txt\<open>NS4\<close>
 apply (metis B_trusts_NS3 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst unique_session_keys)
 done
 
@@ -359,16 +359,16 @@
 apply (erule ns_shared.induct, force)
 apply (drule_tac [4] NS3_msg_in_parts_spies)
 apply (analz_mono_contra, simp_all, blast)
-txt{*NS2*}
+txt\<open>NS2\<close>
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
-txt{*NS5*}
+txt\<open>NS5\<close>
 apply (blast dest!: A_trusts_NS2
              dest: Says_imp_knows_Spy [THEN analz.Inj]
                    unique_session_keys Crypt_Spy_analz_bad)
 done
 
 
-text{*Very strong Oops condition reveals protocol's weakness*}
+text\<open>Very strong Oops condition reveals protocol's weakness\<close>
 lemma B_trusts_NS5:
      "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs);
        Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
@@ -378,9 +378,9 @@
 by (blast intro: B_trusts_NS5_lemma
           dest: B_trusts_NS3 Spy_not_see_encrypted_key)
 
-text{*Unaltered so far wrt original version*}
+text\<open>Unaltered so far wrt original version\<close>
 
-subsection{*Lemmas for reasoning about predicate "Issues"*}
+subsection\<open>Lemmas for reasoning about predicate "Issues"\<close>
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
@@ -414,15 +414,15 @@
 apply (induct_tac "evs")
 apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
-txt{* Resembles @{text"used_subset_append"} in theory Event.*}
+txt\<open>Resembles \<open>used_subset_append\<close> in theory Event.\<close>
 done
 
 lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
 
 
-subsection{*Guarantees of non-injective agreement on the session key, and
+subsection\<open>Guarantees of non-injective agreement on the session key, and
 of key distribution. They also express forms of freshness of certain messages,
-namely that agents were alive after something happened.*}
+namely that agents were alive after something happened.\<close>
 
 lemma B_Issues_A:
      "\<lbrakk> Says B A (Crypt K (Nonce Nb)) \<in> set evs;
@@ -437,24 +437,24 @@
 apply (erule rev_mp)
 apply (erule ns_shared.induct, analz_mono_contra)
 apply (simp_all)
-txt{*fake*}
+txt\<open>fake\<close>
 apply blast
 apply (simp_all add: takeWhile_tail)
-txt{*NS3 remains by pure coincidence!*}
+txt\<open>NS3 remains by pure coincidence!\<close>
 apply (force dest!: A_trusts_NS2 Says_Server_message_form)
-txt{*NS4 would be the non-trivial case can be solved by Nb being used*}
+txt\<open>NS4 would be the non-trivial case can be solved by Nb being used\<close>
 apply (blast dest: parts_spies_takeWhile_mono [THEN subsetD]
                    parts_spies_evs_revD2 [THEN subsetD])
 done
 
-text{*Tells A that B was alive after she sent him the session key.  The
+text\<open>Tells A that B was alive after she sent him the session key.  The
 session key must be assumed confidential for this deduction to be meaningful,
 but that assumption can be relaxed by the appropriate argument.
 
 Precisely, the theorem guarantees (to A) key distribution of the session key
 to B. It also guarantees (to A) non-injective agreement of B with A on the
 session key. Both goals are available to A in the sense of Goal Availability.
-*}
+\<close>
 lemma A_authenticates_and_keydist_to_B:
      "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs);
        Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
@@ -474,13 +474,13 @@
 apply (erule rev_mp)
 apply (erule ns_shared.induct, analz_mono_contra)
 apply (simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*NS2*}
+txt\<open>NS2\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*NS3*}
+txt\<open>NS3\<close>
 apply (metis NS3_msg_in_parts_spies parts_cut_eq)
-txt{*NS5, the most important case, can be solved by unicity*}
+txt\<open>NS5, the most important case, can be solved by unicity\<close>
 apply (metis A_trusts_NS2 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst analz.Snd unique_session_keys)
 done
 
@@ -497,20 +497,20 @@
 apply (erule rev_mp)
 apply (erule ns_shared.induct, analz_mono_contra)
 apply (simp_all)
-txt{*fake*}
+txt\<open>fake\<close>
 apply blast
 apply (simp_all add: takeWhile_tail)
-txt{*NS3 remains by pure coincidence!*}
+txt\<open>NS3 remains by pure coincidence!\<close>
 apply (force dest!: A_trusts_NS2 Says_Server_message_form)
-txt{*NS5 is the non-trivial case and cannot be solved as in @{term B_Issues_A}! because NB is not fresh. We need @{term A_trusts_NS5}, proved for this very purpose*}
+txt\<open>NS5 is the non-trivial case and cannot be solved as in @{term B_Issues_A}! because NB is not fresh. We need @{term A_trusts_NS5}, proved for this very purpose\<close>
 apply (blast dest: A_trusts_NS5 parts_spies_takeWhile_mono [THEN subsetD]
         parts_spies_evs_revD2 [THEN subsetD])
 done
 
-text{*Tells B that A was alive after B issued NB.
+text\<open>Tells B that A was alive after B issued NB.
 
 Precisely, the theorem guarantees (to B) key distribution of the session key to A. It also guarantees (to B) non-injective agreement of A with B on the session key. Both goals are available to B in the sense of Goal Availability.
-*}
+\<close>
 lemma B_authenticates_and_keydist_to_A:
      "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs);
        Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
--- a/src/HOL/Auth/OtwayRees.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/OtwayRees.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,15 +3,15 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Original Otway-Rees Protocol*}
+section\<open>The Original Otway-Rees Protocol\<close>
 
 theory OtwayRees imports Public begin
 
-text{* From page 244 of
+text\<open>From page 244 of
   Burrows, Abadi and Needham (1989).  A Logic of Authentication.
   Proc. Royal Soc. 426
 
-This is the original version, which encrypts Nonce NB.*}
+This is the original version, which encrypts Nonce NB.\<close>
 
 inductive_set otway :: "event list set"
   where
@@ -85,7 +85,7 @@
 declare Fake_parts_insert_in_Un  [dest]
 
 
-text{*A "possibility property": there are traces that reach the end*}
+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.
              Says B A {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key K|}|}
@@ -127,10 +127,10 @@
   some reason proofs work without them!*)
 
 
-text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
-NOBODY sends messages containing X! *}
+text\<open>Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
+NOBODY sends messages containing X!\<close>
 
-text{*Spy never sees a good agent's shared key!*}
+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)"
 by (erule otway.induct, force,
@@ -146,7 +146,7 @@
 by (blast dest: Spy_see_shrK)
 
 
-subsection{*Towards Secrecy: Proofs Involving @{term analz}*}
+subsection\<open>Towards Secrecy: Proofs Involving @{term analz}\<close>
 
 (*Describes the form of K and NA when the Server sends this message.  Also
   for Oops case.*)
@@ -167,9 +167,9 @@
 ****)
 
 
-text{*Session keys are not used to encrypt other session keys*}
+text\<open>Session keys are not used to encrypt other session keys\<close>
 
-text{*The equality makes the induction hypothesis easier to apply*}
+text\<open>The equality makes the induction hypothesis easier to apply\<close>
 lemma analz_image_freshK [rule_format]:
  "evs \<in> otway ==>
    \<forall>K KK. KK <= -(range shrK) -->
@@ -188,7 +188,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{*The Key K uniquely identifies the Server's  message. *}
+text\<open>The Key K uniquely identifies the Server's  message.\<close>
 lemma unique_session_keys:
      "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|}   \<in> set evs;
          Says Server B' {|NA',X',Crypt (shrK B') {|NB',K|}|} \<in> set evs;
@@ -196,13 +196,13 @@
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule otway.induct, simp_all)
-apply blast+  --{*OR3 and OR4*}
+apply blast+  \<comment>\<open>OR3 and OR4\<close>
 done
 
 
-subsection{*Authenticity properties relating to NA*}
+subsection\<open>Authenticity properties relating to NA\<close>
 
-text{*Only OR1 can have caused such a part of a message to appear.*}
+text\<open>Only OR1 can have caused such a part of a message to appear.\<close>
 lemma Crypt_imp_OR1 [rule_format]:
  "[| A \<notin> bad;  evs \<in> otway |]
   ==> Crypt (shrK A) {|NA, Agent A, Agent B|} \<in> parts (knows Spy evs) -->
@@ -222,7 +222,7 @@
 by (blast dest: Crypt_imp_OR1)
 
 
-text{*The Nonce NA uniquely identifies A's message*}
+text\<open>The Nonce NA uniquely identifies A's message\<close>
 lemma unique_NA:
      "[| Crypt (shrK A) {|NA, Agent A, Agent B|} \<in> parts (knows Spy evs);
          Crypt (shrK A) {|NA, Agent A, Agent C|} \<in> parts (knows Spy evs);
@@ -234,9 +234,9 @@
 done
 
 
-text{*It is impossible to re-use a nonce in both OR1 and OR2.  This holds because
+text\<open>It is impossible to re-use a nonce in both OR1 and OR2.  This holds because
   OR2 encrypts Nonce NB.  It prevents the attack that can occur in the
-  over-simplified version of this protocol: see @{text OtwayRees_Bad}.*}
+  over-simplified version of this protocol: see \<open>OtwayRees_Bad\<close>.\<close>
 lemma no_nonce_OR1_OR2:
    "[| Crypt (shrK A) {|NA, Agent A, Agent B|} \<in> parts (knows Spy evs);
        A \<notin> bad;  evs \<in> otway |]
@@ -246,8 +246,8 @@
        drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+)
 done
 
-text{*Crucial property: If the encrypted message appears, and A has used NA
-  to start a run, then it originated with the Server!*}
+text\<open>Crucial property: If the encrypted message appears, and A has used NA
+  to start a run, then it originated with the Server!\<close>
 lemma NA_Crypt_imp_Server_msg [rule_format]:
      "[| A \<notin> bad;  evs \<in> otway |]
       ==> Says A B {|NA, Agent A, Agent B,
@@ -259,16 +259,16 @@
                            Crypt (shrK B) {|NB, Key K|}|} \<in> set evs)"
 apply (erule otway.induct, force,
        drule_tac [4] OR2_parts_knows_Spy, simp_all, blast)
-apply blast  --{*OR1: by freshness*}
-apply (blast dest!: no_nonce_OR1_OR2 intro: unique_NA)  --{*OR3*}
-apply (blast intro!: Crypt_imp_OR1)  --{*OR4*}
+apply blast  \<comment>\<open>OR1: by freshness\<close>
+apply (blast dest!: no_nonce_OR1_OR2 intro: unique_NA)  \<comment>\<open>OR3\<close>
+apply (blast intro!: Crypt_imp_OR1)  \<comment>\<open>OR4\<close>
 done
 
 
-text{*Corollary: if A receives B's OR4 message and the nonce NA agrees
+text\<open>Corollary: if A receives B's OR4 message and the nonce NA agrees
   then the key really did come from the Server!  CANNOT prove this of the
   bad form of this protocol, even though we can prove
-  @{text Spy_not_see_encrypted_key} *}
+  \<open>Spy_not_see_encrypted_key\<close>\<close>
 lemma A_trusts_OR4:
      "[| Says A  B {|NA, Agent A, Agent B,
                      Crypt (shrK A) {|NA, Agent A, Agent B|}|} \<in> set evs;
@@ -282,9 +282,9 @@
 by (blast intro!: NA_Crypt_imp_Server_msg)
 
 
-text{*Crucial secrecy property: Spy does not see the keys sent in msg OR3
+text\<open>Crucial secrecy property: Spy does not see the keys sent in msg OR3
     Does not in itself guarantee security: an attack could violate
-    the premises, e.g. by having @{term "A=Spy"}*}
+    the premises, e.g. by having @{term "A=Spy"}\<close>
 lemma secrecy_lemma:
  "[| A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
   ==> Says Server B
@@ -297,8 +297,8 @@
 apply (drule_tac [6] OR4_analz_knows_Spy)
 apply (drule_tac [4] OR2_analz_knows_Spy)
 apply (simp_all add: analz_insert_eq analz_insert_freshK pushes)
-apply spy_analz  --{*Fake*}
-apply (blast dest: unique_session_keys)+  --{*OR3, OR4, Oops*}
+apply spy_analz  \<comment>\<open>Fake\<close>
+apply (blast dest: unique_session_keys)+  \<comment>\<open>OR3, OR4, Oops\<close>
 done
 
 theorem Spy_not_see_encrypted_key:
@@ -310,13 +310,13 @@
       ==> Key K \<notin> analz (knows Spy evs)"
 by (blast dest: Says_Server_message_form secrecy_lemma)
 
-text{*This form is an immediate consequence of the previous result.  It is 
+text\<open>This form is an immediate consequence of the previous result.  It is 
 similar to the assertions established by other methods.  It is equivalent
 to the previous result in that the Spy already has @{term analz} and
 @{term synth} at his disposal.  However, the conclusion 
 @{term "Key K \<notin> knows Spy evs"} appears not to be inductive: all the cases
 other than Fake are trivial, while Fake requires 
-@{term "Key K \<notin> analz (knows Spy evs)"}. *}
+@{term "Key K \<notin> analz (knows Spy evs)"}.\<close>
 lemma Spy_not_know_encrypted_key:
      "[| Says Server B
           {|NA, Crypt (shrK A) {|NA, Key K|},
@@ -327,8 +327,8 @@
 by (blast dest: Spy_not_see_encrypted_key)
 
 
-text{*A's guarantee.  The Oops premise quantifies over NB because A cannot know
-  what it is.*}
+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 A  B {|NA, Agent A, Agent B,
                      Crypt (shrK A) {|NA, Agent A, Agent B|}|} \<in> set evs;
@@ -339,10 +339,10 @@
 by (blast dest!: A_trusts_OR4 Spy_not_see_encrypted_key)
 
 
-subsection{*Authenticity properties relating to NB*}
+subsection\<open>Authenticity properties relating to NB\<close>
 
-text{*Only OR2 can have caused such a part of a message to appear.  We do not
-  know anything about X: it does NOT have to have the right form.*}
+text\<open>Only OR2 can have caused such a part of a message to appear.  We do not
+  know anything about X: it does NOT have to have the right form.\<close>
 lemma Crypt_imp_OR2:
      "[| Crypt (shrK B) {|NA, NB, Agent A, Agent B|} \<in> parts (knows Spy evs);
          B \<notin> bad;  evs \<in> otway |]
@@ -356,7 +356,7 @@
 done
 
 
-text{*The Nonce NB uniquely identifies B's  message*}
+text\<open>The Nonce NB uniquely identifies B's  message\<close>
 lemma unique_NB:
      "[| Crypt (shrK B) {|NA, NB, Agent A, Agent B|} \<in> parts(knows Spy evs);
          Crypt (shrK B) {|NC, NB, Agent C, Agent B|} \<in> parts(knows Spy evs);
@@ -365,11 +365,11 @@
 apply (erule rev_mp, erule rev_mp)
 apply (erule otway.induct, force,
        drule_tac [4] OR2_parts_knows_Spy, simp_all)
-apply blast+  --{*Fake, OR2*}
+apply blast+  \<comment>\<open>Fake, OR2\<close>
 done
 
-text{*If the encrypted message appears, and B has used Nonce NB,
-  then it originated with the Server!  Quite messy proof.*}
+text\<open>If the encrypted message appears, and B has used Nonce NB,
+  then it originated with the Server!  Quite messy proof.\<close>
 lemma NB_Crypt_imp_Server_msg [rule_format]:
  "[| B \<notin> bad;  evs \<in> otway |]
   ==> Crypt (shrK B) {|NB, Key K|} \<in> parts (knows Spy evs)
@@ -384,15 +384,15 @@
 apply simp
 apply (erule otway.induct, force,
        drule_tac [4] OR2_parts_knows_Spy, simp_all)
-apply blast  --{*Fake*}
-apply blast  --{*OR2*}
-apply (blast dest: unique_NB dest!: no_nonce_OR1_OR2)  --{*OR3*}
-apply (blast dest!: Crypt_imp_OR2)  --{*OR4*}
+apply blast  \<comment>\<open>Fake\<close>
+apply blast  \<comment>\<open>OR2\<close>
+apply (blast dest: unique_NB dest!: no_nonce_OR1_OR2)  \<comment>\<open>OR3\<close>
+apply (blast dest!: Crypt_imp_OR2)  \<comment>\<open>OR4\<close>
 done
 
 
-text{*Guarantee for B: if it gets a message with matching NB then the Server
-  has sent the correct message.*}
+text\<open>Guarantee for B: if it gets a message with matching NB then the Server
+  has sent the correct message.\<close>
 theorem B_trusts_OR3:
      "[| Says B Server {|NA, Agent A, Agent B, X',
                          Crypt (shrK B) {|NA, NB, Agent A, Agent B|} |}
@@ -407,8 +407,8 @@
 by (blast intro!: NB_Crypt_imp_Server_msg)
 
 
-text{*The obvious combination of @{text B_trusts_OR3} with 
-      @{text Spy_not_see_encrypted_key}*}
+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:
      "[| Says B Server {|NA, Agent A, Agent B, X',
                          Crypt (shrK B) {|NA, NB, Agent A, Agent B|} |}
@@ -434,9 +434,9 @@
 done
 
 
-text{*After getting and checking OR4, agent A can trust that B has been active.
+text\<open>After getting and checking OR4, agent A can trust that B has been active.
   We could probably prove that X has the expected form, but that is not
-  strictly necessary for authentication.*}
+  strictly necessary for authentication.\<close>
 theorem A_auths_B:
      "[| Says B' A {|NA, Crypt (shrK A) {|NA, Key K|}|} \<in> set evs;
          Says A  B {|NA, Agent A, Agent B,
--- a/src/HOL/Auth/OtwayReesBella.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/OtwayReesBella.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,19 +2,19 @@
     Author:     Giampaolo Bella, Catania University
 *)
 
-section{*Bella's version of the Otway-Rees protocol*}
+section\<open>Bella's version of the Otway-Rees protocol\<close>
 
 
 theory OtwayReesBella imports Public begin
 
-text{*Bella's modifications to a version of the Otway-Rees protocol taken from
+text\<open>Bella's modifications to a version of the Otway-Rees protocol taken from
 the BAN paper only concern message 7. The updated protocol makes the goal of
 key distribution of the session key available to A. Investigating the
 principle of Goal Availability undermines the BAN claim about the original
 protocol, that "this protocol does not make use of Kab as an encryption key,
 so neither principal can know whether the key is known to the other". The
 updated protocol makes no use of the session key to encrypt but informs A that
-B knows it.*}
+B knows it.\<close>
 
 inductive_set orb :: "event list set"
  where
@@ -76,7 +76,7 @@
 declare Fake_parts_insert_in_Un  [dest]
 
 
-text{*Fragile proof, with backtracking in the possibility call.*}
+text\<open>Fragile proof, with backtracking in the possibility call.\<close>
 lemma possibility_thm: "\<lbrakk>A \<noteq> Server; B \<noteq> Server; Key K \<notin> used[]\<rbrakk>    
       \<Longrightarrow>   \<exists> evs \<in> orb.           
      Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"
@@ -125,15 +125,15 @@
     OR2_analz_knows_Spy [THEN analz_into_parts]
 
 ML
-{*
+\<open>
 fun parts_explicit_tac ctxt i =
     forward_tac ctxt [@{thm Oops_parts_knows_Spy}] (i+7) THEN
     forward_tac ctxt [@{thm OR4_parts_knows_Spy}]  (i+6) THEN
     forward_tac ctxt [@{thm OR2_parts_knows_Spy}]  (i+4)
-*}
+\<close>
  
-method_setup parts_explicit = {*
-    Scan.succeed (SIMPLE_METHOD' o parts_explicit_tac) *}
+method_setup parts_explicit = \<open>
+    Scan.succeed (SIMPLE_METHOD' o parts_explicit_tac)\<close>
   "to explicitly state that some message components belong to parts knows Spy"
 
 
@@ -159,10 +159,10 @@
 
 
 
-subsection{* Proofs involving analz *}
+subsection\<open>Proofs involving analz\<close>
 
-text{*Describes the form of K and NA when the Server sends this message.  Also
-  for Oops case.*}
+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: 
 "\<lbrakk>Says Server B  \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  
      evs \<in> orb\<rbrakk>                                            
@@ -233,7 +233,7 @@
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
 ML
-{*
+\<open>
 structure OtwayReesBella =
 struct
 
@@ -244,23 +244,23 @@
       addsimps @{thms analz_image_freshK_simps})
 
 end
-*}
+\<close>
 
-method_setup analz_freshCryptK = {*
+method_setup analz_freshCryptK = \<open>
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
           REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshCryptK_lemma}),
           ALLGOALS (asm_simp_tac
-            (put_simpset OtwayReesBella.analz_image_freshK_ss ctxt))]))) *}
+            (put_simpset OtwayReesBella.analz_image_freshK_ss ctxt))])))\<close>
   "for proving useful rewrite rule"
 
 
-method_setup disentangle = {*
+method_setup disentangle = \<open>
     Scan.succeed
      (fn ctxt => SIMPLE_METHOD
       (REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE, disjE] 
-                   ORELSE' hyp_subst_tac ctxt))) *}
+                   ORELSE' hyp_subst_tac ctxt)))\<close>
   "for eliminating conjunctions, disjunctions and the like"
 
 
@@ -303,21 +303,21 @@
 apply (frule_tac [7] Gets_Server_message_form)
 apply (frule_tac [9] Says_Server_message_form)
 apply disentangle
-txt{*letting the simplifier solve OR2*}
+txt\<open>letting the simplifier solve OR2\<close>
 apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Snd])
 apply (simp_all (no_asm_simp) add: analz_insert_eq pushes split_ifs)
 apply (spy_analz)
-txt{*OR1*}
+txt\<open>OR1\<close>
 apply blast
-txt{*Oops*}
+txt\<open>Oops\<close>
 prefer 4 apply (blast dest: analz_insert_freshCryptK)
-txt{*OR4 - ii*}
+txt\<open>OR4 - ii\<close>
 prefer 3 apply (blast)
-txt{*OR3*}
+txt\<open>OR3\<close>
 (*adding Gets_imp_ and Says_imp_ for efficiency*)
 apply (blast dest: 
        A_trusts_OR1 unique_Na Key_not_used analz_insert_freshCryptK)
-txt{*OR4 - i *}
+txt\<open>OR4 - i\<close>
 apply clarify
 apply (simp add: pushes split_ifs)
 apply (case_tac "Aaa\<in>bad")
@@ -370,7 +370,7 @@
 done
 
 
-text{*Other properties as for the original protocol*}
+text\<open>Other properties as for the original protocol\<close>
 
 
 end
--- a/src/HOL/Auth/OtwayRees_AN.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/OtwayRees_AN.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Otway-Rees Protocol as Modified by Abadi and Needham*}
+section\<open>The Otway-Rees Protocol as Modified by Abadi and Needham\<close>
 
 theory OtwayRees_AN imports Public begin
 
-text{*
+text\<open>
 This simplified version has minimal encryption and explicit messages.
 
 Note that the formalization does not even assume that nonces are fresh.
@@ -19,36 +19,36 @@
   Abadi and Needham (1996).  
   Prudent Engineering Practice for Cryptographic Protocols.
   IEEE Trans. SE 22 (1)
-*}
+\<close>
 
 inductive_set otway :: "event list set"
   where
-   Nil: --{*The empty trace*}
+   Nil: \<comment>\<open>The empty trace\<close>
         "[] \<in> otway"
 
- | Fake: --{*The Spy may say anything he can say.  The sender field is correct,
-            but agents don't use that information.*}
+ | 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"
 
         
- | Reception: --{*A message that has been sent can be received by the
-                  intended recipient.*}
+ | 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"
 
- | OR1:  --{*Alice initiates a protocol run*}
+ | OR1:  \<comment>\<open>Alice initiates a protocol run\<close>
          "evs1 \<in> otway
           ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
 
- | OR2:  --{*Bob's response to Alice's message.*}
+ | OR2:  \<comment>\<open>Bob's response to Alice's message.\<close>
          "[| evs2 \<in> otway;
              Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
           ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                  # evs2 \<in> otway"
 
- | OR3:  --{*The Server receives Bob's message.  Then he sends a new
-           session key to Bob with a packet for forwarding to Alice.*}
+ | 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;
              Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                \<in>set evs3 |]
@@ -57,17 +57,17 @@
                  Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|}
               # evs3 \<in> otway"
 
- | OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
+ | 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 "B \<noteq> Server"} because we allow messages to self.*}
+             Need @{term "B \<noteq> Server"} because we allow messages to self.\<close>
          "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4;
              Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
                \<in>set evs4 |]
           ==> Says B A X # evs4 \<in> otway"
 
- | Oops: --{*This message models possible leaks of session keys.  The nonces
-             identify the protocol run.*}
+ | Oops: \<comment>\<open>This message models possible leaks of session keys.  The nonces
+             identify the protocol run.\<close>
          "[| evso \<in> otway;
              Says Server B
                       {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|},
@@ -82,7 +82,7 @@
 declare Fake_parts_insert_in_Un  [dest]
 
 
-text{*A "possibility property": there are traces that reach the end*}
+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.
            Says B A (Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|})
@@ -101,7 +101,7 @@
 
 
 
-text{* For reasoning about the encrypted portion of messages *}
+text\<open>For reasoning about the encrypted portion of messages\<close>
 
 lemma OR4_analz_knows_Spy:
      "[| Gets B {|X, Crypt(shrK B) X'|} \<in> set evs;  evs \<in> otway |]
@@ -109,10 +109,10 @@
 by blast
 
 
-text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
-NOBODY sends messages containing X! *}
+text\<open>Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
+NOBODY sends messages containing X!\<close>
 
-text{*Spy never sees a good agent's shared key!*}
+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)"
 by (erule otway.induct, simp_all, blast+)
@@ -126,9 +126,9 @@
 by (blast dest: Spy_see_shrK)
 
 
-subsection{*Proofs involving analz *}
+subsection\<open>Proofs involving analz\<close>
 
-text{*Describes the form of K and NA when the Server sends this message.*}
+text\<open>Describes the form of K and NA when the Server sends this message.\<close>
 lemma Says_Server_message_form:
      "[| Says Server B
             {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
@@ -152,9 +152,9 @@
 ****)
 
 
-text{* Session keys are not used to encrypt other session keys *}
+text\<open>Session keys are not used to encrypt other session keys\<close>
 
-text{*The equality makes the induction hypothesis easier to apply*}
+text\<open>The equality makes the induction hypothesis easier to apply\<close>
 lemma analz_image_freshK [rule_format]:
  "evs \<in> otway ==>
    \<forall>K KK. KK <= -(range shrK) -->
@@ -172,7 +172,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{*The Key K uniquely identifies the Server's message.*}
+text\<open>The Key K uniquely identifies the Server's message.\<close>
 lemma unique_session_keys:
      "[| Says Server B
           {|Crypt (shrK A) {|NA, Agent A, Agent B, K|},
@@ -185,13 +185,13 @@
         evs \<in> otway |]
      ==> A=A' & B=B' & NA=NA' & NB=NB'"
 apply (erule rev_mp, erule rev_mp, erule otway.induct, simp_all)
-apply blast+  --{*OR3 and OR4*}
+apply blast+  \<comment>\<open>OR3 and OR4\<close>
 done
 
 
-subsection{*Authenticity properties relating to NA*}
+subsection\<open>Authenticity properties relating to NA\<close>
 
-text{*If the encrypted message appears then it originated with the Server!*}
+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) {|NA, Agent A, Agent B, Key K|} \<in> parts (knows Spy evs)
@@ -201,12 +201,12 @@
                     \<in> set evs)"
 apply (erule otway.induct, force)
 apply (simp_all add: ex_disj_distrib)
-apply blast+  --{*Fake, OR3*}
+apply blast+  \<comment>\<open>Fake, OR3\<close>
 done
 
 
-text{*Corollary: if A receives B's OR4 message then it originated with the
-      Server. Freshness may be inferred from nonce NA.*}
+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) {|NA, Agent A, Agent B, Key K|}) \<in> set evs;
          A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
@@ -217,9 +217,9 @@
 by (blast intro!: NA_Crypt_imp_Server_msg)
 
 
-text{*Crucial secrecy property: Spy does not see the keys sent in msg OR3
+text\<open>Crucial secrecy property: Spy does not see the keys sent in msg OR3
     Does not in itself guarantee security: an attack could violate
-    the premises, e.g. by having @{term "A=Spy"}*}
+    the premises, e.g. by having @{term "A=Spy"}\<close>
 lemma secrecy_lemma:
      "[| A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
       ==> Says Server B
@@ -232,8 +232,8 @@
 apply (frule_tac [7] Says_Server_message_form)
 apply (drule_tac [6] OR4_analz_knows_Spy)
 apply (simp_all add: analz_insert_eq analz_insert_freshK pushes)
-apply spy_analz  --{*Fake*}
-apply (blast dest: unique_session_keys)+  --{*OR3, OR4, Oops*}
+apply spy_analz  \<comment>\<open>Fake\<close>
+apply (blast dest: unique_session_keys)+  \<comment>\<open>OR3, OR4, Oops\<close>
 done
 
 
@@ -248,8 +248,8 @@
   by (metis secrecy_lemma)
 
 
-text{*A's guarantee.  The Oops premise quantifies over NB because A cannot know
-  what it is.*}
+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) {|NA, Agent A, Agent B, Key K|}) \<in> set evs;
          \<forall>NB. Notes Spy {|NA, NB, Key K|} \<notin> set evs;
@@ -259,9 +259,9 @@
 
 
 
-subsection{*Authenticity properties relating to NB*}
+subsection\<open>Authenticity properties relating to NB\<close>
 
-text{*If the encrypted message appears then it originated with the Server!*}
+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) {|NB, Agent A, Agent B, Key K|} \<in> parts (knows Spy evs)
@@ -270,13 +270,13 @@
                      Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
                    \<in> set evs)"
 apply (erule otway.induct, force, simp_all add: ex_disj_distrib)
-apply blast+  --{*Fake, OR3*}
+apply blast+  \<comment>\<open>Fake, OR3\<close>
 done
 
 
 
-text{*Guarantee for B: if it gets a well-formed certificate then the Server
-  has sent the correct message in round 3.*}
+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 {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
            \<in> set evs;
@@ -288,8 +288,8 @@
 by (blast intro!: NB_Crypt_imp_Server_msg)
 
 
-text{*The obvious combination of @{text B_trusts_OR3} with 
-      @{text Spy_not_see_encrypted_key}*}
+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 {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
           \<in> set evs;
--- a/src/HOL/Auth/OtwayRees_Bad.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/OtwayRees_Bad.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -4,44 +4,44 @@
 *)
 
 
-section{*The Otway-Rees Protocol: The Faulty BAN Version*}
+section\<open>The Otway-Rees Protocol: The Faulty BAN Version\<close>
 
 theory OtwayRees_Bad imports Public begin
 
-text{*The FAULTY version omitting encryption of Nonce NB, as suggested on 
+text\<open>The FAULTY version omitting encryption of Nonce NB, as suggested on 
 page 247 of
   Burrows, Abadi and Needham (1988).  A Logic of Authentication.
   Proc. Royal Soc. 426
 
 This file illustrates the consequences of such errors.  We can still prove
-impressive-looking properties such as @{text Spy_not_see_encrypted_key}, yet
+impressive-looking properties such as \<open>Spy_not_see_encrypted_key\<close>, yet
 the protocol is open to a middleperson attack.  Attempting to prove some key
-lemmas indicates the possibility of this attack.*}
+lemmas indicates the possibility of this attack.\<close>
 
 inductive_set otway :: "event list set"
   where
-   Nil: --{*The empty trace*}
+   Nil: \<comment>\<open>The empty trace\<close>
         "[] \<in> otway"
 
- | Fake: --{*The Spy may say anything he can say.  The sender field is correct,
-            but agents don't use that information.*}
+ | 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"
 
         
- | Reception: --{*A message that has been sent can be received by the
-                  intended recipient.*}
+ | 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"
 
- | OR1:  --{*Alice initiates a protocol run*}
+ | OR1:  \<comment>\<open>Alice initiates a protocol run\<close>
          "[| evs1 \<in> otway;  Nonce NA \<notin> used evs1 |]
           ==> Says A B {|Nonce NA, Agent A, Agent B,
                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
                  # evs1 \<in> otway"
 
- | OR2:  --{*Bob's response to Alice's message.
-             This variant of the protocol does NOT encrypt NB.*}
+ | 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 {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
           ==> Says B Server
@@ -49,9 +49,9 @@
                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
                  # evs2 \<in> otway"
 
- | OR3:  --{*The Server receives Bob's message and checks that the three NAs
+ | 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.*}
+           forwarding to Alice.\<close>
          "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server
                   {|Nonce NA, Agent A, Agent B,
@@ -65,9 +65,9 @@
                     Crypt (shrK B) {|Nonce NB, Key KAB|}|}
                  # evs3 \<in> otway"
 
- | OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
+ | 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 "B \<noteq> Server"} because we allow messages to self.*}
+             Need @{term "B \<noteq> Server"} because we allow messages to self.\<close>
          "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
                              Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
@@ -76,8 +76,8 @@
                \<in> set evs4 |]
           ==> Says B A {|Nonce NA, X|} # evs4 \<in> otway"
 
- | Oops: --{*This message models possible leaks of session keys.  The nonces
-             identify the protocol run.*}
+ | Oops: \<comment>\<open>This message models possible leaks of session keys.  The nonces
+             identify the protocol run.\<close>
          "[| evso \<in> otway;
              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
                \<in> set evso |]
@@ -89,7 +89,7 @@
 declare analz_into_parts [dest]
 declare Fake_parts_insert_in_Un  [dest]
 
-text{*A "possibility property": there are traces that reach the end*}
+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.
             Says B A {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key K|}|}
@@ -109,7 +109,7 @@
 done
 
 
-subsection{*For reasoning about the encrypted portion of messages *}
+subsection\<open>For reasoning about the encrypted portion of messages\<close>
 
 lemma OR2_analz_knows_Spy:
      "[| Gets B {|N, Agent A, Agent B, X|} \<in> set evs;  evs \<in> otway |]
@@ -126,15 +126,15 @@
       ==> K \<in> parts (knows Spy evs)"
 by blast
 
-text{*Forwarding lemma: see comments in OtwayRees.thy*}
+text\<open>Forwarding lemma: see comments in OtwayRees.thy\<close>
 lemmas OR2_parts_knows_Spy =
     OR2_analz_knows_Spy [THEN analz_into_parts]
 
 
-text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
-NOBODY sends messages containing X! *}
+text\<open>Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
+NOBODY sends messages containing X!\<close>
 
-text{*Spy never sees a good agent's shared key!*}
+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)"
 by (erule otway.induct, force,
@@ -150,10 +150,10 @@
 by (blast dest: Spy_see_shrK)
 
 
-subsection{*Proofs involving analz *}
+subsection\<open>Proofs involving analz\<close>
 
-text{*Describes the form of K and NA when the Server sends this message.  Also
-  for Oops case.*}
+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 {|NA, X, Crypt (shrK B) {|NB, Key K|}|} \<in> set evs;
          evs \<in> otway |]
@@ -173,9 +173,9 @@
 ****)
 
 
-text{*Session keys are not used to encrypt other session keys*}
+text\<open>Session keys are not used to encrypt other session keys\<close>
 
-text{*The equality makes the induction hypothesis easier to apply*}
+text\<open>The equality makes the induction hypothesis easier to apply\<close>
 lemma analz_image_freshK [rule_format]:
  "evs \<in> otway ==>
    \<forall>K KK. KK <= -(range shrK) -->
@@ -194,7 +194,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{*The Key K uniquely identifies the Server's  message. *}
+text\<open>The Key K uniquely identifies the Server's  message.\<close>
 lemma unique_session_keys:
      "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|}   \<in> set evs;
          Says Server B' {|NA',X',Crypt (shrK B') {|NB',K|}|} \<in> set evs;
@@ -202,13 +202,13 @@
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule otway.induct, simp_all)
-apply blast+  --{*OR3 and OR4*}
+apply blast+  \<comment>\<open>OR3 and OR4\<close>
 done
 
 
-text{*Crucial secrecy property: Spy does not see the keys sent in msg OR3
+text\<open>Crucial secrecy property: Spy does not see the keys sent in msg OR3
     Does not in itself guarantee security: an attack could violate
-    the premises, e.g. by having @{term "A=Spy"} *}
+    the premises, e.g. by having @{term "A=Spy"}\<close>
 lemma secrecy_lemma:
  "[| A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
   ==> Says Server B
@@ -221,8 +221,8 @@
 apply (drule_tac [6] OR4_analz_knows_Spy)
 apply (drule_tac [4] OR2_analz_knows_Spy)
 apply (simp_all add: analz_insert_eq analz_insert_freshK pushes)
-apply spy_analz  --{*Fake*}
-apply (blast dest: unique_session_keys)+  --{*OR3, OR4, Oops*}
+apply spy_analz  \<comment>\<open>Fake\<close>
+apply (blast dest: unique_session_keys)+  \<comment>\<open>OR3, OR4, Oops\<close>
 done
 
 
@@ -236,11 +236,11 @@
 by (blast dest: Says_Server_message_form secrecy_lemma)
 
 
-subsection{*Attempting to prove stronger properties *}
+subsection\<open>Attempting to prove stronger properties\<close>
 
-text{*Only OR1 can have caused such a part of a message to appear. The premise
+text\<open>Only OR1 can have caused such a part of a message to appear. The premise
   @{term "A \<noteq> B"} prevents OR2's similar-looking cryptogram from being picked 
-  up. Original Otway-Rees doesn't need it.*}
+  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) {|NA, Agent A, Agent B|} \<in> parts (knows Spy evs) -->
@@ -250,11 +250,11 @@
     drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+)
 
 
-text{*Crucial property: If the encrypted message appears, and A has used NA
+text\<open>Crucial property: If the encrypted message appears, and A has used NA
   to start a run, then it originated with the Server!
-  The premise @{term "A \<noteq> B"} allows use of @{text Crypt_imp_OR1}*}
-text{*Only it is FALSE.  Somebody could make a fake message to Server
-          substituting some other nonce NA' for NB.*}
+  The premise @{term "A \<noteq> B"} 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) {|NA, Key K|} \<in> parts (knows Spy evs) -->
            Says A B {|NA, Agent A, Agent B,
@@ -266,12 +266,12 @@
                   Crypt (shrK B) {|NB, Key K|}|}  \<in> set evs)"
 apply (erule otway.induct, force,
        drule_tac [4] OR2_parts_knows_Spy, simp_all)
-apply blast  --{*Fake*}
-apply blast  --{*OR1: it cannot be a new Nonce, contradiction.*}
-txt{*OR3 and OR4*}
+apply blast  \<comment>\<open>Fake\<close>
+apply blast  \<comment>\<open>OR1: it cannot be a new Nonce, contradiction.\<close>
+txt\<open>OR3 and OR4\<close>
 apply (simp_all add: ex_disj_distrib)
- prefer 2 apply (blast intro!: Crypt_imp_OR1)  --{*OR4*}
-txt{*OR3*}
+ prefer 2 apply (blast intro!: Crypt_imp_OR1)  \<comment>\<open>OR4\<close>
+txt\<open>OR3\<close>
 apply clarify
 (*The hypotheses at this point suggest an attack in which nonce NB is used
   in two different roles:
--- a/src/HOL/Auth/Public.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Public.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -14,7 +14,7 @@
 lemma invKey_K: "K \<in> symKeys ==> invKey K = K"
 by (simp add: symKeys_def)
 
-subsection{*Asymmetric Keys*}
+subsection\<open>Asymmetric Keys\<close>
 
 datatype keymode = Signature | Encryption
 
@@ -43,8 +43,8 @@
   "priSK A == privateKey Signature A"
 
 
-text{*These abbreviations give backward compatibility.  They represent the
-simple situation where the signature and encryption keys are the same.*}
+text\<open>These abbreviations give backward compatibility.  They represent the
+simple situation where the signature and encryption keys are the same.\<close>
 
 abbreviation
   pubK :: "agent => key" where
@@ -55,8 +55,8 @@
   "priK A == invKey (pubEK A)"
 
 
-text{*By freeness of agents, no two agents have the same key.  Since
-  @{term "True\<noteq>False"}, no agent has identical signing and encryption keys*}
+text\<open>By freeness of agents, no two agents have the same key.  Since
+  @{term "True\<noteq>False"}, no agent has identical signing and encryption keys\<close>
 specification (publicKey)
   injective_publicKey:
     "publicKey b A = publicKey c A' ==> b=c & A=A'"
@@ -77,7 +77,7 @@
 declare publicKey_neq_privateKey [iff]
 
 
-subsection{*Basic properties of @{term pubK} and @{term priK}*}
+subsection\<open>Basic properties of @{term pubK} and @{term priK}\<close>
 
 lemma publicKey_inject [iff]: "(publicKey b A = publicKey c A') = (b=c & A=A')"
 by (blast dest!: injective_publicKey) 
@@ -104,7 +104,7 @@
 
 
 
-subsection{*"Image" equations that hold for injective functions*}
+subsection\<open>"Image" equations that hold for injective functions\<close>
 
 lemma invKey_image_eq [simp]: "(invKey x \<in> invKey`A) = (x \<in> A)"
 by auto
@@ -125,26 +125,26 @@
 by auto
 
 
-subsection{*Symmetric Keys*}
+subsection\<open>Symmetric Keys\<close>
 
-text{*For some protocols, it is convenient to equip agents with symmetric as
-well as asymmetric keys.  The theory @{text Shared} assumes that all keys
-are symmetric.*}
+text\<open>For some protocols, it is convenient to equip agents with symmetric as
+well as asymmetric keys.  The theory \<open>Shared\<close> assumes that all keys
+are symmetric.\<close>
 
 consts
-  shrK    :: "agent => key"    --{*long-term shared keys*}
+  shrK    :: "agent => key"    \<comment>\<open>long-term shared keys\<close>
 
 specification (shrK)
   inj_shrK: "inj shrK"
-  --{*No two agents have the same long-term key*}
+  \<comment>\<open>No two agents have the same long-term key\<close>
    apply (rule exI [of _ "case_agent 0 (\<lambda>n. n + 2) 1"]) 
    apply (simp add: inj_on_def split: agent.split) 
    done
 
 axiomatization where
-  sym_shrK [iff]: "shrK X \<in> symKeys" --{*All shared keys are symmetric*}
+  sym_shrK [iff]: "shrK X \<in> symKeys" \<comment>\<open>All shared keys are symmetric\<close>
 
-text{*Injectiveness: Agents' long-term keys are distinct.*}
+text\<open>Injectiveness: Agents' long-term keys are distinct.\<close>
 lemmas shrK_injective = inj_shrK [THEN inj_eq]
 declare shrK_injective [iff]
 
@@ -189,15 +189,15 @@
 lemma shrK_image_eq [simp]: "(shrK x \<in> shrK ` AA) = (x \<in> AA)"
 by auto
 
-text{*For some reason, moving this up can make some proofs loop!*}
+text\<open>For some reason, moving this up can make some proofs loop!\<close>
 declare invKey_K [simp]
 
 
-subsection{*Initial States of Agents*}
+subsection\<open>Initial States of Agents\<close>
 
-text{*Note: for all practical purposes, all that matters is the initial
+text\<open>Note: for all practical purposes, all that matters is the initial
 knowledge of the Spy.  All other agents are automata, merely following the
-protocol.*}
+protocol.\<close>
 
 overloading
   initState \<equiv> initState
@@ -224,10 +224,10 @@
 end
 
 
-text{*These lemmas allow reasoning about @{term "used evs"} rather than
+text\<open>These lemmas allow reasoning about @{term "used evs"} rather than
    @{term "knows Spy evs"}, which is useful when there are private Notes. 
    Because they depend upon the definition of @{term initState}, they cannot
-   be moved up.*}
+   be moved up.\<close>
 
 lemma used_parts_subset_parts [rule_format]:
      "\<forall>X \<in> used evs. parts {X} \<subseteq> used evs"
@@ -235,15 +235,15 @@
  prefer 2
  apply (simp add: used_Cons split: event.split)
  apply (metis Un_iff empty_subsetI insert_subset le_supI1 le_supI2 parts_subset_iff)
-txt{*Base case*}
+txt\<open>Base case\<close>
 apply (auto dest!: parts_cut simp add: used_Nil) 
 done
 
 lemma MPair_used_D: "{|X,Y|} \<in> used H ==> X \<in> used H & Y \<in> used H"
 by (drule used_parts_subset_parts, simp, blast)
 
-text{*There was a similar theorem in Event.thy, so perhaps this one can
-  be moved up if proved directly by induction.*}
+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!]:
      "[| {|X,Y|} \<in> used H;
          [| X \<in> used H; Y \<in> used H |] ==> P |] 
@@ -251,8 +251,8 @@
 by (blast dest: MPair_used_D) 
 
 
-text{*Rewrites should not refer to  @{term "initState(Friend i)"} because
-  that expression is not in normal form.*}
+text\<open>Rewrites should not refer to  @{term "initState(Friend i)"} because
+  that expression is not in normal form.\<close>
 
 lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
 apply (unfold keysFor_def)
@@ -293,20 +293,20 @@
 declare neq_shrK [simp]
 
 
-subsection{*Function @{term spies} *}
+subsection\<open>Function @{term spies}\<close>
 
 lemma not_SignatureE [elim!]: "b \<noteq> Signature \<Longrightarrow> b = Encryption"
   by (cases b, auto) 
 
-text{*Agents see their own private keys!*}
+text\<open>Agents see their own private keys!\<close>
 lemma priK_in_initState [iff]: "Key (privateKey b A) \<in> initState A"
   by (cases A, auto)
 
-text{*Agents see all public keys!*}
+text\<open>Agents see all public keys!\<close>
 lemma publicKey_in_initState [iff]: "Key (publicKey b A) \<in> initState B"
   by (cases B, auto) 
 
-text{*All public keys are visible*}
+text\<open>All public keys are visible\<close>
 lemma spies_pubK [iff]: "Key (publicKey b A) \<in> spies evs"
 apply (induct_tac "evs")
 apply (auto simp add: imageI knows_Cons split add: event.split)
@@ -315,14 +315,14 @@
 lemmas analz_spies_pubK = spies_pubK [THEN analz.Inj]
 declare analz_spies_pubK [iff]
 
-text{*Spy sees private keys of bad agents!*}
+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"
 apply (induct_tac "evs")
 apply (auto simp add: imageI knows_Cons split add: event.split)
 done
 
-text{*Spy sees long-term shared keys of bad agents!*}
+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"
 apply (induct_tac "evs")
@@ -346,7 +346,7 @@
 by force
 
 
-subsection{*Fresh Nonces*}
+subsection\<open>Fresh Nonces\<close>
 
 lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)"
 by (induct_tac "B", auto)
@@ -355,9 +355,9 @@
 by (simp add: used_Nil)
 
 
-subsection{*Supply fresh nonces for possibility theorems*}
+subsection\<open>Supply fresh nonces for possibility theorems\<close>
 
-text{*In any trace, there is an upper bound N on the greatest nonce in use*}
+text\<open>In any trace, there is an upper bound N on the greatest nonce in use\<close>
 lemma Nonce_supply_lemma: "EX N. ALL n. N<=n --> Nonce n \<notin> used evs"
 apply (induct_tac "evs")
 apply (rule_tac x = 0 in exI)
@@ -374,7 +374,7 @@
 apply (rule someI, fast)
 done
 
-subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
+subsection\<open>Specialized Rewriting for Theorems About @{term analz} and Image\<close>
 
 lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
 by blast
@@ -386,22 +386,22 @@
 lemma Crypt_imp_keysFor :"[|Crypt K X \<in> H; K \<in> symKeys|] ==> K \<in> keysFor H"
 by (drule Crypt_imp_invKey_keysFor, simp)
 
-text{*Lemma for the trivial direction of the if-and-only-if of the 
-Session Key Compromise Theorem*}
+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)) --> (K \<in> nE | Key K \<in> analz H)  ==>  
          (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])
 
 lemmas analz_image_freshK_simps =
-       simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
+       simp_thms mem_simps \<comment>\<open>these two allow its use with \<open>only:\<close>\<close>
        disj_comms 
        image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset
        analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD]
        insert_Key_singleton 
        Key_not_used insert_Key_image Un_assoc [THEN sym]
 
-ML {*
+ML \<open>
 structure Public =
 struct
 
@@ -428,25 +428,25 @@
      REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
 
 end
-*}
+\<close>
 
-method_setup analz_freshK = {*
+method_setup analz_freshK = \<open>
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
           REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}),
-          ALLGOALS (asm_simp_tac (put_simpset Public.analz_image_freshK_ss ctxt))]))) *}
+          ALLGOALS (asm_simp_tac (put_simpset Public.analz_image_freshK_ss ctxt))])))\<close>
     "for proving the Session Key Compromise theorem"
 
 
-subsection{*Specialized Methods for Possibility Theorems*}
+subsection\<open>Specialized Methods for Possibility Theorems\<close>
 
-method_setup possibility = {*
-    Scan.succeed (SIMPLE_METHOD o Public.possibility_tac) *}
+method_setup possibility = \<open>
+    Scan.succeed (SIMPLE_METHOD o Public.possibility_tac)\<close>
     "for proving possibility theorems"
 
-method_setup basic_possibility = {*
-    Scan.succeed (SIMPLE_METHOD o Public.basic_possibility_tac) *}
+method_setup basic_possibility = \<open>
+    Scan.succeed (SIMPLE_METHOD o Public.basic_possibility_tac)\<close>
     "for proving possibility theorems"
 
 end
--- a/src/HOL/Auth/Recur.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Recur.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Otway-Bull Recursive Authentication Protocol*}
+section\<open>The Otway-Bull Recursive Authentication Protocol\<close>
 
 theory Recur imports Public begin
 
-text{*End marker for message bundles*}
+text\<open>End marker for message bundles\<close>
 abbreviation
   END :: "msg" where
   "END == Number 0"
@@ -117,7 +117,7 @@
 **)
 
 
-text{*Simplest case: Alice goes directly to the server*}
+text\<open>Simplest case: Alice goes directly to the server\<close>
 lemma "Key K \<notin> used [] 
        ==> \<exists>NA. \<exists>evs \<in> recur.
               Says Server A {|Crypt (shrK A) {|Key K, Agent Server, Nonce NA|},
@@ -129,7 +129,7 @@
 done
 
 
-text{*Case two: Alice, Bob and the server*}
+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.
@@ -176,7 +176,7 @@
 apply (auto dest: Key_not_used respond_imp_not_used)
 done
 
-text{*Simple inductive reasoning about responses*}
+text\<open>Simple inductive reasoning about responses\<close>
 lemma respond_imp_responses:
      "(PA,RB,KAB) \<in> respond evs ==> RB \<in> responses evs"
 apply (erule respond.induct)
@@ -210,7 +210,7 @@
 lemma Spy_see_shrK [simp]:
      "evs \<in> recur ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 apply (erule recur.induct, auto)
-txt{*RA3.  It's ugly to call auto twice, but it seems necessary.*}
+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
 
@@ -244,7 +244,7 @@
 done 
 
 
-text{*Version for the protocol.  Proof is easy, thanks to the lemma.*}
+text\<open>Version for the protocol.  Proof is easy, thanks to the lemma.\<close>
 lemma raw_analz_image_freshK:
  "evs \<in> recur ==>
    \<forall>K KK. KK \<subseteq> - (range shrK) -->
@@ -254,7 +254,7 @@
 apply (drule_tac [4] RA2_analz_spies,
        drule_tac [5] respond_imp_responses,
        drule_tac [6] RA4_analz_spies, analz_freshK, spy_analz)
-txt{*RA3*}
+txt\<open>RA3\<close>
 apply (simp_all add: resp_analz_image_freshK_lemma)
 done
 
@@ -276,7 +276,7 @@
          add: analz_image_freshK_simps raw_analz_image_freshK)
 
 
-text{*Everything that's hashed is already in past traffic. *}
+text\<open>Everything that's hashed is already in past traffic.\<close>
 lemma Hash_imp_body:
      "[| Hash {|Key(shrK A), X|} \<in> parts (spies evs);
          evs \<in> recur;  A \<notin> bad |] ==> X \<in> parts (spies evs)"
@@ -285,9 +285,9 @@
        drule_tac [6] RA4_parts_spies,
        drule_tac [5] respond_imp_responses,
        drule_tac [4] RA2_parts_spies)
-txt{*RA3 requires a further induction*}
+txt\<open>RA3 requires a further induction\<close>
 apply (erule_tac [5] responses.induct, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast intro: parts_insertI)
 done
 
@@ -307,10 +307,10 @@
 apply (erule recur.induct,
        drule_tac [5] respond_imp_responses)
 apply (force, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
 apply (erule_tac [3] responses.induct)
-txt{*RA1,2: creation of new Nonce*}
+txt\<open>RA1,2: creation of new Nonce\<close>
 apply simp_all
 apply (blast dest!: Hash_imp_body)+
 done
@@ -339,14 +339,14 @@
 apply (erule rev_mp, erule responses.induct)
 apply (simp_all del: image_insert
              add: analz_image_freshK_simps resp_analz_image_freshK_lemma)
-txt{*Simplification using two distinct treatments of "image"*}
+txt\<open>Simplification using two distinct treatments of "image"\<close>
 apply (simp add: parts_insert2, blast)
 done
 
 lemmas resp_analz_insert =
        resp_analz_insert_lemma [OF _ raw_analz_image_freshK]
 
-text{*The last key returned by respond indeed appears in a certificate*}
+text\<open>The last key returned by respond indeed appears in a certificate\<close>
 lemma respond_certificate:
      "(Hash[Key(shrK A)] {|Agent A, B, NA, P|}, RA, K) \<in> respond evs
       ==> Crypt (shrK A) {|Key K, B, NA|} \<in> parts {RA}"
@@ -392,12 +392,12 @@
 apply (simp_all del: image_insert
                 add: analz_image_freshK_simps split_ifs shrK_in_analz_respond
                      resp_analz_image_freshK parts_insert2)
-txt{*Base case of respond*}
+txt\<open>Base case of respond\<close>
 apply blast
-txt{*Inductive step of respond*}
+txt\<open>Inductive step of respond\<close>
 apply (intro allI conjI impI, simp_all)
-txt{*by unicity, either @{term "B=Aa"} or @{term "B=A'"}, a contradiction
-     if @{term "B \<in> bad"}*}   
+txt\<open>by unicity, either @{term "B=Aa"} or @{term "B=A'"}, a contradiction
+     if @{term "B \<in> bad"}\<close>   
 apply (blast dest: unique_session_keys respond_certificate)
 apply (blast dest!: respond_certificate)
 apply (blast dest!: resp_analz_insert)
@@ -414,21 +414,21 @@
        frule_tac [5] respond_imp_responses,
        drule_tac [6] RA4_analz_spies,
        simp_all add: split_ifs analz_insert_eq analz_insert_freshK)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*RA2*}
+txt\<open>RA2\<close>
 apply blast 
-txt{*RA3*}
+txt\<open>RA3\<close>
 apply (simp add: parts_insert_spies)
 apply (metis Key_in_parts_respond parts.Body parts.Fst resp_analz_insert 
              respond_Spy_not_see_session_key usedI)
-txt{*RA4*}
+txt\<open>RA4\<close>
 apply blast 
 done
 
 (**** Authenticity properties for Agents ****)
 
-text{*The response never contains Hashes*}
+text\<open>The response never contains Hashes\<close>
 lemma Hash_in_parts_respond:
      "[| Hash {|Key (shrK B), M|} \<in> parts (insert RB H);
          (PB,RB,K) \<in> respond evs |]
@@ -437,10 +437,10 @@
 apply (erule respond_imp_responses [THEN responses.induct], auto)
 done
 
-text{*Only RA1 or RA2 can have caused such a part of a message to appear.
+text\<open>Only RA1 or RA2 can have caused such a part of a message to appear.
   This result is of no use to B, who cannot verify the Hash.  Moreover,
   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.*}
+  used to prove B's presence to A at the run's conclusion.\<close>
 lemma Hash_auth_sender [rule_format]:
      "[| Hash {|Key(shrK A), Agent A, Agent B, NA, P|} \<in> parts(spies evs);
          A \<notin> bad;  evs \<in> recur |]
@@ -451,7 +451,7 @@
        drule_tac [6] RA4_parts_spies,
        drule_tac [4] RA2_parts_spies,
        simp_all)
-txt{*Fake, RA3*}
+txt\<open>Fake, RA3\<close>
 apply (blast dest: Hash_in_parts_respond)+
 done
 
@@ -460,23 +460,23 @@
 **)
 
 
-text{*Certificates can only originate with the Server.*}
+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  &
                    Crypt (shrK A) Y \<in> parts {RC}"
 apply (erule rev_mp, erule recur.induct, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*RA1*}
+txt\<open>RA1\<close>
 apply blast
-txt{*RA2: it cannot be a new Nonce, contradiction.*}
+txt\<open>RA2: it cannot be a new Nonce, contradiction.\<close>
 apply blast
-txt{*RA3.  Pity that the proof is so brittle: this step requires the rewriting,
-       which however would break all other steps.*}
+txt\<open>RA3.  Pity that the proof is so brittle: this step requires the rewriting,
+       which however would break all other steps.\<close>
 apply (simp add: parts_insert_spies, blast)
-txt{*RA4*}
+txt\<open>RA4\<close>
 apply blast
 done
 
--- a/src/HOL/Auth/Shared.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Shared.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -16,12 +16,12 @@
 
 specification (shrK)
   inj_shrK: "inj shrK"
-  --{*No two agents have the same long-term key*}
+  \<comment>\<open>No two agents have the same long-term key\<close>
    apply (rule exI [of _ "case_agent 0 (\<lambda>n. n + 2) 1"]) 
    apply (simp add: inj_on_def split: agent.split) 
    done
 
-text{*Server knows all long-term keys; other agents know only their own*}
+text\<open>Server knows all long-term keys; other agents know only their own\<close>
 
 overloading
   initState \<equiv> initState
@@ -35,7 +35,7 @@
 end
 
 
-subsection{*Basic properties of shrK*}
+subsection\<open>Basic properties of shrK\<close>
 
 (*Injectiveness: Agents' long-term keys are distinct.*)
 lemmas shrK_injective = inj_shrK [THEN inj_eq]
@@ -51,12 +51,12 @@
      "[| Crypt K X \<in> analz H;  Key K  \<in> analz H |] ==> X \<in> analz H"
 by auto
 
-text{*Now cancel the @{text dest} attribute given to
- @{text analz.Decrypt} in its declaration.*}
+text\<open>Now cancel the \<open>dest\<close> attribute given to
+ \<open>analz.Decrypt\<close> in its declaration.\<close>
 declare analz.Decrypt [rule del]
 
-text{*Rewrites should not refer to  @{term "initState(Friend i)"} because
-  that expression is not in normal form.*}
+text\<open>Rewrites should not refer to  @{term "initState(Friend i)"} because
+  that expression is not in normal form.\<close>
 
 lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
 apply (unfold keysFor_def)
@@ -73,7 +73,7 @@
 by (metis Crypt_imp_invKey_keysFor invKey_K)
 
 
-subsection{*Function "knows"*}
+subsection\<open>Function "knows"\<close>
 
 (*Spy sees shared keys of agents!*)
 lemma Spy_knows_Spy_bad [intro!]: "A: bad ==> Key (shrK A) \<in> knows Spy evs"
@@ -108,7 +108,7 @@
 declare shrK_sym_neq [simp]
 
 
-subsection{*Fresh nonces*}
+subsection\<open>Fresh nonces\<close>
 
 lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)"
 by (induct_tac "B", auto)
@@ -117,7 +117,7 @@
 by (simp add: used_Nil)
 
 
-subsection{*Supply fresh nonces for possibility theorems.*}
+subsection\<open>Supply fresh nonces for possibility theorems.\<close>
 
 (*In any trace, there is an upper bound N on the greatest nonce in use.*)
 lemma Nonce_supply_lemma: "\<exists>N. ALL n. N<=n --> Nonce n \<notin> used evs"
@@ -152,14 +152,14 @@
 apply (rule someI, blast)
 done
 
-text{*Unlike the corresponding property of nonces, we cannot prove
+text\<open>Unlike the corresponding property of nonces, we cannot prove
     @{term "finite KK ==> \<exists>K. K \<notin> KK & Key K \<notin> used evs"}.
     We have infinitely many agents and there is nothing to stop their
     long-term keys from exhausting all the natural numbers.  Instead,
-    possibility theorems must assume the existence of a few keys.*}
+    possibility theorems must assume the existence of a few keys.\<close>
 
 
-subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
+subsection\<open>Specialized Rewriting for Theorems About @{term analz} and Image\<close>
 
 lemma subset_Compl_range: "A <= - (range shrK) ==> shrK x \<notin> A"
 by blast
@@ -175,7 +175,7 @@
     erase occurrences of forwarded message components (X). **)
 
 lemmas analz_image_freshK_simps =
-       simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
+       simp_thms mem_simps \<comment>\<open>these two allow its use with \<open>only:\<close>\<close>
        disj_comms 
        image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset
        analz_insert_eq Un_upper2 [THEN analz_mono, THEN [2] rev_subsetD]
@@ -189,10 +189,10 @@
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
 
-subsection{*Tactics for possibility theorems*}
+subsection\<open>Tactics for possibility theorems\<close>
 
 ML
-{*
+\<open>
 structure Shared =
 struct
 
@@ -223,7 +223,7 @@
       addsimps @{thms analz_image_freshK_simps})
 
 end
-*}
+\<close>
 
 
 
@@ -234,20 +234,20 @@
 
 (*Specialized methods*)
 
-method_setup analz_freshK = {*
+method_setup analz_freshK = \<open>
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
           REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}),
-          ALLGOALS (asm_simp_tac (put_simpset Shared.analz_image_freshK_ss ctxt))]))) *}
+          ALLGOALS (asm_simp_tac (put_simpset Shared.analz_image_freshK_ss ctxt))])))\<close>
     "for proving the Session Key Compromise theorem"
 
-method_setup possibility = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (Shared.possibility_tac ctxt)) *}
+method_setup possibility = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (Shared.possibility_tac ctxt))\<close>
     "for proving possibility theorems"
 
-method_setup basic_possibility = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (Shared.basic_possibility_tac ctxt)) *}
+method_setup basic_possibility = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (Shared.basic_possibility_tac ctxt))\<close>
     "for proving possibility theorems"
 
 lemma knows_subset_knows_Cons: "knows A evs <= knows A (e # evs)"
--- a/src/HOL/Auth/Smartcard/Auth_Smartcard.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Smartcard/Auth_Smartcard.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -2,7 +2,7 @@
     Copyright   1996  University of Cambridge
 *)
 
-section {* Smartcard protocols: rely on conventional Message and on new EventSC and Smartcard *}
+section \<open>Smartcard protocols: rely on conventional Message and on new EventSC and Smartcard\<close>
 
 theory Auth_Smartcard
 imports
--- a/src/HOL/Auth/Smartcard/EventSC.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Smartcard/EventSC.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -1,4 +1,4 @@
-section{*Theory of Events for Security Protocols that use smartcards*}
+section\<open>Theory of Events for Security Protocols that use smartcards\<close>
 
 theory EventSC
 imports
@@ -11,7 +11,7 @@
 
 datatype card = Card agent
 
-text{*Four new events express the traffic between an agent and his card*}
+text\<open>Four new events express the traffic between an agent and his card\<close>
 datatype  
   event = Says  agent agent msg
         | Notes agent       msg
@@ -32,7 +32,7 @@
   "insecureM == \<not>secureM"
 
 
-text{*Spy has access to his own key for spoof messages, but Server is secure*}
+text\<open>Spy has access to his own key for spoof messages, but Server is secure\<close>
 specification (bad)
   Spy_in_bad     [iff]: "Spy \<in> bad"
   Server_not_bad [iff]: "Server \<notin> bad"
@@ -95,9 +95,9 @@
                   | C_Gets C X   => used evs
                   | Outpts C A X  => parts{X} \<union> (used evs)
                   | A_Gets A X   => used evs)"
-    --{*@{term Gets} always follows @{term Says} in real protocols. 
+    \<comment>\<open>@{term Gets} always follows @{term Says} in real protocols. 
        Likewise, @{term C_Gets} will always have to follow @{term Inputs}
-       and @{term A_Gets} will always have to follow @{term Outpts}*}
+       and @{term A_Gets} will always have to follow @{term Outpts}\<close>
 
 lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs \<longrightarrow> X \<in> used evs"
 apply (induct_tac evs)
@@ -116,7 +116,7 @@
 done
 
 
-subsection{*Function @{term knows}*}
+subsection\<open>Function @{term knows}\<close>
 
 (*Simplifying   
  parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
@@ -127,8 +127,8 @@
      "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
 by simp
 
-text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
-      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
+text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits
+      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}\<close>
 lemma knows_Spy_Notes [simp]:
      "knows Spy (Notes A X # evs) =  
           (if A\<in>bad then insert X (knows Spy evs) else knows Spy evs)"
@@ -192,7 +192,7 @@
 
 
 
-text{*Spy sees what is sent on the traffic*}
+text\<open>Spy sees what is sent on the traffic\<close>
 lemma Says_imp_knows_Spy [rule_format]:
      "Says A B X \<in> set evs \<longrightarrow> X \<in> knows Spy evs"
 apply (induct_tac "evs")
@@ -237,15 +237,15 @@
 
 
 
-text{*Elimination rules: derive contradictions from old Says events containing
-  items known to be fresh*}
+text\<open>Elimination rules: derive contradictions from old Says events containing
+  items known to be fresh\<close>
 lemmas knows_Spy_partsEs =
      Says_imp_knows_Spy [THEN parts.Inj, elim_format]
      parts.Body [elim_format]
 
 
 
-subsection{*Knowledge of Agents*}
+subsection\<open>Knowledge of Agents\<close>
 
 lemma knows_Inputs: "knows A (Inputs A C X # evs) = insert X (knows A evs)"
 by simp
@@ -287,21 +287,21 @@
 by (simp add: subset_insertI)
 
 
-text{*Agents know what they say*}
+text\<open>Agents know what they say\<close>
 lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs \<longrightarrow> X \<in> knows A evs"
 apply (induct_tac "evs")
 apply (simp_all (no_asm_simp) split add: event.split)
 apply blast
 done
 
-text{*Agents know what they note*}
+text\<open>Agents know what they note\<close>
 lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
 apply (induct_tac "evs")
 apply (simp_all (no_asm_simp) split add: event.split)
 apply blast
 done
 
-text{*Agents know what they receive*}
+text\<open>Agents know what they receive\<close>
 lemma Gets_imp_knows_agents [rule_format]:
      "A \<noteq> Spy \<longrightarrow> Gets A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
 apply (induct_tac "evs")
@@ -390,7 +390,7 @@
 
 
 
-text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
+text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close>
 declare knows_Cons [simp del]
         used_Nil [simp del] used_Cons [simp del]
 
@@ -404,7 +404,7 @@
 done
 
 
-text{*For proving @{text new_keys_not_used}*}
+text\<open>For proving \<open>new_keys_not_used\<close>\<close>
 lemma keysFor_parts_insert:
      "\<lbrakk> K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) \<rbrakk>   
       \<Longrightarrow> K \<in> keysFor (parts (G \<union> H)) \<or> Key (invKey K) \<in> parts H" 
--- a/src/HOL/Auth/Smartcard/ShoupRubin.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -1,7 +1,7 @@
 (*  Author:     Giampaolo Bella, Catania University
 *)
 
-section{*Original Shoup-Rubin protocol*}
+section\<open>Original Shoup-Rubin protocol\<close>
 
 theory ShoupRubin imports Smartcard begin
 
@@ -732,7 +732,7 @@
 
 
 ML
-{*
+\<open>
 structure ShoupRubin =
 struct
 
@@ -758,18 +758,18 @@
          REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
 
 end
-*}
+\<close>
 
-method_setup prepare = {*
-    Scan.succeed (SIMPLE_METHOD o ShoupRubin.prepare_tac) *}
+method_setup prepare = \<open>
+    Scan.succeed (SIMPLE_METHOD o ShoupRubin.prepare_tac)\<close>
   "to launch a few simple facts that will help the simplifier"
 
-method_setup parts_prepare = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubin.parts_prepare_tac ctxt)) *}
+method_setup parts_prepare = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubin.parts_prepare_tac ctxt))\<close>
   "additional facts to reason about parts"
 
-method_setup analz_prepare = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubin.analz_prepare_tac ctxt)) *}
+method_setup analz_prepare = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubin.analz_prepare_tac ctxt))\<close>
   "additional facts to reason about analz"
 
 
@@ -813,7 +813,7 @@
 apply auto
 done
 
-method_setup sc_analz_freshK = {*
+method_setup sc_analz_freshK = \<open>
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST
@@ -823,7 +823,7 @@
           addsimps [@{thm knows_Spy_Inputs_secureM_sr_Spy},
                     @{thm knows_Spy_Outpts_secureM_sr_Spy},
                     @{thm shouprubin_assumes_securemeans}, 
-                    @{thm analz_image_Key_Un_Nonce}]))]))) *}
+                    @{thm analz_image_Key_Un_Nonce}]))])))\<close>
     "for proving the Session Key Compromise theorem for smartcard protocols"
 
 
@@ -1320,7 +1320,7 @@
 apply auto
 done
 
-text{*@{term step2_integrity} also is a reliability theorem*}
+text\<open>@{term step2_integrity} also is a reliability theorem\<close>
 lemma Says_Server_message_form: 
      "\<lbrakk> Says Server A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs;  
          evs \<in> sr \<rbrakk>                   
@@ -1333,11 +1333,11 @@
 done
 (*cannot be made useful to A in form of a Gets event*)
 
-text{*
+text\<open>
   step4integrity is @{term Outpts_A_Card_form_4}
 
   step7integrity is @{term Outpts_B_Card_form_7}
-*}
+\<close>
 
 lemma step8_integrity: 
      "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs;  
@@ -1351,10 +1351,10 @@
 done
 
 
-text{*  step9integrity is @{term Inputs_A_Card_form_9}
+text\<open>step9integrity is @{term Inputs_A_Card_form_9}
 
         step10integrity is @{term Outpts_A_Card_form_10}.
-*}
+\<close>
 
 lemma step11_integrity: 
      "\<lbrakk> Says A B (Certificate) \<in> set evs; 
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -1,15 +1,15 @@
 (*  Author:     Giampaolo Bella, Catania University
 *)
 
-section{*Bella's modification of the Shoup-Rubin protocol*}
+section\<open>Bella's modification of the Shoup-Rubin protocol\<close>
 
 theory ShoupRubinBella imports Smartcard begin
 
-text{*The modifications are that message 7 now mentions A, while message 10
+text\<open>The modifications are that message 7 now mentions A, while message 10
 now mentions Nb and B. The lack of explicitness of the original version was
 discovered by investigating adherence to the principle of Goal
 Availability. Only the updated version makes the goals of confidentiality,
-authentication and key distribution available to both peers.*}
+authentication and key distribution available to both peers.\<close>
 
 axiomatization sesK :: "nat*key => key"
 where
@@ -742,7 +742,7 @@
 
 
 ML
-{*
+\<open>
 structure ShoupRubinBella =
 struct
 
@@ -767,18 +767,18 @@
          REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
 
 end
-*}
+\<close>
 
-method_setup prepare = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.prepare_tac ctxt)) *}
+method_setup prepare = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.prepare_tac ctxt))\<close>
   "to launch a few simple facts that will help the simplifier"
 
-method_setup parts_prepare = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.parts_prepare_tac ctxt)) *}
+method_setup parts_prepare = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.parts_prepare_tac ctxt))\<close>
   "additional facts to reason about parts"
 
-method_setup analz_prepare = {*
-    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.analz_prepare_tac ctxt)) *}
+method_setup analz_prepare = \<open>
+    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.analz_prepare_tac ctxt))\<close>
   "additional facts to reason about analz"
 
 
@@ -823,7 +823,7 @@
 apply auto
 done
 
-method_setup sc_analz_freshK = {*
+method_setup sc_analz_freshK = \<open>
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
@@ -832,7 +832,7 @@
               addsimps [@{thm knows_Spy_Inputs_secureM_srb_Spy},
                   @{thm knows_Spy_Outpts_secureM_srb_Spy},
                   @{thm shouprubin_assumes_securemeans},
-                  @{thm analz_image_Key_Un_Nonce}]))]))) *}
+                  @{thm analz_image_Key_Un_Nonce}]))])))\<close>
     "for proving the Session Key Compromise theorem for smartcard protocols"
 
 
@@ -1319,7 +1319,7 @@
 apply auto
 done
 
-text{*@{term step2_integrity} also is a reliability theorem*}
+text\<open>@{term step2_integrity} also is a reliability theorem\<close>
 lemma Says_Server_message_form: 
      "\<lbrakk> Says Server A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs;  
          evs \<in> srb \<rbrakk>                   
@@ -1332,11 +1332,11 @@
 done
 (*cannot be made useful to A in form of a Gets event*)
 
-text{*
+text\<open>
   step4integrity is @{term Outpts_A_Card_form_4}
 
   step7integrity is @{term Outpts_B_Card_form_7}
-*}
+\<close>
 
 lemma step8_integrity: 
      "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs;  
@@ -1350,9 +1350,9 @@
 done
 
 
-text{*  step9integrity is @{term Inputs_A_Card_form_9}
+text\<open>step9integrity is @{term Inputs_A_Card_form_9}
         step10integrity is @{term Outpts_A_Card_form_10}.
-*}
+\<close>
 
 
 lemma step11_integrity: 
--- a/src/HOL/Auth/Smartcard/Smartcard.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -1,20 +1,20 @@
 (* Author:     Giampaolo Bella, Catania University
 *)
 
-section{*Theory of smartcards*}
+section\<open>Theory of smartcards\<close>
 
 theory Smartcard
 imports EventSC "../All_Symmetric"
 begin
 
-text{*  
+text\<open>
 As smartcards handle long-term (symmetric) keys, this theoy extends and 
 supersedes theory Private.thy
 
 An agent is bad if she reveals her PIN to the spy, not the shared key that
 is embedded in her card. An agent's being bad implies nothing about her 
 smartcard, which independently may be stolen or cloned.
-*}
+\<close>
 
 axiomatization
   shrK    :: "agent => key" and  (*long-term keys saved in smart cards*)
@@ -25,9 +25,9 @@
   Pairkey :: "agent * agent => nat" and
   pairK   :: "agent * agent => key"
 where
-  inj_shrK: "inj shrK" and  --{*No two smartcards store the same key*}
-  inj_crdK: "inj crdK" and  --{*Nor do two cards*}
-  inj_pin : "inj pin" and   --{*Nor do two agents have the same pin*}
+  inj_shrK: "inj shrK" and  \<comment>\<open>No two smartcards store the same key\<close>
+  inj_crdK: "inj crdK" and  \<comment>\<open>Nor do two cards\<close>
+  inj_pin : "inj pin" and   \<comment>\<open>Nor do two agents have the same pin\<close>
 
   (*pairK is injective on each component, if we assume encryption to be a PRF
     or at least collision free *)
@@ -49,7 +49,7 @@
   illegalUse_def: "illegalUse (Card A) = ( (Card A \<in> stolen \<and> A \<in> bad)  \<or>  Card A \<in> cloned )"
 
 
-text{*initState must be defined with care*}
+text\<open>initState must be defined with care\<close>
 
 overloading
   initState \<equiv> initState
@@ -76,7 +76,7 @@
 
 end
 
-text{*Still relying on axioms*}
+text\<open>Still relying on axioms\<close>
 axiomatization where
   Key_supply_ax:  "finite KK \<Longrightarrow> \<exists> K. K \<notin> KK & Key K \<notin> used evs" and
 
@@ -89,7 +89,7 @@
 
 
 
-subsection{*Basic properties of shrK*}
+subsection\<open>Basic properties of shrK\<close>
 
 (*Injectiveness: Agents' long-term keys are distinct.*)
 declare inj_shrK [THEN inj_eq, iff]
@@ -106,14 +106,14 @@
      "\<lbrakk> Crypt K X \<in> analz H;  Key K  \<in> analz H \<rbrakk> \<Longrightarrow> X \<in> analz H"
 by auto
 
-text{*Now cancel the @{text dest} attribute given to
- @{text analz.Decrypt} in its declaration.*}
+text\<open>Now cancel the \<open>dest\<close> attribute given to
+ \<open>analz.Decrypt\<close> in its declaration.\<close>
 declare analz.Decrypt [rule del]
 
-text{*Rewrites should not refer to  @{term "initState(Friend i)"} because
-  that expression is not in normal form.*}
+text\<open>Rewrites should not refer to  @{term "initState(Friend i)"} because
+  that expression is not in normal form.\<close>
 
-text{*Added to extend initstate with set of nonces*}
+text\<open>Added to extend initstate with set of nonces\<close>
 lemma parts_image_Nonce [simp]: "parts (Nonce`N) = Nonce`N"
 apply auto
 apply (erule parts.induct)
@@ -135,7 +135,7 @@
 by (drule Crypt_imp_invKey_keysFor, simp)
 
 
-subsection{*Function "knows"*}
+subsection\<open>Function "knows"\<close>
 
 (*Spy knows the pins of bad agents!*)
 lemma Spy_knows_bad [intro!]: "A \<in> bad \<Longrightarrow> Key (pin A) \<in> knows Spy evs"
@@ -260,7 +260,7 @@
 declare pairK_neq [THEN not_sym, simp]
 
 
-subsection{*Fresh nonces*}
+subsection\<open>Fresh nonces\<close>
 
 lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState (Friend i))"
 by auto
@@ -276,7 +276,7 @@
 So, we must use old-style supply fresh nonce theorems relying on the appropriate axiom*)
 
 
-subsection{*Supply fresh nonces for possibility theorems.*}
+subsection\<open>Supply fresh nonces for possibility theorems.\<close>
 
 
 lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs"
@@ -309,14 +309,14 @@
 
 
 
-text{*Unlike the corresponding property of nonces, we cannot prove
+text\<open>Unlike the corresponding property of nonces, we cannot prove
     @{term "finite KK \<Longrightarrow> \<exists>K. K \<notin> KK & Key K \<notin> used evs"}.
     We have infinitely many agents and there is nothing to stop their
     long-term keys from exhausting all the natural numbers.  Instead,
-    possibility theorems must assume the existence of a few keys.*}
+    possibility theorems must assume the existence of a few keys.\<close>
 
 
-subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
+subsection\<open>Specialized Rewriting for Theorems About @{term analz} and Image\<close>
 
 lemma subset_Compl_range_shrK: "A \<subseteq> - (range shrK) \<Longrightarrow> shrK x \<notin> A"
 by blast
@@ -343,7 +343,7 @@
     erase occurrences of forwarded message components (X). **)
 
 lemmas analz_image_freshK_simps =
-       simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
+       simp_thms mem_simps \<comment>\<open>these two allow its use with \<open>only:\<close>\<close>
        disj_comms 
        image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset
        analz_insert_eq Un_upper2 [THEN analz_mono, THEN [2] rev_subsetD]
@@ -358,10 +358,10 @@
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
 
-subsection{*Tactics for possibility theorems*}
+subsection\<open>Tactics for possibility theorems\<close>
 
 ML
-{*
+\<open>
 structure Smartcard =
 struct
 
@@ -390,7 +390,7 @@
                delsimps [@{thm imp_disjL}]    (*reduces blow-up*)
                addsimps @{thms analz_image_freshK_simps})
 end
-*}
+\<close>
 
 
 (*Lets blast_tac perform this step without needing the simplifier*)
@@ -400,22 +400,22 @@
 
 (*Specialized methods*)
 
-method_setup analz_freshK = {*
+method_setup analz_freshK = \<open>
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
           REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}),
-          ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt))]))) *}
+          ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt))])))\<close>
     "for proving the Session Key Compromise theorem"
 
-method_setup possibility = {*
+method_setup possibility = \<open>
     Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD (Smartcard.possibility_tac ctxt)) *}
+        SIMPLE_METHOD (Smartcard.possibility_tac ctxt))\<close>
     "for proving possibility theorems"
 
-method_setup basic_possibility = {*
+method_setup basic_possibility = \<open>
     Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD (Smartcard.basic_possibility_tac ctxt)) *}
+        SIMPLE_METHOD (Smartcard.basic_possibility_tac ctxt))\<close>
     "for proving possibility theorems"
 
 lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)"
--- a/src/HOL/Auth/TLS.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/TLS.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -38,16 +38,16 @@
 Notes A {|Agent B, Nonce PMS|}.
 *)
 
-section{*The TLS Protocol: Transport Layer Security*}
+section\<open>The TLS Protocol: Transport Layer Security\<close>
 
 theory TLS imports Public "~~/src/HOL/Library/Nat_Bijection" begin
 
 definition certificate :: "[agent,key] => msg" where
     "certificate A KA == Crypt (priSK Server) {|Agent A, Key KA|}"
 
-text{*TLS apparently does not require separate keypairs for encryption and
+text\<open>TLS apparently does not require separate keypairs for encryption and
 signature.  Therefore, we formalize signature as encryption using the
-private encryption key.*}
+private encryption key.\<close>
 
 datatype role = ClientRole | ServerRole
 
@@ -72,14 +72,14 @@
 
 specification (PRF)
   inj_PRF: "inj PRF"
-  --{*the pseudo-random function is collision-free*}
+  \<comment>\<open>the pseudo-random function is collision-free\<close>
    apply (rule exI [of _ "%(x,y,z). prod_encode(x, prod_encode(y,z))"])
    apply (simp add: inj_on_def prod_encode_eq)
    done
 
 specification (sessionK)
   inj_sessionK: "inj sessionK"
-  --{*sessionK is collision-free; also, no clientK clashes with any serverK.*}
+  \<comment>\<open>sessionK is collision-free; also, no clientK clashes with any serverK.\<close>
    apply (rule exI [of _ 
          "%((x,y,z), r). prod_encode(case_role 0 1 r, 
                            prod_encode(x, prod_encode(y,z)))"])
@@ -87,66 +87,66 @@
    done
 
 axiomatization where
-  --{*sessionK makes symmetric keys*}
+  \<comment>\<open>sessionK makes symmetric keys\<close>
   isSym_sessionK: "sessionK nonces \<in> symKeys" and
 
-  --{*sessionK never clashes with a long-term symmetric key  
-     (they don't exist in TLS anyway)*}
+  \<comment>\<open>sessionK never clashes with a long-term symmetric key  
+     (they don't exist in TLS anyway)\<close>
   sessionK_neq_shrK [iff]: "sessionK nonces \<noteq> shrK A"
 
 
 inductive_set tls :: "event list set"
   where
-   Nil:  --{*The initial, empty trace*}
+   Nil:  \<comment>\<open>The initial, empty trace\<close>
          "[] \<in> tls"
 
- | Fake: --{*The Spy may say anything he can say.  The sender field is correct,
-          but agents don't use that information.*}
+ | 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"
 
- | SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
-                to available nonces*}
+ | SpyKeys: \<comment>\<open>The spy may apply @{term PRF} and @{term sessionK}
+                to available nonces\<close>
          "[| evsSK \<in> tls;
              {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
           ==> Notes Spy {| Nonce (PRF(M,NA,NB)),
                            Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
 
  | ClientHello:
-         --{*(7.4.1.2)
-           PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
+         \<comment>\<open>(7.4.1.2)
+           PA represents \<open>CLIENT_VERSION\<close>, \<open>CIPHER_SUITES\<close> and \<open>COMPRESSION_METHODS\<close>.
            It is uninterpreted but will be confirmed in the FINISHED messages.
-           NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
+           NA is CLIENT RANDOM, while SID is \<open>SESSION_ID\<close>.
            UNIX TIME is omitted because the protocol doesn't use it.
            May assume @{term "NA \<notin> range PRF"} because CLIENT RANDOM is 
-           28 bytes while MASTER SECRET is 48 bytes*}
+           28 bytes while MASTER SECRET is 48 bytes\<close>
          "[| evsCH \<in> tls;  Nonce NA \<notin> used evsCH;  NA \<notin> range PRF |]
           ==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
                 # evsCH  \<in>  tls"
 
  | ServerHello:
-         --{*7.4.1.3 of the TLS Internet-Draft
-           PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
+         \<comment>\<open>7.4.1.3 of the TLS Internet-Draft
+           PB represents \<open>CLIENT_VERSION\<close>, \<open>CIPHER_SUITE\<close> and \<open>COMPRESSION_METHOD\<close>.
            SERVER CERTIFICATE (7.4.2) is always present.
-           @{text CERTIFICATE_REQUEST} (7.4.4) is implied.*}
+           \<open>CERTIFICATE_REQUEST\<close> (7.4.4) is implied.\<close>
          "[| evsSH \<in> tls;  Nonce NB \<notin> used evsSH;  NB \<notin> range PRF;
              Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
                \<in> set evsSH |]
           ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  \<in>  tls"
 
  | Certificate:
-         --{*SERVER (7.4.2) or CLIENT (7.4.6) 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"
 
  | ClientKeyExch:
-         --{*CLIENT KEY EXCHANGE (7.4.7).
+         \<comment>\<open>CLIENT KEY EXCHANGE (7.4.7).
            The client, A, chooses PMS, the PREMASTER SECRET.
            She encrypts PMS using the supplied KB, which ought to be pubK B.
            We assume @{term "PMS \<notin> range PRF"} because a clash betweem the PMS
            and another MASTER SECRET is highly unlikely (even though
            both items have the same length, 48 bytes).
            The Note event records in the trace that she knows PMS
-               (see REMARK at top). *}
+               (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))
@@ -154,28 +154,28 @@
               # evsCX  \<in>  tls"
 
  | CertVerify:
-        --{*The optional Certificate Verify (7.4.8) message contains the
+        \<comment>\<open>The optional Certificate Verify (7.4.8) message contains the
           specific components listed in the security analysis, F.1.1.2.
           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*}
+          assures B of A's presence\<close>
          "[| evsCV \<in> tls;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCV;
              Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
           ==> Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))
               # evsCV  \<in>  tls"
 
-        --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
+        \<comment>\<open>Finally come the FINISHED messages (7.4.8), confirming PA and PB
           among other things.  The master-secret is PRF(PMS,NA,NB).
-          Either party may send its message first.*}
+          Either party may send its message first.\<close>
 
  | ClientFinished:
-        --{*The occurrence of Notes A {|Agent B, Nonce PMS|} stops the
+        \<comment>\<open>The occurrence of Notes A {|Agent B, Nonce PMS|} stops the
           rule's applying when the Spy has satisfied the "Says A B" by
           repaying messages sent by the true client; in that case, the
           Spy does not know PMS and could not send ClientFinished.  One
           could simply put @{term "A\<noteq>Spy"} into the rule, but one should not
-          expect the spy to be well-behaved.*}
+          expect the spy to be well-behaved.\<close>
          "[| evsCF \<in> tls;
              Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
                \<in> set evsCF;
@@ -189,8 +189,8 @@
               # evsCF  \<in>  tls"
 
  | ServerFinished:
-        --{*Keeping A' and A'' distinct means B cannot even check that the
-          two messages originate from the same source. *}
+        \<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;
              Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
                \<in> set evsSF;
@@ -204,10 +204,10 @@
               # evsSF  \<in>  tls"
 
  | ClientAccepts:
-        --{*Having transmitted ClientFinished and received an identical
+        \<comment>\<open>Having transmitted ClientFinished and received an identical
           message encrypted with serverK, the client stores the parameters
           needed to resume this session.  The "Notes A ..." premise is
-          used to prove @{text Notes_master_imp_Crypt_PMS}.*}
+          used to prove \<open>Notes_master_imp_Crypt_PMS\<close>.\<close>
          "[| evsCA \<in> tls;
              Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
              M = PRF(PMS,NA,NB);
@@ -220,10 +220,10 @@
              Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA  \<in>  tls"
 
  | ServerAccepts:
-        --{*Having transmitted ServerFinished and received an identical
+        \<comment>\<open>Having transmitted ServerFinished and received an identical
           message encrypted with clientK, the server stores the parameters
           needed to resume this session.  The "Says A'' B ..." premise is
-          used to prove @{text Notes_master_imp_Crypt_PMS}.*}
+          used to prove \<open>Notes_master_imp_Crypt_PMS\<close>.\<close>
          "[| evsSA \<in> tls;
              A \<noteq> B;
              Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
@@ -237,8 +237,8 @@
              Notes B {|Number SID, Agent A, Agent B, Nonce M|} # evsSA  \<in>  tls"
 
  | ClientResume:
-         --{*If A recalls the @{text SESSION_ID}, then she sends a FINISHED
-             message using the new nonces and stored MASTER SECRET.*}
+         \<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;
              Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCR;
@@ -250,8 +250,8 @@
               # evsCR  \<in>  tls"
 
  | ServerResume:
-         --{*Resumption (7.3):  If B finds the @{text SESSION_ID} then he can 
-             send a FINISHED message using the recovered MASTER SECRET*}
+         \<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;
              Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
              Says B  A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
@@ -263,11 +263,11 @@
                 \<in>  tls"
 
  | Oops:
-         --{*The most plausible compromise is of an old session key.  Losing
+         \<comment>\<open>The most plausible compromise is of an old session key.  Losing
            the MASTER SECRET or PREMASTER SECRET is more serious but
            rather unlikely.  The assumption @{term "A\<noteq>Spy"} is essential: 
            otherwise the Spy could learn session keys merely by 
-           replaying messages!*}
+           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"
@@ -293,10 +293,10 @@
 declare Fake_parts_insert_in_Un  [dest]
 
 
-text{*Automatically unfold the definition of "certificate"*}
+text\<open>Automatically unfold the definition of "certificate"\<close>
 declare certificate_def [simp]
 
-text{*Injectiveness of key-generating functions*}
+text\<open>Injectiveness of key-generating functions\<close>
 declare inj_PRF [THEN inj_eq, iff]
 declare inj_sessionK [THEN inj_eq, iff]
 declare isSym_sessionK [simp]
@@ -317,10 +317,10 @@
 lemmas keys_distinct = pubK_neq_sessionK priK_neq_sessionK
 
 
-subsection{*Protocol Proofs*}
+subsection\<open>Protocol Proofs\<close>
 
-text{*Possibility properties state that some traces run the protocol to the
-end.  Four paths and 12 rules are considered.*}
+text\<open>Possibility properties state that some traces run the protocol to the
+end.  Four paths and 12 rules are considered.\<close>
 
 
 (** These proofs assume that the Nonce_supply nonces
@@ -330,7 +330,7 @@
 **)
 
 
-text{*Possibility property ending with ClientAccepts.*}
+text\<open>Possibility property ending with ClientAccepts.\<close>
 lemma "[| \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B |]
       ==> \<exists>SID M. \<exists>evs \<in> tls.
             Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs"
@@ -343,7 +343,7 @@
 done
 
 
-text{*And one for ServerAccepts.  Either FINISHED message may come first.*}
+text\<open>And one for ServerAccepts.  Either FINISHED message may come first.\<close>
 lemma "[| \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF; A \<noteq> B |]
       ==> \<exists>SID NA PA NB PB M. \<exists>evs \<in> tls.
            Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs"
@@ -356,7 +356,7 @@
 done
 
 
-text{*Another one, for CertVerify (which is optional)*}
+text\<open>Another one, for CertVerify (which is optional)\<close>
 lemma "[| \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B |]
        ==> \<exists>NB PMS. \<exists>evs \<in> tls.
               Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|})) 
@@ -369,8 +369,8 @@
 done
 
 
-text{*Another one, for session resumption (both ServerResume and ClientResume).
-  NO tls.Nil here: we refer to a previous session, not the empty trace.*}
+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;
           Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs0;
           Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs0;
@@ -389,13 +389,13 @@
 done
 
 
-subsection{*Inductive proofs about tls*}
+subsection\<open>Inductive proofs about tls\<close>
 
 
 (** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY
     sends messages containing X! **)
 
-text{*Spy never sees a good agent's private key!*}
+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)"
 by (erule tls.induct, force, simp_all, blast)
@@ -409,10 +409,10 @@
 by (blast dest: Spy_see_priK)
 
 
-text{*This lemma says that no false certificates exist.  One might extend the
+text\<open>This lemma says that no false certificates exist.  One might extend the
   model to include bogus certificates for the agents, but there seems
   little point in doing so: the loss of their private keys is a worse
-  breach of security.*}
+  breach of security.\<close>
 lemma certificate_valid:
     "[| certificate B KB \<in> parts (spies evs);  evs \<in> tls |] ==> KB = pubK B"
 apply (erule rev_mp)
@@ -422,7 +422,7 @@
 lemmas CX_KB_is_pubKB = Says_imp_spies [THEN parts.Inj, THEN certificate_valid]
 
 
-subsubsection{*Properties of items found in Notes*}
+subsubsection\<open>Properties of items found in Notes\<close>
 
 lemma Notes_Crypt_parts_spies:
      "[| Notes A {|Agent B, X|} \<in> set evs;  evs \<in> tls |]
@@ -433,34 +433,34 @@
 apply (blast intro: parts_insertI)
 done
 
-text{*C may be either A or B*}
+text\<open>C may be either A or B\<close>
 lemma Notes_master_imp_Crypt_PMS:
      "[| Notes C {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \<in> set evs;
          evs \<in> tls |]
       ==> Crypt (pubK B) (Nonce PMS) \<in> parts (spies evs)"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast intro: parts_insertI)
-txt{*Client, Server Accept*}
+txt\<open>Client, Server Accept\<close>
 apply (blast dest!: Notes_Crypt_parts_spies)+
 done
 
-text{*Compared with the theorem above, both premise and conclusion are stronger*}
+text\<open>Compared with the theorem above, both premise and conclusion are stronger\<close>
 lemma Notes_master_imp_Notes_PMS:
      "[| Notes A {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \<in> set evs;
          evs \<in> tls |]
       ==> Notes A {|Agent B, Nonce PMS|} \<in> set evs"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all)
-txt{*ServerAccepts*}
+txt\<open>ServerAccepts\<close>
 apply blast
 done
 
 
-subsubsection{*Protocol goal: if B receives CertVerify, then A sent it*}
+subsubsection\<open>Protocol goal: if B receives CertVerify, then A sent it\<close>
 
-text{*B can check A's signature if he has received A's certificate.*}
+text\<open>B can check A's signature if he has received A's certificate.\<close>
 lemma TrustCertVerify_lemma:
      "[| X \<in> parts (spies evs);
          X = Crypt (priK A) (Hash{|nb, Agent B, pms|});
@@ -470,7 +470,7 @@
 apply (erule tls.induct, force, simp_all, blast)
 done
 
-text{*Final version: B checks X using the distributed KA instead of priK A*}
+text\<open>Final version: B checks X using the distributed KA instead of priK A\<close>
 lemma TrustCertVerify:
      "[| X \<in> parts (spies evs);
          X = Crypt (invKey KA) (Hash{|nb, Agent B, pms|});
@@ -480,7 +480,7 @@
 by (blast dest!: certificate_valid intro!: TrustCertVerify_lemma)
 
 
-text{*If CertVerify is present then A has chosen PMS.*}
+text\<open>If CertVerify is present then A has chosen PMS.\<close>
 lemma UseCertVerify_lemma:
      "[| Crypt (priK A) (Hash{|nb, Agent B, Nonce PMS|}) \<in> parts (spies evs);
          evs \<in> tls;  A \<notin> bad |]
@@ -489,7 +489,7 @@
 apply (erule tls.induct, force, simp_all, blast)
 done
 
-text{*Final version using the distributed KA instead of priK A*}
+text\<open>Final version using the distributed KA instead of priK A\<close>
 lemma UseCertVerify:
      "[| Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|})
            \<in> parts (spies evs);
@@ -502,7 +502,7 @@
 lemma no_Notes_A_PRF [simp]:
      "evs \<in> tls ==> Notes A {|Agent B, Nonce (PRF x)|} \<notin> set evs"
 apply (erule tls.induct, force, simp_all)
-txt{*ClientKeyExch: PMS is assumed to differ from any PRF.*}
+txt\<open>ClientKeyExch: PMS is assumed to differ from any PRF.\<close>
 apply blast
 done
 
@@ -512,18 +512,18 @@
       ==> Nonce PMS \<in> parts (spies evs)"
 apply (erule rev_mp)
 apply (erule tls.induct, force, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast intro: parts_insertI)
-txt{*Easy, e.g. by freshness*}
+txt\<open>Easy, e.g. by freshness\<close>
 apply (blast dest: Notes_Crypt_parts_spies)+
 done
 
 
 
 
-subsubsection{*Unicity results for PMS, the pre-master-secret*}
+subsubsection\<open>Unicity results for PMS, the pre-master-secret\<close>
 
-text{*PMS determines B.*}
+text\<open>PMS determines B.\<close>
 lemma Crypt_unique_PMS:
      "[| Crypt(pubK B)  (Nonce PMS) \<in> parts (spies evs);
          Crypt(pubK B') (Nonce PMS) \<in> parts (spies evs);
@@ -532,7 +532,7 @@
       ==> B=B'"
 apply (erule rev_mp, erule rev_mp, erule rev_mp)
 apply (erule tls.induct, analz_mono_contra, force, simp_all (no_asm_simp))
-txt{*Fake, ClientKeyExch*}
+txt\<open>Fake, ClientKeyExch\<close>
 apply blast+
 done
 
@@ -543,7 +543,7 @@
     determines B alone, and only if PMS is secret.
 **)
 
-text{*In A's internal Note, PMS determines A and B.*}
+text\<open>In A's internal Note, PMS determines A and B.\<close>
 lemma Notes_unique_PMS:
      "[| Notes A  {|Agent B,  Nonce PMS|} \<in> set evs;
          Notes A' {|Agent B', Nonce PMS|} \<in> set evs;
@@ -551,15 +551,15 @@
       ==> A=A' & B=B'"
 apply (erule rev_mp, erule rev_mp)
 apply (erule tls.induct, force, simp_all)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: Notes_Crypt_parts_spies)
 done
 
 
-subsection{*Secrecy Theorems*}
+subsection\<open>Secrecy Theorems\<close>
 
-text{*Key compromise lemma needed to prove @{term analz_image_keys}.
-  No collection of keys can help the spy get new private keys.*}
+text\<open>Key compromise lemma needed to prove @{term analz_image_keys}.
+  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 Un (spies evs))) =
@@ -569,18 +569,18 @@
                 del: image_insert
                 add: image_Un [THEN sym]
                      insert_Key_image Un_assoc [THEN sym])
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
 done
 
 
-text{*slightly speeds up the big simplification below*}
+text\<open>slightly speeds up the big simplification below\<close>
 lemma range_sessionkeys_not_priK:
      "KK <= range sessionK ==> priK B \<notin> KK"
 by blast
 
 
-text{*Lemma for the trivial direction of the if-and-only-if*}
+text\<open>Lemma for the trivial direction of the if-and-only-if\<close>
 lemma analz_image_keys_lemma:
      "(X \<in> analz (G Un H)) --> (X \<in> analz H)  ==>
       (X \<in> analz (G Un H))  =  (X \<in> analz H)"
@@ -605,11 +605,11 @@
                      insert_Key_singleton
                      range_sessionkeys_not_priK analz_image_priK)
 apply (simp_all add: insert_absorb)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
 done
 
-text{*Knowing some session keys is no help in getting new nonces*}
+text\<open>Knowing some session keys is no help in getting new nonces\<close>
 lemma analz_insert_key [simp]:
      "evs \<in> tls ==>
       (Nonce N \<in> analz (insert (Key (sessionK z)) (spies evs))) =
@@ -618,15 +618,15 @@
          add: insert_Key_singleton analz_image_keys)
 
 
-subsubsection{*Protocol goal: serverK(Na,Nb,M) and clientK(Na,Nb,M) remain secure*}
+subsubsection\<open>Protocol goal: serverK(Na,Nb,M) and clientK(Na,Nb,M) remain secure\<close>
 
 (** Some lemmas about session keys, comprising clientK and serverK **)
 
 
-text{*Lemma: session keys are never used if PMS is fresh.
+text\<open>Lemma: session keys are never used if PMS is fresh.
   Nonces don't have to agree, allowing session resumption.
   Converse doesn't hold; revealing PMS doesn't force the keys to be sent.
-  THEY ARE NOT SUITABLE AS SAFE ELIM RULES.*}
+  THEY ARE NOT SUITABLE AS SAFE ELIM RULES.\<close>
 lemma PMS_lemma:
      "[| Nonce PMS \<notin> parts (spies evs);
          K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role);
@@ -635,11 +635,11 @@
 apply (erule rev_mp, erule ssubst)
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) 
 apply (force, simp_all (no_asm_simp))
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast intro: parts_insertI)
-txt{*SpyKeys*}
+txt\<open>SpyKeys\<close>
 apply blast
-txt{*Many others*}
+txt\<open>Many others\<close>
 apply (force dest!: Notes_Crypt_parts_spies Notes_master_imp_Crypt_PMS)+
 done
 
@@ -655,11 +655,11 @@
       ==> Nonce PMS \<in> parts (spies evs)"
 by (blast dest: PMS_lemma)
 
-text{*Write keys are never sent if M (MASTER SECRET) is secure.
+text\<open>Write keys are never sent if M (MASTER SECRET) is secure.
   Converse fails; betraying M doesn't force the keys to be sent!
   The strong Oops condition can be weakened later by unicity reasoning,
   with some effort.
-  NO LONGER USED: see @{text clientK_not_spied} and @{text serverK_not_spied}*}
+  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 |]
@@ -667,57 +667,57 @@
 apply (erule rev_mp, erule rev_mp)
 apply (erule tls.induct, analz_mono_contra)
 apply (force, simp_all (no_asm_simp))
-txt{*Fake, SpyKeys*}
+txt\<open>Fake, SpyKeys\<close>
 apply blast+
 done
 
 
-text{*If A sends ClientKeyExch to an honest B, then the PMS will stay secret.*}
+text\<open>If A sends ClientKeyExch to an honest B, then the PMS will stay secret.\<close>
 lemma Spy_not_see_PMS:
      "[| Notes A {|Agent B, Nonce PMS|} \<in> set evs;
          evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
       ==> 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{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*SpyKeys*}
+txt\<open>SpyKeys\<close>
 apply force
 apply (simp_all add: insert_absorb) 
-txt{*ClientHello, ServerHello, ClientKeyExch: mostly freshness reasoning*}
+txt\<open>ClientHello, ServerHello, ClientKeyExch: mostly freshness reasoning\<close>
 apply (blast dest: Notes_Crypt_parts_spies)
 apply (blast dest: Notes_Crypt_parts_spies)
 apply (blast dest: Notes_Crypt_parts_spies)
-txt{*ClientAccepts and ServerAccepts: because @{term "PMS \<notin> range PRF"}*}
+txt\<open>ClientAccepts and ServerAccepts: because @{term "PMS \<notin> range PRF"}\<close>
 apply force+
 done
 
 
-text{*If A sends ClientKeyExch to an honest B, then the MASTER SECRET
-  will stay secret.*}
+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 {|Agent B, Nonce PMS|} \<in> set evs;
          evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
       ==> 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{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*SpyKeys: by secrecy of the PMS, Spy cannot make the MS*}
+txt\<open>SpyKeys: by secrecy of the PMS, Spy cannot make the MS\<close>
 apply (blast dest!: Spy_not_see_PMS)
 apply (simp_all add: insert_absorb)
-txt{*ClientAccepts and ServerAccepts: because PMS was already visible;
-  others, freshness etc.*}
+txt\<open>ClientAccepts and ServerAccepts: because PMS was already visible;
+  others, freshness etc.\<close>
 apply (blast dest: Notes_Crypt_parts_spies Spy_not_see_PMS 
                    Notes_imp_knows_Spy [THEN analz.Inj])+
 done
 
 
 
-subsubsection{*Weakening the Oops conditions for leakage of clientK*}
+subsubsection\<open>Weakening the Oops conditions for leakage of clientK\<close>
 
-text{*If A created PMS then nobody else (except the Spy in replays)
-  would send a message using a clientK generated from that PMS.*}
+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;
          Notes A {|Agent B, Nonce PMS|} \<in> set evs;
@@ -726,16 +726,16 @@
 apply (erule rev_mp, erule rev_mp)
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (force, simp_all)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: PMS_Crypt_sessionK_not_spied)
-txt{*ClientFinished, ClientResume: by unicity of PMS*}
+txt\<open>ClientFinished, ClientResume: by unicity of PMS\<close>
 apply (blast dest!: Notes_master_imp_Notes_PMS 
              intro: Notes_unique_PMS [THEN conjunct1])+
 done
 
 
-text{*If A created PMS and has not leaked her clientK to the Spy,
-  then it is completely secure: not even in parts!*}
+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 {|Agent B, Nonce PMS|} \<in> set evs;
          Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) \<notin> set evs;
@@ -745,21 +745,21 @@
 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))
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply blast 
-txt{*SpyKeys*}
+txt\<open>SpyKeys\<close>
 apply (blast dest!: Spy_not_see_MS)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: PMS_sessionK_not_spied)
-txt{*Oops*}
+txt\<open>Oops\<close>
 apply (blast intro: Says_clientK_unique)
 done
 
 
-subsubsection{*Weakening the Oops conditions for leakage of serverK*}
+subsubsection\<open>Weakening the Oops conditions for leakage of serverK\<close>
 
-text{*If A created PMS for B, then nobody other than B or the Spy would
-  send a message using a serverK generated from that PMS.*}
+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;
          Notes A {|Agent B, Nonce PMS|} \<in> set evs;
@@ -768,16 +768,16 @@
 apply (erule rev_mp, erule rev_mp)
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (force, simp_all)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: PMS_Crypt_sessionK_not_spied)
-txt{*ServerResume, ServerFinished: by unicity of PMS*}
+txt\<open>ServerResume, ServerFinished: by unicity of PMS\<close>
 apply (blast dest!: Notes_master_imp_Crypt_PMS 
              dest: Spy_not_see_PMS Notes_Crypt_parts_spies Crypt_unique_PMS)+
 done
 
 
-text{*If A created PMS for B, and B has not leaked his serverK to the Spy,
-  then it is completely secure: not even in parts!*}
+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 {|Agent B, Nonce PMS|} \<in> set evs;
          Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) \<notin> set evs;
@@ -786,22 +786,22 @@
 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))
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast 
-txt{*SpyKeys*}
+txt\<open>SpyKeys\<close>
 apply (blast dest!: Spy_not_see_MS)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: PMS_sessionK_not_spied)
-txt{*Oops*}
+txt\<open>Oops\<close>
 apply (blast intro: Says_serverK_unique)
 done
 
 
-subsubsection{*Protocol goals: if A receives ServerFinished, then B is present
+subsubsection\<open>Protocol goals: if A receives ServerFinished, then B is present
      and has used the quoted values PA, PB, etc.  Note that it is up to A
-     to compare PA with what she originally sent.*}
+     to compare PA with what she originally sent.\<close>
 
-text{*The mention of her name (A) in X assures A that B knows who she is.*}
+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))
                (Hash{|Number SID, Nonce M,
@@ -815,17 +815,17 @@
 apply (erule ssubst)+
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (force, simp_all (no_asm_simp))
-txt{*Fake: the Spy doesn't have the critical session key!*}
+txt\<open>Fake: the Spy doesn't have the critical session key!\<close>
 apply (blast dest: serverK_not_spied)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: PMS_Crypt_sessionK_not_spied)
 done
 
-text{*This version refers not to ServerFinished but to any message from B.
+text\<open>This version refers not to ServerFinished but to any message from B.
   We don't assume B has received CertVerify, and an intruder could
   have changed A's identity in all other messages, so we can't be sure
   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.*}
+  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 -->
@@ -835,22 +835,22 @@
 apply (erule ssubst)
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (force, simp_all (no_asm_simp) add: ex_disj_distrib)
-txt{*Fake: the Spy doesn't have the critical session key!*}
+txt\<open>Fake: the Spy doesn't have the critical session key!\<close>
 apply (blast dest: serverK_not_spied)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (clarify, blast dest!: PMS_Crypt_sessionK_not_spied)
-txt{*ServerResume, ServerFinished: by unicity of PMS*}
+txt\<open>ServerResume, ServerFinished: by unicity of PMS\<close>
 apply (blast dest!: Notes_master_imp_Crypt_PMS 
              dest: Spy_not_see_PMS Notes_Crypt_parts_spies Crypt_unique_PMS)+
 done
 
 
-subsubsection{*Protocol goal: if B receives any message encrypted with clientK
-      then A has sent it*}
+subsubsection\<open>Protocol goal: if B receives any message encrypted with clientK
+      then A has sent it\<close>
 
-text{*ASSUMING that A chose PMS.  Authentication is
+text\<open>ASSUMING that A chose PMS.  Authentication is
      assumed here; B cannot verify it.  But if the message is
-     ClientFinished, then B can then check the quoted values PA, PB, etc.*}
+     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 |]
@@ -861,18 +861,18 @@
 apply (erule ssubst)
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (force, simp_all (no_asm_simp))
-txt{*Fake: the Spy doesn't have the critical session key!*}
+txt\<open>Fake: the Spy doesn't have the critical session key!\<close>
 apply (blast dest: clientK_not_spied)
-txt{*ClientKeyExch*}
+txt\<open>ClientKeyExch\<close>
 apply (blast dest!: PMS_Crypt_sessionK_not_spied)
-txt{*ClientFinished, ClientResume: by unicity of PMS*}
+txt\<open>ClientFinished, ClientResume: by unicity of PMS\<close>
 apply (blast dest!: Notes_master_imp_Notes_PMS dest: Notes_unique_PMS)+
 done
 
 
-subsubsection{*Protocol goal: if B receives ClientFinished, and if B is able to
+subsubsection\<open>Protocol goal: if B receives ClientFinished, and if B is able to
      check a CertVerify from A, then A has used the quoted
-     values PA, PB, etc.  Even this one requires A to be uncompromised.*}
+     values PA, PB, etc.  Even this one requires A to be uncompromised.\<close>
 lemma AuthClientFinished:
      "[| M = PRF(PMS,NA,NB);
          Says A Spy (Key(clientK(Na,Nb,M))) \<notin> set evs;
--- a/src/HOL/Auth/WooLam.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/WooLam.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Woo-Lam Protocol*}
+section\<open>The Woo-Lam Protocol\<close>
 
 theory WooLam imports Public begin
 
-text{*Simplified version from page 11 of
+text\<open>Simplified version from page 11 of
   Abadi and Needham (1996). 
   Prudent Engineering Practice for Cryptographic Protocols.
   IEEE Trans. S.E. 22(1), pages 6-15.
@@ -15,7 +15,7 @@
 Note: this differs from the Woo-Lam protocol discussed by Lowe (1996):
   Some New Attacks upon Security Protocols.
   Computer Security Foundations Workshop
-*}
+\<close>
 
 inductive_set woolam :: "event list set"
   where
--- a/src/HOL/Auth/Yahalom.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Yahalom.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,16 +3,16 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Yahalom Protocol*}
+section\<open>The Yahalom Protocol\<close>
 
 theory Yahalom imports Public begin
 
-text{*From page 257 of
+text\<open>From page 257 of
   Burrows, Abadi and Needham (1989).  A Logic of Authentication.
   Proc. Royal Soc. 426
 
 This theory has the prototypical example of a secrecy relation, KeyCryptNonce.
-*}
+\<close>
 
 inductive_set yahalom :: "event list set"
   where
@@ -53,10 +53,10 @@
                 # evs3 \<in> yahalom"
 
  | YM4:  
-       --{*Alice receives the Server's (?) message, checks her Nonce, and
+       \<comment>\<open>Alice receives the Server's (?) message, checks her Nonce, and
            uses the new session key to send Bob his Nonce.  The premise
-           @{term "A \<noteq> Server"} is needed for @{text Says_Server_not_range}.
-           Alice can check that K is symmetric by its length.*}
+           @{term "A \<noteq> Server"} is needed for \<open>Says_Server_not_range\<close>.
+           Alice can check that K is symmetric by its length.\<close>
          "[| evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
              Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
                 \<in> set evs4;
@@ -85,7 +85,7 @@
 declare Fake_parts_insert_in_Un  [dest]
 declare analz_into_parts [dest]
 
-text{*A "possibility property": there are traces that reach the end*}
+text\<open>A "possibility property": there are traces that reach the end\<close>
 lemma "[| A \<noteq> Server; K \<in> symKeys; Key K \<notin> used [] |]
       ==> \<exists>X NB. \<exists>evs \<in> yahalom.
              Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs"
@@ -99,13 +99,13 @@
 done
 
 
-subsection{*Regularity Lemmas for Yahalom*}
+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"
 by (erule rev_mp, erule yahalom.induct, auto)
 
-text{*Must be proved separately for each protocol*}
+text\<open>Must be proved separately for each protocol\<close>
 lemma Gets_imp_knows_Spy:
      "[| Gets B X \<in> set evs; evs \<in> yahalom |]  ==> X \<in> knows Spy evs"
 by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
@@ -114,7 +114,7 @@
 declare Gets_imp_analz_Spy [dest]
 
 
-text{*Lets us treat YM4 using a similar argument as for the Fake case.*}
+text\<open>Lets us treat YM4 using a similar argument as for the Fake case.\<close>
 lemma YM4_analz_knows_Spy:
      "[| Gets A {|Crypt (shrK A) Y, X|} \<in> set evs;  evs \<in> yahalom |]
       ==> X \<in> analz (knows Spy evs)"
@@ -123,16 +123,16 @@
 lemmas YM4_parts_knows_Spy =
        YM4_analz_knows_Spy [THEN analz_into_parts]
 
-text{*For Oops*}
+text\<open>For Oops\<close>
 lemma YM4_Key_parts_knows_Spy:
      "Says Server A {|Crypt (shrK A) {|B,K,NA,NB|}, X|} \<in> set evs
       ==> K \<in> parts (knows Spy evs)"
   by (metis parts.Body parts.Fst parts.Snd  Says_imp_knows_Spy parts.Inj)
 
-text{*Theorems of the form @{term "X \<notin> parts (knows Spy evs)"} imply 
-that NOBODY sends messages containing X! *}
+text\<open>Theorems of the form @{term "X \<notin> parts (knows Spy evs)"} imply 
+that NOBODY sends messages containing X!\<close>
 
-text{*Spy never sees a good agent's shared key!*}
+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)"
 by (erule yahalom.induct, force,
@@ -146,29 +146,29 @@
      "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> yahalom|] ==> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
-text{*Nobody can have used non-existent keys!
-    Needed to apply @{text analz_insert_Key}*}
+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))"
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert, auto)
 done
 
 
-text{*Earlier, all protocol proofs declared this theorem.
-  But only a few proofs need it, e.g. Yahalom and Kerberos IV.*}
+text\<open>Earlier, all protocol proofs declared this theorem.
+  But only a few proofs need it, e.g. Yahalom and Kerberos IV.\<close>
 lemma new_keys_not_analzd:
  "[|K \<in> symKeys; evs \<in> yahalom; Key K \<notin> used evs|]
   ==> K \<notin> keysFor (analz (knows Spy evs))"
 by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD])
 
 
-text{*Describes the form of K when the Server sends this message.  Useful for
-  Oops as well as main secrecy property.*}
+text\<open>Describes the form of K when the Server sends this message.  Useful for
+  Oops as well as main secrecy property.\<close>
 lemma Says_Server_not_range [simp]:
      "[| Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|}
            \<in> set evs;   evs \<in> yahalom |]
@@ -176,7 +176,7 @@
 by (erule rev_mp, erule yahalom.induct, simp_all)
 
 
-subsection{*Secrecy Theorems*}
+subsection\<open>Secrecy Theorems\<close>
 
 (****
  The following is to prove theorems of the form
@@ -187,7 +187,7 @@
  A more general formula must be proved inductively.
 ****)
 
-text{* Session keys are not used to encrypt other session keys *}
+text\<open>Session keys are not used to encrypt other session keys\<close>
 
 lemma analz_image_freshK [rule_format]:
  "evs \<in> yahalom ==>
@@ -207,7 +207,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{*The Key K uniquely identifies the Server's  message.*}
+text\<open>The Key K uniquely identifies the Server's  message.\<close>
 lemma unique_session_keys:
      "[| Says Server A
           {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \<in> set evs;
@@ -217,12 +217,12 @@
      ==> A=A' & B=B' & na=na' & nb=nb'"
 apply (erule rev_mp, erule rev_mp)
 apply (erule yahalom.induct, simp_all)
-txt{*YM3, by freshness, and YM4*}
+txt\<open>YM3, by freshness, and YM4\<close>
 apply blast+
 done
 
 
-text{*Crucial secrecy property: Spy does not see the keys sent in msg YM3*}
+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
@@ -233,11 +233,11 @@
           Key K \<notin> analz (knows Spy evs)"
 apply (erule yahalom.induct, force,
        drule_tac [6] YM4_analz_knows_Spy)
-apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz)   --{*Fake*}
-apply (blast dest: unique_session_keys)+  --{*YM3, Oops*}
+apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz)   \<comment>\<open>Fake\<close>
+apply (blast dest: unique_session_keys)+  \<comment>\<open>YM3, Oops\<close>
 done
 
-text{*Final version*}
+text\<open>Final version\<close>
 lemma Spy_not_see_encrypted_key:
      "[| Says Server A
             {|Crypt (shrK A) {|Agent B, Key K, na, nb|},
@@ -249,9 +249,9 @@
 by (blast dest: secrecy_lemma)
 
 
-subsubsection{* Security Guarantee for A upon receiving YM3 *}
+subsubsection\<open>Security Guarantee for A upon receiving YM3\<close>
 
-text{*If the encrypted message appears then it originated with the Server*}
+text\<open>If the encrypted message appears then it originated with the Server\<close>
 lemma A_trusts_YM3:
      "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs);
          A \<notin> bad;  evs \<in> yahalom |]
@@ -262,12 +262,12 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast+
 done
 
-text{*The obvious combination of @{text A_trusts_YM3} with
-  @{text Spy_not_see_encrypted_key}*}
+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) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs);
          Notes Spy {|na, nb, Key K|} \<notin> set evs;
@@ -276,10 +276,10 @@
   by (metis A_trusts_YM3 secrecy_lemma)
 
 
-subsubsection{* Security Guarantees for B upon receiving YM4 *}
+subsubsection\<open>Security Guarantees for B upon receiving YM4\<close>
 
-text{*B knows, by the first part of A's message, that the Server distributed
-  the key for A and B.  But this part says nothing about nonces.*}
+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) {|Agent A, Key K|} \<in> parts (knows Spy evs);
          B \<notin> bad;  evs \<in> yahalom |]
@@ -291,15 +291,15 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast+
 done
 
-text{*B knows, by the second part of A's message, that the Server
+text\<open>B knows, by the second part of A's message, that the Server
   distributed the key quoting nonce NB.  This part says nothing about
   agent names.  Secrecy of NB is crucial.  Note that @{term "Nonce NB
   \<notin> analz(knows Spy evs)"} must be the FIRST antecedent of the
-  induction formula.*}
+  induction formula.\<close>
 
 lemma B_trusts_YM4_newK [rule_format]:
      "[|Crypt K (Nonce NB) \<in> parts (knows Spy evs);
@@ -312,20 +312,20 @@
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy)
 apply (analz_mono_contra, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast
 apply blast
-txt{*YM4.  A is uncompromised because NB is secure
-  A's certificate guarantees the existence of the Server message*}
+txt\<open>YM4.  A is uncompromised because NB is secure
+  A's certificate guarantees the existence of the Server message\<close>
 apply (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad
              dest: Says_imp_spies
                    parts.Inj [THEN parts.Fst, THEN A_trusts_YM3])
 done
 
 
-subsubsection{* Towards proving secrecy of Nonce NB *}
+subsubsection\<open>Towards proving secrecy of Nonce NB\<close>
 
-text{*Lemmas about the predicate KeyWithNonce*}
+text\<open>Lemmas about the predicate KeyWithNonce\<close>
 
 lemma KeyWithNonceI:
  "Says Server A
@@ -349,14 +349,14 @@
    "KeyWithNonce K NB (Gets A X # evs) = KeyWithNonce K NB evs"
 by (simp add: KeyWithNonce_def)
 
-text{*A fresh key cannot be associated with any nonce
-  (with respect to a given trace). *}
+text\<open>A fresh key cannot be associated with any nonce
+  (with respect to a given trace).\<close>
 lemma fresh_not_KeyWithNonce:
      "Key K \<notin> used evs ==> ~ KeyWithNonce K NB evs"
 by (unfold KeyWithNonce_def, blast)
 
-text{*The Server message associates K with NB' and therefore not with any
-  other nonce NB.*}
+text\<open>The Server message associates K with NB' and therefore not with any
+  other nonce NB.\<close>
 lemma Says_Server_KeyWithNonce:
  "[| Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB'|}, X|}
        \<in> set evs;
@@ -365,13 +365,13 @@
 by (unfold KeyWithNonce_def, blast dest: unique_session_keys)
 
 
-text{*The only nonces that can be found with the help of session keys are
+text\<open>The only nonces that can be found with the help of session keys are
   those distributed as nonce NB by the Server.  The form of the theorem
-  recalls @{text analz_image_freshK}, but it is much more complicated.*}
+  recalls \<open>analz_image_freshK\<close>, but it is much more complicated.\<close>
 
 
-text{*As with @{text analz_image_freshK}, we take some pains to express the 
-  property as a logical equivalence so that the simplifier can apply it.*}
+text\<open>As with \<open>analz_image_freshK\<close>, we take some pains to express the 
+  property as a logical equivalence so that the simplifier can apply it.\<close>
 lemma Nonce_secrecy_lemma:
      "P --> (X \<in> analz (G Un H)) --> (X \<in> analz H)  ==>
       P --> (X \<in> analz (G Un H)) = (X \<in> analz H)"
@@ -392,29 +392,29 @@
             analz_image_freshK fresh_not_KeyWithNonce
             imp_disj_not1               (*Moves NBa\<noteq>NB to the front*)
             Says_Server_KeyWithNonce)
-txt{*For Oops, simplification proves @{prop "NBa\<noteq>NB"}.  By
+txt\<open>For Oops, simplification proves @{prop "NBa\<noteq>NB"}.  By
   @{term Says_Server_KeyWithNonce}, we get @{prop "~ KeyWithNonce K NB
   evs"}; then simplification can apply the induction hypothesis with
-  @{term "KK = {K}"}.*}
-txt{*Fake*}
+  @{term "KK = {K}"}.\<close>
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*YM2*}
+txt\<open>YM2\<close>
 apply blast
-txt{*YM3*}
+txt\<open>YM3\<close>
 apply blast
-txt{*YM4*}
+txt\<open>YM4\<close>
 apply (erule_tac V = "\<forall>KK. P KK" for P in thin_rl, clarify)
-txt{*If @{prop "A \<in> bad"} then @{term NBa} is known, therefore
+txt\<open>If @{prop "A \<in> bad"} then @{term NBa} is known, therefore
   @{prop "NBa \<noteq> NB"}.  Previous two steps make the next step
-  faster.*}
+  faster.\<close>
 apply (metis A_trusts_YM3 Gets_imp_analz_Spy Gets_imp_knows_Spy KeyWithNonce_def
       Spy_analz_shrK analz.Fst analz.Snd analz_shrK_Decrypt parts.Fst parts.Inj)
 done
 
 
-text{*Version required below: if NB can be decrypted using a session key then
+text\<open>Version required below: if NB can be decrypted using a session key then
    it was distributed with that key.  The more general form above is required
-   for the induction to carry through.*}
+   for the induction to carry through.\<close>
 lemma single_Nonce_secrecy:
      "[| Says Server A
           {|Crypt (shrK A) {|Agent B, Key KAB, na, Nonce NB'|}, X|}
@@ -427,7 +427,7 @@
                   Nonce_secrecy Says_Server_KeyWithNonce)
 
 
-subsubsection{* The Nonce NB uniquely identifies B's message. *}
+subsubsection\<open>The Nonce NB uniquely identifies B's message.\<close>
 
 lemma unique_NB:
      "[| Crypt (shrK B) {|Agent A, Nonce NA, nb|} \<in> parts (knows Spy evs);
@@ -437,13 +437,13 @@
 apply (erule rev_mp, erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, and YM2 by freshness*}
+txt\<open>Fake, and YM2 by freshness\<close>
 apply blast+
 done
 
 
-text{*Variant useful for proving secrecy of NB.  Because nb is assumed to be
-  secret, we no longer must assume B, B' not bad.*}
+text\<open>Variant useful for proving secrecy of NB.  Because nb is assumed to be
+  secret, we no longer must assume B, B' not bad.\<close>
 lemma Says_unique_NB:
      "[| Says C S   {|X,  Crypt (shrK B) {|Agent A, Nonce NA, nb|}|}
            \<in> set evs;
@@ -455,7 +455,7 @@
           dest: Says_imp_spies unique_NB parts.Inj analz.Inj)
 
 
-subsubsection{* A nonce value is never used both as NA and as NB *}
+subsubsection\<open>A nonce value is never used both as NA and as NB\<close>
 
 lemma no_nonce_YM1_YM2:
      "[|Crypt (shrK B') {|Agent A', Nonce NB, nb'|} \<in> parts(knows Spy evs);
@@ -465,11 +465,11 @@
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy)
 apply (analz_mono_contra, simp_all)
-txt{*Fake, YM2*}
+txt\<open>Fake, YM2\<close>
 apply blast+
 done
 
-text{*The Server sends YM3 only in response to YM2.*}
+text\<open>The Server sends YM3 only in response to YM2.\<close>
 lemma Says_Server_imp_YM2:
      "[| Says Server A {|Crypt (shrK A) {|Agent B, k, na, nb|}, X|} \<in> set evs;
          evs \<in> yahalom |]
@@ -477,7 +477,7 @@
              \<in> set evs"
 by (erule rev_mp, erule yahalom.induct, auto)
 
-text{*A vital theorem for B, that nonce NB remains secure from the Spy.*}
+text\<open>A vital theorem for B, that nonce NB remains secure from the Spy.\<close>
 lemma Spy_not_see_NB :
      "[| Says B Server
                 {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
@@ -490,36 +490,36 @@
        frule_tac [6] YM4_analz_knows_Spy)
 apply (simp_all add: split_ifs pushes new_keys_not_analzd analz_insert_eq
                      analz_insert_freshK)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply spy_analz
-txt{*YM1: NB=NA is impossible anyway, but NA is secret because it is fresh!*}
+txt\<open>YM1: NB=NA is impossible anyway, but NA is secret because it is fresh!\<close>
 apply blast
-txt{*YM2*}
+txt\<open>YM2\<close>
 apply blast
-txt{*Prove YM3 by showing that no NB can also be an NA*}
+txt\<open>Prove YM3 by showing that no NB can also be an NA\<close>
 apply (blast dest!: no_nonce_YM1_YM2 dest: Gets_imp_Says Says_unique_NB)
-txt{*LEVEL 7: YM4 and Oops remain*}
+txt\<open>LEVEL 7: YM4 and Oops remain\<close>
 apply (clarify, simp add: all_conj_distrib)
-txt{*YM4: key K is visible to Spy, contradicting session key secrecy theorem*}
-txt{*Case analysis on Aa:bad; PROOF FAILED problems
-  use @{text Says_unique_NB} to identify message components: @{term "Aa=A"}, @{term "Ba=B"}*}
+txt\<open>YM4: key K is visible to Spy, contradicting session key secrecy theorem\<close>
+txt\<open>Case analysis on Aa:bad; PROOF FAILED problems
+  use \<open>Says_unique_NB\<close> to identify message components: @{term "Aa=A"}, @{term "Ba=B"}\<close>
 apply (blast dest!: Says_unique_NB analz_shrK_Decrypt
                     parts.Inj [THEN parts.Fst, THEN A_trusts_YM3]
              dest: Gets_imp_Says Says_imp_spies Says_Server_imp_YM2
                    Spy_not_see_encrypted_key)
-txt{*Oops case: if the nonce is betrayed now, show that the Oops event is
-  covered by the quantified Oops assumption.*}
+txt\<open>Oops case: if the nonce is betrayed now, show that the Oops event is
+  covered by the quantified Oops assumption.\<close>
 apply (clarify, simp add: all_conj_distrib)
 apply (frule Says_Server_imp_YM2, assumption)
 apply (metis Gets_imp_Says Says_Server_not_range Says_unique_NB no_nonce_YM1_YM2 parts.Snd single_Nonce_secrecy spies_partsEs(1))
 done
 
 
-text{*B's session key guarantee from YM4.  The two certificates contribute to a
+text\<open>B's session key guarantee from YM4.  The two certificates contribute to a
   single conclusion about the Server's message.  Note that the "Notes Spy"
-  assumption must quantify over @{text \<forall>} POSSIBLE keys instead of our particular K.
+  assumption must quantify over \<open>\<forall>\<close> POSSIBLE keys instead of our particular K.
   If this run is broken and the spy substitutes a certificate containing an
-  old key, B has no means of telling.*}
+  old key, B has no means of telling.\<close>
 lemma B_trusts_YM4:
      "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},
                   Crypt K (Nonce NB)|} \<in> set evs;
@@ -538,8 +538,8 @@
 
 
 
-text{*The obvious combination of @{text B_trusts_YM4} with 
-  @{text Spy_not_see_encrypted_key}*}
+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 {|Crypt (shrK B) {|Agent A, Key K|},
                   Crypt K (Nonce NB)|} \<in> set evs;
@@ -552,9 +552,9 @@
   by (metis B_trusts_YM4 Spy_not_see_encrypted_key)
 
 
-subsection{*Authenticating B to A*}
+subsection\<open>Authenticating B to A\<close>
 
-text{*The encryption in message YM2 tells us it cannot be faked.*}
+text\<open>The encryption in message YM2 tells us it cannot be faked.\<close>
 lemma B_Said_YM2 [rule_format]:
      "[|Crypt (shrK B) {|Agent A, Nonce NA, nb|} \<in> parts (knows Spy evs);
         evs \<in> yahalom|]
@@ -563,11 +563,11 @@
             \<in> set evs"
 apply (erule rev_mp, erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
 done
 
-text{*If the server sends YM3 then B sent YM2*}
+text\<open>If the server sends YM3 then B sent YM2\<close>
 lemma YM3_auth_B_to_A_lemma:
      "[|Says Server A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, nb|}, X|}
        \<in> set evs;  evs \<in> yahalom|]
@@ -575,11 +575,11 @@
           Says B Server {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, nb|}|}
             \<in> set evs"
 apply (erule rev_mp, erule yahalom.induct, simp_all)
-txt{*YM3, YM4*}
+txt\<open>YM3, YM4\<close>
 apply (blast dest!: B_Said_YM2)+
 done
 
-text{*If A receives YM3 then B has used nonce NA (and therefore is alive)*}
+text\<open>If A receives YM3 then B has used nonce NA (and therefore is alive)\<close>
 lemma YM3_auth_B_to_A:
      "[| Gets A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, nb|}, X|}
            \<in> set evs;
@@ -590,12 +590,12 @@
          not_parts_not_analz)
 
 
-subsection{*Authenticating A to B using the certificate 
-  @{term "Crypt K (Nonce NB)"}*}
+subsection\<open>Authenticating A to B using the certificate 
+  @{term "Crypt K (Nonce NB)"}\<close>
 
-text{*Assuming the session key is secure, if both certificates are present then
+text\<open>Assuming the session key is secure, if both certificates are present then
   A has said NB.  We can't be sure about the rest of A's message, but only
-  NB matters for freshness.*}
+  NB matters for freshness.\<close>
 lemma A_Said_YM3_lemma [rule_format]:
      "evs \<in> yahalom
       ==> Key K \<notin> analz (knows Spy evs) -->
@@ -606,23 +606,23 @@
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy)
 apply (analz_mono_contra, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*YM3: by @{text new_keys_not_used}, the message
-   @{term "Crypt K (Nonce NB)"} could not exist*}
+txt\<open>YM3: by \<open>new_keys_not_used\<close>, the message
+   @{term "Crypt K (Nonce NB)"} could not exist\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*YM4: was @{term "Crypt K (Nonce NB)"} the very last message?
-    If not, use the induction hypothesis*}
+txt\<open>YM4: was @{term "Crypt K (Nonce NB)"} the very last message?
+    If not, use the induction hypothesis\<close>
 apply (simp add: ex_disj_distrib)
-txt{*yes: apply unicity of session keys*}
+txt\<open>yes: apply unicity of session keys\<close>
 apply (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK
                     Crypt_Spy_analz_bad
              dest: Says_imp_knows_Spy [THEN parts.Inj] unique_session_keys)
 done
 
-text{*If B receives YM4 then A has used nonce NB (and therefore is alive).
+text\<open>If B receives YM4 then A has used nonce NB (and therefore is alive).
   Moreover, A associates K with NB (thus is talking about the same run).
-  Other premises guarantee secrecy of K.*}
+  Other premises guarantee secrecy of K.\<close>
 lemma YM4_imp_A_Said_YM3 [rule_format]:
      "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},
                   Crypt K (Nonce NB)|} \<in> set evs;
--- a/src/HOL/Auth/Yahalom2.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Yahalom2.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,11 +3,11 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Yahalom Protocol, Variant 2*}
+section\<open>The Yahalom Protocol, Variant 2\<close>
 
 theory Yahalom2 imports Public begin
 
-text{*
+text\<open>
 This version trades encryption of NB for additional explicitness in YM3.
 Also in YM3, care is taken to make the two certificates distinct.
 
@@ -16,7 +16,7 @@
   Proc. Royal Soc. 426
 
 This theory has the prototypical example of a secrecy relation, KeyCryptNonce.
-*}
+\<close>
 
 inductive_set yahalom :: "event list set"
   where
@@ -81,7 +81,7 @@
 declare Fake_parts_insert_in_Un  [dest]
 declare analz_into_parts [dest]
 
-text{*A "possibility property": there are traces that reach the end*}
+text\<open>A "possibility property": there are traces that reach the end\<close>
 lemma "Key K \<notin> used []
        ==> \<exists>X NB. \<exists>evs \<in> yahalom.
              Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs"
@@ -98,7 +98,7 @@
      "[| Gets B X \<in> set evs; evs \<in> yahalom |] ==> \<exists>A. Says A B X \<in> set evs"
 by (erule rev_mp, erule yahalom.induct, auto)
 
-text{*Must be proved separately for each protocol*}
+text\<open>Must be proved separately for each protocol\<close>
 lemma Gets_imp_knows_Spy:
      "[| Gets B X \<in> set evs; evs \<in> yahalom |]  ==> X \<in> knows Spy evs"
 by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
@@ -106,10 +106,10 @@
 declare Gets_imp_knows_Spy [THEN analz.Inj, dest]
 
 
-subsection{*Inductive Proofs*}
+subsection\<open>Inductive Proofs\<close>
 
-text{*Result for reasoning about the encrypted portion of messages.
-Lets us treat YM4 using a similar argument as for the Fake case.*}
+text\<open>Result for reasoning about the encrypted portion of messages.
+Lets us treat YM4 using a similar argument as for the Fake case.\<close>
 lemma YM4_analz_knows_Spy:
      "[| Gets A {|NB, Crypt (shrK A) Y, X|} \<in> set evs;  evs \<in> yahalom |]
       ==> X \<in> analz (knows Spy evs)"
@@ -122,7 +122,7 @@
 (** Theorems of the form X \<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
-text{*Spy never sees a good agent's shared key!*}
+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)"
 by (erule yahalom.induct, force,
@@ -136,26 +136,26 @@
      "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> yahalom|] ==> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
-text{*Nobody can have used non-existent keys!  
-    Needed to apply @{text analz_insert_Key}*}
+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))"
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert)
-txt{*YM3*}
+txt\<open>YM3\<close>
 apply blast
-txt{*YM4*}
+txt\<open>YM4\<close>
 apply auto
 apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj])
 done
 
 
-text{*Describes the form of K when the Server sends this message.  Useful for
-  Oops as well as main secrecy property.*}
+text\<open>Describes the form of K when the Server sends this message.  Useful for
+  Oops as well as main secrecy property.\<close>
 lemma Says_Server_message_form:
      "[| Says Server A {|nb', Crypt (shrK A) {|Agent B, Key K, na|}, X|}
           \<in> set evs;  evs \<in> yahalom |]
@@ -191,7 +191,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{*The Key K uniquely identifies the Server's  message*}
+text\<open>The Key K uniquely identifies the Server's  message\<close>
 lemma unique_session_keys:
      "[| Says Server A
           {|nb, Crypt (shrK A) {|Agent B, Key K, na|}, X|} \<in> set evs;
@@ -201,12 +201,12 @@
      ==> A=A' & B=B' & na=na' & nb=nb'"
 apply (erule rev_mp, erule rev_mp)
 apply (erule yahalom.induct, simp_all)
-txt{*YM3, by freshness*}
+txt\<open>YM3, by freshness\<close>
 apply blast
 done
 
 
-subsection{*Crucial Secrecy Property: Spy Does Not See Key @{term KAB}*}
+subsection\<open>Crucial Secrecy Property: Spy Does Not See Key @{term KAB}\<close>
 
 lemma secrecy_lemma:
      "[| A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
@@ -223,7 +223,7 @@
 done
 
 
-text{*Final version*}
+text\<open>Final version\<close>
 lemma Spy_not_see_encrypted_key:
      "[| Says Server A
             {|nb, Crypt (shrK A) {|Agent B, Key K, na|},
@@ -236,13 +236,13 @@
 
 
 
-text{*This form is an immediate consequence of the previous result.  It is
+text\<open>This form is an immediate consequence of the previous result.  It is
 similar to the assertions established by other methods.  It is equivalent
 to the previous result in that the Spy already has @{term analz} and
 @{term synth} at his disposal.  However, the conclusion
 @{term "Key K \<notin> knows Spy evs"} appears not to be inductive: all the cases
 other than Fake are trivial, while Fake requires
-@{term "Key K \<notin> analz (knows Spy evs)"}. *}
+@{term "Key K \<notin> analz (knows Spy evs)"}.\<close>
 lemma Spy_not_know_encrypted_key:
      "[| Says Server A
             {|nb, Crypt (shrK A) {|Agent B, Key K, na|},
@@ -254,10 +254,10 @@
 by (blast dest: Spy_not_see_encrypted_key)
 
 
-subsection{*Security Guarantee for A upon receiving YM3*}
+subsection\<open>Security Guarantee for A upon receiving YM3\<close>
 
-text{*If the encrypted message appears then it originated with the Server.
-  May now apply @{text Spy_not_see_encrypted_key}, subject to its conditions.*}
+text\<open>If the encrypted message appears then it originated with the Server.
+  May now apply \<open>Spy_not_see_encrypted_key\<close>, subject to its conditions.\<close>
 lemma A_trusts_YM3:
      "[| Crypt (shrK A) {|Agent B, Key K, na|} \<in> parts (knows Spy evs);
          A \<notin> bad;  evs \<in> yahalom |]
@@ -268,12 +268,12 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast+
 done
 
-text{*The obvious combination of @{text A_trusts_YM3} with 
-@{text Spy_not_see_encrypted_key}*}
+text\<open>The obvious combination of \<open>A_trusts_YM3\<close> with 
+\<open>Spy_not_see_encrypted_key\<close>\<close>
 theorem A_gets_good_key:
      "[| Crypt (shrK A) {|Agent B, Key K, na|} \<in> parts (knows Spy evs);
          \<forall>nb. Notes Spy {|na, nb, Key K|} \<notin> set evs;
@@ -282,10 +282,10 @@
 by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key)
 
 
-subsection{*Security Guarantee for B upon receiving YM4*}
+subsection\<open>Security Guarantee for B upon receiving YM4\<close>
 
-text{*B knows, by the first part of A's message, that the Server distributed
-  the key for A and B, and has associated it with NB.*}
+text\<open>B knows, by the first part of A's message, that the Server distributed
+  the key for A and B, and has associated it with NB.\<close>
 lemma B_trusts_YM4_shrK:
      "[| Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}
            \<in> parts (knows Spy evs);
@@ -298,16 +298,16 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast+
 done
 
 
-text{*With this protocol variant, we don't need the 2nd part of YM4 at all:
-  Nonce NB is available in the first part.*}
+text\<open>With this protocol variant, we don't need the 2nd part of YM4 at all:
+  Nonce NB is available in the first part.\<close>
 
-text{*What can B deduce from receipt of YM4?  Stronger and simpler than Yahalom
-  because we do not have to show that NB is secret. *}
+text\<open>What can B deduce from receipt of YM4?  Stronger and simpler than Yahalom
+  because we do not have to show that NB is secret.\<close>
 lemma B_trusts_YM4:
      "[| Gets B {|Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|},  X|}
            \<in> set evs;
@@ -320,8 +320,8 @@
 by (blast dest!: B_trusts_YM4_shrK)
 
 
-text{*The obvious combination of @{text B_trusts_YM4} with 
-@{text Spy_not_see_encrypted_key}*}
+text\<open>The obvious combination of \<open>B_trusts_YM4\<close> with 
+\<open>Spy_not_see_encrypted_key\<close>\<close>
 theorem B_gets_good_key:
      "[| Gets B {|Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}, X|}
            \<in> set evs;
@@ -331,9 +331,9 @@
 by (blast dest!: B_trusts_YM4 Spy_not_see_encrypted_key)
 
 
-subsection{*Authenticating B to A*}
+subsection\<open>Authenticating B to A\<close>
 
-text{*The encryption in message YM2 tells us it cannot be faked.*}
+text\<open>The encryption in message YM2 tells us it cannot be faked.\<close>
 lemma B_Said_YM2:
      "[| Crypt (shrK B) {|Agent A, Nonce NA|} \<in> parts (knows Spy evs);
          B \<notin> bad;  evs \<in> yahalom |]
@@ -343,12 +343,12 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM2*}
+txt\<open>Fake, YM2\<close>
 apply blast+
 done
 
 
-text{*If the server sends YM3 then B sent YM2, perhaps with a different NB*}
+text\<open>If the server sends YM3 then B sent YM2, perhaps with a different NB\<close>
 lemma YM3_auth_B_to_A_lemma:
      "[| Says Server A {|nb, Crypt (shrK A) {|Agent B, Key K, Nonce NA|}, X|}
            \<in> set evs;
@@ -358,11 +358,11 @@
                        \<in> set evs"
 apply (erule rev_mp)
 apply (erule yahalom.induct, simp_all)
-txt{*Fake, YM2, YM3*}
+txt\<open>Fake, YM2, YM3\<close>
 apply (blast dest!: B_Said_YM2)+
 done
 
-text{*If A receives YM3 then B has used nonce NA (and therefore is alive)*}
+text\<open>If A receives YM3 then B has used nonce NA (and therefore is alive)\<close>
 theorem YM3_auth_B_to_A:
      "[| Gets A {|nb, Crypt (shrK A) {|Agent B, Key K, Nonce NA|}, X|}
            \<in> set evs;
@@ -373,17 +373,17 @@
 by (blast dest!: A_trusts_YM3 YM3_auth_B_to_A_lemma)
 
 
-subsection{*Authenticating A to B*}
+subsection\<open>Authenticating A to B\<close>
 
-text{*using the certificate @{term "Crypt K (Nonce NB)"}*}
+text\<open>using the certificate @{term "Crypt K (Nonce NB)"}\<close>
 
-text{*Assuming the session key is secure, if both certificates are present then
+text\<open>Assuming the session key is secure, if both certificates are present then
   A has said NB.  We can't be sure about the rest of A's message, but only
   NB matters for freshness.  Note that @{term "Key K \<notin> analz (knows Spy evs)"}
-  must be the FIRST antecedent of the induction formula.*}
+  must be the FIRST antecedent of the induction formula.\<close>
 
-text{*This lemma allows a use of @{text unique_session_keys} in the next proof,
-  which otherwise is extremely slow.*}
+text\<open>This lemma allows a use of \<open>unique_session_keys\<close> in the next proof,
+  which otherwise is extremely slow.\<close>
 lemma secure_unique_session_keys:
      "[| Crypt (shrK A) {|Agent B, Key K, na|} \<in> analz (spies evs);
          Crypt (shrK A') {|Agent B', Key K, na'|} \<in> analz (spies evs);
@@ -404,20 +404,20 @@
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy)
 apply (analz_mono_contra, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*YM3: by @{text new_keys_not_used}, the message
-   @{term "Crypt K (Nonce NB)"} could not exist*}
+txt\<open>YM3: by \<open>new_keys_not_used\<close>, the message
+   @{term "Crypt K (Nonce NB)"} could not exist\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*YM4: was   @{term "Crypt K (Nonce NB)"} the very last message?  If so, 
-    apply unicity of session keys; if not, use the induction hypothesis*}
+txt\<open>YM4: was   @{term "Crypt K (Nonce NB)"} the very last message?  If so, 
+    apply unicity of session keys; if not, use the induction hypothesis\<close>
 apply (blast dest!: B_trusts_YM4_shrK dest: secure_unique_session_keys)
 done
 
 
-text{*If B receives YM4 then A has used nonce NB (and therefore is alive).
+text\<open>If B receives YM4 then A has used nonce NB (and therefore is alive).
   Moreover, A associates K with NB (thus is talking about the same run).
-  Other premises guarantee secrecy of K.*}
+  Other premises guarantee secrecy of K.\<close>
 theorem YM4_imp_A_Said_YM3 [rule_format]:
      "[| Gets B {|Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|},
                   Crypt K (Nonce NB)|} \<in> set evs;
--- a/src/HOL/Auth/Yahalom_Bad.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/Yahalom_Bad.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -3,15 +3,15 @@
     Copyright   1996  University of Cambridge
 *)
 
-section{*The Yahalom Protocol: A Flawed Version*}
+section\<open>The Yahalom Protocol: A Flawed Version\<close>
 
 theory Yahalom_Bad imports Public begin
 
-text{*
+text\<open>
 Demonstrates of why Oops is necessary.  This protocol can be attacked because
 it doesn't keep NB secret, but without Oops it can be "verified" anyway.
 The issues are discussed in lcp's LICS 2000 invited lecture.
-*}
+\<close>
 
 inductive_set yahalom :: "event list set"
   where
@@ -67,7 +67,7 @@
 declare analz_into_parts [dest]
 
 
-text{*A "possibility property": there are traces that reach the end*}
+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.
               Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs"
@@ -80,7 +80,7 @@
 apply (possibility, simp add: used_Cons) 
 done
 
-subsection{*Regularity Lemmas for Yahalom*}
+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"
@@ -94,9 +94,9 @@
 declare Gets_imp_knows_Spy [THEN analz.Inj, dest]
 
 
-subsection{* For reasoning about the encrypted portion of messages *}
+subsection\<open>For reasoning about the encrypted portion of messages\<close>
 
-text{*Lets us treat YM4 using a similar argument as for the Fake case.*}
+text\<open>Lets us treat YM4 using a similar argument as for the Fake case.\<close>
 lemma YM4_analz_knows_Spy:
      "[| Gets A {|Crypt (shrK A) Y, X|} \<in> set evs;  evs \<in> yahalom |]
       ==> X \<in> analz (knows Spy evs)"
@@ -106,10 +106,10 @@
        YM4_analz_knows_Spy [THEN analz_into_parts]
 
 
-text{*Theorems of the form @{term "X \<notin> parts (knows Spy evs)"} imply 
-            that NOBODY sends messages containing X!*}
+text\<open>Theorems of the form @{term "X \<notin> parts (knows Spy evs)"} imply 
+            that NOBODY sends messages containing X!\<close>
 
-text{*Spy never sees a good agent's shared key!*}
+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)"
 apply (erule yahalom.induct, force,
@@ -124,20 +124,20 @@
      "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> yahalom|] ==> A \<in> bad"
 by (blast dest: Spy_see_shrK)
 
-text{*Nobody can have used non-existent keys!
-    Needed to apply @{text analz_insert_Key}*}
+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))"
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (force dest!: keysFor_parts_insert, auto)
 done
 
 
-subsection{*Secrecy Theorems*}
+subsection\<open>Secrecy Theorems\<close>
 
 (****
  The following is to prove theorems of the form
@@ -148,7 +148,7 @@
  A more general formula must be proved inductively.
 ****)
 
-subsection{* Session keys are not used to encrypt other session keys *}
+subsection\<open>Session keys are not used to encrypt other session keys\<close>
 
 lemma analz_image_freshK [rule_format]:
  "evs \<in> yahalom ==>
@@ -165,7 +165,7 @@
 by (simp only: analz_image_freshK analz_image_freshK_simps)
 
 
-text{*The Key K uniquely identifies the Server's  message.*}
+text\<open>The Key K uniquely identifies the Server's  message.\<close>
 lemma unique_session_keys:
      "[| Says Server A
           {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \<in> set evs;
@@ -175,12 +175,12 @@
      ==> A=A' & B=B' & na=na' & nb=nb'"
 apply (erule rev_mp, erule rev_mp)
 apply (erule yahalom.induct, simp_all)
-txt{*YM3, by freshness, and YM4*}
+txt\<open>YM3, by freshness, and YM4\<close>
 apply blast+
 done
 
 
-text{* Crucial secrecy property: Spy does not see the keys sent in msg YM3 *}
+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
@@ -193,7 +193,7 @@
 apply (blast dest: unique_session_keys)  (*YM3*)
 done
 
-text{*Final version*}
+text\<open>Final version\<close>
 lemma Spy_not_see_encrypted_key:
      "[| Says Server A
             {|Crypt (shrK A) {|Agent B, Key K, na, nb|},
@@ -204,9 +204,9 @@
 by (blast dest: secrecy_lemma)
 
 
-subsection{* Security Guarantee for A upon receiving YM3 *}
+subsection\<open>Security Guarantee for A upon receiving YM3\<close>
 
-text{*If the encrypted message appears then it originated with the Server*}
+text\<open>If the encrypted message appears then it originated with the Server\<close>
 lemma A_trusts_YM3:
      "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs);
          A \<notin> bad;  evs \<in> yahalom |]
@@ -217,22 +217,22 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast+
 done
 
-text{*The obvious combination of @{text A_trusts_YM3} with
-  @{text Spy_not_see_encrypted_key}*}
+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) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs);
          A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
       ==> Key K \<notin> analz (knows Spy evs)"
 by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key)
 
-subsection{* Security Guarantees for B upon receiving YM4 *}
+subsection\<open>Security Guarantees for B upon receiving YM4\<close>
 
-text{*B knows, by the first part of A's message, that the Server distributed
-  the key for A and B.  But this part says nothing about nonces.*}
+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) {|Agent A, Key K|} \<in> parts (knows Spy evs);
          B \<notin> bad;  evs \<in> yahalom |]
@@ -243,21 +243,21 @@
 apply (erule rev_mp)
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy, simp_all)
-txt{*Fake, YM3*}
+txt\<open>Fake, YM3\<close>
 apply blast+
 done
 
-subsection{*The Flaw in the Model*}
+subsection\<open>The Flaw in the Model\<close>
 
-text{* Up to now, the reasoning is similar to standard Yahalom.  Now the
+text\<open>Up to now, the reasoning is similar to standard Yahalom.  Now the
     doubtful reasoning occurs.  We should not be assuming that an unknown
     key is secure, but the model allows us to: there is no Oops rule to
-    let session keys become compromised.*}
+    let session keys become compromised.\<close>
 
-text{*B knows, by the second part of A's message, that the Server distributed
+text\<open>B knows, by the second part of A's message, that the Server distributed
   the key quoting nonce NB.  This part says nothing about agent names.
   Secrecy of K is assumed; the valid Yahalom proof uses (and later proves)
-  the secrecy of NB.*}
+  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) -->
@@ -270,20 +270,20 @@
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy)
 apply (analz_mono_contra, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*YM3*}
+txt\<open>YM3\<close>
 apply blast
-txt{*A is uncompromised because NB is secure
-  A's certificate guarantees the existence of the Server message*}
+txt\<open>A is uncompromised because NB is secure
+  A's certificate guarantees the existence of the Server message\<close>
 apply (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad
              dest: Says_imp_spies
                    parts.Inj [THEN parts.Fst, THEN A_trusts_YM3])
 done
 
 
-text{*B's session key guarantee from YM4.  The two certificates contribute to a
-  single conclusion about the Server's message. *}
+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 {|Crypt (shrK B) {|Agent A, Key K|},
                   Crypt K (Nonce NB)|} \<in> set evs;
@@ -299,8 +299,8 @@
                 unique_session_keys)
 
 
-text{*The obvious combination of @{text B_trusts_YM4} with 
-  @{text Spy_not_see_encrypted_key}*}
+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 {|Crypt (shrK B) {|Agent A, Key K|},
                   Crypt K (Nonce NB)|} \<in> set evs;
@@ -318,9 +318,9 @@
 
 (*** Authenticating A to B using the certificate Crypt K (Nonce NB) ***)
 
-text{*Assuming the session key is secure, if both certificates are present then
+text\<open>Assuming the session key is secure, if both certificates are present then
   A has said NB.  We can't be sure about the rest of A's message, but only
-  NB matters for freshness.*}
+  NB matters for freshness.\<close>
 lemma A_Said_YM3_lemma [rule_format]:
      "evs \<in> yahalom
       ==> Key K \<notin> analz (knows Spy evs) -->
@@ -331,23 +331,23 @@
 apply (erule yahalom.induct, force,
        frule_tac [6] YM4_parts_knows_Spy)
 apply (analz_mono_contra, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast
-txt{*YM3: by @{text new_keys_not_used}, the message
-   @{term "Crypt K (Nonce NB)"} could not exist*}
+txt\<open>YM3: by \<open>new_keys_not_used\<close>, the message
+   @{term "Crypt K (Nonce NB)"} could not exist\<close>
 apply (force dest!: Crypt_imp_keysFor)
-txt{*YM4: was @{term "Crypt K (Nonce NB)"} the very last message?
-    If not, use the induction hypothesis*}
+txt\<open>YM4: was @{term "Crypt K (Nonce NB)"} the very last message?
+    If not, use the induction hypothesis\<close>
 apply (simp add: ex_disj_distrib)
-txt{*yes: apply unicity of session keys*}
+txt\<open>yes: apply unicity of session keys\<close>
 apply (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK
                     Crypt_Spy_analz_bad
              dest: Says_imp_knows_Spy [THEN parts.Inj] unique_session_keys)
 done
 
-text{*If B receives YM4 then A has used nonce NB (and therefore is alive).
+text\<open>If B receives YM4 then A has used nonce NB (and therefore is alive).
   Moreover, A associates K with NB (thus is talking about the same run).
-  Other premises guarantee secrecy of K.*}
+  Other premises guarantee secrecy of K.\<close>
 lemma YM4_imp_A_Said_YM3 [rule_format]:
      "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},
                   Crypt K (Nonce NB)|} \<in> set evs;
--- a/src/HOL/Auth/ZhouGollmann.thy	Thu Dec 10 21:31:24 2015 +0100
+++ b/src/HOL/Auth/ZhouGollmann.thy	Thu Dec 10 21:39:33 2015 +0100
@@ -21,8 +21,8 @@
 
 
 definition broken :: "agent set" where    
-    --{*the compromised honest agents; TTP is included as it's not allowed to
-        use the protocol*}
+    \<comment>\<open>the compromised honest agents; TTP is included as it's not allowed to
+        use the protocol\<close>
    "broken == bad - {Spy}"
 
 declare broken_def [simp]
@@ -89,7 +89,7 @@
 declare symKey_neq_priEK [THEN not_sym, simp]
 
 
-text{*A "possibility property": there are traces that reach the end*}
+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|] ==>
      \<exists>L. \<exists>evs \<in> zg.
            Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K,
@@ -104,7 +104,7 @@
 apply (basic_possibility, auto)
 done
 
-subsection {*Basic Lemmas*}
+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"
@@ -117,8 +117,8 @@
 by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
 
 
-text{*Lets us replace proofs about @{term "used evs"} by simpler proofs 
-about @{term "parts (spies evs)"}.*}
+text\<open>Lets us replace proofs about @{term "used evs"} by simpler proofs 
+about @{term "parts (spies evs)"}.\<close>
 lemma Crypt_used_imp_spies:
      "[| Crypt K X \<in> used evs; evs \<in> zg |]
       ==> Crypt K X \<in> parts (spies evs)"
@@ -137,7 +137,7 @@
 apply (erule zg.induct, auto)
 done
 
-text{*For reasoning about C, which is encrypted in message ZG2*}
+text\<open>For reasoning about C, which is encrypted in message ZG2\<close>
 lemma ZG2_msg_in_parts_spies:
      "[|Gets B {|F, B', L, C, X|} \<in> set evs; evs \<in> zg|]
       ==> C \<in> parts (spies evs)"
@@ -150,7 +150,7 @@
 apply (frule_tac [5] ZG2_msg_in_parts_spies, auto)
 done
 
-text{*So that blast can use it too*}
+text\<open>So that blast can use it too\<close>
 declare  Spy_see_priK [THEN [2] rev_iffD1, dest!]
 
 lemma Spy_analz_priK [simp]:
@@ -158,12 +158,12 @@
 by auto 
 
 
-subsection{*About NRO: Validity for @{term B}*}
+subsection\<open>About NRO: Validity for @{term B}\<close>
 
-text{*Below we prove that if @{term NRO} exists then @{term A} definitely
-sent it, provided @{term A} is not broken.*}
+text\<open>Below we prove that if @{term NRO} exists then @{term A} definitely
+sent it, provided @{term A} is not broken.\<close>
 
-text{*Strong conclusion for a good agent*}
+text\<open>Strong conclusion for a good agent\<close>
 lemma NRO_validity_good:
      "[|NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
         NRO \<in> parts (spies evs);
@@ -182,7 +182,7 @@
 apply (erule zg.induct, simp_all)
 done
 
-text{*Holds also for @{term "A = Spy"}!*}
+text\<open>Holds also for @{term "A = Spy"}!\<close>
 theorem NRO_validity:
      "[|Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs;
         NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
@@ -191,19 +191,19 @@
 apply (drule Gets_imp_Says, assumption) 
 apply clarify 
 apply (frule NRO_sender, auto)
-txt{*We are left with the case where the sender is @{term Spy} and not
+txt\<open>We are left with the case where the sender is @{term Spy} and not
   equal to @{term A}, because @{term "A \<notin> bad"}. 
-  Thus theorem @{text NRO_validity_good} applies.*}
+  Thus theorem \<open>NRO_validity_good\<close> applies.\<close>
 apply (blast dest: NRO_validity_good [OF refl])
 done
 
 
-subsection{*About NRR: Validity for @{term A}*}
+subsection\<open>About NRR: Validity for @{term A}\<close>
 
-text{*Below we prove that if @{term NRR} exists then @{term B} definitely
-sent it, provided @{term B} is not broken.*}
+text\<open>Below we prove that if @{term NRR} exists then @{term B} definitely
+sent it, provided @{term B} is not broken.\<close>
 
-text{*Strong conclusion for a good agent*}
+text\<open>Strong conclusion for a good agent\<close>
 lemma NRR_validity_good:
      "[|NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
         NRR \<in> parts (spies evs);
@@ -222,7 +222,7 @@
 apply (erule zg.induct, simp_all)
 done
 
-text{*Holds also for @{term "B = Spy"}!*}
+text\<open>Holds also for @{term "B = Spy"}!\<close>
 theorem NRR_validity:
      "[|Says B' A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs;
         NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
@@ -230,18 +230,18 @@
     ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs"
 apply clarify 
 apply (frule NRR_sender, auto)
-txt{*We are left with the case where @{term "B' = Spy"} and  @{term "B' \<noteq> B"},
-  i.e. @{term "B \<notin> bad"}, when we can apply @{text NRR_validity_good}.*}
+txt\<open>We are left with the case where @{term "B' = Spy"} and  @{term "B' \<noteq> B"},
+  i.e. @{term "B \<notin> bad"}, when we can apply \<open>NRR_validity_good\<close>.\<close>
  apply (blast dest: NRR_validity_good [OF refl])
 done
 
 
-subsection{*Proofs About @{term sub_K}*}
+subsection\<open>Proofs About @{term sub_K}\<close>
 
-text{*Below we prove that if @{term sub_K} exists then @{term A} definitely
-sent it, provided @{term A} is not broken.  *}
+text\<open>Below we prove that if @{term sub_K} exists then @{term A} definitely
+sent it, provided @{term A} is not broken.\<close>
 
-text{*Strong conclusion for a good agent*}
+text\<open>Strong conclusion for a good agent\<close>
 lemma sub_K_validity_good:
      "[|sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
         sub_K \<in> parts (spies evs);
@@ -251,7 +251,7 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all)
-txt{*Fake*} 
+txt\<open>Fake\<close> 
 apply (blast dest!: Fake_parts_sing_imp_Un)
 done
 
@@ -262,7 +262,7 @@
 apply (erule zg.induct, simp_all)
 done
 
-text{*Holds also for @{term "A = Spy"}!*}
+text\<open>Holds also for @{term "A = Spy"}!\<close>
 theorem sub_K_validity:
      "[|Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs;
         sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
@@ -271,19 +271,19 @@
 apply (drule Gets_imp_Says, assumption) 
 apply clarify 
 apply (frule sub_K_sender, auto)
-txt{*We are left with the case where the sender is @{term Spy} and not
+txt\<open>We are left with the case where the sender is @{term Spy} and not
   equal to @{term A}, because @{term "A \<notin> bad"}. 
-  Thus theorem @{text sub_K_validity_good} applies.*}
+  Thus theorem \<open>sub_K_validity_good\<close> applies.\<close>
 apply (blast dest: sub_K_validity_good [OF refl])
 done
 
 
 
-subsection{*Proofs About @{term con_K}*}
+subsection\<open>Proofs About @{term con_K}\<close>
 
-text{*Below we prove that if @{term con_K} exists, then @{term TTP} has it,
+text\<open>Below we prove that if @{term con_K} exists, then @{term TTP} has it,
 and therefore @{term A} and @{term B}) can get it too.  Moreover, we know
-that @{term A} sent @{term sub_K}*}
+that @{term A} sent @{term sub_K}\<close>
 
 lemma con_K_validity:
      "[|con_K \<in> used evs;
@@ -296,15 +296,15 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest!: Fake_parts_sing_imp_Un)
-txt{*ZG2*} 
+txt\<open>ZG2\<close> 
 apply (blast dest: parts_cut)
 done
 
-text{*If @{term TTP} holds @{term con_K} then @{term A} sent
+text\<open>If @{term TTP} holds @{term con_K} then @{term A} sent
  @{term sub_K}.  We assume that @{term A} is not broken.  Importantly, nothing
-  needs to be assumed about the form of @{term con_K}!*}
+  needs to be assumed about the form of @{term con_K}!\<close>
 lemma Notes_TTP_imp_Says_A:
      "[|Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
            \<in> set evs;
@@ -315,13 +315,13 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all)
-txt{*ZG4*}
+txt\<open>ZG4\<close>
 apply clarify 
 apply (rule sub_K_validity, auto) 
 done
 
-text{*If @{term con_K} exists, then @{term A} sent @{term sub_K}.  We again
-   assume that @{term A} is not broken. *}
+text\<open>If @{term con_K} exists, then @{term A} sent @{term sub_K}.  We again
+   assume that @{term A} is not broken.\<close>
 theorem B_sub_K_validity:
      "[|con_K \<in> used evs;
         con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
@@ -332,13 +332,13 @@
 by (blast dest: con_K_validity Notes_TTP_imp_Says_A)
 
 
-subsection{*Proving fairness*}
+subsection\<open>Proving fairness\<close>
 
-text{*Cannot prove that, if @{term B} has NRO, then  @{term A} has her NRR.
+text\<open>Cannot prove that, if @{term B} has NRO, then  @{term A} has her NRR.
 It would appear that @{term B} has a small advantage, though it is
-useless to win disputes: @{term B} needs to present @{term con_K} as well.*}
+useless to win disputes: @{term B} needs to present @{term con_K} as well.\<close>
 
-text{*Strange: unicity of the label protects @{term A}?*}
+text\<open>Strange: unicity of the label protects @{term A}?\<close>
 lemma A_unicity: 
      "[|NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|};
         NRO \<in> parts (spies evs);
@@ -351,13 +351,13 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) 
-txt{*ZG1: freshness*}
+txt\<open>ZG1: freshness\<close>
 apply (blast dest: parts.Body) 
 done
 
 
-text{*Fairness lemma: if @{term sub_K} exists, then @{term A} holds 
-NRR.  Relies on unicity of labels.*}
+text\<open>Fairness lemma: if @{term sub_K} exists, then @{term A} holds 
+NRR.  Relies on unicity of labels.\<close>
 lemma sub_K_implies_NRR:
      "[| NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|};
          NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, Crypt K M|};
@@ -372,11 +372,11 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply blast 
-txt{*ZG1: freshness*}
+txt\<open>ZG1: freshness\<close>
 apply (blast dest: parts.Body) 
-txt{*ZG3*} 
+txt\<open>ZG3\<close> 
 apply (blast dest: A_unicity [OF refl]) 
 done
 
@@ -386,16 +386,16 @@
       ==> L \<in> used evs"
 apply (erule rev_mp)
 apply (erule zg.induct, auto)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest!: Fake_parts_sing_imp_Un)
-txt{*ZG2: freshness*}
+txt\<open>ZG2: freshness\<close>
 apply (blast dest: parts.Body) 
 done
 
 
-text{*Fairness for @{term A}: if @{term con_K} and @{term NRO} exist, 
+text\<open>Fairness for @{term A}: if @{term con_K} and @{term NRO} exist, 
 then @{term A} holds NRR.  @{term A} must be uncompromised, but there is no
-assumption about @{term B}.*}
+assumption about @{term B}.\<close>
 theorem A_fairness_NRO:
      "[|con_K \<in> used evs;
         NRO \<in> parts (spies evs);
@@ -410,21 +410,21 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all)
-   txt{*Fake*}
+   txt\<open>Fake\<close>
    apply (simp add: parts_insert_knows_A) 
    apply (blast dest: Fake_parts_sing_imp_Un) 
-  txt{*ZG1*}
+  txt\<open>ZG1\<close>
   apply (blast dest: Crypt_used_imp_L_used) 
- txt{*ZG2*}
+ txt\<open>ZG2\<close>
  apply (blast dest: parts_cut)
-txt{*ZG4*} 
+txt\<open>ZG4\<close> 
 apply (blast intro: sub_K_implies_NRR [OF refl] 
              dest: Gets_imp_knows_Spy [THEN parts.Inj])
 done
 
-text{*Fairness for @{term B}: NRR exists at all, then @{term B} holds NRO.
+text\<open>Fairness for @{term B}: NRR exists at all, then @{term B} holds NRO.
 @{term B} must be uncompromised, but there is no assumption about @{term
-A}. *}
+A}.\<close>
 theorem B_fairness_NRR:
      "[|NRR \<in> used evs;
         NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
@@ -435,16 +435,15 @@
 apply (erule rev_mp)
 apply (erule zg.induct)
 apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all)
-txt{*Fake*}
+txt\<open>Fake\<close>
 apply (blast dest!: Fake_parts_sing_imp_Un)
-txt{*ZG2*}
+txt\<open>ZG2\<close>
 apply (blast dest: parts_cut)
 done
 
 
-text{*If @{term con_K} exists at all, then @{term B} can get it, by @{text
-con_K_validity}.  Cannot conclude that also NRO is available to @{term B},
+text\<open>If @{term con_K} exists at all, then @{term B} can get it, by \<open>con_K_validity\<close>.  Cannot conclude that also NRO is available to @{term B},
 because if @{term A} were unfair, @{term A} could build message 3 without
-building message 1, which contains NRO. *}
+building message 1, which contains NRO.\<close>
 
 end