src/HOL/Metis_Examples/Message.thy
changeset 63167 0909deb8059b
parent 62390 842917225d56
child 67443 3abf6a722518
equal deleted inserted replaced
63166:143f58bb34f9 63167:0909deb8059b
     3     Author:     Jasmin Blanchette, TU Muenchen
     3     Author:     Jasmin Blanchette, TU Muenchen
     4 
     4 
     5 Metis example featuring message authentication.
     5 Metis example featuring message authentication.
     6 *)
     6 *)
     7 
     7 
     8 section {* Metis Example Featuring Message Authentication *}
     8 section \<open>Metis Example Featuring Message Authentication\<close>
     9 
     9 
    10 theory Message
    10 theory Message
    11 imports Main
    11 imports Main
    12 begin
    12 begin
    13 
    13 
    17 by (metis Un_commute Un_left_absorb)
    17 by (metis Un_commute Un_left_absorb)
    18 
    18 
    19 type_synonym key = nat
    19 type_synonym key = nat
    20 
    20 
    21 consts
    21 consts
    22   all_symmetric :: bool        --{*true if all keys are symmetric*}
    22   all_symmetric :: bool        \<comment>\<open>true if all keys are symmetric\<close>
    23   invKey        :: "key=>key"  --{*inverse of a symmetric key*}
    23   invKey        :: "key=>key"  \<comment>\<open>inverse of a symmetric key\<close>
    24 
    24 
    25 specification (invKey)
    25 specification (invKey)
    26   invKey [simp]: "invKey (invKey K) = K"
    26   invKey [simp]: "invKey (invKey K) = K"
    27   invKey_symmetric: "all_symmetric --> invKey = id"
    27   invKey_symmetric: "all_symmetric --> invKey = id"
    28 by (metis id_apply)
    28 by (metis id_apply)
    29 
    29 
    30 
    30 
    31 text{*The inverse of a symmetric key is itself; that of a public key
    31 text\<open>The inverse of a symmetric key is itself; that of a public key
    32       is the private key and vice versa*}
    32       is the private key and vice versa\<close>
    33 
    33 
    34 definition symKeys :: "key set" where
    34 definition symKeys :: "key set" where
    35   "symKeys == {K. invKey K = K}"
    35   "symKeys == {K. invKey K = K}"
    36 
    36 
    37 datatype  --{*We allow any number of friendly agents*}
    37 datatype  \<comment>\<open>We allow any number of friendly agents\<close>
    38   agent = Server | Friend nat | Spy
    38   agent = Server | Friend nat | Spy
    39 
    39 
    40 datatype
    40 datatype
    41      msg = Agent  agent     --{*Agent names*}
    41      msg = Agent  agent     \<comment>\<open>Agent names\<close>
    42          | Number nat       --{*Ordinary integers, timestamps, ...*}
    42          | Number nat       \<comment>\<open>Ordinary integers, timestamps, ...\<close>
    43          | Nonce  nat       --{*Unguessable nonces*}
    43          | Nonce  nat       \<comment>\<open>Unguessable nonces\<close>
    44          | Key    key       --{*Crypto keys*}
    44          | Key    key       \<comment>\<open>Crypto keys\<close>
    45          | Hash   msg       --{*Hashing*}
    45          | Hash   msg       \<comment>\<open>Hashing\<close>
    46          | MPair  msg msg   --{*Compound messages*}
    46          | MPair  msg msg   \<comment>\<open>Compound messages\<close>
    47          | Crypt  key msg   --{*Encryption, public- or shared-key*}
    47          | Crypt  key msg   \<comment>\<open>Encryption, public- or shared-key\<close>
    48 
    48 
    49 
    49 
    50 text{*Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...*}
    50 text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
    51 syntax
    51 syntax
    52   "_MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
    52   "_MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
    53 translations
    53 translations
    54   "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
    54   "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
    55   "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
    55   "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
    56 
    56 
    57 
    57 
    58 definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
    58 definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
    59     --{*Message Y paired with a MAC computed with the help of X*}
    59     \<comment>\<open>Message Y paired with a MAC computed with the help of X\<close>
    60     "Hash[X] Y == \<lbrace> Hash\<lbrace>X,Y\<rbrace>, Y\<rbrace>"
    60     "Hash[X] Y == \<lbrace> Hash\<lbrace>X,Y\<rbrace>, Y\<rbrace>"
    61 
    61 
    62 definition keysFor :: "msg set => key set" where
    62 definition keysFor :: "msg set => key set" where
    63     --{*Keys useful to decrypt elements of a message set*}
    63     \<comment>\<open>Keys useful to decrypt elements of a message set\<close>
    64   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
    64   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
    65 
    65 
    66 
    66 
    67 subsubsection{*Inductive Definition of All Parts" of a Message*}
    67 subsubsection\<open>Inductive Definition of All Parts" of a Message\<close>
    68 
    68 
    69 inductive_set
    69 inductive_set
    70   parts :: "msg set => msg set"
    70   parts :: "msg set => msg set"
    71   for H :: "msg set"
    71   for H :: "msg set"
    72   where
    72   where
    81    apply (metis parts.Inj set_rev_mp)
    81    apply (metis parts.Inj set_rev_mp)
    82   apply (metis parts.Fst)
    82   apply (metis parts.Fst)
    83  apply (metis parts.Snd)
    83  apply (metis parts.Snd)
    84 by (metis parts.Body)
    84 by (metis parts.Body)
    85 
    85 
    86 text{*Equations hold because constructors are injective.*}
    86 text\<open>Equations hold because constructors are injective.\<close>
    87 lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
    87 lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
    88 by (metis agent.inject image_iff)
    88 by (metis agent.inject image_iff)
    89 
    89 
    90 lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x \<in> A)"
    90 lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x \<in> A)"
    91 by (metis image_iff msg.inject(4))
    91 by (metis image_iff msg.inject(4))
    92 
    92 
    93 lemma Nonce_Key_image_eq [simp]: "Nonce x \<notin> Key`A"
    93 lemma Nonce_Key_image_eq [simp]: "Nonce x \<notin> Key`A"
    94 by (metis image_iff msg.distinct(23))
    94 by (metis image_iff msg.distinct(23))
    95 
    95 
    96 
    96 
    97 subsubsection{*Inverse of keys *}
    97 subsubsection\<open>Inverse of keys\<close>
    98 
    98 
    99 lemma invKey_eq [simp]: "(invKey K = invKey K') = (K = K')"
    99 lemma invKey_eq [simp]: "(invKey K = invKey K') = (K = K')"
   100 by (metis invKey)
   100 by (metis invKey)
   101 
   101 
   102 
   102 
   103 subsection{*keysFor operator*}
   103 subsection\<open>keysFor operator\<close>
   104 
   104 
   105 lemma keysFor_empty [simp]: "keysFor {} = {}"
   105 lemma keysFor_empty [simp]: "keysFor {} = {}"
   106 by (unfold keysFor_def, blast)
   106 by (unfold keysFor_def, blast)
   107 
   107 
   108 lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
   108 lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
   109 by (unfold keysFor_def, blast)
   109 by (unfold keysFor_def, blast)
   110 
   110 
   111 lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
   111 lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
   112 by (unfold keysFor_def, blast)
   112 by (unfold keysFor_def, blast)
   113 
   113 
   114 text{*Monotonicity*}
   114 text\<open>Monotonicity\<close>
   115 lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
   115 lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
   116 by (unfold keysFor_def, blast)
   116 by (unfold keysFor_def, blast)
   117 
   117 
   118 lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
   118 lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
   119 by (unfold keysFor_def, auto)
   119 by (unfold keysFor_def, auto)
   142 
   142 
   143 lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
   143 lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
   144 by (unfold keysFor_def, blast)
   144 by (unfold keysFor_def, blast)
   145 
   145 
   146 
   146 
   147 subsection{*Inductive relation "parts"*}
   147 subsection\<open>Inductive relation "parts"\<close>
   148 
   148 
   149 lemma MPair_parts:
   149 lemma MPair_parts:
   150      "[| \<lbrace>X,Y\<rbrace> \<in> parts H;
   150      "[| \<lbrace>X,Y\<rbrace> \<in> parts H;
   151          [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
   151          [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
   152 by (blast dest: parts.Fst parts.Snd)
   152 by (blast dest: parts.Fst parts.Snd)
   153 
   153 
   154 declare MPair_parts [elim!] parts.Body [dest!]
   154 declare MPair_parts [elim!] parts.Body [dest!]
   155 text{*NB These two rules are UNSAFE in the formal sense, as they discard the
   155 text\<open>NB These two rules are UNSAFE in the formal sense, as they discard the
   156      compound message.  They work well on THIS FILE.
   156      compound message.  They work well on THIS FILE.
   157   @{text MPair_parts} is left as SAFE because it speeds up proofs.
   157   \<open>MPair_parts\<close> is left as SAFE because it speeds up proofs.
   158   The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
   158   The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\<close>
   159 
   159 
   160 lemma parts_increasing: "H \<subseteq> parts(H)"
   160 lemma parts_increasing: "H \<subseteq> parts(H)"
   161 by blast
   161 by blast
   162 
   162 
   163 lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD]
   163 lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD]
   169 done
   169 done
   170 
   170 
   171 lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
   171 lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
   172 by simp
   172 by simp
   173 
   173 
   174 text{*WARNING: loops if H = {Y}, therefore must not be repeated!*}
   174 text\<open>WARNING: loops if H = {Y}, therefore must not be repeated!\<close>
   175 lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
   175 lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
   176 apply (erule parts.induct)
   176 apply (erule parts.induct)
   177 apply fast+
   177 apply fast+
   178 done
   178 done
   179 
   179 
   180 
   180 
   181 subsubsection{*Unions *}
   181 subsubsection\<open>Unions\<close>
   182 
   182 
   183 lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
   183 lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
   184 by (intro Un_least parts_mono Un_upper1 Un_upper2)
   184 by (intro Un_least parts_mono Un_upper1 Un_upper2)
   185 
   185 
   186 lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
   186 lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
   210 done
   210 done
   211 
   211 
   212 lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
   212 lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
   213 by (intro equalityI parts_UN_subset1 parts_UN_subset2)
   213 by (intro equalityI parts_UN_subset1 parts_UN_subset2)
   214 
   214 
   215 text{*Added to simplify arguments to parts, analz and synth.
   215 text\<open>Added to simplify arguments to parts, analz and synth.
   216   NOTE: the UN versions are no longer used!*}
   216   NOTE: the UN versions are no longer used!\<close>
   217 
   217 
   218 
   218 
   219 text{*This allows @{text blast} to simplify occurrences of
   219 text\<open>This allows \<open>blast\<close> to simplify occurrences of
   220   @{term "parts(G\<union>H)"} in the assumption.*}
   220   @{term "parts(G\<union>H)"} in the assumption.\<close>
   221 lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE]
   221 lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE]
   222 declare in_parts_UnE [elim!]
   222 declare in_parts_UnE [elim!]
   223 
   223 
   224 lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
   224 lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
   225 by (blast intro: parts_mono [THEN [2] rev_subsetD])
   225 by (blast intro: parts_mono [THEN [2] rev_subsetD])
   226 
   226 
   227 subsubsection{*Idempotence and transitivity *}
   227 subsubsection\<open>Idempotence and transitivity\<close>
   228 
   228 
   229 lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
   229 lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
   230 by (erule parts.induct, blast+)
   230 by (erule parts.induct, blast+)
   231 
   231 
   232 lemma parts_idem [simp]: "parts (parts H) = parts H"
   232 lemma parts_idem [simp]: "parts (parts H) = parts H"
   243 
   243 
   244 lemma parts_cut: "[|Y\<in> parts (insert X G);  X\<in> parts H|] ==> Y\<in> parts(G \<union> H)"
   244 lemma parts_cut: "[|Y\<in> parts (insert X G);  X\<in> parts H|] ==> Y\<in> parts(G \<union> H)"
   245 by (metis (no_types) Un_insert_left Un_insert_right insert_absorb le_supE
   245 by (metis (no_types) Un_insert_left Un_insert_right insert_absorb le_supE
   246           parts_Un parts_idem parts_increasing parts_trans)
   246           parts_Un parts_idem parts_increasing parts_trans)
   247 
   247 
   248 subsubsection{*Rewrite rules for pulling out atomic messages *}
   248 subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
   249 
   249 
   250 lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
   250 lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
   251 
   251 
   252 
   252 
   253 lemma parts_insert_Agent [simp]:
   253 lemma parts_insert_Agent [simp]:
   308 apply (simp_all add: parts_insert2)
   308 apply (simp_all add: parts_insert2)
   309 apply (metis Suc_n_not_le_n)
   309 apply (metis Suc_n_not_le_n)
   310 apply (metis le_trans linorder_linear)
   310 apply (metis le_trans linorder_linear)
   311 done
   311 done
   312 
   312 
   313 subsection{*Inductive relation "analz"*}
   313 subsection\<open>Inductive relation "analz"\<close>
   314 
   314 
   315 text{*Inductive definition of "analz" -- what can be broken down from a set of
   315 text\<open>Inductive definition of "analz" -- what can be broken down from a set of
   316     messages, including keys.  A form of downward closure.  Pairs can
   316     messages, including keys.  A form of downward closure.  Pairs can
   317     be taken apart; messages decrypted with known keys.  *}
   317     be taken apart; messages decrypted with known keys.\<close>
   318 
   318 
   319 inductive_set
   319 inductive_set
   320   analz :: "msg set => msg set"
   320   analz :: "msg set => msg set"
   321   for H :: "msg set"
   321   for H :: "msg set"
   322   where
   322   where
   325   | Snd:     "\<lbrace>X,Y\<rbrace> \<in> analz H ==> Y \<in> analz H"
   325   | Snd:     "\<lbrace>X,Y\<rbrace> \<in> analz H ==> Y \<in> analz H"
   326   | Decrypt [dest]:
   326   | Decrypt [dest]:
   327              "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
   327              "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
   328 
   328 
   329 
   329 
   330 text{*Monotonicity; Lemma 1 of Lowe's paper*}
   330 text\<open>Monotonicity; Lemma 1 of Lowe's paper\<close>
   331 lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
   331 lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
   332 apply auto
   332 apply auto
   333 apply (erule analz.induct)
   333 apply (erule analz.induct)
   334 apply (auto dest: analz.Fst analz.Snd)
   334 apply (auto dest: analz.Fst analz.Snd)
   335 done
   335 done
   336 
   336 
   337 text{*Making it safe speeds up proofs*}
   337 text\<open>Making it safe speeds up proofs\<close>
   338 lemma MPair_analz [elim!]:
   338 lemma MPair_analz [elim!]:
   339      "[| \<lbrace>X,Y\<rbrace> \<in> analz H;
   339      "[| \<lbrace>X,Y\<rbrace> \<in> analz H;
   340              [| X \<in> analz H; Y \<in> analz H |] ==> P
   340              [| X \<in> analz H; Y \<in> analz H |] ==> P
   341           |] ==> P"
   341           |] ==> P"
   342 by (blast dest: analz.Fst analz.Snd)
   342 by (blast dest: analz.Fst analz.Snd)
   365 apply (erule analz.induct, auto)
   365 apply (erule analz.induct, auto)
   366 done
   366 done
   367 
   367 
   368 lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
   368 lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
   369 
   369 
   370 subsubsection{*General equational properties *}
   370 subsubsection\<open>General equational properties\<close>
   371 
   371 
   372 lemma analz_empty [simp]: "analz{} = {}"
   372 lemma analz_empty [simp]: "analz{} = {}"
   373 apply safe
   373 apply safe
   374 apply (erule analz.induct, blast+)
   374 apply (erule analz.induct, blast+)
   375 done
   375 done
   376 
   376 
   377 text{*Converse fails: we can analz more from the union than from the
   377 text\<open>Converse fails: we can analz more from the union than from the
   378   separate parts, as a key in one might decrypt a message in the other*}
   378   separate parts, as a key in one might decrypt a message in the other\<close>
   379 lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
   379 lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
   380 by (intro Un_least analz_mono Un_upper1 Un_upper2)
   380 by (intro Un_least analz_mono Un_upper1 Un_upper2)
   381 
   381 
   382 lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
   382 lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
   383 by (blast intro: analz_mono [THEN [2] rev_subsetD])
   383 by (blast intro: analz_mono [THEN [2] rev_subsetD])
   384 
   384 
   385 subsubsection{*Rewrite rules for pulling out atomic messages *}
   385 subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
   386 
   386 
   387 lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
   387 lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
   388 
   388 
   389 lemma analz_insert_Agent [simp]:
   389 lemma analz_insert_Agent [simp]:
   390      "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
   390      "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
   408      "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
   408      "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
   409 apply (rule analz_insert_eq_I)
   409 apply (rule analz_insert_eq_I)
   410 apply (erule analz.induct, auto)
   410 apply (erule analz.induct, auto)
   411 done
   411 done
   412 
   412 
   413 text{*Can only pull out Keys if they are not needed to decrypt the rest*}
   413 text\<open>Can only pull out Keys if they are not needed to decrypt the rest\<close>
   414 lemma analz_insert_Key [simp]:
   414 lemma analz_insert_Key [simp]:
   415     "K \<notin> keysFor (analz H) ==>
   415     "K \<notin> keysFor (analz H) ==>
   416           analz (insert (Key K) H) = insert (Key K) (analz H)"
   416           analz (insert (Key K) H) = insert (Key K) (analz H)"
   417 apply (unfold keysFor_def)
   417 apply (unfold keysFor_def)
   418 apply (rule analz_insert_eq_I)
   418 apply (rule analz_insert_eq_I)
   427 apply (erule analz.induct, auto)
   427 apply (erule analz.induct, auto)
   428 apply (erule analz.induct)
   428 apply (erule analz.induct)
   429 apply (blast intro: analz.Fst analz.Snd)+
   429 apply (blast intro: analz.Fst analz.Snd)+
   430 done
   430 done
   431 
   431 
   432 text{*Can pull out enCrypted message if the Key is not known*}
   432 text\<open>Can pull out enCrypted message if the Key is not known\<close>
   433 lemma analz_insert_Crypt:
   433 lemma analz_insert_Crypt:
   434      "Key (invKey K) \<notin> analz H
   434      "Key (invKey K) \<notin> analz H
   435       ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
   435       ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
   436 apply (rule analz_insert_eq_I)
   436 apply (rule analz_insert_eq_I)
   437 apply (erule analz.induct, auto)
   437 apply (erule analz.induct, auto)
   457      "Key (invKey K) \<in> analz H ==>
   457      "Key (invKey K) \<in> analz H ==>
   458                analz (insert (Crypt K X) H) =
   458                analz (insert (Crypt K X) H) =
   459                insert (Crypt K X) (analz (insert X H))"
   459                insert (Crypt K X) (analz (insert X H))"
   460 by (intro equalityI lemma1 lemma2)
   460 by (intro equalityI lemma1 lemma2)
   461 
   461 
   462 text{*Case analysis: either the message is secure, or it is not! Effective,
   462 text\<open>Case analysis: either the message is secure, or it is not! Effective,
   463 but can cause subgoals to blow up! Use with @{text "if_split"}; apparently
   463 but can cause subgoals to blow up! Use with \<open>if_split\<close>; apparently
   464 @{text "split_tac"} does not cope with patterns such as @{term"analz (insert
   464 \<open>split_tac\<close> does not cope with patterns such as @{term"analz (insert
   465 (Crypt K X) H)"} *}
   465 (Crypt K X) H)"}\<close>
   466 lemma analz_Crypt_if [simp]:
   466 lemma analz_Crypt_if [simp]:
   467      "analz (insert (Crypt K X) H) =
   467      "analz (insert (Crypt K X) H) =
   468           (if (Key (invKey K) \<in> analz H)
   468           (if (Key (invKey K) \<in> analz H)
   469            then insert (Crypt K X) (analz (insert X H))
   469            then insert (Crypt K X) (analz (insert X H))
   470            else insert (Crypt K X) (analz H))"
   470            else insert (Crypt K X) (analz H))"
   471 by (simp add: analz_insert_Crypt analz_insert_Decrypt)
   471 by (simp add: analz_insert_Crypt analz_insert_Decrypt)
   472 
   472 
   473 
   473 
   474 text{*This rule supposes "for the sake of argument" that we have the key.*}
   474 text\<open>This rule supposes "for the sake of argument" that we have the key.\<close>
   475 lemma analz_insert_Crypt_subset:
   475 lemma analz_insert_Crypt_subset:
   476      "analz (insert (Crypt K X) H) \<subseteq>
   476      "analz (insert (Crypt K X) H) \<subseteq>
   477            insert (Crypt K X) (analz (insert X H))"
   477            insert (Crypt K X) (analz (insert X H))"
   478 apply (rule subsetI)
   478 apply (rule subsetI)
   479 apply (erule analz.induct, auto)
   479 apply (erule analz.induct, auto)
   484 apply auto
   484 apply auto
   485 apply (erule analz.induct, auto)
   485 apply (erule analz.induct, auto)
   486 done
   486 done
   487 
   487 
   488 
   488 
   489 subsubsection{*Idempotence and transitivity *}
   489 subsubsection\<open>Idempotence and transitivity\<close>
   490 
   490 
   491 lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
   491 lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
   492 by (erule analz.induct, blast+)
   492 by (erule analz.induct, blast+)
   493 
   493 
   494 lemma analz_idem [simp]: "analz (analz H) = analz H"
   494 lemma analz_idem [simp]: "analz (analz H) = analz H"
   507 declare analz_trans[intro]
   507 declare analz_trans[intro]
   508 
   508 
   509 lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
   509 lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
   510 by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset)
   510 by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset)
   511 
   511 
   512 text{*This rewrite rule helps in the simplification of messages that involve
   512 text\<open>This rewrite rule helps in the simplification of messages that involve
   513   the forwarding of unknown components (X).  Without it, removing occurrences
   513   the forwarding of unknown components (X).  Without it, removing occurrences
   514   of X can be very complicated. *}
   514   of X can be very complicated.\<close>
   515 lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
   515 lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
   516 by (blast intro: analz_cut analz_insertI)
   516 by (blast intro: analz_cut analz_insertI)
   517 
   517 
   518 
   518 
   519 text{*A congruence rule for "analz" *}
   519 text\<open>A congruence rule for "analz"\<close>
   520 
   520 
   521 lemma analz_subset_cong:
   521 lemma analz_subset_cong:
   522      "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |]
   522      "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |]
   523       ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
   523       ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
   524 apply simp
   524 apply simp
   533 
   533 
   534 lemma analz_insert_cong:
   534 lemma analz_insert_cong:
   535      "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
   535      "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
   536 by (force simp only: insert_def intro!: analz_cong)
   536 by (force simp only: insert_def intro!: analz_cong)
   537 
   537 
   538 text{*If there are no pairs or encryptions then analz does nothing*}
   538 text\<open>If there are no pairs or encryptions then analz does nothing\<close>
   539 lemma analz_trivial:
   539 lemma analz_trivial:
   540      "[| \<forall>X Y. \<lbrace>X,Y\<rbrace> \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
   540      "[| \<forall>X Y. \<lbrace>X,Y\<rbrace> \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
   541 apply safe
   541 apply safe
   542 apply (erule analz.induct, blast+)
   542 apply (erule analz.induct, blast+)
   543 done
   543 done
   544 
   544 
   545 text{*These two are obsolete (with a single Spy) but cost little to prove...*}
   545 text\<open>These two are obsolete (with a single Spy) but cost little to prove...\<close>
   546 lemma analz_UN_analz_lemma:
   546 lemma analz_UN_analz_lemma:
   547      "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
   547      "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
   548 apply (erule analz.induct)
   548 apply (erule analz.induct)
   549 apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
   549 apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
   550 done
   550 done
   551 
   551 
   552 lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
   552 lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
   553 by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
   553 by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
   554 
   554 
   555 
   555 
   556 subsection{*Inductive relation "synth"*}
   556 subsection\<open>Inductive relation "synth"\<close>
   557 
   557 
   558 text{*Inductive definition of "synth" -- what can be built up from a set of
   558 text\<open>Inductive definition of "synth" -- what can be built up from a set of
   559     messages.  A form of upward closure.  Pairs can be built, messages
   559     messages.  A form of upward closure.  Pairs can be built, messages
   560     encrypted with known keys.  Agent names are public domain.
   560     encrypted with known keys.  Agent names are public domain.
   561     Numbers can be guessed, but Nonces cannot be.  *}
   561     Numbers can be guessed, but Nonces cannot be.\<close>
   562 
   562 
   563 inductive_set
   563 inductive_set
   564   synth :: "msg set => msg set"
   564   synth :: "msg set => msg set"
   565   for H :: "msg set"
   565   for H :: "msg set"
   566   where
   566   where
   569   | Number [intro]:   "Number n  \<in> synth H"
   569   | Number [intro]:   "Number n  \<in> synth H"
   570   | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
   570   | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
   571   | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> \<lbrace>X,Y\<rbrace> \<in> synth H"
   571   | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> \<lbrace>X,Y\<rbrace> \<in> synth H"
   572   | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
   572   | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
   573 
   573 
   574 text{*Monotonicity*}
   574 text\<open>Monotonicity\<close>
   575 lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
   575 lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
   576   by (auto, erule synth.induct, auto)
   576   by (auto, erule synth.induct, auto)
   577 
   577 
   578 text{*NO @{text Agent_synth}, as any Agent name can be synthesized.
   578 text\<open>NO \<open>Agent_synth\<close>, as any Agent name can be synthesized.
   579   The same holds for @{term Number}*}
   579   The same holds for @{term Number}\<close>
   580 inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
   580 inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
   581 inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
   581 inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
   582 inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
   582 inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
   583 inductive_cases MPair_synth [elim!]: "\<lbrace>X,Y\<rbrace> \<in> synth H"
   583 inductive_cases MPair_synth [elim!]: "\<lbrace>X,Y\<rbrace> \<in> synth H"
   584 inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
   584 inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
   585 
   585 
   586 
   586 
   587 lemma synth_increasing: "H \<subseteq> synth(H)"
   587 lemma synth_increasing: "H \<subseteq> synth(H)"
   588 by blast
   588 by blast
   589 
   589 
   590 subsubsection{*Unions *}
   590 subsubsection\<open>Unions\<close>
   591 
   591 
   592 text{*Converse fails: we can synth more from the union than from the
   592 text\<open>Converse fails: we can synth more from the union than from the
   593   separate parts, building a compound message using elements of each.*}
   593   separate parts, building a compound message using elements of each.\<close>
   594 lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
   594 lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
   595 by (intro Un_least synth_mono Un_upper1 Un_upper2)
   595 by (intro Un_least synth_mono Un_upper1 Un_upper2)
   596 
   596 
   597 lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
   597 lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
   598 by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono)
   598 by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono)
   599 
   599 
   600 subsubsection{*Idempotence and transitivity *}
   600 subsubsection\<open>Idempotence and transitivity\<close>
   601 
   601 
   602 lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
   602 lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
   603 by (erule synth.induct, blast+)
   603 by (erule synth.induct, blast+)
   604 
   604 
   605 lemma synth_idem: "synth (synth H) = synth H"
   605 lemma synth_idem: "synth (synth H) = synth H"
   637 lemma keysFor_synth [simp]:
   637 lemma keysFor_synth [simp]:
   638     "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
   638     "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
   639 by (unfold keysFor_def, blast)
   639 by (unfold keysFor_def, blast)
   640 
   640 
   641 
   641 
   642 subsubsection{*Combinations of parts, analz and synth *}
   642 subsubsection\<open>Combinations of parts, analz and synth\<close>
   643 
   643 
   644 lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
   644 lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
   645 apply (rule equalityI)
   645 apply (rule equalityI)
   646 apply (rule subsetI)
   646 apply (rule subsetI)
   647 apply (erule parts.induct)
   647 apply (erule parts.induct)
   679   hence "\<forall>x\<^sub>1. analz x\<^sub>1 \<union> synth x\<^sub>1 = analz (synth x\<^sub>1)" by (metis Un_commute)
   679   hence "\<forall>x\<^sub>1. analz x\<^sub>1 \<union> synth x\<^sub>1 = analz (synth x\<^sub>1)" by (metis Un_commute)
   680   thus "analz (synth H) = analz H \<union> synth H" by metis
   680   thus "analz (synth H) = analz H \<union> synth H" by metis
   681 qed
   681 qed
   682 
   682 
   683 
   683 
   684 subsubsection{*For reasoning about the Fake rule in traces *}
   684 subsubsection\<open>For reasoning about the Fake rule in traces\<close>
   685 
   685 
   686 lemma parts_insert_subset_Un: "X \<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
   686 lemma parts_insert_subset_Un: "X \<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
   687 proof -
   687 proof -
   688   assume "X \<in> G"
   688   assume "X \<in> G"
   689   hence "\<forall>x\<^sub>1. G \<subseteq> x\<^sub>1 \<longrightarrow> X \<in> x\<^sub>1 " by auto
   689   hence "\<forall>x\<^sub>1. G \<subseteq> x\<^sub>1 \<longrightarrow> X \<in> x\<^sub>1 " by auto