1(*  Title:      HOL/Metis_Examples/Message.thy
2    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
3    Author:     Jasmin Blanchette, TU Muenchen
4
5Metis example featuring message authentication.
6*)
7
8section \<open>Metis Example Featuring Message Authentication\<close>
9
10theory Message
11imports Main
12begin
13
14declare [[metis_new_skolem]]
15
16lemma strange_Un_eq [simp]: "A \<union> (B \<union> A) = B \<union> A"
17by (metis Un_commute Un_left_absorb)
18
19type_synonym key = nat
20
21consts
22  all_symmetric :: bool        \<comment> \<open>true if all keys are symmetric\<close>
23  invKey        :: "key=>key"  \<comment> \<open>inverse of a symmetric key\<close>
24
25specification (invKey)
26  invKey [simp]: "invKey (invKey K) = K"
27  invKey_symmetric: "all_symmetric --> invKey = id"
28by (metis id_apply)
29
30
31text\<open>The inverse of a symmetric key is itself; that of a public key
32      is the private key and vice versa\<close>
33
34definition symKeys :: "key set" where
35  "symKeys == {K. invKey K = K}"
36
37datatype  \<comment> \<open>We allow any number of friendly agents\<close>
38  agent = Server | Friend nat | Spy
39
40datatype
41     msg = Agent  agent     \<comment> \<open>Agent names\<close>
42         | Number nat       \<comment> \<open>Ordinary integers, timestamps, ...\<close>
43         | Nonce  nat       \<comment> \<open>Unguessable nonces\<close>
44         | Key    key       \<comment> \<open>Crypto keys\<close>
45         | Hash   msg       \<comment> \<open>Hashing\<close>
46         | MPair  msg msg   \<comment> \<open>Compound messages\<close>
47         | Crypt  key msg   \<comment> \<open>Encryption, public- or shared-key\<close>
48
49
50text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
51syntax
52  "_MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
53translations
54  "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
55  "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
56
57
58definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
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>"
61
62definition keysFor :: "msg set => key set" where
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}"
65
66
67subsubsection\<open>Inductive Definition of All Parts" of a Message\<close>
68
69inductive_set
70  parts :: "msg set => msg set"
71  for H :: "msg set"
72  where
73    Inj [intro]:               "X \<in> H ==> X \<in> parts H"
74  | Fst:         "\<lbrace>X,Y\<rbrace>   \<in> parts H ==> X \<in> parts H"
75  | Snd:         "\<lbrace>X,Y\<rbrace>   \<in> parts H ==> Y \<in> parts H"
76  | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
77
78lemma parts_mono: "G \<subseteq> H ==> parts(G) \<subseteq> parts(H)"
79apply auto
80apply (erule parts.induct)
81   apply (metis parts.Inj rev_subsetD)
82  apply (metis parts.Fst)
83 apply (metis parts.Snd)
84by (metis parts.Body)
85
86text\<open>Equations hold because constructors are injective.\<close>
87lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x\<in>A)"
88by (metis agent.inject image_iff)
89
90lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x \<in> A)"
91by (metis image_iff msg.inject(4))
92
93lemma Nonce_Key_image_eq [simp]: "Nonce x \<notin> Key`A"
94by (metis image_iff msg.distinct(23))
95
96
97subsubsection\<open>Inverse of keys\<close>
98
99lemma invKey_eq [simp]: "(invKey K = invKey K') = (K = K')"
100by (metis invKey)
101
102
103subsection\<open>keysFor operator\<close>
104
105lemma keysFor_empty [simp]: "keysFor {} = {}"
106by (unfold keysFor_def, blast)
107
108lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
109by (unfold keysFor_def, blast)
110
111lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
112by (unfold keysFor_def, blast)
113
114text\<open>Monotonicity\<close>
115lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
116by (unfold keysFor_def, blast)
117
118lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
119by (unfold keysFor_def, auto)
120
121lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
122by (unfold keysFor_def, auto)
123
124lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
125by (unfold keysFor_def, auto)
126
127lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
128by (unfold keysFor_def, auto)
129
130lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
131by (unfold keysFor_def, auto)
132
133lemma keysFor_insert_MPair [simp]: "keysFor (insert \<lbrace>X,Y\<rbrace> H) = keysFor H"
134by (unfold keysFor_def, auto)
135
136lemma keysFor_insert_Crypt [simp]:
137    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
138by (unfold keysFor_def, auto)
139
140lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
141by (unfold keysFor_def, auto)
142
143lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
144by (unfold keysFor_def, blast)
145
146
147subsection\<open>Inductive relation "parts"\<close>
148
149lemma MPair_parts:
150     "[| \<lbrace>X,Y\<rbrace> \<in> parts H;
151         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
152by (blast dest: parts.Fst parts.Snd)
153
154declare MPair_parts [elim!] parts.Body [dest!]
155text\<open>NB These two rules are UNSAFE in the formal sense, as they discard the
156     compound message.  They work well on THIS FILE.
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.\<close>
159
160lemma parts_increasing: "H \<subseteq> parts(H)"
161by blast
162
163lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD]
164
165lemma parts_empty [simp]: "parts{} = {}"
166apply safe
167apply (erule parts.induct)
168apply blast+
169done
170
171lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
172by simp
173
174text\<open>WARNING: loops if H = {Y}, therefore must not be repeated!\<close>
175lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
176apply (erule parts.induct)
177apply fast+
178done
179
180
181subsubsection\<open>Unions\<close>
182
183lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
184by (intro Un_least parts_mono Un_upper1 Un_upper2)
185
186lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
187apply (rule subsetI)
188apply (erule parts.induct, blast+)
189done
190
191lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
192by (intro equalityI parts_Un_subset1 parts_Un_subset2)
193
194lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
195apply (subst insert_is_Un [of _ H])
196apply (simp only: parts_Un)
197done
198
199lemma parts_insert2:
200     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
201by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un)
202
203
204lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
205by (intro UN_least parts_mono UN_upper)
206
207lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
208apply (rule subsetI)
209apply (erule parts.induct, blast+)
210done
211
212lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
213by (intro equalityI parts_UN_subset1 parts_UN_subset2)
214
215text\<open>Added to simplify arguments to parts, analz and synth.
216  NOTE: the UN versions are no longer used!\<close>
217
218
219text\<open>This allows \<open>blast\<close> to simplify occurrences of
220  \<^term>\<open>parts(G\<union>H)\<close> in the assumption.\<close>
221lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE]
222declare in_parts_UnE [elim!]
223
224lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
225by (blast intro: parts_mono [THEN [2] rev_subsetD])
226
227subsubsection\<open>Idempotence and transitivity\<close>
228
229lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
230by (erule parts.induct, blast+)
231
232lemma parts_idem [simp]: "parts (parts H) = parts H"
233by blast
234
235lemma parts_subset_iff [simp]: "(parts G \<subseteq> parts H) = (G \<subseteq> parts H)"
236apply (rule iffI)
237apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing)
238apply (metis parts_idem parts_mono)
239done
240
241lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
242by (blast dest: parts_mono)
243
244lemma parts_cut: "[|Y\<in> parts (insert X G);  X\<in> parts H|] ==> Y\<in> parts(G \<union> H)"
245by (metis (no_types) Un_insert_left Un_insert_right insert_absorb le_supE
246          parts_Un parts_idem parts_increasing parts_trans)
247
248subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
249
250lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
251
252
253lemma parts_insert_Agent [simp]:
254     "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
255apply (rule parts_insert_eq_I)
256apply (erule parts.induct, auto)
257done
258
259lemma parts_insert_Nonce [simp]:
260     "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
261apply (rule parts_insert_eq_I)
262apply (erule parts.induct, auto)
263done
264
265lemma parts_insert_Number [simp]:
266     "parts (insert (Number N) H) = insert (Number N) (parts H)"
267apply (rule parts_insert_eq_I)
268apply (erule parts.induct, auto)
269done
270
271lemma parts_insert_Key [simp]:
272     "parts (insert (Key K) H) = insert (Key K) (parts H)"
273apply (rule parts_insert_eq_I)
274apply (erule parts.induct, auto)
275done
276
277lemma parts_insert_Hash [simp]:
278     "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
279apply (rule parts_insert_eq_I)
280apply (erule parts.induct, auto)
281done
282
283lemma parts_insert_Crypt [simp]:
284     "parts (insert (Crypt K X) H) =
285          insert (Crypt K X) (parts (insert X H))"
286apply (rule equalityI)
287apply (rule subsetI)
288apply (erule parts.induct, auto)
289apply (blast intro: parts.Body)
290done
291
292lemma parts_insert_MPair [simp]:
293     "parts (insert \<lbrace>X,Y\<rbrace> H) =
294          insert \<lbrace>X,Y\<rbrace> (parts (insert X (insert Y H)))"
295apply (rule equalityI)
296apply (rule subsetI)
297apply (erule parts.induct, auto)
298apply (blast intro: parts.Fst parts.Snd)+
299done
300
301lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
302apply auto
303apply (erule parts.induct, auto)
304done
305
306lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
307apply (induct_tac "msg")
308apply (simp_all add: parts_insert2)
309apply (metis Suc_n_not_le_n)
310apply (metis le_trans linorder_linear)
311done
312
313subsection\<open>Inductive relation "analz"\<close>
314
315text\<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
317    be taken apart; messages decrypted with known keys.\<close>
318
319inductive_set
320  analz :: "msg set => msg set"
321  for H :: "msg set"
322  where
323    Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
324  | Fst:     "\<lbrace>X,Y\<rbrace> \<in> analz H ==> X \<in> analz H"
325  | Snd:     "\<lbrace>X,Y\<rbrace> \<in> analz H ==> Y \<in> analz H"
326  | Decrypt [dest]:
327             "[|Crypt K X \<in> analz H; Key(invKey K) \<in> analz H|] ==> X \<in> analz H"
328
329
330text\<open>Monotonicity; Lemma 1 of Lowe's paper\<close>
331lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
332apply auto
333apply (erule analz.induct)
334apply (auto dest: analz.Fst analz.Snd)
335done
336
337text\<open>Making it safe speeds up proofs\<close>
338lemma MPair_analz [elim!]:
339     "[| \<lbrace>X,Y\<rbrace> \<in> analz H;
340             [| X \<in> analz H; Y \<in> analz H |] ==> P
341          |] ==> P"
342by (blast dest: analz.Fst analz.Snd)
343
344lemma analz_increasing: "H \<subseteq> analz(H)"
345by blast
346
347lemma analz_subset_parts: "analz H \<subseteq> parts H"
348apply (rule subsetI)
349apply (erule analz.induct, blast+)
350done
351
352lemmas analz_into_parts = analz_subset_parts [THEN subsetD]
353
354lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD]
355
356lemma parts_analz [simp]: "parts (analz H) = parts H"
357apply (rule equalityI)
358apply (metis analz_subset_parts parts_subset_iff)
359apply (metis analz_increasing parts_mono)
360done
361
362
363lemma analz_parts [simp]: "analz (parts H) = parts H"
364apply auto
365apply (erule analz.induct, auto)
366done
367
368lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
369
370subsubsection\<open>General equational properties\<close>
371
372lemma analz_empty [simp]: "analz{} = {}"
373apply safe
374apply (erule analz.induct, blast+)
375done
376
377text\<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\<close>
379lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
380by (intro Un_least analz_mono Un_upper1 Un_upper2)
381
382lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
383by (blast intro: analz_mono [THEN [2] rev_subsetD])
384
385subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
386
387lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
388
389lemma analz_insert_Agent [simp]:
390     "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
391apply (rule analz_insert_eq_I)
392apply (erule analz.induct, auto)
393done
394
395lemma analz_insert_Nonce [simp]:
396     "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
397apply (rule analz_insert_eq_I)
398apply (erule analz.induct, auto)
399done
400
401lemma analz_insert_Number [simp]:
402     "analz (insert (Number N) H) = insert (Number N) (analz H)"
403apply (rule analz_insert_eq_I)
404apply (erule analz.induct, auto)
405done
406
407lemma analz_insert_Hash [simp]:
408     "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
409apply (rule analz_insert_eq_I)
410apply (erule analz.induct, auto)
411done
412
413text\<open>Can only pull out Keys if they are not needed to decrypt the rest\<close>
414lemma analz_insert_Key [simp]:
415    "K \<notin> keysFor (analz H) ==>
416          analz (insert (Key K) H) = insert (Key K) (analz H)"
417apply (unfold keysFor_def)
418apply (rule analz_insert_eq_I)
419apply (erule analz.induct, auto)
420done
421
422lemma analz_insert_MPair [simp]:
423     "analz (insert \<lbrace>X,Y\<rbrace> H) =
424          insert \<lbrace>X,Y\<rbrace> (analz (insert X (insert Y H)))"
425apply (rule equalityI)
426apply (rule subsetI)
427apply (erule analz.induct, auto)
428apply (erule analz.induct)
429apply (blast intro: analz.Fst analz.Snd)+
430done
431
432text\<open>Can pull out enCrypted message if the Key is not known\<close>
433lemma analz_insert_Crypt:
434     "Key (invKey K) \<notin> analz H
435      ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
436apply (rule analz_insert_eq_I)
437apply (erule analz.induct, auto)
438
439done
440
441lemma lemma1: "Key (invKey K) \<in> analz H ==>
442               analz (insert (Crypt K X) H) \<subseteq>
443               insert (Crypt K X) (analz (insert X H))"
444apply (rule subsetI)
445apply (erule_tac x = x in analz.induct, auto)
446done
447
448lemma lemma2: "Key (invKey K) \<in> analz H ==>
449               insert (Crypt K X) (analz (insert X H)) \<subseteq>
450               analz (insert (Crypt K X) H)"
451apply auto
452apply (erule_tac x = x in analz.induct, auto)
453apply (blast intro: analz_insertI analz.Decrypt)
454done
455
456lemma analz_insert_Decrypt:
457     "Key (invKey K) \<in> analz H ==>
458               analz (insert (Crypt K X) H) =
459               insert (Crypt K X) (analz (insert X H))"
460by (intro equalityI lemma1 lemma2)
461
462text\<open>Case analysis: either the message is secure, or it is not! Effective,
463but can cause subgoals to blow up! Use with \<open>if_split\<close>; apparently
464\<open>split_tac\<close> does not cope with patterns such as \<^term>\<open>analz (insert
465(Crypt K X) H)\<close>\<close>
466lemma analz_Crypt_if [simp]:
467     "analz (insert (Crypt K X) H) =
468          (if (Key (invKey K) \<in> analz H)
469           then insert (Crypt K X) (analz (insert X H))
470           else insert (Crypt K X) (analz H))"
471by (simp add: analz_insert_Crypt analz_insert_Decrypt)
472
473
474text\<open>This rule supposes "for the sake of argument" that we have the key.\<close>
475lemma analz_insert_Crypt_subset:
476     "analz (insert (Crypt K X) H) \<subseteq>
477           insert (Crypt K X) (analz (insert X H))"
478apply (rule subsetI)
479apply (erule analz.induct, auto)
480done
481
482
483lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
484apply auto
485apply (erule analz.induct, auto)
486done
487
488
489subsubsection\<open>Idempotence and transitivity\<close>
490
491lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
492by (erule analz.induct, blast+)
493
494lemma analz_idem [simp]: "analz (analz H) = analz H"
495by blast
496
497lemma analz_subset_iff [simp]: "(analz G \<subseteq> analz H) = (G \<subseteq> analz H)"
498apply (rule iffI)
499apply (iprover intro: subset_trans analz_increasing)
500apply (frule analz_mono, simp)
501done
502
503lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
504by (drule analz_mono, blast)
505
506
507declare analz_trans[intro]
508
509lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
510by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset)
511
512text\<open>This rewrite rule helps in the simplification of messages that involve
513  the forwarding of unknown components (X).  Without it, removing occurrences
514  of X can be very complicated.\<close>
515lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
516by (blast intro: analz_cut analz_insertI)
517
518
519text\<open>A congruence rule for "analz"\<close>
520
521lemma analz_subset_cong:
522     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |]
523      ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
524apply simp
525apply (metis Un_absorb2 Un_commute Un_subset_iff Un_upper1 Un_upper2 analz_mono)
526done
527
528
529lemma analz_cong:
530     "[| analz G = analz G'; analz H = analz H'
531               |] ==> analz (G \<union> H) = analz (G' \<union> H')"
532by (intro equalityI analz_subset_cong, simp_all)
533
534lemma analz_insert_cong:
535     "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
536by (force simp only: insert_def intro!: analz_cong)
537
538text\<open>If there are no pairs or encryptions then analz does nothing\<close>
539lemma analz_trivial:
540     "[| \<forall>X Y. \<lbrace>X,Y\<rbrace> \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
541apply safe
542apply (erule analz.induct, blast+)
543done
544
545text\<open>These two are obsolete (with a single Spy) but cost little to prove...\<close>
546lemma analz_UN_analz_lemma:
547     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
548apply (erule analz.induct)
549apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
550done
551
552lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
553by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
554
555
556subsection\<open>Inductive relation "synth"\<close>
557
558text\<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
560    encrypted with known keys.  Agent names are public domain.
561    Numbers can be guessed, but Nonces cannot be.\<close>
562
563inductive_set
564  synth :: "msg set => msg set"
565  for H :: "msg set"
566  where
567    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
568  | Agent  [intro]:   "Agent agt \<in> synth H"
569  | Number [intro]:   "Number n  \<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"
572  | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
573
574text\<open>Monotonicity\<close>
575lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
576  by (auto, erule synth.induct, auto)
577
578text\<open>NO \<open>Agent_synth\<close>, as any Agent name can be synthesized.
579  The same holds for \<^term>\<open>Number\<close>\<close>
580inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
581inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
582inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
583inductive_cases MPair_synth [elim!]: "\<lbrace>X,Y\<rbrace> \<in> synth H"
584inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
585
586
587lemma synth_increasing: "H \<subseteq> synth(H)"
588by blast
589
590subsubsection\<open>Unions\<close>
591
592text\<open>Converse fails: we can synth more from the union than from the
593  separate parts, building a compound message using elements of each.\<close>
594lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
595by (intro Un_least synth_mono Un_upper1 Un_upper2)
596
597lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
598by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono)
599
600subsubsection\<open>Idempotence and transitivity\<close>
601
602lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
603by (erule synth.induct, blast+)
604
605lemma synth_idem: "synth (synth H) = synth H"
606by blast
607
608lemma synth_subset_iff [simp]: "(synth G \<subseteq> synth H) = (G \<subseteq> synth H)"
609apply (rule iffI)
610apply (iprover intro: subset_trans synth_increasing)
611apply (frule synth_mono, simp add: synth_idem)
612done
613
614lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
615by (drule synth_mono, blast)
616
617lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
618by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono)
619
620lemma Agent_synth [simp]: "Agent A \<in> synth H"
621by blast
622
623lemma Number_synth [simp]: "Number n \<in> synth H"
624by blast
625
626lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
627by blast
628
629lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
630by blast
631
632lemma Crypt_synth_eq [simp]:
633     "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
634by blast
635
636
637lemma keysFor_synth [simp]:
638    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
639by (unfold keysFor_def, blast)
640
641
642subsubsection\<open>Combinations of parts, analz and synth\<close>
643
644lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
645apply (rule equalityI)
646apply (rule subsetI)
647apply (erule parts.induct)
648apply (metis UnCI)
649apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Fst parts_increasing)
650apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Snd parts_increasing)
651apply (metis Body Crypt_synth UnCI UnE insert_absorb insert_subset parts_increasing)
652apply (metis Un_subset_iff parts_increasing parts_mono synth_increasing)
653done
654
655lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
656apply (rule equalityI)
657apply (metis analz_idem analz_subset_cong order_eq_refl)
658apply (metis analz_increasing analz_subset_cong order_eq_refl)
659done
660
661declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro]
662
663lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
664apply (rule equalityI)
665apply (rule subsetI)
666apply (erule analz.induct)
667apply (metis UnCI UnE Un_commute analz.Inj)
668apply (metis MPair_synth UnCI UnE Un_commute analz.Fst analz.Inj)
669apply (metis MPair_synth UnCI UnE Un_commute analz.Inj analz.Snd)
670apply (blast intro: analz.Decrypt)
671apply blast
672done
673
674lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
675proof -
676  have "\<forall>x\<^sub>2 x\<^sub>1. synth x\<^sub>1 \<union> analz (x\<^sub>1 \<union> x\<^sub>2) = analz (synth x\<^sub>1 \<union> x\<^sub>2)" by (metis Un_commute analz_synth_Un)
677  hence "\<forall>x\<^sub>1. synth x\<^sub>1 \<union> analz x\<^sub>1 = analz (synth x\<^sub>1 \<union> {})" by (metis Un_empty_right)
678  hence "\<forall>x\<^sub>1. synth x\<^sub>1 \<union> analz x\<^sub>1 = analz (synth x\<^sub>1)" by (metis Un_empty_right)
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
681qed
682
683
684subsubsection\<open>For reasoning about the Fake rule in traces\<close>
685
686lemma parts_insert_subset_Un: "X \<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
687proof -
688  assume "X \<in> G"
689  hence "\<forall>x\<^sub>1. G \<subseteq> x\<^sub>1 \<longrightarrow> X \<in> x\<^sub>1 " by auto
690  hence "\<forall>x\<^sub>1. X \<in> G \<union> x\<^sub>1" by (metis Un_upper1)
691  hence "insert X H \<subseteq> G \<union> H" by (metis Un_upper2 insert_subset)
692  hence "parts (insert X H) \<subseteq> parts (G \<union> H)" by (metis parts_mono)
693  thus "parts (insert X H) \<subseteq> parts G \<union> parts H" by (metis parts_Un)
694qed
695
696lemma Fake_parts_insert:
697     "X \<in> synth (analz H) ==>
698      parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
699proof -
700  assume A1: "X \<in> synth (analz H)"
701  have F1: "\<forall>x\<^sub>1. analz x\<^sub>1 \<union> synth (analz x\<^sub>1) = analz (synth (analz x\<^sub>1))"
702    by (metis analz_idem analz_synth)
703  have F2: "\<forall>x\<^sub>1. parts x\<^sub>1 \<union> synth (analz x\<^sub>1) = parts (synth (analz x\<^sub>1))"
704    by (metis parts_analz parts_synth)
705  have F3: "X \<in> synth (analz H)" using A1 by metis
706  have "\<forall>x\<^sub>2 x\<^sub>1::msg set. x\<^sub>1 \<le> sup x\<^sub>1 x\<^sub>2" by (metis inf_sup_ord(3))
707  hence F4: "\<forall>x\<^sub>1. analz x\<^sub>1 \<subseteq> analz (synth x\<^sub>1)" by (metis analz_synth)
708  have F5: "X \<in> synth (analz H)" using F3 by metis
709  have "\<forall>x\<^sub>1. analz x\<^sub>1 \<subseteq> synth (analz x\<^sub>1)
710         \<longrightarrow> analz (synth (analz x\<^sub>1)) = synth (analz x\<^sub>1)"
711    using F1 by (metis subset_Un_eq)
712  hence F6: "\<forall>x\<^sub>1. analz (synth (analz x\<^sub>1)) = synth (analz x\<^sub>1)"
713    by (metis synth_increasing)
714  have "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> analz (synth x\<^sub>1)" using F4 by (metis analz_subset_iff)
715  hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> analz (synth (analz x\<^sub>1))" by (metis analz_subset_iff)
716  hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> synth (analz x\<^sub>1)" using F6 by metis
717  hence "H \<subseteq> synth (analz H)" by metis
718  hence "H \<subseteq> synth (analz H) \<and> X \<in> synth (analz H)" using F5 by metis
719  hence "insert X H \<subseteq> synth (analz H)" by (metis insert_subset)
720  hence "parts (insert X H) \<subseteq> parts (synth (analz H))" by (metis parts_mono)
721  hence "parts (insert X H) \<subseteq> parts H \<union> synth (analz H)" using F2 by metis
722  thus "parts (insert X H) \<subseteq> synth (analz H) \<union> parts H" by (metis Un_commute)
723qed
724
725lemma Fake_parts_insert_in_Un:
726     "[|Z \<in> parts (insert X H);  X \<in> synth (analz H)|]
727      ==> Z \<in>  synth (analz H) \<union> parts H"
728by (blast dest: Fake_parts_insert [THEN subsetD, dest])
729
730declare synth_mono [intro]
731
732lemma Fake_analz_insert:
733     "X \<in> synth (analz G) ==>
734      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
735by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un
736          analz_mono analz_synth_Un insert_absorb)
737
738lemma Fake_analz_insert_simpler:
739     "X \<in> synth (analz G) ==>
740      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
741apply (rule subsetI)
742apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
743apply (metis Un_commute analz_analz_Un analz_synth_Un)
744by (metis Un_upper1 Un_upper2 analz_mono insert_absorb insert_subset)
745
746end
747