Theory NS_Public
section‹The Needham-Schroeder Public-Key Protocol›
text ‹Flawed version, vulnerable to Lowe's attack.
From Burrows, Abadi and Needham. A Logic of Authentication.
Proc. Royal Soc. 426 (1989), p. 260›
theory NS_Public imports Public begin
inductive_set ns_public :: "event list set"
where
Nil: "[] ∈ ns_public"
| Fake: "⟦evsf ∈ ns_public; X ∈ synth (analz (spies evsf))⟧
⟹ Says Spy B X # evsf ∈ ns_public"
| NS1: "⟦evs1 ∈ ns_public; Nonce NA ∉ used evs1⟧
⟹ Says A B (Crypt (pubEK B) ⦃Nonce NA, Agent A⦄)
# evs1 ∈ ns_public"
| NS2: "⟦evs2 ∈ ns_public; Nonce NB ∉ used evs2;
Says A' B (Crypt (pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs2⟧
⟹ Says B A (Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄)
# evs2 ∈ ns_public"
| NS3: "⟦evs3 ∈ ns_public;
Says A B (Crypt (pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs3;
Says B' A (Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs3⟧
⟹ Says A B (Crypt (pubEK B) (Nonce NB)) # evs3 ∈ ns_public"
declare knows_Spy_partsEs [elim]
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un [dest]
text ‹A "possibility property": there are traces that reach the end›
lemma "∃NB. ∃evs ∈ ns_public. Says A B (Crypt (pubEK B) (Nonce NB)) ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, THEN ns_public.NS3])
by possibility
subsection ‹Inductive proofs about @{term ns_public}›
text ‹Spy never sees another agent's private key! (unless it's bad at start)›
lemma Spy_see_priEK [simp]:
"evs ∈ ns_public ⟹ (Key (priEK A) ∈ parts (spies evs)) = (A ∈ bad)"
by (erule ns_public.induct, auto)
lemma Spy_analz_priEK [simp]:
"evs ∈ ns_public ⟹ (Key (priEK A) ∈ analz (spies evs)) = (A ∈ bad)"
by auto
subsection ‹Authenticity properties obtained from {term NS1}›
text ‹It is impossible to re-use a nonce in both {term NS1} and {term NS2}, provided the nonce
is secret. (Honest users generate fresh nonces.)›
lemma no_nonce_NS1_NS2:
"⟦evs ∈ ns_public;
Crypt (pubEK C) ⦃NA', Nonce NA, Agent D⦄ ∈ parts (spies evs);
Crypt (pubEK B) ⦃Nonce NA, Agent A⦄ ∈ parts (spies evs)⟧
⟹ Nonce NA ∈ analz (spies evs)"
by (induct rule: ns_public.induct) (auto intro: analz_insertI)
text ‹Unicity for {term NS1}: nonce {term NA} identifies agents {term A} and {term B}›
lemma unique_NA:
assumes NA: "Crypt(pubEK B) ⦃Nonce NA, Agent A ⦄ ∈ parts(spies evs)"
"Crypt(pubEK B') ⦃Nonce NA, Agent A'⦄ ∈ parts(spies evs)"
"Nonce NA ∉ analz (spies evs)"
and evs: "evs ∈ ns_public"
shows "A=A' ∧ B=B'"
using evs NA
by (induction rule: ns_public.induct) (auto intro!: analz_insertI split: if_split_asm)
text ‹Secrecy: Spy does not see the nonce sent in msg {term NS1} if {term A} and {term B} are secure
The major premise "Says A B ..." makes it a dest-rule, hence the given assumption order. ›
theorem Spy_not_see_NA:
assumes NA: "Says A B (Crypt(pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs"
"A ∉ bad" "B ∉ bad"
and evs: "evs ∈ ns_public"
shows "Nonce NA ∉ analz (spies evs)"
using evs NA
proof (induction rule: ns_public.induct)
case (Fake evsf X B)
then show ?case
by spy_analz
next
case (NS2 evs2 NB A' B NA A)
then show ?case
by simp (metis Says_imp_analz_Spy analz_into_parts parts.simps unique_NA usedI)
next
case (NS3 evs3 A B NA B' NB)
then show ?case
by simp (meson Says_imp_analz_Spy analz_into_parts no_nonce_NS1_NS2)
qed auto
text ‹Authentication for {term A}: if she receives message 2 and has used {term NA}
to start a run, then {term B} has sent message 2.›
lemma A_trusts_NS2_lemma:
"⟦evs ∈ ns_public;
Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄ ∈ parts (spies evs);
Says A B (Crypt(pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs;
A ∉ bad; B ∉ bad⟧
⟹ Says B A (Crypt(pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs"
by (induct rule: ns_public.induct) (auto dest: Spy_not_see_NA unique_NA)
theorem A_trusts_NS2:
"⟦Says A B (Crypt(pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs;
Says B' A (Crypt(pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs;
A ∉ bad; B ∉ bad; evs ∈ ns_public⟧
⟹ Says B A (Crypt(pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs"
by (blast intro: A_trusts_NS2_lemma)
text ‹If the encrypted message appears then it originated with Alice in {term NS1}›
lemma B_trusts_NS1:
"⟦evs ∈ ns_public;
Crypt (pubEK B) ⦃Nonce NA, Agent A⦄ ∈ parts (spies evs);
Nonce NA ∉ analz (spies evs)⟧
⟹ Says A B (Crypt (pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs"
by (induct evs rule: ns_public.induct) (use analz_insertI in ‹auto split: if_split_asm›)
subsection ‹Authenticity properties obtained from {term NS2}›
text ‹Unicity for {term NS2}: nonce {term NB} identifies nonce {term NA} and agent {term A}
[proof closely follows that for @{thm [source] unique_NA}]›
lemma unique_NB [dest]:
assumes NB: "Crypt(pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄ ∈ parts(spies evs)"
"Crypt(pubEK A') ⦃Nonce NA', Nonce NB, Agent B'⦄ ∈ parts(spies evs)"
"Nonce NB ∉ analz (spies evs)"
and evs: "evs ∈ ns_public"
shows "A=A' ∧ NA=NA' ∧ B=B'"
using evs NB
by (induction rule: ns_public.induct) (auto intro!: analz_insertI split: if_split_asm)
text ‹{term NB} remains secret›
theorem Spy_not_see_NB [dest]:
assumes NB: "Says B A (Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs"
"A ∉ bad" "B ∉ bad"
and evs: "evs ∈ ns_public"
shows "Nonce NB ∉ analz (spies evs)"
using evs NB evs
proof (induction rule: ns_public.induct)
case Fake
then show ?case by spy_analz
next
case NS2
then show ?case
by (auto intro!: no_nonce_NS1_NS2)
qed auto
text ‹Authentication for {term B}: if he receives message 3 and has used {term NB}
in message 2, then {term A} has sent message 3.›
lemma B_trusts_NS3_lemma:
"⟦evs ∈ ns_public;
Crypt (pubEK B) (Nonce NB) ∈ parts (spies evs);
Says B A (Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs;
A ∉ bad; B ∉ bad⟧
⟹ Says A B (Crypt (pubEK B) (Nonce NB)) ∈ set evs"
proof (induction rule: ns_public.induct)
case (NS3 evs3 A B NA B' NB)
then show ?case
by simp (blast intro: no_nonce_NS1_NS2)
qed auto
theorem B_trusts_NS3:
"⟦Says B A (Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs;
Says A' B (Crypt (pubEK B) (Nonce NB)) ∈ set evs;
A ∉ bad; B ∉ bad; evs ∈ ns_public⟧
⟹ Says A B (Crypt (pubEK B) (Nonce NB)) ∈ set evs"
by (blast intro: B_trusts_NS3_lemma)
subsection‹Overall guarantee for {term B}›
text ‹If NS3 has been sent and the nonce NB agrees with the nonce B joined with
NA, then A initiated the run using NA.›
theorem B_trusts_protocol:
"⟦A ∉ bad; B ∉ bad; evs ∈ ns_public⟧ ⟹
Crypt (pubEK B) (Nonce NB) ∈ parts (spies evs) ⟶
Says B A (Crypt (pubEK A) ⦃Nonce NA, Nonce NB, Agent B⦄) ∈ set evs ⟶
Says A B (Crypt (pubEK B) ⦃Nonce NA, Agent A⦄) ∈ set evs"
by (erule ns_public.induct, auto)
end