13508
|
1 |
(******************************************************************************
|
|
2 |
date: march 2002
|
|
3 |
author: Frederic Blanqui
|
|
4 |
email: blanqui@lri.fr
|
|
5 |
webpage: http://www.lri.fr/~blanqui/
|
|
6 |
|
|
7 |
University of Cambridge, Computer Laboratory
|
|
8 |
William Gates Building, JJ Thomson Avenue
|
|
9 |
Cambridge CB3 0FD, United Kingdom
|
|
10 |
******************************************************************************)
|
|
11 |
|
|
12 |
header{*Otway-Rees Protocol*}
|
|
13 |
|
|
14 |
theory OtwayRees = Guard_Shared:
|
|
15 |
|
|
16 |
subsection{*messages used in the protocol*}
|
|
17 |
|
|
18 |
syntax nil :: "msg"
|
|
19 |
|
|
20 |
translations "nil" == "Number 0"
|
|
21 |
|
|
22 |
syntax or1 :: "agent => agent => nat => event"
|
|
23 |
|
|
24 |
translations "or1 A B NA"
|
|
25 |
=> "Says A B {|Nonce NA, Agent A, Agent B,
|
|
26 |
Ciph A {|Nonce NA, Agent A, Agent B|}|}"
|
|
27 |
|
|
28 |
syntax or1' :: "agent => agent => agent => nat => msg => event"
|
|
29 |
|
|
30 |
translations "or1' A' A B NA X"
|
|
31 |
=> "Says A' B {|Nonce NA, Agent A, Agent B, X|}"
|
|
32 |
|
|
33 |
syntax or2 :: "agent => agent => nat => nat => msg => event"
|
|
34 |
|
|
35 |
translations "or2 A B NA NB X"
|
|
36 |
=> "Says B Server {|Nonce NA, Agent A, Agent B, X,
|
|
37 |
Ciph B {|Nonce NA, Nonce NB, Agent A, Agent B|}|}"
|
|
38 |
|
|
39 |
syntax or2' :: "agent => agent => agent => nat => nat => event"
|
|
40 |
|
|
41 |
translations "or2' B' A B NA NB"
|
|
42 |
=> "Says B' Server {|Nonce NA, Agent A, Agent B,
|
|
43 |
Ciph A {|Nonce NA, Agent A, Agent B|},
|
|
44 |
Ciph B {|Nonce NA, Nonce NB, Agent A, Agent B|}|}"
|
|
45 |
|
|
46 |
syntax or3 :: "agent => agent => nat => nat => key => event"
|
|
47 |
|
|
48 |
translations "or3 A B NA NB K"
|
|
49 |
=> "Says Server B {|Nonce NA, Ciph A {|Nonce NA, Key K|},
|
|
50 |
Ciph B {|Nonce NB, Key K|}|}"
|
|
51 |
|
|
52 |
syntax or3':: "agent => msg => agent => agent => nat => nat => key => event"
|
|
53 |
|
|
54 |
translations "or3' S Y A B NA NB K"
|
|
55 |
=> "Says S B {|Nonce NA, Y, Ciph B {|Nonce NB, Key K|}|}"
|
|
56 |
|
|
57 |
syntax or4 :: "agent => agent => nat => msg => event"
|
|
58 |
|
|
59 |
translations "or4 A B NA X" => "Says B A {|Nonce NA, X, nil|}"
|
|
60 |
|
|
61 |
syntax or4' :: "agent => agent => nat => msg => event"
|
|
62 |
|
|
63 |
translations "or4' B' A NA K" =>
|
|
64 |
"Says B' A {|Nonce NA, Ciph A {|Nonce NA, Key K|}, nil|}"
|
|
65 |
|
|
66 |
subsection{*definition of the protocol*}
|
|
67 |
|
|
68 |
consts or :: "event list set"
|
|
69 |
|
|
70 |
inductive or
|
|
71 |
intros
|
|
72 |
|
|
73 |
Nil: "[]:or"
|
|
74 |
|
|
75 |
Fake: "[| evs:or; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs:or"
|
|
76 |
|
|
77 |
OR1: "[| evs1:or; Nonce NA ~:used evs1 |] ==> or1 A B NA # evs1:or"
|
|
78 |
|
|
79 |
OR2: "[| evs2:or; or1' A' A B NA X:set evs2; Nonce NB ~:used evs2 |]
|
|
80 |
==> or2 A B NA NB X # evs2:or"
|
|
81 |
|
|
82 |
OR3: "[| evs3:or; or2' B' A B NA NB:set evs3; Key K ~:used evs3 |]
|
|
83 |
==> or3 A B NA NB K # evs3:or"
|
|
84 |
|
|
85 |
OR4: "[| evs4:or; or2 A B NA NB X:set evs4; or3' S Y A B NA NB K:set evs4 |]
|
|
86 |
==> or4 A B NA X # evs4:or"
|
|
87 |
|
|
88 |
subsection{*declarations for tactics*}
|
|
89 |
|
|
90 |
declare knows_Spy_partsEs [elim]
|
|
91 |
declare Fake_parts_insert [THEN subsetD, dest]
|
|
92 |
declare initState.simps [simp del]
|
|
93 |
|
|
94 |
subsection{*general properties of or*}
|
|
95 |
|
|
96 |
lemma or_has_no_Gets: "evs:or ==> ALL A X. Gets A X ~:set evs"
|
|
97 |
by (erule or.induct, auto)
|
|
98 |
|
|
99 |
lemma or_is_Gets_correct [iff]: "Gets_correct or"
|
|
100 |
by (auto simp: Gets_correct_def dest: or_has_no_Gets)
|
|
101 |
|
|
102 |
lemma or_is_one_step [iff]: "one_step or"
|
|
103 |
by (unfold one_step_def, clarify, ind_cases "ev#evs:or", auto)
|
|
104 |
|
|
105 |
lemma or_has_only_Says' [rule_format]: "evs:or ==>
|
|
106 |
ev:set evs --> (EX A B X. ev=Says A B X)"
|
|
107 |
by (erule or.induct, auto)
|
|
108 |
|
|
109 |
lemma or_has_only_Says [iff]: "has_only_Says or"
|
|
110 |
by (auto simp: has_only_Says_def dest: or_has_only_Says')
|
|
111 |
|
|
112 |
subsection{*or is regular*}
|
|
113 |
|
|
114 |
lemma or1'_parts_spies [dest]: "or1' A' A B NA X:set evs
|
|
115 |
==> X:parts (spies evs)"
|
|
116 |
by blast
|
|
117 |
|
|
118 |
lemma or2_parts_spies [dest]: "or2 A B NA NB X:set evs
|
|
119 |
==> X:parts (spies evs)"
|
|
120 |
by blast
|
|
121 |
|
|
122 |
lemma or3_parts_spies [dest]: "Says S B {|NA, Y, Ciph B {|NB, K|}|}:set evs
|
|
123 |
==> K:parts (spies evs)"
|
|
124 |
by blast
|
|
125 |
|
|
126 |
lemma or_is_regular [iff]: "regular or"
|
|
127 |
apply (simp only: regular_def, clarify)
|
|
128 |
apply (erule or.induct, simp_all add: initState.simps knows.simps)
|
|
129 |
by (auto dest: parts_sub)
|
|
130 |
|
|
131 |
subsection{*guardedness of KAB*}
|
|
132 |
|
|
133 |
lemma Guard_KAB [rule_format]: "[| evs:or; A ~:bad; B ~:bad |] ==>
|
|
134 |
or3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)"
|
|
135 |
apply (erule or.induct)
|
|
136 |
(* Nil *)
|
|
137 |
apply simp_all
|
|
138 |
(* Fake *)
|
|
139 |
apply (clarify, erule in_synth_GuardK, erule GuardK_analz, simp)
|
|
140 |
(* OR1 *)
|
|
141 |
apply blast
|
|
142 |
(* OR2 *)
|
|
143 |
apply safe
|
|
144 |
apply (blast dest: Says_imp_spies, blast)
|
|
145 |
(* OR3 *)
|
|
146 |
apply blast
|
|
147 |
apply (drule_tac A=Server in Key_neq, simp+, rule No_Key, simp)
|
|
148 |
apply (drule_tac A=Server in Key_neq, simp+, rule No_Key, simp)
|
|
149 |
(* OR4 *)
|
|
150 |
by (blast dest: Says_imp_spies in_GuardK_kparts)
|
|
151 |
|
|
152 |
subsection{*guardedness of NB*}
|
|
153 |
|
|
154 |
lemma Guard_NB [rule_format]: "[| evs:or; B ~:bad |] ==>
|
|
155 |
or2 A B NA NB X:set evs --> Guard NB {shrK B} (spies evs)"
|
|
156 |
apply (erule or.induct)
|
|
157 |
(* Nil *)
|
|
158 |
apply simp_all
|
|
159 |
(* Fake *)
|
|
160 |
apply safe
|
|
161 |
apply (erule in_synth_Guard, erule Guard_analz, simp)
|
|
162 |
(* OR1 *)
|
|
163 |
apply (drule_tac n=NB in Nonce_neq, simp+, rule No_Nonce, simp)
|
|
164 |
apply (drule_tac n=NB in Nonce_neq, simp+, rule No_Nonce, simp)
|
|
165 |
(* OR2 *)
|
|
166 |
apply blast
|
|
167 |
apply (drule_tac n=NA in Nonce_neq, simp+, rule No_Nonce, simp)
|
|
168 |
apply (blast intro!: No_Nonce dest: used_parts)
|
|
169 |
apply (drule_tac n=NA in Nonce_neq, simp+, rule No_Nonce, simp)
|
|
170 |
apply (blast intro!: No_Nonce dest: used_parts)
|
|
171 |
apply (blast dest: Says_imp_spies)
|
|
172 |
apply (blast dest: Says_imp_spies)
|
|
173 |
apply (case_tac "Ba=B", clarsimp)
|
|
174 |
apply (drule_tac n=NB and A=B in Nonce_neq, simp+)
|
|
175 |
apply (drule Says_imp_spies)
|
|
176 |
apply (drule_tac n'=NAa in in_Guard_kparts_neq, simp+, rule No_Nonce, simp)
|
|
177 |
(* OR3 *)
|
|
178 |
apply (drule Says_imp_spies)
|
|
179 |
apply (frule_tac n'=NAa in in_Guard_kparts_neq, simp+, rule No_Nonce, simp)
|
|
180 |
apply (case_tac "Aa=B", clarsimp)
|
|
181 |
apply (case_tac "NAa=NB", clarsimp)
|
|
182 |
apply (drule Says_imp_spies)
|
|
183 |
apply (drule_tac Y="{|Nonce NB, Agent Aa, Agent Ba|}"
|
|
184 |
and K="shrK Aa" in in_Guard_kparts_Crypt, simp+)
|
|
185 |
apply (simp add: No_Nonce)
|
|
186 |
apply (case_tac "Ba=B", clarsimp)
|
|
187 |
apply (case_tac "NBa=NB", clarify)
|
|
188 |
apply (drule Says_imp_spies)
|
|
189 |
apply (drule_tac Y="{|Nonce NAa, Nonce NB, Agent Aa, Agent Ba|}"
|
|
190 |
and K="shrK Ba" in in_Guard_kparts_Crypt, simp+)
|
|
191 |
apply (simp add: No_Nonce)
|
|
192 |
(* OR4 *)
|
|
193 |
by (blast dest: Says_imp_spies)+
|
|
194 |
|
|
195 |
end
|