src/HOL/Auth/OtwayRees_AN.thy
changeset 61830 4f5ab843cf5b
parent 58889 5b7a9633cfa8
child 61956 38b73f7940af
--- 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;