author | paulson <lp15@cam.ac.uk> |
Tue, 17 Sep 2019 12:36:04 +0100 | |
changeset 70721 | 47258727fa42 |
parent 69597 | ff784d5a5bfb |
child 76287 | cdc14f94c754 |
permissions | -rw-r--r-- |
41775 | 1 |
(* Title: HOL/Auth/Guard/Guard_Public.thy |
2 |
Author: Frederic Blanqui, University of Cambridge Computer Laboratory |
|
3 |
Copyright 2002 University of Cambridge |
|
13508 | 4 |
|
41775 | 5 |
Lemmas on guarded messages for public protocols. |
6 |
*) |
|
13508 | 7 |
|
48260 | 8 |
theory Guard_Public imports Guard "../Public" Extensions begin |
13508 | 9 |
|
61830 | 10 |
subsection\<open>Extensions to Theory \<open>Public\<close>\<close> |
13508 | 11 |
|
12 |
declare initState.simps [simp del] |
|
13 |
||
61830 | 14 |
subsubsection\<open>signature\<close> |
13508 | 15 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
20217
diff
changeset
|
16 |
definition sign :: "agent => msg => msg" where |
61956 | 17 |
"sign A X == \<lbrace>Agent A, X, Crypt (priK A) (Hash X)\<rbrace>" |
13508 | 18 |
|
19 |
lemma sign_inj [iff]: "(sign A X = sign A' X') = (A=A' & X=X')" |
|
20 |
by (auto simp: sign_def) |
|
21 |
||
61830 | 22 |
subsubsection\<open>agent associated to a key\<close> |
13508 | 23 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
20217
diff
changeset
|
24 |
definition agt :: "key => agent" where |
67613 | 25 |
"agt K == SOME A. K = priK A | K = pubK A" |
13508 | 26 |
|
27 |
lemma agt_priK [simp]: "agt (priK A) = A" |
|
28 |
by (simp add: agt_def) |
|
29 |
||
30 |
lemma agt_pubK [simp]: "agt (pubK A) = A" |
|
31 |
by (simp add: agt_def) |
|
32 |
||
69597 | 33 |
subsubsection\<open>basic facts about \<^term>\<open>initState\<close>\<close> |
13508 | 34 |
|
67613 | 35 |
lemma no_Crypt_in_parts_init [simp]: "Crypt K X \<notin> parts (initState A)" |
13508 | 36 |
by (cases A, auto simp: initState.simps) |
37 |
||
67613 | 38 |
lemma no_Crypt_in_analz_init [simp]: "Crypt K X \<notin> analz (initState A)" |
13508 | 39 |
by auto |
40 |
||
67613 | 41 |
lemma no_priK_in_analz_init [simp]: "A \<notin> bad |
42 |
\<Longrightarrow> Key (priK A) \<notin> analz (initState Spy)" |
|
13508 | 43 |
by (auto simp: initState.simps) |
44 |
||
67613 | 45 |
lemma priK_notin_initState_Friend [simp]: "A \<noteq> Friend C |
46 |
\<Longrightarrow> Key (priK A) \<notin> parts (initState (Friend C))" |
|
13508 | 47 |
by (auto simp: initState.simps) |
48 |
||
49 |
lemma keyset_init [iff]: "keyset (initState A)" |
|
50 |
by (cases A, auto simp: keyset_def initState.simps) |
|
51 |
||
61830 | 52 |
subsubsection\<open>sets of private keys\<close> |
13508 | 53 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
20217
diff
changeset
|
54 |
definition priK_set :: "key set => bool" where |
67613 | 55 |
"priK_set Ks \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> (\<exists>A. K = priK A)" |
13508 | 56 |
|
67613 | 57 |
lemma in_priK_set: "[| priK_set Ks; K \<in> Ks |] ==> \<exists>A. K = priK A" |
13508 | 58 |
by (simp add: priK_set_def) |
59 |
||
60 |
lemma priK_set1 [iff]: "priK_set {priK A}" |
|
61 |
by (simp add: priK_set_def) |
|
62 |
||
63 |
lemma priK_set2 [iff]: "priK_set {priK A, priK B}" |
|
64 |
by (simp add: priK_set_def) |
|
65 |
||
61830 | 66 |
subsubsection\<open>sets of good keys\<close> |
13508 | 67 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
20217
diff
changeset
|
68 |
definition good :: "key set => bool" where |
67613 | 69 |
"good Ks == \<forall>K. K \<in> Ks \<longrightarrow> agt K \<notin> bad" |
13508 | 70 |
|
67613 | 71 |
lemma in_good: "[| good Ks; K \<in> Ks |] ==> agt K \<notin> bad" |
13508 | 72 |
by (simp add: good_def) |
73 |
||
67613 | 74 |
lemma good1 [simp]: "A \<notin> bad \<Longrightarrow> good {priK A}" |
13508 | 75 |
by (simp add: good_def) |
76 |
||
67613 | 77 |
lemma good2 [simp]: "[| A \<notin> bad; B \<notin> bad |] ==> good {priK A, priK B}" |
13508 | 78 |
by (simp add: good_def) |
79 |
||
61830 | 80 |
subsubsection\<open>greatest nonce used in a trace, 0 if there is no nonce\<close> |
13508 | 81 |
|
35418 | 82 |
primrec greatest :: "event list => nat" |
83 |
where |
|
84 |
"greatest [] = 0" |
|
85 |
| "greatest (ev # evs) = max (greatest_msg (msg ev)) (greatest evs)" |
|
13508 | 86 |
|
67613 | 87 |
lemma greatest_is_greatest: "Nonce n \<in> used evs \<Longrightarrow> n \<le> greatest evs" |
13508 | 88 |
apply (induct evs, auto simp: initState.simps) |
89 |
apply (drule used_sub_parts_used, safe) |
|
90 |
apply (drule greatest_msg_is_greatest, arith) |
|
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
16417
diff
changeset
|
91 |
by simp |
13508 | 92 |
|
61830 | 93 |
subsubsection\<open>function giving a new nonce\<close> |
13508 | 94 |
|
67613 | 95 |
definition new :: "event list \<Rightarrow> nat" where |
96 |
"new evs \<equiv> Suc (greatest evs)" |
|
13508 | 97 |
|
67613 | 98 |
lemma new_isnt_used [iff]: "Nonce (new evs) \<notin> used evs" |
13508 | 99 |
by (clarify, drule greatest_is_greatest, auto simp: new_def) |
100 |
||
61830 | 101 |
subsection\<open>Proofs About Guarded Messages\<close> |
13508 | 102 |
|
61830 | 103 |
subsubsection\<open>small hack necessary because priK is defined as the inverse of pubK\<close> |
13508 | 104 |
|
105 |
lemma pubK_is_invKey_priK: "pubK A = invKey (priK A)" |
|
106 |
by simp |
|
107 |
||
108 |
lemmas pubK_is_invKey_priK_substI = pubK_is_invKey_priK [THEN ssubst] |
|
109 |
||
110 |
lemmas invKey_invKey_substI = invKey [THEN ssubst] |
|
111 |
||
67613 | 112 |
lemma "Nonce n \<in> parts {X} \<Longrightarrow> Crypt (pubK A) X \<in> guard n {priK A}" |
13508 | 113 |
apply (rule pubK_is_invKey_priK_substI, rule invKey_invKey_substI) |
114 |
by (rule Guard_Nonce, simp+) |
|
115 |
||
61830 | 116 |
subsubsection\<open>guardedness results\<close> |
13508 | 117 |
|
67613 | 118 |
lemma sign_guard [intro]: "X \<in> guard n Ks \<Longrightarrow> sign A X \<in> guard n Ks" |
13508 | 119 |
by (auto simp: sign_def) |
120 |
||
121 |
lemma Guard_init [iff]: "Guard n Ks (initState B)" |
|
122 |
by (induct B, auto simp: Guard_def initState.simps) |
|
123 |
||
124 |
lemma Guard_knows_max': "Guard n Ks (knows_max' C evs) |
|
125 |
==> Guard n Ks (knows_max C evs)" |
|
126 |
by (simp add: knows_max_def) |
|
127 |
||
67613 | 128 |
lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \<notin> used evs |
129 |
\<Longrightarrow> Guard n Ks (spies evs)" |
|
13508 | 130 |
by (auto simp: Guard_def dest: not_used_not_known parts_sub) |
131 |
||
67613 | 132 |
lemma Nonce_not_used_Guard [dest]: "[| evs \<in> p; Nonce n \<notin> used evs; |
13508 | 133 |
Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" |
134 |
by (auto simp: Guard_def dest: known_used parts_trans) |
|
135 |
||
67613 | 136 |
lemma Nonce_not_used_Guard_max [dest]: "[| evs \<in> p; Nonce n \<notin> used evs; |
13508 | 137 |
Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" |
138 |
by (auto simp: Guard_def dest: known_max_used parts_trans) |
|
139 |
||
67613 | 140 |
lemma Nonce_not_used_Guard_max' [dest]: "[| evs \<in> p; Nonce n \<notin> used evs; |
13508 | 141 |
Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" |
142 |
apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) |
|
143 |
by (auto simp: knows_max_def) |
|
144 |
||
61830 | 145 |
subsubsection\<open>regular protocols\<close> |
13508 | 146 |
|
67613 | 147 |
definition regular :: "event list set \<Rightarrow> bool" where |
148 |
"regular p \<equiv> \<forall>evs A. evs \<in> p \<longrightarrow> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)" |
|
13508 | 149 |
|
67613 | 150 |
lemma priK_parts_iff_bad [simp]: "[| evs \<in> p; regular p |] ==> |
151 |
(Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)" |
|
13508 | 152 |
by (auto simp: regular_def) |
153 |
||
67613 | 154 |
lemma priK_analz_iff_bad [simp]: "[| evs \<in> p; regular p |] ==> |
155 |
(Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)" |
|
13508 | 156 |
by auto |
157 |
||
67613 | 158 |
lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \<in> p; |
159 |
priK_set Ks; good Ks; regular p |] ==> Nonce n \<notin> analz (spies evs)" |
|
13508 | 160 |
apply (clarify, simp only: knows_decomp) |
161 |
apply (drule Guard_invKey_keyset, simp+, safe) |
|
162 |
apply (drule in_good, simp) |
|
163 |
apply (drule in_priK_set, simp+, clarify) |
|
164 |
apply (frule_tac A=A in priK_analz_iff_bad) |
|
165 |
by (simp add: knows_decomp)+ |
|
166 |
||
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
16417
diff
changeset
|
167 |
end |