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