author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 76288 | b82ac7ef65ec |
permissions | -rw-r--r-- |
41775 | 1 |
(* Title: HOL/Auth/Guard/P1.thy |
2 |
Author: Frederic Blanqui, University of Cambridge Computer Laboratory |
|
3 |
Copyright 2002 University of Cambridge |
|
13508 | 4 |
|
41775 | 5 |
From G. Karjoth, N. Asokan and C. Gulcu |
6 |
"Protecting the computation results of free-roaming agents" |
|
7 |
Mobiles Agents 1998, LNCS 1477. |
|
8 |
*) |
|
13508 | 9 |
|
61830 | 10 |
section\<open>Protocol P1\<close> |
13508 | 11 |
|
39216 | 12 |
theory P1 imports "../Public" Guard_Public List_Msg begin |
13508 | 13 |
|
61830 | 14 |
subsection\<open>Protocol Definition\<close> |
13508 | 15 |
|
16 |
(****************************************************************************** |
|
17 |
||
18 |
the contents of the messages are not completely specified in the paper |
|
19 |
we assume that the user sends his request and his itinerary in the clear |
|
20 |
||
61956 | 21 |
we will adopt the following format for messages: \<lbrace>A,r,I,L\<rbrace> |
13508 | 22 |
A: originator (agent) |
23 |
r: request (number) |
|
24 |
I: next shops (agent list) |
|
25 |
L: collected offers (offer list) |
|
26 |
||
27 |
in the paper, the authors use nonces r_i to add redundancy in the offer |
|
28 |
in order to make it safer against dictionary attacks |
|
29 |
it is not necessary in our modelization since crypto is assumed to be strong |
|
30 |
(Crypt in injective) |
|
31 |
******************************************************************************) |
|
32 |
||
61830 | 33 |
subsubsection\<open>offer chaining: |
34 |
B chains his offer for A with the head offer of L for sending it to C\<close> |
|
13508 | 35 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
36 |
definition chain :: "agent => nat => agent => msg => agent => msg" where |
13508 | 37 |
"chain B ofr A L C == |
38 |
let m1= Crypt (pubK A) (Nonce ofr) in |
|
61956 | 39 |
let m2= Hash \<lbrace>head L, Agent C\<rbrace> in |
40 |
sign B \<lbrace>m1,m2\<rbrace>" |
|
13508 | 41 |
|
42 |
declare Let_def [simp] |
|
43 |
||
44 |
lemma chain_inj [iff]: "(chain B ofr A L C = chain B' ofr' A' L' C') |
|
45 |
= (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')" |
|
46 |
by (auto simp: chain_def Let_def) |
|
47 |
||
67613 | 48 |
lemma Nonce_in_chain [iff]: "Nonce ofr \<in> parts {chain B ofr A L C}" |
13508 | 49 |
by (auto simp: chain_def sign_def) |
50 |
||
61830 | 51 |
subsubsection\<open>agent whose key is used to sign an offer\<close> |
13508 | 52 |
|
35418 | 53 |
fun shop :: "msg => msg" where |
61956 | 54 |
"shop \<lbrace>B,X,Crypt K H\<rbrace> = Agent (agt K)" |
13508 | 55 |
|
56 |
lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B" |
|
57 |
by (simp add: chain_def sign_def) |
|
58 |
||
61830 | 59 |
subsubsection\<open>nonce used in an offer\<close> |
13508 | 60 |
|
35418 | 61 |
fun nonce :: "msg => msg" where |
61956 | 62 |
"nonce \<lbrace>B,\<lbrace>Crypt K ofr,m2\<rbrace>,CryptH\<rbrace> = ofr" |
13508 | 63 |
|
64 |
lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr" |
|
65 |
by (simp add: chain_def sign_def) |
|
66 |
||
61830 | 67 |
subsubsection\<open>next shop\<close> |
13508 | 68 |
|
35418 | 69 |
fun next_shop :: "msg => agent" where |
61956 | 70 |
"next_shop \<lbrace>B,\<lbrace>m1,Hash\<lbrace>headL,Agent C\<rbrace>\<rbrace>,CryptH\<rbrace> = C" |
13508 | 71 |
|
72 |
lemma next_shop_chain [iff]: "next_shop (chain B ofr A L C) = C" |
|
73 |
by (simp add: chain_def sign_def) |
|
74 |
||
61830 | 75 |
subsubsection\<open>anchor of the offer list\<close> |
13508 | 76 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
77 |
definition anchor :: "agent => nat => agent => msg" where |
13508 | 78 |
"anchor A n B == chain A n A (cons nil nil) B" |
79 |
||
80 |
lemma anchor_inj [iff]: "(anchor A n B = anchor A' n' B') |
|
81 |
= (A=A' & n=n' & B=B')" |
|
82 |
by (auto simp: anchor_def) |
|
83 |
||
67613 | 84 |
lemma Nonce_in_anchor [iff]: "Nonce n \<in> parts {anchor A n B}" |
13508 | 85 |
by (auto simp: anchor_def) |
86 |
||
87 |
lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A" |
|
88 |
by (simp add: anchor_def) |
|
89 |
||
90 |
lemma nonce_anchor [simp]: "nonce (anchor A n B) = Nonce n" |
|
91 |
by (simp add: anchor_def) |
|
92 |
||
93 |
lemma next_shop_anchor [iff]: "next_shop (anchor A n B) = B" |
|
94 |
by (simp add: anchor_def) |
|
95 |
||
61830 | 96 |
subsubsection\<open>request event\<close> |
13508 | 97 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
98 |
definition reqm :: "agent => nat => nat => msg => agent => msg" where |
61956 | 99 |
"reqm A r n I B == \<lbrace>Agent A, Number r, cons (Agent A) (cons (Agent B) I), |
100 |
cons (anchor A n B) nil\<rbrace>" |
|
13508 | 101 |
|
102 |
lemma reqm_inj [iff]: "(reqm A r n I B = reqm A' r' n' I' B') |
|
103 |
= (A=A' & r=r' & n=n' & I=I' & B=B')" |
|
104 |
by (auto simp: reqm_def) |
|
105 |
||
67613 | 106 |
lemma Nonce_in_reqm [iff]: "Nonce n \<in> parts {reqm A r n I B}" |
13508 | 107 |
by (auto simp: reqm_def) |
108 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
109 |
definition req :: "agent => nat => nat => msg => agent => event" where |
13508 | 110 |
"req A r n I B == Says A B (reqm A r n I B)" |
111 |
||
112 |
lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B') |
|
113 |
= (A=A' & r=r' & n=n' & I=I' & B=B')" |
|
114 |
by (auto simp: req_def) |
|
115 |
||
61830 | 116 |
subsubsection\<open>propose event\<close> |
13508 | 117 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
118 |
definition prom :: "agent => nat => agent => nat => msg => msg => |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
119 |
msg => agent => msg" where |
61956 | 120 |
"prom B ofr A r I L J C == \<lbrace>Agent A, Number r, |
121 |
app (J, del (Agent B, I)), cons (chain B ofr A L C) L\<rbrace>" |
|
13508 | 122 |
|
123 |
lemma prom_inj [dest]: "prom B ofr A r I L J C |
|
124 |
= prom B' ofr' A' r' I' L' J' C' |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
125 |
\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" |
13508 | 126 |
by (auto simp: prom_def) |
127 |
||
67613 | 128 |
lemma Nonce_in_prom [iff]: "Nonce ofr \<in> parts {prom B ofr A r I L J C}" |
13508 | 129 |
by (auto simp: prom_def) |
130 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
131 |
definition pro :: "agent => nat => agent => nat => msg => msg => |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
132 |
msg => agent => event" where |
13508 | 133 |
"pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" |
134 |
||
135 |
lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
136 |
\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" |
13508 | 137 |
by (auto simp: pro_def dest: prom_inj) |
138 |
||
61830 | 139 |
subsubsection\<open>protocol\<close> |
13508 | 140 |
|
23746 | 141 |
inductive_set p1 :: "event list set" |
142 |
where |
|
13508 | 143 |
|
67613 | 144 |
Nil: "[] \<in> p1" |
13508 | 145 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
146 |
| Fake: "\<lbrakk>evsf \<in> p1; X \<in> synth (analz (spies evsf))\<rbrakk> \<Longrightarrow> Says Spy B X # evsf \<in> p1" |
13508 | 147 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
148 |
| Request: "\<lbrakk>evsr \<in> p1; Nonce n \<notin> used evsr; I \<in> agl\<rbrakk> \<Longrightarrow> req A r n I B # evsr \<in> p1" |
13508 | 149 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
150 |
| Propose: "\<lbrakk>evsp \<in> p1; Says A' B \<lbrace>Agent A,Number r,I,cons M L\<rbrace> \<in> set evsp; |
67613 | 151 |
I \<in> agl; J \<in> agl; isin (Agent C, app (J, del (Agent B, I))); |
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
152 |
Nonce ofr \<notin> used evsp\<rbrakk> \<Longrightarrow> pro B ofr A r I (cons M L) J C # evsp \<in> p1" |
13508 | 153 |
|
61830 | 154 |
subsubsection\<open>Composition of Traces\<close> |
13508 | 155 |
|
67613 | 156 |
lemma "evs' \<in> p1 \<Longrightarrow> |
157 |
evs \<in> p1 \<and> (\<forall>n. Nonce n \<in> used evs' \<longrightarrow> Nonce n \<notin> used evs) \<longrightarrow> |
|
158 |
evs' @ evs \<in> p1" |
|
13508 | 159 |
apply (erule p1.induct, safe) |
160 |
apply (simp_all add: used_ConsI) |
|
161 |
apply (erule p1.Fake, erule synth_sub, rule analz_mono, rule knows_sub_app) |
|
162 |
apply (erule p1.Request, safe, simp_all add: req_def, force) |
|
163 |
apply (erule_tac A'=A' in p1.Propose, simp_all) |
|
164 |
apply (drule_tac x=ofr in spec, simp add: pro_def, blast) |
|
165 |
apply (erule_tac A'=A' in p1.Propose, auto simp: pro_def) |
|
166 |
done |
|
167 |
||
61830 | 168 |
subsubsection\<open>Valid Offer Lists\<close> |
13508 | 169 |
|
23746 | 170 |
inductive_set |
67613 | 171 |
valid :: "agent \<Rightarrow> nat \<Rightarrow> agent \<Rightarrow> msg set" |
23746 | 172 |
for A :: agent and n :: nat and B :: agent |
173 |
where |
|
67613 | 174 |
Request [intro]: "cons (anchor A n B) nil \<in> valid A n B" |
13508 | 175 |
|
67613 | 176 |
| Propose [intro]: "L \<in> valid A n B |
177 |
\<Longrightarrow> cons (chain (next_shop (head L)) ofr A L C) L \<in> valid A n B" |
|
13508 | 178 |
|
61830 | 179 |
subsubsection\<open>basic properties of valid\<close> |
13508 | 180 |
|
67613 | 181 |
lemma valid_not_empty: "L \<in> valid A n B \<Longrightarrow> \<exists>M L'. L = cons M L'" |
13508 | 182 |
by (erule valid.cases, auto) |
183 |
||
67613 | 184 |
lemma valid_pos_len: "L \<in> valid A n B \<Longrightarrow> 0 < len L" |
13508 | 185 |
by (erule valid.induct, auto) |
186 |
||
61830 | 187 |
subsubsection\<open>offers of an offer list\<close> |
13508 | 188 |
|
67613 | 189 |
definition offer_nonces :: "msg \<Rightarrow> msg set" where |
190 |
"offer_nonces L \<equiv> {X. X \<in> parts {L} \<and> (\<exists>n. X = Nonce n)}" |
|
13508 | 191 |
|
61830 | 192 |
subsubsection\<open>the originator can get the offers\<close> |
13508 | 193 |
|
67613 | 194 |
lemma "L \<in> valid A n B \<Longrightarrow> offer_nonces L \<subseteq> analz (insert L (initState A))" |
13508 | 195 |
by (erule valid.induct, auto simp: anchor_def chain_def sign_def |
196 |
offer_nonces_def initState.simps) |
|
197 |
||
61830 | 198 |
subsubsection\<open>list of offers\<close> |
13508 | 199 |
|
35418 | 200 |
fun offers :: "msg => msg" where |
61956 | 201 |
"offers (cons M L) = cons \<lbrace>shop M, nonce M\<rbrace> (offers L)" | |
13508 | 202 |
"offers other = nil" |
203 |
||
61830 | 204 |
subsubsection\<open>list of agents whose keys are used to sign a list of offers\<close> |
13508 | 205 |
|
35418 | 206 |
fun shops :: "msg => msg" where |
207 |
"shops (cons M L) = cons (shop M) (shops L)" | |
|
13508 | 208 |
"shops other = other" |
209 |
||
67613 | 210 |
lemma shops_in_agl: "L \<in> valid A n B \<Longrightarrow> shops L \<in> agl" |
13508 | 211 |
by (erule valid.induct, auto simp: anchor_def chain_def sign_def) |
212 |
||
61830 | 213 |
subsubsection\<open>builds a trace from an itinerary\<close> |
13508 | 214 |
|
67613 | 215 |
fun offer_list :: "agent \<times> nat \<times> agent \<times> msg \<times> nat \<Rightarrow> msg" where |
35418 | 216 |
"offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil" | |
13508 | 217 |
"offer_list (A,n,B,cons (Agent C) I,ofr) = ( |
218 |
let L = offer_list (A,n,B,I,Suc ofr) in |
|
219 |
cons (chain (next_shop (head L)) ofr A L C) L)" |
|
220 |
||
67613 | 221 |
lemma "I \<in> agl \<Longrightarrow> \<forall>ofr. offer_list (A,n,B,I,ofr) \<in> valid A n B" |
13508 | 222 |
by (erule agl.induct, auto) |
223 |
||
67613 | 224 |
fun trace :: "agent \<times> nat \<times> agent \<times> nat \<times> msg \<times> msg \<times> msg |
225 |
\<Rightarrow> event list" where |
|
35418 | 226 |
"trace (B,ofr,A,r,I,L,nil) = []" | |
13508 | 227 |
"trace (B,ofr,A,r,I,L,cons (Agent D) K) = ( |
228 |
let C = (if K=nil then B else agt_nb (head K)) in |
|
229 |
let I' = (if K=nil then cons (Agent A) (cons (Agent B) I) |
|
230 |
else cons (Agent A) (app (I, cons (head K) nil))) in |
|
231 |
let I'' = app (I, cons (head K) nil) in |
|
232 |
pro C (Suc ofr) A r I' L nil D |
|
233 |
# trace (B,Suc ofr,A,r,I'',tail L,K))" |
|
234 |
||
67613 | 235 |
definition trace' :: "agent \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> msg \<Rightarrow> agent \<Rightarrow> nat \<Rightarrow> event list" where |
236 |
"trace' A r n I B ofr \<equiv> ( |
|
13508 | 237 |
let AI = cons (Agent A) I in |
238 |
let L = offer_list (A,n,B,AI,ofr) in |
|
239 |
trace (B,ofr,A,r,nil,L,AI))" |
|
240 |
||
241 |
declare trace'_def [simp] |
|
242 |
||
61830 | 243 |
subsubsection\<open>there is a trace in which the originator receives a valid answer\<close> |
13508 | 244 |
|
67613 | 245 |
lemma p1_not_empty: "evs \<in> p1 \<Longrightarrow> req A r n I B \<in> set evs \<longrightarrow> |
246 |
(\<exists>evs'. evs' @ evs \<in> p1 \<and> pro B' ofr A r I' L J A \<in> set evs' \<and> L \<in> valid A n B)" |
|
13508 | 247 |
oops |
248 |
||
249 |
||
61830 | 250 |
subsection\<open>properties of protocol P1\<close> |
13508 | 251 |
|
61830 | 252 |
text\<open>publicly verifiable forward integrity: |
253 |
anyone can verify the validity of an offer list\<close> |
|
13508 | 254 |
|
61830 | 255 |
subsubsection\<open>strong forward integrity: |
256 |
except the last one, no offer can be modified\<close> |
|
13508 | 257 |
|
67613 | 258 |
lemma strong_forward_integrity: "\<forall>L. Suc i < len L |
259 |
\<longrightarrow> L \<in> valid A n B \<and> repl (L,Suc i,M) \<in> valid A n B \<longrightarrow> M = ith (L,Suc i)" |
|
13508 | 260 |
apply (induct i) |
261 |
(* i = 0 *) |
|
262 |
apply clarify |
|
263 |
apply (frule len_not_empty, clarsimp) |
|
264 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 265 |
apply (ind_cases "\<lbrace>x,xa,l'a\<rbrace> \<in> valid A n B" for x xa l'a) |
266 |
apply (ind_cases "\<lbrace>x,M,l'a\<rbrace> \<in> valid A n B" for x l'a) |
|
13508 | 267 |
apply (simp add: chain_def) |
268 |
(* i > 0 *) |
|
269 |
apply clarify |
|
270 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 271 |
apply (ind_cases "\<lbrace>x,repl(l',Suc na,M)\<rbrace> \<in> valid A n B" for x l' na) |
13508 | 272 |
apply (frule len_not_empty, clarsimp) |
67613 | 273 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 274 |
by (drule_tac x=l' in spec, simp, blast) |
275 |
||
61830 | 276 |
subsubsection\<open>insertion resilience: |
277 |
except at the beginning, no offer can be inserted\<close> |
|
13508 | 278 |
|
67613 | 279 |
lemma chain_isnt_head [simp]: "L \<in> valid A n B \<Longrightarrow> |
280 |
head L \<noteq> chain (next_shop (head L)) ofr A L C" |
|
13508 | 281 |
by (erule valid.induct, auto simp: chain_def sign_def anchor_def) |
282 |
||
67613 | 283 |
lemma insertion_resilience: "\<forall>L. L \<in> valid A n B \<longrightarrow> Suc i < len L |
284 |
\<longrightarrow> ins (L,Suc i,M) \<notin> valid A n B" |
|
71989
bad75618fb82
extraction of equations x = t from premises beneath meta-all
haftmann
parents:
67613
diff
changeset
|
285 |
supply [[simproc del: defined_all]] |
13508 | 286 |
apply (induct i) |
287 |
(* i = 0 *) |
|
288 |
apply clarify |
|
289 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 290 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l', simp) |
291 |
apply (ind_cases "\<lbrace>x,M,l'\<rbrace> \<in> valid A n B" for x l', clarsimp) |
|
292 |
apply (ind_cases "\<lbrace>head l',l'\<rbrace> \<in> valid A n B" for l', simp, simp) |
|
13508 | 293 |
(* i > 0 *) |
294 |
apply clarify |
|
295 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 296 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 297 |
apply (frule len_not_empty, clarsimp) |
67613 | 298 |
apply (ind_cases "\<lbrace>x,ins(l',Suc na,M)\<rbrace> \<in> valid A n B" for x l' na) |
13508 | 299 |
apply (frule len_not_empty, clarsimp) |
300 |
by (drule_tac x=l' in spec, clarsimp) |
|
301 |
||
61830 | 302 |
subsubsection\<open>truncation resilience: |
303 |
only shop i can truncate at offer i\<close> |
|
13508 | 304 |
|
67613 | 305 |
lemma truncation_resilience: "\<forall>L. L \<in> valid A n B \<longrightarrow> Suc i < len L |
306 |
\<longrightarrow> cons M (trunc (L,Suc i)) \<in> valid A n B \<longrightarrow> shop M = shop (ith (L,i))" |
|
13508 | 307 |
apply (induct i) |
308 |
(* i = 0 *) |
|
309 |
apply clarify |
|
310 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 311 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 312 |
apply (frule len_not_empty, clarsimp) |
67613 | 313 |
apply (ind_cases "\<lbrace>M,l'\<rbrace> \<in> valid A n B" for l') |
13508 | 314 |
apply (frule len_not_empty, clarsimp, simp) |
315 |
(* i > 0 *) |
|
316 |
apply clarify |
|
317 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 318 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 319 |
apply (frule len_not_empty, clarsimp) |
320 |
by (drule_tac x=l' in spec, clarsimp) |
|
321 |
||
61830 | 322 |
subsubsection\<open>declarations for tactics\<close> |
13508 | 323 |
|
324 |
declare knows_Spy_partsEs [elim] |
|
325 |
declare Fake_parts_insert [THEN subsetD, dest] |
|
326 |
declare initState.simps [simp del] |
|
327 |
||
61830 | 328 |
subsubsection\<open>get components of a message\<close> |
13508 | 329 |
|
67613 | 330 |
lemma get_ML [dest]: "Says A' B \<lbrace>A,r,I,M,L\<rbrace> \<in> set evs \<Longrightarrow> |
331 |
M \<in> parts (spies evs) \<and> L \<in> parts (spies evs)" |
|
13508 | 332 |
by blast |
333 |
||
61830 | 334 |
subsubsection\<open>general properties of p1\<close> |
13508 | 335 |
|
336 |
lemma reqm_neq_prom [iff]: |
|
67613 | 337 |
"reqm A r n I B \<noteq> prom B' ofr A' r' I' (cons M L) J C" |
13508 | 338 |
by (auto simp: reqm_def prom_def) |
339 |
||
340 |
lemma prom_neq_reqm [iff]: |
|
67613 | 341 |
"prom B' ofr A' r' I' (cons M L) J C \<noteq> reqm A r n I B" |
13508 | 342 |
by (auto simp: reqm_def prom_def) |
343 |
||
67613 | 344 |
lemma req_neq_pro [iff]: "req A r n I B \<noteq> pro B' ofr A' r' I' (cons M L) J C" |
13508 | 345 |
by (auto simp: req_def pro_def) |
346 |
||
67613 | 347 |
lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C \<noteq> req A r n I B" |
13508 | 348 |
by (auto simp: req_def pro_def) |
349 |
||
67613 | 350 |
lemma p1_has_no_Gets: "evs \<in> p1 \<Longrightarrow> \<forall>A X. Gets A X \<notin> set evs" |
13508 | 351 |
by (erule p1.induct, auto simp: req_def pro_def) |
352 |
||
353 |
lemma p1_is_Gets_correct [iff]: "Gets_correct p1" |
|
354 |
by (auto simp: Gets_correct_def dest: p1_has_no_Gets) |
|
355 |
||
356 |
lemma p1_is_one_step [iff]: "one_step p1" |
|
76288
b82ac7ef65ec
Removal of the "unfold" method in favour of "unfolding"
paulson <lp15@cam.ac.uk>
parents:
76287
diff
changeset
|
357 |
unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> p1" for ev evs, auto) |
13508 | 358 |
|
67613 | 359 |
lemma p1_has_only_Says' [rule_format]: "evs \<in> p1 \<Longrightarrow> |
360 |
ev \<in> set evs \<longrightarrow> (\<exists>A B X. ev=Says A B X)" |
|
13508 | 361 |
by (erule p1.induct, auto simp: req_def pro_def) |
362 |
||
363 |
lemma p1_has_only_Says [iff]: "has_only_Says p1" |
|
364 |
by (auto simp: has_only_Says_def dest: p1_has_only_Says') |
|
365 |
||
366 |
lemma p1_is_regular [iff]: "regular p1" |
|
367 |
apply (simp only: regular_def, clarify) |
|
368 |
apply (erule_tac p1.induct) |
|
369 |
apply (simp_all add: initState.simps knows.simps pro_def prom_def |
|
370 |
req_def reqm_def anchor_def chain_def sign_def) |
|
371 |
by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans) |
|
372 |
||
61830 | 373 |
subsubsection\<open>private keys are safe\<close> |
13508 | 374 |
|
375 |
lemma priK_parts_Friend_imp_bad [rule_format,dest]: |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
376 |
"\<lbrakk>evs \<in> p1; Friend B \<noteq> A\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
377 |
\<Longrightarrow> (Key (priK A) \<in> parts (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)" |
13508 | 378 |
apply (erule p1.induct) |
379 |
apply (simp_all add: initState.simps knows.simps pro_def prom_def |
|
17778 | 380 |
req_def reqm_def anchor_def chain_def sign_def) |
13508 | 381 |
apply (blast dest: no_Key_in_agl) |
382 |
apply (auto del: parts_invKey disjE dest: parts_trans |
|
383 |
simp add: no_Key_in_appdel) |
|
384 |
done |
|
385 |
||
386 |
lemma priK_analz_Friend_imp_bad [rule_format,dest]: |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
387 |
"\<lbrakk>evs \<in> p1; Friend B \<noteq> A\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
388 |
\<Longrightarrow> (Key (priK A) \<in> analz (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)" |
13508 | 389 |
by auto |
390 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
391 |
lemma priK_notin_knows_max_Friend: "\<lbrakk>evs \<in> p1; A \<notin> bad; A \<noteq> Friend C\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
392 |
\<Longrightarrow> Key (priK A) \<notin> analz (knows_max (Friend C) evs)" |
13508 | 393 |
apply (rule not_parts_not_analz, simp add: knows_max_def, safe) |
394 |
apply (drule_tac H="spies' evs" in parts_sub) |
|
395 |
apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) |
|
396 |
apply (drule_tac H="spies evs" in parts_sub) |
|
397 |
by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend) |
|
398 |
||
61830 | 399 |
subsubsection\<open>general guardedness properties\<close> |
13508 | 400 |
|
67613 | 401 |
lemma agl_guard [intro]: "I \<in> agl \<Longrightarrow> I \<in> guard n Ks" |
13508 | 402 |
by (erule agl.induct, auto) |
403 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
404 |
lemma Says_to_knows_max'_guard: "\<lbrakk>Says A' C \<lbrace>A'',r,I,L\<rbrace> \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
405 |
Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks" |
13508 | 406 |
by (auto dest: Says_to_knows_max') |
407 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
408 |
lemma Says_from_knows_max'_guard: "\<lbrakk>Says C A' \<lbrace>A'',r,I,L\<rbrace> \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
409 |
Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks" |
13508 | 410 |
by (auto dest: Says_from_knows_max') |
411 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
412 |
lemma Says_Nonce_not_used_guard: "\<lbrakk>Says A' B \<lbrace>A'',r,I,L\<rbrace> \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
413 |
Nonce n \<notin> used evs\<rbrakk> \<Longrightarrow> L \<in> guard n Ks" |
13508 | 414 |
by (drule not_used_not_parts, auto) |
415 |
||
61830 | 416 |
subsubsection\<open>guardedness of messages\<close> |
13508 | 417 |
|
67613 | 418 |
lemma chain_guard [iff]: "chain B ofr A L C \<in> guard n {priK A}" |
13508 | 419 |
by (case_tac "ofr=n", auto simp: chain_def sign_def) |
420 |
||
67613 | 421 |
lemma chain_guard_Nonce_neq [intro]: "n \<noteq> ofr |
422 |
\<Longrightarrow> chain B ofr A' L C \<in> guard n {priK A}" |
|
13508 | 423 |
by (auto simp: chain_def sign_def) |
424 |
||
67613 | 425 |
lemma anchor_guard [iff]: "anchor A n' B \<in> guard n {priK A}" |
13508 | 426 |
by (case_tac "n'=n", auto simp: anchor_def) |
427 |
||
67613 | 428 |
lemma anchor_guard_Nonce_neq [intro]: "n \<noteq> n' |
429 |
\<Longrightarrow> anchor A' n' B \<in> guard n {priK A}" |
|
13508 | 430 |
by (auto simp: anchor_def) |
431 |
||
67613 | 432 |
lemma reqm_guard [intro]: "I \<in> agl \<Longrightarrow> reqm A r n' I B \<in> guard n {priK A}" |
13508 | 433 |
by (case_tac "n'=n", auto simp: reqm_def) |
434 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
435 |
lemma reqm_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> n'; I \<in> agl\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
436 |
\<Longrightarrow> reqm A' r n' I B \<in> guard n {priK A}" |
13508 | 437 |
by (auto simp: reqm_def) |
438 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
439 |
lemma prom_guard [intro]: "\<lbrakk>I \<in> agl; J \<in> agl; L \<in> guard n {priK A}\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
440 |
\<Longrightarrow> prom B ofr A r I L J C \<in> guard n {priK A}" |
13508 | 441 |
by (auto simp: prom_def) |
442 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
443 |
lemma prom_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> ofr; I \<in> agl; J \<in> agl; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
444 |
L \<in> guard n {priK A}\<rbrakk> \<Longrightarrow> prom B ofr A' r I L J C \<in> guard n {priK A}" |
13508 | 445 |
by (auto simp: prom_def) |
446 |
||
61830 | 447 |
subsubsection\<open>Nonce uniqueness\<close> |
13508 | 448 |
|
67613 | 449 |
lemma uniq_Nonce_in_chain [dest]: "Nonce k \<in> parts {chain B ofr A L C} \<Longrightarrow> k=ofr" |
13508 | 450 |
by (auto simp: chain_def sign_def) |
451 |
||
67613 | 452 |
lemma uniq_Nonce_in_anchor [dest]: "Nonce k \<in> parts {anchor A n B} \<Longrightarrow> k=n" |
13508 | 453 |
by (auto simp: anchor_def chain_def sign_def) |
454 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
455 |
lemma uniq_Nonce_in_reqm [dest]: "\<lbrakk>Nonce k \<in> parts {reqm A r n I B}; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
456 |
I \<in> agl\<rbrakk> \<Longrightarrow> k=n" |
13508 | 457 |
by (auto simp: reqm_def dest: no_Nonce_in_agl) |
458 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
459 |
lemma uniq_Nonce_in_prom [dest]: "\<lbrakk>Nonce k \<in> parts {prom B ofr A r I L J C}; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
460 |
I \<in> agl; J \<in> agl; Nonce k \<notin> parts {L}\<rbrakk> \<Longrightarrow> k=ofr" |
13508 | 461 |
by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) |
462 |
||
61830 | 463 |
subsubsection\<open>requests are guarded\<close> |
13508 | 464 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
465 |
lemma req_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow> |
67613 | 466 |
req A r n I B \<in> set evs \<longrightarrow> Guard n {priK A} (spies evs)" |
13508 | 467 |
apply (erule p1.induct, simp) |
468 |
apply (simp add: req_def knows.simps, safe) |
|
469 |
apply (erule in_synth_Guard, erule Guard_analz, simp) |
|
470 |
by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) |
|
471 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
472 |
lemma req_imp_Guard_Friend: "\<lbrakk>evs \<in> p1; A \<notin> bad; req A r n I B \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
473 |
\<Longrightarrow> Guard n {priK A} (knows_max (Friend C) evs)" |
13508 | 474 |
apply (rule Guard_knows_max') |
475 |
apply (rule_tac H="spies evs" in Guard_mono) |
|
476 |
apply (rule req_imp_Guard, simp+) |
|
477 |
apply (rule_tac B="spies' evs" in subset_trans) |
|
478 |
apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) |
|
479 |
by (rule knows'_sub_knows) |
|
480 |
||
61830 | 481 |
subsubsection\<open>propositions are guarded\<close> |
13508 | 482 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
483 |
lemma pro_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad\<rbrakk> \<Longrightarrow> |
67613 | 484 |
pro B ofr A r I (cons M L) J C \<in> set evs \<longrightarrow> Guard ofr {priK A} (spies evs)" |
71989
bad75618fb82
extraction of equations x = t from premises beneath meta-all
haftmann
parents:
67613
diff
changeset
|
485 |
supply [[simproc del: defined_all]] |
13508 | 486 |
apply (erule p1.induct) (* +3 subgoals *) |
487 |
(* Nil *) |
|
488 |
apply simp |
|
489 |
(* Fake *) |
|
490 |
apply (simp add: pro_def, safe) (* +4 subgoals *) |
|
491 |
(* 1 *) |
|
492 |
apply (erule in_synth_Guard, drule Guard_analz, simp, simp) |
|
493 |
(* 2 *) |
|
494 |
apply simp |
|
495 |
(* 3 *) |
|
496 |
apply (simp, simp add: req_def pro_def, blast) |
|
497 |
(* 4 *) |
|
498 |
apply (simp add: pro_def) |
|
499 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard) |
|
500 |
(* 5 *) |
|
501 |
apply simp |
|
502 |
apply safe (* +1 subgoal *) |
|
503 |
apply (simp add: pro_def) |
|
504 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard) |
|
505 |
(* 6 *) |
|
506 |
apply (simp add: pro_def) |
|
507 |
apply (blast dest: Says_imp_knows_Spy) |
|
508 |
(* Request *) |
|
509 |
apply (simp add: pro_def) |
|
510 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard) |
|
511 |
(* Propose *) |
|
512 |
apply simp |
|
513 |
apply safe (* +1 subgoal *) |
|
514 |
(* 1 *) |
|
515 |
apply (simp add: pro_def) |
|
516 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard) |
|
517 |
(* 2 *) |
|
518 |
apply (simp add: pro_def) |
|
519 |
by (blast dest: Says_imp_knows_Spy) |
|
520 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
521 |
lemma pro_imp_Guard_Friend: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
522 |
pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
523 |
\<Longrightarrow> Guard ofr {priK A} (knows_max (Friend D) evs)" |
13508 | 524 |
apply (rule Guard_knows_max') |
525 |
apply (rule_tac H="spies evs" in Guard_mono) |
|
526 |
apply (rule pro_imp_Guard, simp+) |
|
527 |
apply (rule_tac B="spies' evs" in subset_trans) |
|
528 |
apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) |
|
529 |
by (rule knows'_sub_knows) |
|
530 |
||
61830 | 531 |
subsubsection\<open>data confidentiality: |
532 |
no one other than the originator can decrypt the offers\<close> |
|
13508 | 533 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
534 |
lemma Nonce_req_notin_spies: "\<lbrakk>evs \<in> p1; req A r n I B \<in> set evs; A \<notin> bad\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
535 |
\<Longrightarrow> Nonce n \<notin> analz (spies evs)" |
13508 | 536 |
by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) |
537 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
538 |
lemma Nonce_req_notin_knows_max_Friend: "\<lbrakk>evs \<in> p1; req A r n I B \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
539 |
A \<notin> bad; A \<noteq> Friend C\<rbrakk> \<Longrightarrow> Nonce n \<notin> analz (knows_max (Friend C) evs)" |
13508 | 540 |
apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) |
541 |
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) |
|
542 |
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) |
|
543 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
544 |
lemma Nonce_pro_notin_spies: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
545 |
pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> \<Longrightarrow> Nonce ofr \<notin> analz (spies evs)" |
13508 | 546 |
by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) |
547 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
548 |
lemma Nonce_pro_notin_knows_max_Friend: "\<lbrakk>evs \<in> p1; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
549 |
A \<noteq> Friend D; pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
550 |
\<Longrightarrow> Nonce ofr \<notin> analz (knows_max (Friend D) evs)" |
13508 | 551 |
apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) |
552 |
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) |
|
553 |
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) |
|
554 |
||
61830 | 555 |
subsubsection\<open>non repudiability: |
556 |
an offer signed by B has been sent by B\<close> |
|
13508 | 557 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
558 |
lemma Crypt_reqm: "\<lbrakk>Crypt (priK A) X \<in> parts {reqm A' r n I B}; I \<in> agl\<rbrakk> \<Longrightarrow> A=A'" |
13508 | 559 |
by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) |
560 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
561 |
lemma Crypt_prom: "\<lbrakk>Crypt (priK A) X \<in> parts {prom B ofr A' r I L J C}; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
562 |
I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow> A=B \<or> Crypt (priK A) X \<in> parts {L}" |
13508 | 563 |
apply (simp add: prom_def anchor_def chain_def sign_def) |
564 |
by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) |
|
565 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
566 |
lemma Crypt_safeness: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow> Crypt (priK A) X \<in> parts (spies evs) |
67613 | 567 |
\<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> Crypt (priK A) X \<in> parts {Y})" |
13508 | 568 |
apply (erule p1.induct) |
569 |
(* Nil *) |
|
570 |
apply simp |
|
571 |
(* Fake *) |
|
572 |
apply clarsimp |
|
67613 | 573 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) X \<in> G" in parts_insert_substD, simp) |
13508 | 574 |
apply (erule disjE) |
575 |
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) |
|
576 |
(* Request *) |
|
577 |
apply (simp add: req_def, clarify) |
|
67613 | 578 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) X \<in> G" in parts_insert_substD, simp) |
13508 | 579 |
apply (erule disjE) |
580 |
apply (frule Crypt_reqm, simp, clarify) |
|
581 |
apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast) |
|
582 |
(* Propose *) |
|
583 |
apply (simp add: pro_def, clarify) |
|
67613 | 584 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) X \<in> G" in parts_insert_substD, simp) |
13508 | 585 |
apply (rotate_tac -1, erule disjE) |
586 |
apply (frule Crypt_prom, simp, simp) |
|
587 |
apply (rotate_tac -1, erule disjE) |
|
588 |
apply (rule_tac x=C in exI) |
|
589 |
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast) |
|
67613 | 590 |
apply (subgoal_tac "cons M L \<in> parts (spies evsp)") |
13508 | 591 |
apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast) |
592 |
apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj) |
|
593 |
apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) |
|
594 |
by auto |
|
595 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
596 |
lemma Crypt_Hash_imp_sign: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow> |
67613 | 597 |
Crypt (priK A) (Hash X) \<in> parts (spies evs) |
598 |
\<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})" |
|
13508 | 599 |
apply (erule p1.induct) |
600 |
(* Nil *) |
|
601 |
apply simp |
|
602 |
(* Fake *) |
|
603 |
apply clarsimp |
|
67613 | 604 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) (Hash X) \<in> G" in parts_insert_substD) |
13508 | 605 |
apply simp |
606 |
apply (erule disjE) |
|
607 |
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) |
|
608 |
(* Request *) |
|
609 |
apply (simp add: req_def, clarify) |
|
67613 | 610 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) (Hash X) \<in> G" in parts_insert_substD) |
13508 | 611 |
apply simp |
612 |
apply (erule disjE) |
|
613 |
apply (frule Crypt_reqm, simp+) |
|
614 |
apply (rule_tac x=B in exI, rule_tac x="reqm Aa r n I B" in exI) |
|
615 |
apply (simp add: reqm_def sign_def anchor_def no_Crypt_in_agl) |
|
616 |
apply (simp add: chain_def sign_def, blast) |
|
617 |
(* Propose *) |
|
618 |
apply (simp add: pro_def, clarify) |
|
67613 | 619 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) (Hash X) \<in> G" in parts_insert_substD) |
13508 | 620 |
apply simp |
621 |
apply (rotate_tac -1, erule disjE) |
|
622 |
apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel) |
|
623 |
apply (simp add: chain_def sign_def) |
|
624 |
apply (rotate_tac -1, erule disjE) |
|
625 |
apply (rule_tac x=C in exI) |
|
626 |
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI) |
|
627 |
apply (simp add: prom_def chain_def sign_def) |
|
628 |
apply (erule impE) |
|
629 |
apply (blast dest: get_ML parts_sub) |
|
630 |
apply (blast del: MPair_parts)+ |
|
631 |
done |
|
632 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
633 |
lemma sign_safeness: "\<lbrakk>evs \<in> p1; A \<notin> bad\<rbrakk> \<Longrightarrow> sign A X \<in> parts (spies evs) |
67613 | 634 |
\<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})" |
13508 | 635 |
apply (clarify, simp add: sign_def, frule parts.Snd) |
636 |
apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) |
|
637 |
done |
|
638 |
||
62390 | 639 |
end |