author | wenzelm |
Mon, 13 Apr 2020 22:08:14 +0200 | |
changeset 71751 | abf3e80bd815 |
parent 69597 | ff784d5a5bfb |
child 76287 | cdc14f94c754 |
permissions | -rw-r--r-- |
41775 | 1 |
(* Title: HOL/Auth/Guard/Guard_Shared.thy |
2 |
Author: Frederic Blanqui, University of Cambridge Computer Laboratory |
|
3 |
Copyright 2002 University of Cambridge |
|
4 |
*) |
|
13508 | 5 |
|
61830 | 6 |
section\<open>lemmas on guarded messages for protocols with symmetric keys\<close> |
13508 | 7 |
|
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
35416
diff
changeset
|
8 |
theory Guard_Shared imports Guard GuardK "../Shared" begin |
13508 | 9 |
|
61830 | 10 |
subsection\<open>Extensions to Theory \<open>Shared\<close>\<close> |
13508 | 11 |
|
12 |
declare initState.simps [simp del] |
|
13 |
||
61830 | 14 |
subsubsection\<open>a little abbreviation\<close> |
13508 | 15 |
|
19363 | 16 |
abbreviation |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19363
diff
changeset
|
17 |
Ciph :: "agent => msg => msg" where |
19363 | 18 |
"Ciph A X == Crypt (shrK A) X" |
13508 | 19 |
|
61830 | 20 |
subsubsection\<open>agent associated to a key\<close> |
13508 | 21 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
21404
diff
changeset
|
22 |
definition agt :: "key => agent" where |
67613 | 23 |
"agt K == SOME A. K = shrK A" |
13508 | 24 |
|
25 |
lemma agt_shrK [simp]: "agt (shrK A) = A" |
|
26 |
by (simp add: agt_def) |
|
27 |
||
69597 | 28 |
subsubsection\<open>basic facts about \<^term>\<open>initState\<close>\<close> |
13508 | 29 |
|
67613 | 30 |
lemma no_Crypt_in_parts_init [simp]: "Crypt K X \<notin> parts (initState A)" |
13508 | 31 |
by (cases A, auto simp: initState.simps) |
32 |
||
67613 | 33 |
lemma no_Crypt_in_analz_init [simp]: "Crypt K X \<notin> analz (initState A)" |
13508 | 34 |
by auto |
35 |
||
67613 | 36 |
lemma no_shrK_in_analz_init [simp]: "A \<notin> bad |
37 |
\<Longrightarrow> Key (shrK A) \<notin> analz (initState Spy)" |
|
13508 | 38 |
by (auto simp: initState.simps) |
39 |
||
67613 | 40 |
lemma shrK_notin_initState_Friend [simp]: "A \<noteq> Friend C |
41 |
\<Longrightarrow> Key (shrK A) \<notin> parts (initState (Friend C))" |
|
13508 | 42 |
by (auto simp: initState.simps) |
43 |
||
44 |
lemma keyset_init [iff]: "keyset (initState A)" |
|
45 |
by (cases A, auto simp: keyset_def initState.simps) |
|
46 |
||
61830 | 47 |
subsubsection\<open>sets of symmetric keys\<close> |
13508 | 48 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
21404
diff
changeset
|
49 |
definition shrK_set :: "key set => bool" where |
67613 | 50 |
"shrK_set Ks \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> (\<exists>A. K = shrK A)" |
13508 | 51 |
|
67613 | 52 |
lemma in_shrK_set: "[| shrK_set Ks; K \<in> Ks |] ==> \<exists>A. K = shrK A" |
13508 | 53 |
by (simp add: shrK_set_def) |
54 |
||
55 |
lemma shrK_set1 [iff]: "shrK_set {shrK A}" |
|
56 |
by (simp add: shrK_set_def) |
|
57 |
||
58 |
lemma shrK_set2 [iff]: "shrK_set {shrK A, shrK B}" |
|
59 |
by (simp add: shrK_set_def) |
|
60 |
||
61830 | 61 |
subsubsection\<open>sets of good keys\<close> |
13508 | 62 |
|
67613 | 63 |
definition good :: "key set \<Rightarrow> bool" where |
64 |
"good Ks \<equiv> \<forall>K. K \<in> Ks \<longrightarrow> agt K \<notin> bad" |
|
13508 | 65 |
|
67613 | 66 |
lemma in_good: "[| good Ks; K \<in> Ks |] ==> agt K \<notin> bad" |
13508 | 67 |
by (simp add: good_def) |
68 |
||
67613 | 69 |
lemma good1 [simp]: "A \<notin> bad \<Longrightarrow> good {shrK A}" |
13508 | 70 |
by (simp add: good_def) |
71 |
||
67613 | 72 |
lemma good2 [simp]: "[| A \<notin> bad; B \<notin> bad |] ==> good {shrK A, shrK B}" |
13508 | 73 |
by (simp add: good_def) |
74 |
||
75 |
||
61830 | 76 |
subsection\<open>Proofs About Guarded Messages\<close> |
13508 | 77 |
|
61830 | 78 |
subsubsection\<open>small hack\<close> |
13508 | 79 |
|
80 |
lemma shrK_is_invKey_shrK: "shrK A = invKey (shrK A)" |
|
81 |
by simp |
|
82 |
||
83 |
lemmas shrK_is_invKey_shrK_substI = shrK_is_invKey_shrK [THEN ssubst] |
|
84 |
||
85 |
lemmas invKey_invKey_substI = invKey [THEN ssubst] |
|
86 |
||
67613 | 87 |
lemma "Nonce n \<in> parts {X} \<Longrightarrow> Crypt (shrK A) X \<in> guard n {shrK A}" |
13508 | 88 |
apply (rule shrK_is_invKey_shrK_substI, rule invKey_invKey_substI) |
89 |
by (rule Guard_Nonce, simp+) |
|
90 |
||
61830 | 91 |
subsubsection\<open>guardedness results on nonces\<close> |
13508 | 92 |
|
67613 | 93 |
lemma guard_ciph [simp]: "shrK A \<in> Ks \<Longrightarrow> Ciph A X \<in> guard n Ks" |
13508 | 94 |
by (rule Guard_Nonce, simp) |
95 |
||
67613 | 96 |
lemma guardK_ciph [simp]: "shrK A \<in> Ks \<Longrightarrow> Ciph A X \<in> guardK n Ks" |
13508 | 97 |
by (rule Guard_Key, simp) |
98 |
||
99 |
lemma Guard_init [iff]: "Guard n Ks (initState B)" |
|
100 |
by (induct B, auto simp: Guard_def initState.simps) |
|
101 |
||
102 |
lemma Guard_knows_max': "Guard n Ks (knows_max' C evs) |
|
103 |
==> Guard n Ks (knows_max C evs)" |
|
104 |
by (simp add: knows_max_def) |
|
105 |
||
67613 | 106 |
lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \<notin> used evs |
107 |
\<Longrightarrow> Guard n Ks (spies evs)" |
|
13508 | 108 |
by (auto simp: Guard_def dest: not_used_not_known parts_sub) |
109 |
||
67613 | 110 |
lemma Nonce_not_used_Guard [dest]: "[| evs \<in> p; Nonce n \<notin> used evs; |
13508 | 111 |
Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" |
112 |
by (auto simp: Guard_def dest: known_used parts_trans) |
|
113 |
||
67613 | 114 |
lemma Nonce_not_used_Guard_max [dest]: "[| evs \<in> p; Nonce n \<notin> used evs; |
13508 | 115 |
Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" |
116 |
by (auto simp: Guard_def dest: known_max_used parts_trans) |
|
117 |
||
67613 | 118 |
lemma Nonce_not_used_Guard_max' [dest]: "[| evs \<in> p; Nonce n \<notin> used evs; |
13508 | 119 |
Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" |
120 |
apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) |
|
121 |
by (auto simp: knows_max_def) |
|
122 |
||
61830 | 123 |
subsubsection\<open>guardedness results on keys\<close> |
13508 | 124 |
|
67613 | 125 |
lemma GuardK_init [simp]: "n \<notin> range shrK \<Longrightarrow> GuardK n Ks (initState B)" |
13508 | 126 |
by (induct B, auto simp: GuardK_def initState.simps) |
127 |
||
67613 | 128 |
lemma GuardK_knows_max': "[| GuardK n A (knows_max' C evs); n \<notin> range shrK |] |
13508 | 129 |
==> GuardK n A (knows_max C evs)" |
130 |
by (simp add: knows_max_def) |
|
131 |
||
67613 | 132 |
lemma Key_not_used_GuardK_spies [dest]: "Key n \<notin> used evs |
133 |
\<Longrightarrow> GuardK n A (spies evs)" |
|
13508 | 134 |
by (auto simp: GuardK_def dest: not_used_not_known parts_sub) |
135 |
||
67613 | 136 |
lemma Key_not_used_GuardK [dest]: "[| evs \<in> p; Key n \<notin> used evs; |
13508 | 137 |
Gets_correct p; one_step p |] ==> GuardK n A (knows (Friend C) evs)" |
138 |
by (auto simp: GuardK_def dest: known_used parts_trans) |
|
139 |
||
67613 | 140 |
lemma Key_not_used_GuardK_max [dest]: "[| evs \<in> p; Key n \<notin> used evs; |
13508 | 141 |
Gets_correct p; one_step p |] ==> GuardK n A (knows_max (Friend C) evs)" |
142 |
by (auto simp: GuardK_def dest: known_max_used parts_trans) |
|
143 |
||
67613 | 144 |
lemma Key_not_used_GuardK_max' [dest]: "[| evs \<in> p; Key n \<notin> used evs; |
13508 | 145 |
Gets_correct p; one_step p |] ==> GuardK n A (knows_max' (Friend C) evs)" |
146 |
apply (rule_tac H="knows_max (Friend C) evs" in GuardK_mono) |
|
147 |
by (auto simp: knows_max_def) |
|
148 |
||
61830 | 149 |
subsubsection\<open>regular protocols\<close> |
13508 | 150 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
21404
diff
changeset
|
151 |
definition regular :: "event list set => bool" where |
67613 | 152 |
"regular p \<equiv> \<forall>evs A. evs \<in> p \<longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" |
13508 | 153 |
|
67613 | 154 |
lemma shrK_parts_iff_bad [simp]: "[| evs \<in> p; regular p |] ==> |
155 |
(Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" |
|
13508 | 156 |
by (auto simp: regular_def) |
157 |
||
67613 | 158 |
lemma shrK_analz_iff_bad [simp]: "[| evs \<in> p; regular p |] ==> |
159 |
(Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" |
|
13508 | 160 |
by auto |
161 |
||
67613 | 162 |
lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \<in> p; |
163 |
shrK_set Ks; good Ks; regular p |] ==> Nonce n \<notin> analz (spies evs)" |
|
13508 | 164 |
apply (clarify, simp only: knows_decomp) |
165 |
apply (drule Guard_invKey_keyset, simp+, safe) |
|
166 |
apply (drule in_good, simp) |
|
167 |
apply (drule in_shrK_set, simp+, clarify) |
|
168 |
apply (frule_tac A=A in shrK_analz_iff_bad) |
|
169 |
by (simp add: knows_decomp)+ |
|
170 |
||
61928 | 171 |
lemma GuardK_Key_analz: |
172 |
assumes "GuardK n Ks (spies evs)" "evs \<in> p" "shrK_set Ks" |
|
173 |
"good Ks" "regular p" "n \<notin> range shrK" |
|
174 |
shows "Key n \<notin> analz (spies evs)" |
|
175 |
proof (rule ccontr) |
|
176 |
assume "\<not> Key n \<notin> analz (knows Spy evs)" |
|
177 |
then have *: "Key n \<in> analz (spies' evs \<union> initState Spy)" |
|
178 |
by (simp add: knows_decomp) |
|
179 |
from \<open>GuardK n Ks (spies evs)\<close> |
|
180 |
have "GuardK n Ks (spies' evs \<union> initState Spy)" |
|
181 |
by (simp add: knows_decomp) |
|
182 |
then have "GuardK n Ks (spies' evs)" |
|
183 |
and "finite (spies' evs)" "keyset (initState Spy)" |
|
184 |
by simp_all |
|
185 |
moreover have "Key n \<notin> initState Spy" |
|
186 |
using \<open>n \<notin> range shrK\<close> by (simp add: image_iff initState_Spy) |
|
187 |
ultimately obtain K |
|
188 |
where "K \<in> Ks" and **: "Key K \<in> analz (spies' evs \<union> initState Spy)" |
|
189 |
using * by (auto dest: GuardK_invKey_keyset) |
|
190 |
from \<open>K \<in> Ks\<close> and \<open>good Ks\<close> have "agt K \<notin> bad" |
|
191 |
by (auto dest: in_good) |
|
192 |
from \<open>K \<in> Ks\<close> \<open>shrK_set Ks\<close> obtain A |
|
193 |
where "K = shrK A" |
|
194 |
by (auto dest: in_shrK_set) |
|
195 |
then have "agt K \<in> bad" |
|
196 |
using ** \<open>evs \<in> p\<close> \<open>regular p\<close> shrK_analz_iff_bad [of evs p "agt K"] |
|
197 |
by (simp add: knows_decomp) |
|
198 |
with \<open>agt K \<notin> bad\<close> show False by simp |
|
199 |
qed |
|
13508 | 200 |
|
62390 | 201 |
end |