src/HOL/SET-Protocol/MessageSET.thy
changeset 14199 d3b8d972a488
child 14218 db95d1c2f51b
equal deleted inserted replaced
14198:8ea2645b8132 14199:d3b8d972a488
       
     1 (*  Title:      HOL/Auth/SET/MessageSET
       
     2     ID:         $Id$
       
     3     Authors:     Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
       
     4 *)
       
     5 
       
     6 header{*The Message Theory, Modified for SET*}
       
     7 
       
     8 theory MessageSET = NatPair:
       
     9 
       
    10 subsection{*General Lemmas*}
       
    11 
       
    12 text{*Needed occasionally with @{text spy_analz_tac}, e.g. in
       
    13      @{text analz_insert_Key_newK}*}
       
    14 
       
    15 lemma Un_absorb3 [simp] : "A \<union> (B \<union> A) = B \<union> A"
       
    16 by blast
       
    17 
       
    18 text{*Collapses redundant cases in the huge protocol proofs*}
       
    19 lemmas disj_simps = disj_comms disj_left_absorb disj_assoc 
       
    20 
       
    21 text{*Effective with assumptions like @{term "K \<notin> range pubK"} and 
       
    22    @{term "K \<notin> invKey`range pubK"}*}
       
    23 lemma notin_image_iff: "(y \<notin> f`I) = (\<forall>i\<in>I. f i \<noteq> y)"
       
    24 by blast
       
    25 
       
    26 text{*Effective with the assumption @{term "KK \<subseteq> - (range(invKey o pubK))"} *}
       
    27 lemma disjoint_image_iff: "(A <= - (f`I)) = (\<forall>i\<in>I. f i \<notin> A)"
       
    28 by blast
       
    29 
       
    30 
       
    31 
       
    32 types
       
    33   key = nat
       
    34 
       
    35 consts
       
    36   all_symmetric :: bool        --{*true if all keys are symmetric*}
       
    37   invKey        :: "key=>key"  --{*inverse of a symmetric key*}
       
    38 
       
    39 specification (invKey)
       
    40   invKey [simp]: "invKey (invKey K) = K"
       
    41   invKey_symmetric: "all_symmetric --> invKey = id"
       
    42     by (rule exI [of _ id], auto)
       
    43 
       
    44 
       
    45 text{*The inverse of a symmetric key is itself; that of a public key
       
    46       is the private key and vice versa*}
       
    47 
       
    48 constdefs
       
    49   symKeys :: "key set"
       
    50   "symKeys == {K. invKey K = K}"
       
    51 
       
    52 text{*Agents. We allow any number of certification authorities, cardholders
       
    53             merchants, and payment gateways.*}
       
    54 datatype
       
    55   agent = CA nat | Cardholder nat | Merchant nat | PG nat | Spy
       
    56 
       
    57 text{*Messages*}
       
    58 datatype
       
    59      msg = Agent  agent	    --{*Agent names*}
       
    60          | Number nat       --{*Ordinary integers, timestamps, ...*}
       
    61          | Nonce  nat       --{*Unguessable nonces*}
       
    62          | Pan    nat       --{*Unguessable Primary Account Numbers (??)*}
       
    63          | Key    key       --{*Crypto keys*}
       
    64 	 | Hash   msg       --{*Hashing*}
       
    65 	 | MPair  msg msg   --{*Compound messages*}
       
    66 	 | Crypt  key msg   --{*Encryption, public- or shared-key*}
       
    67 
       
    68 
       
    69 (*Concrete syntax: messages appear as {|A,B,NA|}, etc...*)
       
    70 syntax
       
    71   "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
       
    72 
       
    73 syntax (xsymbols)
       
    74   "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
       
    75 
       
    76 translations
       
    77   "{|x, y, z|}"   == "{|x, {|y, z|}|}"
       
    78   "{|x, y|}"      == "MPair x y"
       
    79 
       
    80 
       
    81 constdefs
       
    82   nat_of_agent :: "agent => nat"
       
    83    "nat_of_agent == agent_case (curry nat2_to_nat 0)
       
    84 			       (curry nat2_to_nat 1)
       
    85 			       (curry nat2_to_nat 2)
       
    86 			       (curry nat2_to_nat 3)
       
    87 			       (nat2_to_nat (4,0))"
       
    88     --{*maps each agent to a unique natural number, for specifications*}
       
    89 
       
    90 text{*The function is indeed injective*}
       
    91 lemma inj_nat_of_agent: "inj nat_of_agent"
       
    92 by (simp add: nat_of_agent_def inj_on_def curry_def
       
    93               nat2_to_nat_inj [THEN inj_eq]  split: agent.split) 
       
    94 
       
    95 
       
    96 constdefs
       
    97   (*Keys useful to decrypt elements of a message set*)
       
    98   keysFor :: "msg set => key set"
       
    99   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
       
   100 
       
   101 subsubsection{*Inductive definition of all "parts" of a message.*}
       
   102 
       
   103 consts  parts   :: "msg set => msg set"
       
   104 inductive "parts H"
       
   105   intros
       
   106     Inj [intro]:               "X \<in> H ==> X \<in> parts H"
       
   107     Fst:         "{|X,Y|}   \<in> parts H ==> X \<in> parts H"
       
   108     Snd:         "{|X,Y|}   \<in> parts H ==> Y \<in> parts H"
       
   109     Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
       
   110 
       
   111 
       
   112 (*Monotonicity*)
       
   113 lemma parts_mono: "G<=H ==> parts(G) <= parts(H)"
       
   114 apply auto
       
   115 apply (erule parts.induct)
       
   116 apply (auto dest: Fst Snd Body)
       
   117 done
       
   118 
       
   119 
       
   120 subsubsection{*Inverse of keys*}
       
   121 
       
   122 (*Equations hold because constructors are injective; cannot prove for all f*)
       
   123 lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
       
   124 by auto
       
   125 
       
   126 lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
       
   127 by auto
       
   128 
       
   129 lemma Cardholder_image_eq [simp]: "(Cardholder x \<in> Cardholder`A) = (x \<in> A)"
       
   130 by auto
       
   131 
       
   132 lemma CA_image_eq [simp]: "(CA x \<in> CA`A) = (x \<in> A)"
       
   133 by auto
       
   134 
       
   135 lemma Pan_image_eq [simp]: "(Pan x \<in> Pan`A) = (x \<in> A)"
       
   136 by auto
       
   137 
       
   138 lemma Pan_Key_image_eq [simp]: "(Pan x \<notin> Key`A)"
       
   139 by auto
       
   140 
       
   141 lemma Nonce_Pan_image_eq [simp]: "(Nonce x \<notin> Pan`A)"
       
   142 by auto
       
   143 
       
   144 lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
       
   145 apply safe
       
   146 apply (drule_tac f = invKey in arg_cong, simp)
       
   147 done
       
   148 
       
   149 
       
   150 subsection{*keysFor operator*}
       
   151 
       
   152 lemma keysFor_empty [simp]: "keysFor {} = {}"
       
   153 by (unfold keysFor_def, blast)
       
   154 
       
   155 lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
       
   156 by (unfold keysFor_def, blast)
       
   157 
       
   158 lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
       
   159 by (unfold keysFor_def, blast)
       
   160 
       
   161 (*Monotonicity*)
       
   162 lemma keysFor_mono: "G\<subseteq>H ==> keysFor(G) \<subseteq> keysFor(H)"
       
   163 by (unfold keysFor_def, blast)
       
   164 
       
   165 lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
       
   166 by (unfold keysFor_def, auto)
       
   167 
       
   168 lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
       
   169 by (unfold keysFor_def, auto)
       
   170 
       
   171 lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
       
   172 by (unfold keysFor_def, auto)
       
   173 
       
   174 lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
       
   175 by (unfold keysFor_def, auto)
       
   176 
       
   177 lemma keysFor_insert_Pan [simp]: "keysFor (insert (Pan A) H) = keysFor H"
       
   178 by (unfold keysFor_def, auto)
       
   179 
       
   180 lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
       
   181 by (unfold keysFor_def, auto)
       
   182 
       
   183 lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H"
       
   184 by (unfold keysFor_def, auto)
       
   185 
       
   186 lemma keysFor_insert_Crypt [simp]:
       
   187     "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
       
   188 by (unfold keysFor_def, auto)
       
   189 
       
   190 lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
       
   191 by (unfold keysFor_def, auto)
       
   192 
       
   193 lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
       
   194 by (unfold keysFor_def, blast)
       
   195 
       
   196 
       
   197 subsection{*Inductive relation "parts"*}
       
   198 
       
   199 lemma MPair_parts:
       
   200      "[| {|X,Y|} \<in> parts H;
       
   201          [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
       
   202 by (blast dest: parts.Fst parts.Snd)
       
   203 
       
   204 declare MPair_parts [elim!]  parts.Body [dest!]
       
   205 text{*NB These two rules are UNSAFE in the formal sense, as they discard the
       
   206      compound message.  They work well on THIS FILE.
       
   207   @{text MPair_parts} is left as SAFE because it speeds up proofs.
       
   208   The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
       
   209 
       
   210 lemma parts_increasing: "H \<subseteq> parts(H)"
       
   211 by blast
       
   212 
       
   213 lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard]
       
   214 
       
   215 lemma parts_empty [simp]: "parts{} = {}"
       
   216 apply safe
       
   217 apply (erule parts.induct, blast+)
       
   218 done
       
   219 
       
   220 lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
       
   221 by simp
       
   222 
       
   223 (*WARNING: loops if H = {Y}, therefore must not be repeated!*)
       
   224 lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
       
   225 by (erule parts.induct, blast+)
       
   226 
       
   227 
       
   228 subsubsection{*Unions*}
       
   229 
       
   230 lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
       
   231 by (intro Un_least parts_mono Un_upper1 Un_upper2)
       
   232 
       
   233 lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
       
   234 apply (rule subsetI)
       
   235 apply (erule parts.induct, blast+)
       
   236 done
       
   237 
       
   238 lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
       
   239 by (intro equalityI parts_Un_subset1 parts_Un_subset2)
       
   240 
       
   241 lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
       
   242 apply (subst insert_is_Un [of _ H])
       
   243 apply (simp only: parts_Un)
       
   244 done
       
   245 
       
   246 (*TWO inserts to avoid looping.  This rewrite is better than nothing.
       
   247   Not suitable for Addsimps: its behaviour can be strange.*)
       
   248 lemma parts_insert2:
       
   249      "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
       
   250 apply (simp add: Un_assoc)
       
   251 apply (simp add: parts_insert [symmetric])
       
   252 done
       
   253 
       
   254 lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
       
   255 by (intro UN_least parts_mono UN_upper)
       
   256 
       
   257 lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
       
   258 apply (rule subsetI)
       
   259 apply (erule parts.induct, blast+)
       
   260 done
       
   261 
       
   262 lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
       
   263 by (intro equalityI parts_UN_subset1 parts_UN_subset2)
       
   264 
       
   265 (*Added to simplify arguments to parts, analz and synth.
       
   266   NOTE: the UN versions are no longer used!*)
       
   267 
       
   268 
       
   269 text{*This allows @{text blast} to simplify occurrences of
       
   270   @{term "parts(G\<union>H)"} in the assumption.*}
       
   271 declare parts_Un [THEN equalityD1, THEN subsetD, THEN UnE, elim!]
       
   272 
       
   273 
       
   274 lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
       
   275 by (blast intro: parts_mono [THEN [2] rev_subsetD])
       
   276 
       
   277 subsubsection{*Idempotence and transitivity*}
       
   278 
       
   279 lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
       
   280 by (erule parts.induct, blast+)
       
   281 
       
   282 lemma parts_idem [simp]: "parts (parts H) = parts H"
       
   283 by blast
       
   284 
       
   285 lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
       
   286 by (drule parts_mono, blast)
       
   287 
       
   288 (*Cut*)
       
   289 lemma parts_cut:
       
   290      "[| Y\<in> parts (insert X G);  X\<in> parts H |] ==> Y\<in> parts (G \<union> H)"
       
   291 by (erule parts_trans, auto)
       
   292 
       
   293 lemma parts_cut_eq [simp]: "X\<in> parts H ==> parts (insert X H) = parts H"
       
   294 by (force dest!: parts_cut intro: parts_insertI)
       
   295 
       
   296 
       
   297 subsubsection{*Rewrite rules for pulling out atomic messages*}
       
   298 
       
   299 lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
       
   300 
       
   301 
       
   302 lemma parts_insert_Agent [simp]:
       
   303      "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
       
   304 apply (rule parts_insert_eq_I)
       
   305 apply (erule parts.induct, auto)
       
   306 done
       
   307 
       
   308 lemma parts_insert_Nonce [simp]:
       
   309      "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
       
   310 apply (rule parts_insert_eq_I)
       
   311 apply (erule parts.induct, auto)
       
   312 done
       
   313 
       
   314 lemma parts_insert_Number [simp]:
       
   315      "parts (insert (Number N) H) = insert (Number N) (parts H)"
       
   316 apply (rule parts_insert_eq_I)
       
   317 apply (erule parts.induct, auto)
       
   318 done
       
   319 
       
   320 lemma parts_insert_Key [simp]:
       
   321      "parts (insert (Key K) H) = insert (Key K) (parts H)"
       
   322 apply (rule parts_insert_eq_I)
       
   323 apply (erule parts.induct, auto)
       
   324 done
       
   325 
       
   326 lemma parts_insert_Pan [simp]:
       
   327      "parts (insert (Pan A) H) = insert (Pan A) (parts H)"
       
   328 apply (rule parts_insert_eq_I)
       
   329 apply (erule parts.induct, auto)
       
   330 done
       
   331 
       
   332 lemma parts_insert_Hash [simp]:
       
   333      "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
       
   334 apply (rule parts_insert_eq_I)
       
   335 apply (erule parts.induct, auto)
       
   336 done
       
   337 
       
   338 lemma parts_insert_Crypt [simp]:
       
   339      "parts (insert (Crypt K X) H) =
       
   340           insert (Crypt K X) (parts (insert X H))"
       
   341 apply (rule equalityI)
       
   342 apply (rule subsetI)
       
   343 apply (erule parts.induct, auto)
       
   344 apply (erule parts.induct)
       
   345 apply (blast intro: parts.Body)+
       
   346 done
       
   347 
       
   348 lemma parts_insert_MPair [simp]:
       
   349      "parts (insert {|X,Y|} H) =
       
   350           insert {|X,Y|} (parts (insert X (insert Y H)))"
       
   351 apply (rule equalityI)
       
   352 apply (rule subsetI)
       
   353 apply (erule parts.induct, auto)
       
   354 apply (erule parts.induct)
       
   355 apply (blast intro: parts.Fst parts.Snd)+
       
   356 done
       
   357 
       
   358 lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
       
   359 apply auto
       
   360 apply (erule parts.induct, auto)
       
   361 done
       
   362 
       
   363 lemma parts_image_Pan [simp]: "parts (Pan`A) = Pan`A"
       
   364 apply auto
       
   365 apply (erule parts.induct, auto)
       
   366 done
       
   367 
       
   368 
       
   369 (*In any message, there is an upper bound N on its greatest nonce.*)
       
   370 lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
       
   371 apply (induct_tac "msg")
       
   372 apply (simp_all (no_asm_simp) add: exI parts_insert2)
       
   373 (*MPair case: blast_tac works out the necessary sum itself!*)
       
   374 prefer 2 apply (blast elim!: add_leE)
       
   375 (*Nonce case*)
       
   376 apply (rule_tac x = "N + Suc nat" in exI)
       
   377 apply (auto elim!: add_leE)
       
   378 done
       
   379 
       
   380 (* Ditto, for numbers.*)
       
   381 lemma msg_Number_supply: "\<exists>N. \<forall>n. N<=n --> Number n \<notin> parts {msg}"
       
   382 apply (induct_tac "msg")
       
   383 apply (simp_all (no_asm_simp) add: exI parts_insert2)
       
   384 prefer 2 apply (blast elim!: add_leE)
       
   385 apply (rule_tac x = "N + Suc nat" in exI, auto)
       
   386 done
       
   387 
       
   388 subsection{*Inductive relation "analz"*}
       
   389 
       
   390 text{*Inductive definition of "analz" -- what can be broken down from a set of
       
   391     messages, including keys.  A form of downward closure.  Pairs can
       
   392     be taken apart; messages decrypted with known keys.*}
       
   393 
       
   394 consts  analz   :: "msg set => msg set"
       
   395 inductive "analz H"
       
   396   intros
       
   397     Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
       
   398     Fst:     "{|X,Y|} \<in> analz H ==> X \<in> analz H"
       
   399     Snd:     "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
       
   400     Decrypt [dest]:
       
   401              "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
       
   402 
       
   403 
       
   404 (*Monotonicity; Lemma 1 of Lowe's paper*)
       
   405 lemma analz_mono: "G<=H ==> analz(G) <= analz(H)"
       
   406 apply auto
       
   407 apply (erule analz.induct)
       
   408 apply (auto dest: Fst Snd)
       
   409 done
       
   410 
       
   411 text{*Making it safe speeds up proofs*}
       
   412 lemma MPair_analz [elim!]:
       
   413      "[| {|X,Y|} \<in> analz H;
       
   414              [| X \<in> analz H; Y \<in> analz H |] ==> P
       
   415           |] ==> P"
       
   416 by (blast dest: analz.Fst analz.Snd)
       
   417 
       
   418 lemma analz_increasing: "H \<subseteq> analz(H)"
       
   419 by blast
       
   420 
       
   421 lemma analz_subset_parts: "analz H \<subseteq> parts H"
       
   422 apply (rule subsetI)
       
   423 apply (erule analz.induct, blast+)
       
   424 done
       
   425 
       
   426 lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard]
       
   427 
       
   428 lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard]
       
   429 
       
   430 
       
   431 lemma parts_analz [simp]: "parts (analz H) = parts H"
       
   432 apply (rule equalityI)
       
   433 apply (rule analz_subset_parts [THEN parts_mono, THEN subset_trans], simp)
       
   434 apply (blast intro: analz_increasing [THEN parts_mono, THEN subsetD])
       
   435 done
       
   436 
       
   437 lemma analz_parts [simp]: "analz (parts H) = parts H"
       
   438 apply auto
       
   439 apply (erule analz.induct, auto)
       
   440 done
       
   441 
       
   442 lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard]
       
   443 
       
   444 subsubsection{*General equational properties*}
       
   445 
       
   446 lemma analz_empty [simp]: "analz{} = {}"
       
   447 apply safe
       
   448 apply (erule analz.induct, blast+)
       
   449 done
       
   450 
       
   451 (*Converse fails: we can analz more from the union than from the
       
   452   separate parts, as a key in one might decrypt a message in the other*)
       
   453 lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
       
   454 by (intro Un_least analz_mono Un_upper1 Un_upper2)
       
   455 
       
   456 lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
       
   457 by (blast intro: analz_mono [THEN [2] rev_subsetD])
       
   458 
       
   459 subsubsection{*Rewrite rules for pulling out atomic messages*}
       
   460 
       
   461 lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
       
   462 
       
   463 lemma analz_insert_Agent [simp]:
       
   464      "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
       
   465 apply (rule analz_insert_eq_I)
       
   466 apply (erule analz.induct, auto)
       
   467 done
       
   468 
       
   469 lemma analz_insert_Nonce [simp]:
       
   470      "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
       
   471 apply (rule analz_insert_eq_I)
       
   472 apply (erule analz.induct, auto)
       
   473 done
       
   474 
       
   475 lemma analz_insert_Number [simp]:
       
   476      "analz (insert (Number N) H) = insert (Number N) (analz H)"
       
   477 apply (rule analz_insert_eq_I)
       
   478 apply (erule analz.induct, auto)
       
   479 done
       
   480 
       
   481 lemma analz_insert_Hash [simp]:
       
   482      "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
       
   483 apply (rule analz_insert_eq_I)
       
   484 apply (erule analz.induct, auto)
       
   485 done
       
   486 
       
   487 (*Can only pull out Keys if they are not needed to decrypt the rest*)
       
   488 lemma analz_insert_Key [simp]:
       
   489     "K \<notin> keysFor (analz H) ==>
       
   490           analz (insert (Key K) H) = insert (Key K) (analz H)"
       
   491 apply (unfold keysFor_def)
       
   492 apply (rule analz_insert_eq_I)
       
   493 apply (erule analz.induct, auto)
       
   494 done
       
   495 
       
   496 lemma analz_insert_MPair [simp]:
       
   497      "analz (insert {|X,Y|} H) =
       
   498           insert {|X,Y|} (analz (insert X (insert Y H)))"
       
   499 apply (rule equalityI)
       
   500 apply (rule subsetI)
       
   501 apply (erule analz.induct, auto)
       
   502 apply (erule analz.induct)
       
   503 apply (blast intro: analz.Fst analz.Snd)+
       
   504 done
       
   505 
       
   506 (*Can pull out enCrypted message if the Key is not known*)
       
   507 lemma analz_insert_Crypt:
       
   508      "Key (invKey K) \<notin> analz H
       
   509       ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
       
   510 apply (rule analz_insert_eq_I)
       
   511 apply (erule analz.induct, auto)
       
   512 done
       
   513 
       
   514 lemma analz_insert_Pan [simp]:
       
   515      "analz (insert (Pan A) H) = insert (Pan A) (analz H)"
       
   516 apply (rule analz_insert_eq_I)
       
   517 apply (erule analz.induct, auto)
       
   518 done
       
   519 
       
   520 lemma lemma1: "Key (invKey K) \<in> analz H ==>
       
   521                analz (insert (Crypt K X) H) \<subseteq>
       
   522                insert (Crypt K X) (analz (insert X H))"
       
   523 apply (rule subsetI)
       
   524 apply (erule_tac xa = x in analz.induct, auto)
       
   525 done
       
   526 
       
   527 lemma lemma2: "Key (invKey K) \<in> analz H ==>
       
   528                insert (Crypt K X) (analz (insert X H)) \<subseteq>
       
   529                analz (insert (Crypt K X) H)"
       
   530 apply auto
       
   531 apply (erule_tac xa = x in analz.induct, auto)
       
   532 apply (blast intro: analz_insertI analz.Decrypt)
       
   533 done
       
   534 
       
   535 lemma analz_insert_Decrypt:
       
   536      "Key (invKey K) \<in> analz H ==>
       
   537                analz (insert (Crypt K X) H) =
       
   538                insert (Crypt K X) (analz (insert X H))"
       
   539 by (intro equalityI lemma1 lemma2)
       
   540 
       
   541 (*Case analysis: either the message is secure, or it is not!
       
   542   Effective, but can cause subgoals to blow up!
       
   543   Use with split_if;  apparently split_tac does not cope with patterns
       
   544   such as "analz (insert (Crypt K X) H)" *)
       
   545 lemma analz_Crypt_if [simp]:
       
   546      "analz (insert (Crypt K X) H) =
       
   547           (if (Key (invKey K) \<in> analz H)
       
   548            then insert (Crypt K X) (analz (insert X H))
       
   549            else insert (Crypt K X) (analz H))"
       
   550 by (simp add: analz_insert_Crypt analz_insert_Decrypt)
       
   551 
       
   552 
       
   553 (*This rule supposes "for the sake of argument" that we have the key.*)
       
   554 lemma analz_insert_Crypt_subset:
       
   555      "analz (insert (Crypt K X) H) \<subseteq>
       
   556            insert (Crypt K X) (analz (insert X H))"
       
   557 apply (rule subsetI)
       
   558 apply (erule analz.induct, auto)
       
   559 done
       
   560 
       
   561 lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
       
   562 apply auto
       
   563 apply (erule analz.induct, auto)
       
   564 done
       
   565 
       
   566 lemma analz_image_Pan [simp]: "analz (Pan`A) = Pan`A"
       
   567 apply auto
       
   568 apply (erule analz.induct, auto)
       
   569 done
       
   570 
       
   571 
       
   572 subsubsection{*Idempotence and transitivity*}
       
   573 
       
   574 lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
       
   575 by (erule analz.induct, blast+)
       
   576 
       
   577 lemma analz_idem [simp]: "analz (analz H) = analz H"
       
   578 by blast
       
   579 
       
   580 lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
       
   581 by (drule analz_mono, blast)
       
   582 
       
   583 (*Cut; Lemma 2 of Lowe*)
       
   584 lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
       
   585 by (erule analz_trans, blast)
       
   586 
       
   587 (*Cut can be proved easily by induction on
       
   588    "Y: analz (insert X H) ==> X: analz H --> Y: analz H"
       
   589 *)
       
   590 
       
   591 (*This rewrite rule helps in the simplification of messages that involve
       
   592   the forwarding of unknown components (X).  Without it, removing occurrences
       
   593   of X can be very complicated. *)
       
   594 lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
       
   595 by (blast intro: analz_cut analz_insertI)
       
   596 
       
   597 
       
   598 text{*A congruence rule for "analz"*}
       
   599 
       
   600 lemma analz_subset_cong:
       
   601      "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H'
       
   602                |] ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
       
   603 apply clarify
       
   604 apply (erule analz.induct)
       
   605 apply (best intro: analz_mono [THEN subsetD])+
       
   606 done
       
   607 
       
   608 lemma analz_cong:
       
   609      "[| analz G = analz G'; analz H = analz H'
       
   610                |] ==> analz (G \<union> H) = analz (G' \<union> H')"
       
   611 by (intro equalityI analz_subset_cong, simp_all)
       
   612 
       
   613 lemma analz_insert_cong:
       
   614      "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
       
   615 by (force simp only: insert_def intro!: analz_cong)
       
   616 
       
   617 (*If there are no pairs or encryptions then analz does nothing*)
       
   618 lemma analz_trivial:
       
   619      "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
       
   620 apply safe
       
   621 apply (erule analz.induct, blast+)
       
   622 done
       
   623 
       
   624 (*These two are obsolete (with a single Spy) but cost little to prove...*)
       
   625 lemma analz_UN_analz_lemma:
       
   626      "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
       
   627 apply (erule analz.induct)
       
   628 apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
       
   629 done
       
   630 
       
   631 lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
       
   632 by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
       
   633 
       
   634 
       
   635 subsection{*Inductive relation "synth"*}
       
   636 
       
   637 text{*Inductive definition of "synth" -- what can be built up from a set of
       
   638     messages.  A form of upward closure.  Pairs can be built, messages
       
   639     encrypted with known keys.  Agent names are public domain.
       
   640     Numbers can be guessed, but Nonces cannot be.*}
       
   641 
       
   642 consts  synth   :: "msg set => msg set"
       
   643 inductive "synth H"
       
   644   intros
       
   645     Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
       
   646     Agent  [intro]:   "Agent agt \<in> synth H"
       
   647     Number [intro]:   "Number n  \<in> synth H"
       
   648     Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
       
   649     MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
       
   650     Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
       
   651 
       
   652 (*Monotonicity*)
       
   653 lemma synth_mono: "G<=H ==> synth(G) <= synth(H)"
       
   654 apply auto
       
   655 apply (erule synth.induct)
       
   656 apply (auto dest: Fst Snd Body)
       
   657 done
       
   658 
       
   659 (*NO Agent_synth, as any Agent name can be synthesized.  Ditto for Number*)
       
   660 inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
       
   661 inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
       
   662 inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
       
   663 inductive_cases MPair_synth [elim!]: "{|X,Y|} \<in> synth H"
       
   664 inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
       
   665 inductive_cases Pan_synth   [elim!]: "Pan A \<in> synth H"
       
   666 
       
   667 
       
   668 lemma synth_increasing: "H \<subseteq> synth(H)"
       
   669 by blast
       
   670 
       
   671 subsubsection{*Unions*}
       
   672 
       
   673 (*Converse fails: we can synth more from the union than from the
       
   674   separate parts, building a compound message using elements of each.*)
       
   675 lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
       
   676 by (intro Un_least synth_mono Un_upper1 Un_upper2)
       
   677 
       
   678 lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
       
   679 by (blast intro: synth_mono [THEN [2] rev_subsetD])
       
   680 
       
   681 subsubsection{*Idempotence and transitivity*}
       
   682 
       
   683 lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
       
   684 by (erule synth.induct, blast+)
       
   685 
       
   686 lemma synth_idem: "synth (synth H) = synth H"
       
   687 by blast
       
   688 
       
   689 lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
       
   690 by (drule synth_mono, blast)
       
   691 
       
   692 (*Cut; Lemma 2 of Lowe*)
       
   693 lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
       
   694 by (erule synth_trans, blast)
       
   695 
       
   696 lemma Agent_synth [simp]: "Agent A \<in> synth H"
       
   697 by blast
       
   698 
       
   699 lemma Number_synth [simp]: "Number n \<in> synth H"
       
   700 by blast
       
   701 
       
   702 lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
       
   703 by blast
       
   704 
       
   705 lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
       
   706 by blast
       
   707 
       
   708 lemma Crypt_synth_eq [simp]: "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
       
   709 by blast
       
   710 
       
   711 lemma Pan_synth_eq [simp]: "(Pan A \<in> synth H) = (Pan A \<in> H)"
       
   712 by blast
       
   713 
       
   714 lemma keysFor_synth [simp]:
       
   715     "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
       
   716 by (unfold keysFor_def, blast)
       
   717 
       
   718 
       
   719 subsubsection{*Combinations of parts, analz and synth*}
       
   720 
       
   721 lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
       
   722 apply (rule equalityI)
       
   723 apply (rule subsetI)
       
   724 apply (erule parts.induct)
       
   725 apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD]
       
   726                     parts.Fst parts.Snd parts.Body)+
       
   727 done
       
   728 
       
   729 lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
       
   730 apply (intro equalityI analz_subset_cong)+
       
   731 apply simp_all
       
   732 done
       
   733 
       
   734 lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
       
   735 apply (rule equalityI)
       
   736 apply (rule subsetI)
       
   737 apply (erule analz.induct)
       
   738 prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD])
       
   739 apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+
       
   740 done
       
   741 
       
   742 lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
       
   743 apply (cut_tac H = "{}" in analz_synth_Un)
       
   744 apply (simp (no_asm_use))
       
   745 done
       
   746 
       
   747 
       
   748 subsubsection{*For reasoning about the Fake rule in traces*}
       
   749 
       
   750 lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
       
   751 by (rule subset_trans [OF parts_mono parts_Un_subset2], blast)
       
   752 
       
   753 (*More specifically for Fake.  Very occasionally we could do with a version
       
   754   of the form  parts{X} \<subseteq> synth (analz H) \<union> parts H *)
       
   755 lemma Fake_parts_insert: "X \<in> synth (analz H) ==>
       
   756       parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
       
   757 apply (drule parts_insert_subset_Un)
       
   758 apply (simp (no_asm_use))
       
   759 apply blast
       
   760 done
       
   761 
       
   762 lemma Fake_parts_insert_in_Un:
       
   763      "[|Z \<in> parts (insert X H);  X: synth (analz H)|] 
       
   764       ==> Z \<in>  synth (analz H) \<union> parts H";
       
   765 by (blast dest: Fake_parts_insert [THEN subsetD, dest])
       
   766 
       
   767 (*H is sometimes (Key ` KK \<union> spies evs), so can't put G=H*)
       
   768 lemma Fake_analz_insert:
       
   769      "X\<in> synth (analz G) ==>
       
   770       analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
       
   771 apply (rule subsetI)
       
   772 apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
       
   773 prefer 2 apply (blast intro: analz_mono [THEN [2] rev_subsetD] analz_mono [THEN synth_mono, THEN [2] rev_subsetD])
       
   774 apply (simp (no_asm_use))
       
   775 apply blast
       
   776 done
       
   777 
       
   778 lemma analz_conj_parts [simp]:
       
   779      "(X \<in> analz H & X \<in> parts H) = (X \<in> analz H)"
       
   780 by (blast intro: analz_subset_parts [THEN subsetD])
       
   781 
       
   782 lemma analz_disj_parts [simp]:
       
   783      "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
       
   784 by (blast intro: analz_subset_parts [THEN subsetD])
       
   785 
       
   786 (*Without this equation, other rules for synth and analz would yield
       
   787   redundant cases*)
       
   788 lemma MPair_synth_analz [iff]:
       
   789      "({|X,Y|} \<in> synth (analz H)) =
       
   790       (X \<in> synth (analz H) & Y \<in> synth (analz H))"
       
   791 by blast
       
   792 
       
   793 lemma Crypt_synth_analz:
       
   794      "[| Key K \<in> analz H;  Key (invKey K) \<in> analz H |]
       
   795        ==> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
       
   796 by blast
       
   797 
       
   798 
       
   799 lemma Hash_synth_analz [simp]:
       
   800      "X \<notin> synth (analz H)
       
   801       ==> (Hash{|X,Y|} \<in> synth (analz H)) = (Hash{|X,Y|} \<in> analz H)"
       
   802 by blast
       
   803 
       
   804 
       
   805 (*We do NOT want Crypt... messages broken up in protocols!!*)
       
   806 declare parts.Body [rule del]
       
   807 
       
   808 
       
   809 text{*Rewrites to push in Key and Crypt messages, so that other messages can
       
   810     be pulled out using the @{text analz_insert} rules*}
       
   811 ML
       
   812 {*
       
   813 fun insComm x y = inst "x" x (inst "y" y insert_commute);
       
   814 
       
   815 bind_thms ("pushKeys",
       
   816            map (insComm "Key ?K")
       
   817                    ["Agent ?C", "Nonce ?N", "Number ?N", "Pan ?PAN",
       
   818 		    "Hash ?X", "MPair ?X ?Y", "Crypt ?X ?K'"]);
       
   819 
       
   820 bind_thms ("pushCrypts",
       
   821            map (insComm "Crypt ?X ?K")
       
   822                      ["Agent ?C", "Nonce ?N", "Number ?N", "Pan ?PAN",
       
   823 		      "Hash ?X'", "MPair ?X' ?Y"]);
       
   824 *}
       
   825 
       
   826 text{*Cannot be added with @{text "[simp]"} -- messages should not always be
       
   827   re-ordered.*}
       
   828 lemmas pushes = pushKeys pushCrypts
       
   829 
       
   830 
       
   831 subsection{*Tactics useful for many protocol proofs*}
       
   832 ML
       
   833 {*
       
   834 val invKey = thm "invKey"
       
   835 val keysFor_def = thm "keysFor_def"
       
   836 val symKeys_def = thm "symKeys_def"
       
   837 val parts_mono = thm "parts_mono";
       
   838 val analz_mono = thm "analz_mono";
       
   839 val Key_image_eq = thm "Key_image_eq";
       
   840 val Nonce_Key_image_eq = thm "Nonce_Key_image_eq";
       
   841 val keysFor_Un = thm "keysFor_Un";
       
   842 val keysFor_mono = thm "keysFor_mono";
       
   843 val keysFor_image_Key = thm "keysFor_image_Key";
       
   844 val Crypt_imp_invKey_keysFor = thm "Crypt_imp_invKey_keysFor";
       
   845 val MPair_parts = thm "MPair_parts";
       
   846 val parts_increasing = thm "parts_increasing";
       
   847 val parts_insertI = thm "parts_insertI";
       
   848 val parts_empty = thm "parts_empty";
       
   849 val parts_emptyE = thm "parts_emptyE";
       
   850 val parts_singleton = thm "parts_singleton";
       
   851 val parts_Un_subset1 = thm "parts_Un_subset1";
       
   852 val parts_Un_subset2 = thm "parts_Un_subset2";
       
   853 val parts_insert = thm "parts_insert";
       
   854 val parts_insert2 = thm "parts_insert2";
       
   855 val parts_UN_subset1 = thm "parts_UN_subset1";
       
   856 val parts_UN_subset2 = thm "parts_UN_subset2";
       
   857 val parts_UN = thm "parts_UN";
       
   858 val parts_insert_subset = thm "parts_insert_subset";
       
   859 val parts_partsD = thm "parts_partsD";
       
   860 val parts_trans = thm "parts_trans";
       
   861 val parts_cut = thm "parts_cut";
       
   862 val parts_cut_eq = thm "parts_cut_eq";
       
   863 val parts_insert_eq_I = thm "parts_insert_eq_I";
       
   864 val parts_image_Key = thm "parts_image_Key";
       
   865 val MPair_analz = thm "MPair_analz";
       
   866 val analz_increasing = thm "analz_increasing";
       
   867 val analz_subset_parts = thm "analz_subset_parts";
       
   868 val not_parts_not_analz = thm "not_parts_not_analz";
       
   869 val parts_analz = thm "parts_analz";
       
   870 val analz_parts = thm "analz_parts";
       
   871 val analz_insertI = thm "analz_insertI";
       
   872 val analz_empty = thm "analz_empty";
       
   873 val analz_Un = thm "analz_Un";
       
   874 val analz_insert_Crypt_subset = thm "analz_insert_Crypt_subset";
       
   875 val analz_image_Key = thm "analz_image_Key";
       
   876 val analz_analzD = thm "analz_analzD";
       
   877 val analz_trans = thm "analz_trans";
       
   878 val analz_cut = thm "analz_cut";
       
   879 val analz_insert_eq = thm "analz_insert_eq";
       
   880 val analz_subset_cong = thm "analz_subset_cong";
       
   881 val analz_cong = thm "analz_cong";
       
   882 val analz_insert_cong = thm "analz_insert_cong";
       
   883 val analz_trivial = thm "analz_trivial";
       
   884 val analz_UN_analz = thm "analz_UN_analz";
       
   885 val synth_mono = thm "synth_mono";
       
   886 val synth_increasing = thm "synth_increasing";
       
   887 val synth_Un = thm "synth_Un";
       
   888 val synth_insert = thm "synth_insert";
       
   889 val synth_synthD = thm "synth_synthD";
       
   890 val synth_trans = thm "synth_trans";
       
   891 val synth_cut = thm "synth_cut";
       
   892 val Agent_synth = thm "Agent_synth";
       
   893 val Number_synth = thm "Number_synth";
       
   894 val Nonce_synth_eq = thm "Nonce_synth_eq";
       
   895 val Key_synth_eq = thm "Key_synth_eq";
       
   896 val Crypt_synth_eq = thm "Crypt_synth_eq";
       
   897 val keysFor_synth = thm "keysFor_synth";
       
   898 val parts_synth = thm "parts_synth";
       
   899 val analz_analz_Un = thm "analz_analz_Un";
       
   900 val analz_synth_Un = thm "analz_synth_Un";
       
   901 val analz_synth = thm "analz_synth";
       
   902 val parts_insert_subset_Un = thm "parts_insert_subset_Un";
       
   903 val Fake_parts_insert = thm "Fake_parts_insert";
       
   904 val Fake_analz_insert = thm "Fake_analz_insert";
       
   905 val analz_conj_parts = thm "analz_conj_parts";
       
   906 val analz_disj_parts = thm "analz_disj_parts";
       
   907 val MPair_synth_analz = thm "MPair_synth_analz";
       
   908 val Crypt_synth_analz = thm "Crypt_synth_analz";
       
   909 val Hash_synth_analz = thm "Hash_synth_analz";
       
   910 val pushes = thms "pushes";
       
   911 
       
   912 
       
   913 (*Prove base case (subgoal i) and simplify others.  A typical base case
       
   914   concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
       
   915   alone.*)
       
   916 fun prove_simple_subgoals_tac i =
       
   917     force_tac (claset(), simpset() addsimps [image_eq_UN]) i THEN
       
   918     ALLGOALS Asm_simp_tac
       
   919 
       
   920 (*Analysis of Fake cases.  Also works for messages that forward unknown parts,
       
   921   but this application is no longer necessary if analz_insert_eq is used.
       
   922   Abstraction over i is ESSENTIAL: it delays the dereferencing of claset
       
   923   DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *)
       
   924 
       
   925 (*Apply rules to break down assumptions of the form
       
   926   Y \<in> parts(insert X H)  and  Y \<in> analz(insert X H)
       
   927 *)
       
   928 val Fake_insert_tac =
       
   929     dresolve_tac [impOfSubs Fake_analz_insert,
       
   930                   impOfSubs Fake_parts_insert] THEN'
       
   931     eresolve_tac [asm_rl, thm"synth.Inj"];
       
   932 
       
   933 fun Fake_insert_simp_tac ss i =
       
   934     REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i;
       
   935 
       
   936 fun atomic_spy_analz_tac (cs,ss) = SELECT_GOAL
       
   937     (Fake_insert_simp_tac ss 1
       
   938      THEN
       
   939      IF_UNSOLVED (Blast.depth_tac
       
   940 		  (cs addIs [analz_insertI,
       
   941 				   impOfSubs analz_subset_parts]) 4 1))
       
   942 
       
   943 (*The explicit claset and simpset arguments help it work with Isar*)
       
   944 fun gen_spy_analz_tac (cs,ss) i =
       
   945   DETERM
       
   946    (SELECT_GOAL
       
   947      (EVERY
       
   948       [  (*push in occurrences of X...*)
       
   949        (REPEAT o CHANGED)
       
   950            (res_inst_tac [("x1","X")] (insert_commute RS ssubst) 1),
       
   951        (*...allowing further simplifications*)
       
   952        simp_tac ss 1,
       
   953        REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
       
   954        DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
       
   955 
       
   956 fun spy_analz_tac i = gen_spy_analz_tac (claset(), simpset()) i
       
   957 *}
       
   958 
       
   959 (*By default only o_apply is built-in.  But in the presence of eta-expansion
       
   960   this means that some terms displayed as (f o g) will be rewritten, and others
       
   961   will not!*)
       
   962 declare o_def [simp]
       
   963 
       
   964 
       
   965 lemma Crypt_notin_image_Key [simp]: "Crypt K X \<notin> Key ` A"
       
   966 by auto
       
   967 
       
   968 lemma Hash_notin_image_Key [simp] :"Hash X \<notin> Key ` A"
       
   969 by auto
       
   970 
       
   971 lemma synth_analz_mono: "G<=H ==> synth (analz(G)) <= synth (analz(H))"
       
   972 by (simp add: synth_mono analz_mono)
       
   973 
       
   974 lemma Fake_analz_eq [simp]:
       
   975      "X \<in> synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)"
       
   976 apply (drule Fake_analz_insert[of _ _ "H"])
       
   977 apply (simp add: synth_increasing[THEN Un_absorb2])
       
   978 apply (drule synth_mono)
       
   979 apply (simp add: synth_idem)
       
   980 apply (blast intro: synth_analz_mono [THEN [2] rev_subsetD])
       
   981 done
       
   982 
       
   983 text{*Two generalizations of @{text analz_insert_eq}*}
       
   984 lemma gen_analz_insert_eq [rule_format]:
       
   985      "X \<in> analz H ==> ALL G. H \<subseteq> G --> analz (insert X G) = analz G";
       
   986 by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
       
   987 
       
   988 lemma synth_analz_insert_eq [rule_format]:
       
   989      "X \<in> synth (analz H)
       
   990       ==> ALL G. H \<subseteq> G --> (Key K \<in> analz (insert X G)) = (Key K \<in> analz G)";
       
   991 apply (erule synth.induct)
       
   992 apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI])
       
   993 done
       
   994 
       
   995 lemma Fake_parts_sing:
       
   996      "X \<in> synth (analz H) ==> parts{X} \<subseteq> synth (analz H) \<union> parts H";
       
   997 apply (rule subset_trans)
       
   998  apply (erule_tac [2] Fake_parts_insert)
       
   999 apply (simp add: parts_mono)
       
  1000 done
       
  1001 
       
  1002 lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
       
  1003 
       
  1004 method_setup spy_analz = {*
       
  1005     Method.ctxt_args (fn ctxt =>
       
  1006         Method.METHOD (fn facts =>
       
  1007             gen_spy_analz_tac (Classical.get_local_claset ctxt,
       
  1008                                Simplifier.get_local_simpset ctxt) 1))*}
       
  1009     "for proving the Fake case when analz is involved"
       
  1010 
       
  1011 method_setup atomic_spy_analz = {*
       
  1012     Method.ctxt_args (fn ctxt =>
       
  1013         Method.METHOD (fn facts =>
       
  1014             atomic_spy_analz_tac (Classical.get_local_claset ctxt,
       
  1015                                   Simplifier.get_local_simpset ctxt) 1))*}
       
  1016     "for debugging spy_analz"
       
  1017 
       
  1018 method_setup Fake_insert_simp = {*
       
  1019     Method.ctxt_args (fn ctxt =>
       
  1020         Method.METHOD (fn facts =>
       
  1021             Fake_insert_simp_tac (Simplifier.get_local_simpset ctxt) 1))*}
       
  1022     "for debugging spy_analz"
       
  1023 
       
  1024 
       
  1025 end