author | bulwahn |
Fri, 13 Jul 2012 08:44:42 +0200 | |
changeset 48254 | 63e0ca00b952 |
parent 37936 | 1e4c5015a72e |
child 56073 | 29e308b56d23 |
permissions | -rw-r--r-- |
37936 | 1 |
(* Title: HOL/Auth/ZhouGollmann.thy |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
2 |
Author: Giampaolo Bella and L C Paulson, Cambridge Univ Computer Lab |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
3 |
Copyright 2003 University of Cambridge |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
4 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
5 |
The protocol of |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
6 |
Jianying Zhou and Dieter Gollmann, |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
7 |
A Fair Non-Repudiation Protocol, |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
8 |
Security and Privacy 1996 (Oakland) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
9 |
55-61 |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
10 |
*) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
11 |
|
16417 | 12 |
theory ZhouGollmann imports Public begin |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
13 |
|
20768 | 14 |
abbreviation |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
15 |
TTP :: agent where "TTP == Server" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
16 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
17 |
abbreviation f_sub :: nat where "f_sub == 5" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
18 |
abbreviation f_nro :: nat where "f_nro == 2" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
19 |
abbreviation f_nrr :: nat where "f_nrr == 3" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
20 |
abbreviation f_con :: nat where "f_con == 4" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
21 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
22 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
23 |
definition broken :: "agent set" where |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
24 |
--{*the compromised honest agents; TTP is included as it's not allowed to |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
25 |
use the protocol*} |
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
26 |
"broken == bad - {Spy}" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
27 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
28 |
declare broken_def [simp] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
29 |
|
23746 | 30 |
inductive_set zg :: "event list set" |
31 |
where |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
32 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
33 |
Nil: "[] \<in> zg" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
34 |
|
23746 | 35 |
| Fake: "[| evsf \<in> zg; X \<in> synth (analz (spies evsf)) |] |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
36 |
==> Says Spy B X # evsf \<in> zg" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
37 |
|
23746 | 38 |
| Reception: "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
39 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
40 |
(*L is fresh for honest agents. |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
41 |
We don't require K to be fresh because we don't bother to prove secrecy! |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
42 |
We just assume that the protocol's objective is to deliver K fairly, |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
43 |
rather than to keep M secret.*) |
23746 | 44 |
| ZG1: "[| evs1 \<in> zg; Nonce L \<notin> used evs1; C = Crypt K (Number m); |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
45 |
K \<in> symKeys; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
46 |
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|] |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
47 |
==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
48 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
49 |
(*B must check that NRO is A's signature to learn the sender's name*) |
23746 | 50 |
| ZG2: "[| evs2 \<in> zg; |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
51 |
Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
52 |
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
53 |
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|] |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
54 |
==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} # evs2 \<in> zg" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
55 |
|
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
56 |
(*A must check that NRR is B's signature to learn the sender's name; |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
57 |
without spy, the matching label would be enough*) |
23746 | 58 |
| ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys; |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
59 |
Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
60 |
Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
61 |
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
62 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|] |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
63 |
==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
64 |
# evs3 \<in> zg" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
65 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
66 |
(*TTP checks that sub_K is A's signature to learn who issued K, then |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
67 |
gives credentials to A and B. The Notes event models the availability of |
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
68 |
the credentials, but the act of fetching them is not modelled. We also |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
69 |
give con_K to the Spy. This makes the threat model more dangerous, while |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
70 |
also allowing lemma @{text Crypt_used_imp_spies} to omit the condition |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
71 |
@{term "K \<noteq> priK TTP"}. *) |
23746 | 72 |
| ZG4: "[| evs4 \<in> zg; K \<in> symKeys; |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
73 |
Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
74 |
\<in> set evs4; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
75 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
76 |
con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
77 |
Nonce L, Key K|}|] |
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
78 |
==> Says TTP Spy con_K |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
79 |
# |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
80 |
Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23746
diff
changeset
|
81 |
# evs4 \<in> zg" |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
82 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
83 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
84 |
declare Says_imp_knows_Spy [THEN analz.Inj, dest] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
85 |
declare Fake_parts_insert_in_Un [dest] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
86 |
declare analz_into_parts [dest] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
87 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
88 |
declare symKey_neq_priEK [simp] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
89 |
declare symKey_neq_priEK [THEN not_sym, simp] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
90 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
91 |
|
14146 | 92 |
text{*A "possibility property": there are traces that reach the end*} |
93 |
lemma "[|A \<noteq> B; TTP \<noteq> A; TTP \<noteq> B; K \<in> symKeys|] ==> |
|
94 |
\<exists>L. \<exists>evs \<in> zg. |
|
95 |
Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, |
|
96 |
Crypt (priK TTP) {|Number f_con, Agent A, Agent B, Nonce L, Key K|} |} |
|
97 |
\<in> set evs" |
|
98 |
apply (intro exI bexI) |
|
99 |
apply (rule_tac [2] zg.Nil |
|
100 |
[THEN zg.ZG1, THEN zg.Reception [of _ A B], |
|
101 |
THEN zg.ZG2, THEN zg.Reception [of _ B A], |
|
102 |
THEN zg.ZG3, THEN zg.Reception [of _ A TTP], |
|
103 |
THEN zg.ZG4]) |
|
104 |
apply (possibility, auto) |
|
105 |
done |
|
106 |
||
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
107 |
subsection {*Basic Lemmas*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
108 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
109 |
lemma Gets_imp_Says: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
110 |
"[| Gets B X \<in> set evs; evs \<in> zg |] ==> \<exists>A. Says A B X \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
111 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
112 |
apply (erule zg.induct, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
113 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
114 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
115 |
lemma Gets_imp_knows_Spy: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
116 |
"[| Gets B X \<in> set evs; evs \<in> zg |] ==> X \<in> spies evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
117 |
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
118 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
119 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
120 |
text{*Lets us replace proofs about @{term "used evs"} by simpler proofs |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
121 |
about @{term "parts (spies evs)"}.*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
122 |
lemma Crypt_used_imp_spies: |
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
123 |
"[| Crypt K X \<in> used evs; evs \<in> zg |] |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
124 |
==> Crypt K X \<in> parts (spies evs)" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
125 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
126 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
127 |
apply (simp_all add: parts_insert_knows_A) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
128 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
129 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
130 |
lemma Notes_TTP_imp_Gets: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
131 |
"[|Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K |} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
132 |
\<in> set evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
133 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
134 |
evs \<in> zg|] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
135 |
==> Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
136 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
137 |
apply (erule zg.induct, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
138 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
139 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
140 |
text{*For reasoning about C, which is encrypted in message ZG2*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
141 |
lemma ZG2_msg_in_parts_spies: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
142 |
"[|Gets B {|F, B', L, C, X|} \<in> set evs; evs \<in> zg|] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
143 |
==> C \<in> parts (spies evs)" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
144 |
by (blast dest: Gets_imp_Says) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
145 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
146 |
(*classical regularity lemma on priK*) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
147 |
lemma Spy_see_priK [simp]: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
148 |
"evs \<in> zg ==> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
149 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
150 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
151 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
152 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
153 |
text{*So that blast can use it too*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
154 |
declare Spy_see_priK [THEN [2] rev_iffD1, dest!] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
155 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
156 |
lemma Spy_analz_priK [simp]: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
157 |
"evs \<in> zg ==> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
158 |
by auto |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
159 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
160 |
|
14741 | 161 |
subsection{*About NRO: Validity for @{term B}*} |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
162 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
163 |
text{*Below we prove that if @{term NRO} exists then @{term A} definitely |
14741 | 164 |
sent it, provided @{term A} is not broken.*} |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
165 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
166 |
text{*Strong conclusion for a good agent*} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
167 |
lemma NRO_validity_good: |
14741 | 168 |
"[|NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}; |
169 |
NRO \<in> parts (spies evs); |
|
170 |
A \<notin> bad; evs \<in> zg |] |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
171 |
==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
172 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
173 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
174 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
175 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
176 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
177 |
|
14741 | 178 |
lemma NRO_sender: |
179 |
"[|Says A' B {|n, b, l, C, Crypt (priK A) X|} \<in> set evs; evs \<in> zg|] |
|
180 |
==> A' \<in> {A,Spy}" |
|
181 |
apply (erule rev_mp) |
|
182 |
apply (erule zg.induct, simp_all) |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
183 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
184 |
|
14741 | 185 |
text{*Holds also for @{term "A = Spy"}!*} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
186 |
theorem NRO_validity: |
15047 | 187 |
"[|Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs; |
14741 | 188 |
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}; |
189 |
A \<notin> broken; evs \<in> zg |] |
|
190 |
==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs" |
|
15047 | 191 |
apply (drule Gets_imp_Says, assumption) |
14741 | 192 |
apply clarify |
193 |
apply (frule NRO_sender, auto) |
|
15047 | 194 |
txt{*We are left with the case where the sender is @{term Spy} and not |
195 |
equal to @{term A}, because @{term "A \<notin> bad"}. |
|
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
196 |
Thus theorem @{text NRO_validity_good} applies.*} |
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
197 |
apply (blast dest: NRO_validity_good [OF refl]) |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
198 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
199 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
200 |
|
14741 | 201 |
subsection{*About NRR: Validity for @{term A}*} |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
202 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
203 |
text{*Below we prove that if @{term NRR} exists then @{term B} definitely |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
204 |
sent it, provided @{term B} is not broken.*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
205 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
206 |
text{*Strong conclusion for a good agent*} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
207 |
lemma NRR_validity_good: |
14741 | 208 |
"[|NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}; |
209 |
NRR \<in> parts (spies evs); |
|
210 |
B \<notin> bad; evs \<in> zg |] |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
211 |
==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
212 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
213 |
apply (erule rev_mp) |
14741 | 214 |
apply (erule zg.induct) |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
215 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
216 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
217 |
|
14741 | 218 |
lemma NRR_sender: |
219 |
"[|Says B' A {|n, a, l, Crypt (priK B) X|} \<in> set evs; evs \<in> zg|] |
|
220 |
==> B' \<in> {B,Spy}" |
|
221 |
apply (erule rev_mp) |
|
222 |
apply (erule zg.induct, simp_all) |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
223 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
224 |
|
14741 | 225 |
text{*Holds also for @{term "B = Spy"}!*} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
226 |
theorem NRR_validity: |
14741 | 227 |
"[|Says B' A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs; |
228 |
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}; |
|
229 |
B \<notin> broken; evs \<in> zg|] |
|
230 |
==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs" |
|
231 |
apply clarify |
|
232 |
apply (frule NRR_sender, auto) |
|
233 |
txt{*We are left with the case where @{term "B' = Spy"} and @{term "B' \<noteq> B"}, |
|
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
234 |
i.e. @{term "B \<notin> bad"}, when we can apply @{text NRR_validity_good}.*} |
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
235 |
apply (blast dest: NRR_validity_good [OF refl]) |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
236 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
237 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
238 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
239 |
subsection{*Proofs About @{term sub_K}*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
240 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
241 |
text{*Below we prove that if @{term sub_K} exists then @{term A} definitely |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
242 |
sent it, provided @{term A} is not broken. *} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
243 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
244 |
text{*Strong conclusion for a good agent*} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
245 |
lemma sub_K_validity_good: |
14741 | 246 |
"[|sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
247 |
sub_K \<in> parts (spies evs); |
|
248 |
A \<notin> bad; evs \<in> zg |] |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
249 |
==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs" |
14741 | 250 |
apply clarify |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
251 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
252 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
253 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
254 |
txt{*Fake*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
255 |
apply (blast dest!: Fake_parts_sing_imp_Un) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
256 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
257 |
|
14741 | 258 |
lemma sub_K_sender: |
259 |
"[|Says A' TTP {|n, b, l, k, Crypt (priK A) X|} \<in> set evs; evs \<in> zg|] |
|
260 |
==> A' \<in> {A,Spy}" |
|
261 |
apply (erule rev_mp) |
|
262 |
apply (erule zg.induct, simp_all) |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
263 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
264 |
|
14741 | 265 |
text{*Holds also for @{term "A = Spy"}!*} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
266 |
theorem sub_K_validity: |
15047 | 267 |
"[|Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs; |
14741 | 268 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
269 |
A \<notin> broken; evs \<in> zg |] |
|
270 |
==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs" |
|
15047 | 271 |
apply (drule Gets_imp_Says, assumption) |
14741 | 272 |
apply clarify |
273 |
apply (frule sub_K_sender, auto) |
|
15047 | 274 |
txt{*We are left with the case where the sender is @{term Spy} and not |
275 |
equal to @{term A}, because @{term "A \<notin> bad"}. |
|
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
276 |
Thus theorem @{text sub_K_validity_good} applies.*} |
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
277 |
apply (blast dest: sub_K_validity_good [OF refl]) |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
278 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
279 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
280 |
|
14741 | 281 |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
282 |
subsection{*Proofs About @{term con_K}*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
283 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
284 |
text{*Below we prove that if @{term con_K} exists, then @{term TTP} has it, |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
285 |
and therefore @{term A} and @{term B}) can get it too. Moreover, we know |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
286 |
that @{term A} sent @{term sub_K}*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
287 |
|
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
288 |
lemma con_K_validity: |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
289 |
"[|con_K \<in> used evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
290 |
con_K = Crypt (priK TTP) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
291 |
{|Number f_con, Agent A, Agent B, Nonce L, Key K|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
292 |
evs \<in> zg |] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
293 |
==> Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
294 |
\<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
295 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
296 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
297 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
298 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
299 |
txt{*Fake*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
300 |
apply (blast dest!: Fake_parts_sing_imp_Un) |
14741 | 301 |
txt{*ZG2*} |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
302 |
apply (blast dest: parts_cut) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
303 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
304 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
305 |
text{*If @{term TTP} holds @{term con_K} then @{term A} sent |
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
306 |
@{term sub_K}. We assume that @{term A} is not broken. Importantly, nothing |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
307 |
needs to be assumed about the form of @{term con_K}!*} |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
308 |
lemma Notes_TTP_imp_Says_A: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
309 |
"[|Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
310 |
\<in> set evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
311 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
312 |
A \<notin> broken; evs \<in> zg|] |
14741 | 313 |
==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs" |
314 |
apply clarify |
|
315 |
apply (erule rev_mp) |
|
316 |
apply (erule zg.induct) |
|
317 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) |
|
318 |
txt{*ZG4*} |
|
15047 | 319 |
apply clarify |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
320 |
apply (rule sub_K_validity, auto) |
14741 | 321 |
done |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
322 |
|
14736
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
323 |
text{*If @{term con_K} exists, then @{term A} sent @{term sub_K}. We again |
7104394df99a
broken no longer includes TTP, and other minor changes
paulson
parents:
14207
diff
changeset
|
324 |
assume that @{term A} is not broken. *} |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
325 |
theorem B_sub_K_validity: |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
326 |
"[|con_K \<in> used evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
327 |
con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B, |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
328 |
Nonce L, Key K|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
329 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
14741 | 330 |
A \<notin> broken; evs \<in> zg|] |
331 |
==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs" |
|
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
332 |
by (blast dest: con_K_validity Notes_TTP_imp_Says_A) |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
333 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
334 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
335 |
subsection{*Proving fairness*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
336 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
337 |
text{*Cannot prove that, if @{term B} has NRO, then @{term A} has her NRR. |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
338 |
It would appear that @{term B} has a small advantage, though it is |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
339 |
useless to win disputes: @{term B} needs to present @{term con_K} as well.*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
340 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
341 |
text{*Strange: unicity of the label protects @{term A}?*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
342 |
lemma A_unicity: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
343 |
"[|NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
344 |
NRO \<in> parts (spies evs); |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
345 |
Says A B {|Number f_nro, Agent B, Nonce L, Crypt K M', NRO'|} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
346 |
\<in> set evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
347 |
A \<notin> bad; evs \<in> zg |] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
348 |
==> M'=M" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
349 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
350 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
351 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
352 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
353 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
354 |
txt{*ZG1: freshness*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
355 |
apply (blast dest: parts.Body) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
356 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
357 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
358 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
359 |
text{*Fairness lemma: if @{term sub_K} exists, then @{term A} holds |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
360 |
NRR. Relies on unicity of labels.*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
361 |
lemma sub_K_implies_NRR: |
14741 | 362 |
"[| NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|}; |
363 |
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, Crypt K M|}; |
|
364 |
sub_K \<in> parts (spies evs); |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
365 |
NRO \<in> parts (spies evs); |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
366 |
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
367 |
A \<notin> bad; evs \<in> zg |] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
368 |
==> Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
369 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
370 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
371 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
372 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
373 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
374 |
txt{*Fake*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
375 |
apply blast |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
376 |
txt{*ZG1: freshness*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
377 |
apply (blast dest: parts.Body) |
14741 | 378 |
txt{*ZG3*} |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
379 |
apply (blast dest: A_unicity [OF refl]) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
380 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
381 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
382 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
383 |
lemma Crypt_used_imp_L_used: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
384 |
"[| Crypt (priK TTP) {|F, A, B, L, K|} \<in> used evs; evs \<in> zg |] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
385 |
==> L \<in> used evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
386 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
387 |
apply (erule zg.induct, auto) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
388 |
txt{*Fake*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
389 |
apply (blast dest!: Fake_parts_sing_imp_Un) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
390 |
txt{*ZG2: freshness*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
391 |
apply (blast dest: parts.Body) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
392 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
393 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
394 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
395 |
text{*Fairness for @{term A}: if @{term con_K} and @{term NRO} exist, |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
396 |
then @{term A} holds NRR. @{term A} must be uncompromised, but there is no |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
397 |
assumption about @{term B}.*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
398 |
theorem A_fairness_NRO: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
399 |
"[|con_K \<in> used evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
400 |
NRO \<in> parts (spies evs); |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
401 |
con_K = Crypt (priK TTP) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
402 |
{|Number f_con, Agent A, Agent B, Nonce L, Key K|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
403 |
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
404 |
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, Crypt K M|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
405 |
A \<notin> bad; evs \<in> zg |] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
406 |
==> Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
407 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
408 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
409 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
410 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
411 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
412 |
txt{*Fake*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
413 |
apply (simp add: parts_insert_knows_A) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
414 |
apply (blast dest: Fake_parts_sing_imp_Un) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
415 |
txt{*ZG1*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
416 |
apply (blast dest: Crypt_used_imp_L_used) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
417 |
txt{*ZG2*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
418 |
apply (blast dest: parts_cut) |
14741 | 419 |
txt{*ZG4*} |
420 |
apply (blast intro: sub_K_implies_NRR [OF refl] |
|
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
421 |
dest: Gets_imp_knows_Spy [THEN parts.Inj]) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
422 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
423 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
424 |
text{*Fairness for @{term B}: NRR exists at all, then @{term B} holds NRO. |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
425 |
@{term B} must be uncompromised, but there is no assumption about @{term |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
426 |
A}. *} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
427 |
theorem B_fairness_NRR: |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
428 |
"[|NRR \<in> used evs; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
429 |
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
430 |
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}; |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
431 |
B \<notin> bad; evs \<in> zg |] |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
432 |
==> Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs" |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
433 |
apply clarify |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
434 |
apply (erule rev_mp) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
435 |
apply (erule zg.induct) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
436 |
apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
437 |
txt{*Fake*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
438 |
apply (blast dest!: Fake_parts_sing_imp_Un) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
439 |
txt{*ZG2*} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
440 |
apply (blast dest: parts_cut) |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
441 |
done |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
442 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
443 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
444 |
text{*If @{term con_K} exists at all, then @{term B} can get it, by @{text |
15068
58d216b32199
minor tweaks to go with the new version of the Accountability paper
paulson
parents:
15047
diff
changeset
|
445 |
con_K_validity}. Cannot conclude that also NRO is available to @{term B}, |
14145
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
446 |
because if @{term A} were unfair, @{term A} could build message 3 without |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
447 |
building message 1, which contains NRO. *} |
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
448 |
|
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents:
diff
changeset
|
449 |
end |