Theory TLS
section‹The TLS Protocol: Transport Layer Security›
theory TLS imports Public "HOL-Library.Nat_Bijection" begin
definition certificate :: "[agent,key] ⇒ msg" where
"certificate A KA == Crypt (priSK Server) ⦃Agent A, Key KA⦄"
text‹TLS apparently does not require separate keypairs for encryption and
signature. Therefore, we formalize signature as encryption using the
private encryption key.›
datatype role = ClientRole | ServerRole
consts
PRF :: "nat*nat*nat ⇒ nat"
sessionK :: "(nat*nat*nat) * role ⇒ key"
abbreviation
clientK :: "nat*nat*nat ⇒ key" where
"clientK X == sessionK(X, ClientRole)"
abbreviation
serverK :: "nat*nat*nat ⇒ key" where
"serverK X == sessionK(X, ServerRole)"
specification (PRF)
inj_PRF: "inj PRF"
apply (rule exI [of _ "λ(x,y,z). prod_encode(x, prod_encode(y,z))"])
apply (simp add: inj_on_def prod_encode_eq)
done
specification (sessionK)
inj_sessionK: "inj sessionK"
apply (rule exI [of _
"λ((x,y,z), r). prod_encode(case_role 0 1 r,
prod_encode(x, prod_encode(y,z)))"])
apply (simp add: inj_on_def prod_encode_eq split: role.split)
done
axiomatization where
isSym_sessionK: "sessionK nonces ∈ symKeys" and
sessionK_neq_shrK [iff]: "sessionK nonces ≠ shrK A"
inductive_set tls :: "event list set"
where
Nil:
"[] ∈ tls"
| Fake:
"⟦evsf ∈ tls; X ∈ synth (analz (spies evsf))⟧
⟹ Says Spy B X # evsf ∈ tls"
| SpyKeys:
"⟦evsSK ∈ tls;
{Nonce NA, Nonce NB, Nonce M} ⊆ analz (spies evsSK)⟧
⟹ Notes Spy ⦃ Nonce (PRF(M,NA,NB)),
Key (sessionK((NA,NB,M),role))⦄ # evsSK ∈ tls"
| ClientHello:
"⟦evsCH ∈ tls; Nonce NA ∉ used evsCH; NA ∉ range PRF⟧
⟹ Says A B ⦃Agent A, Nonce NA, Number SID, Number PA⦄
# evsCH ∈ tls"
| ServerHello:
"⟦evsSH ∈ tls; Nonce NB ∉ used evsSH; NB ∉ range PRF;
Says A' B ⦃Agent A, Nonce NA, Number SID, Number PA⦄
∈ set evsSH⟧
⟹ Says B A ⦃Nonce NB, Number SID, Number PB⦄ # evsSH ∈ tls"
| Certificate:
"evsC ∈ tls ⟹ Says B A (certificate B (pubK B)) # evsC ∈ tls"
| ClientKeyExch:
"⟦evsCX ∈ tls; Nonce PMS ∉ used evsCX; PMS ∉ range PRF;
Says B' A (certificate B KB) ∈ set evsCX⟧
⟹ Says A B (Crypt KB (Nonce PMS))
# Notes A ⦃Agent B, Nonce PMS⦄
# evsCX ∈ tls"
| CertVerify:
"⟦evsCV ∈ tls;
Says B' A ⦃Nonce NB, Number SID, Number PB⦄ ∈ set evsCV;
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evsCV⟧
⟹ Says A B (Crypt (priK A) (Hash⦃Nonce NB, Agent B, Nonce PMS⦄))
# evsCV ∈ tls"
| ClientFinished:
"⟦evsCF ∈ tls;
Says A B ⦃Agent A, Nonce NA, Number SID, Number PA⦄
∈ set evsCF;
Says B' A ⦃Nonce NB, Number SID, Number PB⦄ ∈ set evsCF;
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evsCF;
M = PRF(PMS,NA,NB)⟧
⟹ Says A B (Crypt (clientK(NA,NB,M))
(Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄))
# evsCF ∈ tls"
| ServerFinished:
"⟦evsSF ∈ tls;
Says A' B ⦃Agent A, Nonce NA, Number SID, Number PA⦄
∈ set evsSF;
Says B A ⦃Nonce NB, Number SID, Number PB⦄ ∈ set evsSF;
Says A'' B (Crypt (pubK B) (Nonce PMS)) ∈ set evsSF;
M = PRF(PMS,NA,NB)⟧
⟹ Says B A (Crypt (serverK(NA,NB,M))
(Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄))
# evsSF ∈ tls"
| ClientAccepts:
"⟦evsCA ∈ tls;
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evsCA;
M = PRF(PMS,NA,NB);
X = Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄;
Says A B (Crypt (clientK(NA,NB,M)) X) ∈ set evsCA;
Says B' A (Crypt (serverK(NA,NB,M)) X) ∈ set evsCA⟧
⟹
Notes A ⦃Number SID, Agent A, Agent B, Nonce M⦄ # evsCA ∈ tls"
| ServerAccepts:
"⟦evsSA ∈ tls;
A ≠ B;
Says A'' B (Crypt (pubK B) (Nonce PMS)) ∈ set evsSA;
M = PRF(PMS,NA,NB);
X = Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄;
Says B A (Crypt (serverK(NA,NB,M)) X) ∈ set evsSA;
Says A' B (Crypt (clientK(NA,NB,M)) X) ∈ set evsSA⟧
⟹
Notes B ⦃Number SID, Agent A, Agent B, Nonce M⦄ # evsSA ∈ tls"
| ClientResume:
"⟦evsCR ∈ tls;
Says A B ⦃Agent A, Nonce NA, Number SID, Number PA⦄ ∈ set evsCR;
Says B' A ⦃Nonce NB, Number SID, Number PB⦄ ∈ set evsCR;
Notes A ⦃Number SID, Agent A, Agent B, Nonce M⦄ ∈ set evsCR⟧
⟹ Says A B (Crypt (clientK(NA,NB,M))
(Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄))
# evsCR ∈ tls"
| ServerResume:
"⟦evsSR ∈ tls;
Says A' B ⦃Agent A, Nonce NA, Number SID, Number PA⦄ ∈ set evsSR;
Says B A ⦃Nonce NB, Number SID, Number PB⦄ ∈ set evsSR;
Notes B ⦃Number SID, Agent A, Agent B, Nonce M⦄ ∈ set evsSR⟧
⟹ Says B A (Crypt (serverK(NA,NB,M))
(Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄)) # evsSR
∈ tls"
| Oops:
"⟦evso ∈ tls; A ≠ Spy;
Says A B (Crypt (sessionK((NA,NB,M),role)) X) ∈ set evso⟧
⟹ Says A Spy (Key (sessionK((NA,NB,M),role))) # evso ∈ tls"
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
declare parts.Body [dest]
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un [dest]
text‹Automatically unfold the definition of "certificate"›
declare certificate_def [simp]
text‹Injectiveness of key-generating functions›
declare inj_PRF [THEN inj_eq, iff]
declare inj_sessionK [THEN inj_eq, iff]
declare isSym_sessionK [simp]
lemma pubK_neq_sessionK [iff]: "publicKey b A ≠ sessionK arg"
by (simp add: symKeys_neq_imp_neq)
declare pubK_neq_sessionK [THEN not_sym, iff]
lemma priK_neq_sessionK [iff]: "invKey (publicKey b A) ≠ sessionK arg"
by (simp add: symKeys_neq_imp_neq)
declare priK_neq_sessionK [THEN not_sym, iff]
lemmas keys_distinct = pubK_neq_sessionK priK_neq_sessionK
subsection‹Protocol Proofs›
text‹Possibility properties state that some traces run the protocol to the
end. Four paths and 12 rules are considered.›
text‹Possibility property ending with ClientAccepts.›
lemma "⟦∀evs. (SOME N. Nonce N ∉ used evs) ∉ range PRF; A ≠ B⟧
⟹ ∃SID M. ∃evs ∈ tls.
Notes A ⦃Number SID, Agent A, Agent B, Nonce M⦄ ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] tls.Nil
[THEN tls.ClientHello, THEN tls.ServerHello,
THEN tls.Certificate, THEN tls.ClientKeyExch,
THEN tls.ClientFinished, THEN tls.ServerFinished,
THEN tls.ClientAccepts], possibility, blast+)
done
text‹And one for ServerAccepts. Either FINISHED message may come first.›
lemma "⟦∀evs. (SOME N. Nonce N ∉ used evs) ∉ range PRF; A ≠ B⟧
⟹ ∃SID NA PA NB PB M. ∃evs ∈ tls.
Notes B ⦃Number SID, Agent A, Agent B, Nonce M⦄ ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] tls.Nil
[THEN tls.ClientHello, THEN tls.ServerHello,
THEN tls.Certificate, THEN tls.ClientKeyExch,
THEN tls.ServerFinished, THEN tls.ClientFinished,
THEN tls.ServerAccepts], possibility, blast+)
done
text‹Another one, for CertVerify (which is optional)›
lemma "⟦∀evs. (SOME N. Nonce N ∉ used evs) ∉ range PRF; A ≠ B⟧
⟹ ∃NB PMS. ∃evs ∈ tls.
Says A B (Crypt (priK A) (Hash⦃Nonce NB, Agent B, Nonce PMS⦄))
∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] tls.Nil
[THEN tls.ClientHello, THEN tls.ServerHello,
THEN tls.Certificate, THEN tls.ClientKeyExch,
THEN tls.CertVerify], possibility, blast+)
done
text‹Another one, for session resumption (both ServerResume and ClientResume).
NO tls.Nil here: we refer to a previous session, not the empty trace.›
lemma "⟦evs0 ∈ tls;
Notes A ⦃Number SID, Agent A, Agent B, Nonce M⦄ ∈ set evs0;
Notes B ⦃Number SID, Agent A, Agent B, Nonce M⦄ ∈ set evs0;
∀evs. (SOME N. Nonce N ∉ used evs) ∉ range PRF;
A ≠ B⟧
⟹ ∃NA PA NB PB X. ∃evs ∈ tls.
X = Hash⦃Number SID, Nonce M,
Nonce NA, Number PA, Agent A,
Nonce NB, Number PB, Agent B⦄ ∧
Says A B (Crypt (clientK(NA,NB,M)) X) ∈ set evs ∧
Says B A (Crypt (serverK(NA,NB,M)) X) ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] tls.ClientHello
[THEN tls.ServerHello,
THEN tls.ServerResume, THEN tls.ClientResume], possibility, blast+)
done
subsection‹Inductive proofs about tls›
text‹Spy never sees a good agent's private key!›
lemma Spy_see_priK [simp]:
"evs ∈ tls ⟹ (Key (privateKey b A) ∈ parts (spies evs)) = (A ∈ bad)"
by (erule tls.induct, force, simp_all, blast)
lemma Spy_analz_priK [simp]:
"evs ∈ tls ⟹ (Key (privateKey b A) ∈ analz (spies evs)) = (A ∈ bad)"
by auto
lemma Spy_see_priK_D [dest!]:
"⟦Key (privateKey b A) ∈ parts (knows Spy evs); evs ∈ tls⟧ ⟹ A ∈ bad"
by (blast dest: Spy_see_priK)
text‹This lemma says that no false certificates exist. One might extend the
model to include bogus certificates for the agents, but there seems
little point in doing so: the loss of their private keys is a worse
breach of security.›
lemma certificate_valid:
"⟦certificate B KB ∈ parts (spies evs); evs ∈ tls⟧ ⟹ KB = pubK B"
apply (erule rev_mp)
apply (erule tls.induct, force, simp_all, blast)
done
lemmas CX_KB_is_pubKB = Says_imp_spies [THEN parts.Inj, THEN certificate_valid]
subsubsection‹Properties of items found in Notes›
lemma Notes_Crypt_parts_spies:
"⟦Notes A ⦃Agent B, X⦄ ∈ set evs; evs ∈ tls⟧
⟹ Crypt (pubK B) X ∈ parts (spies evs)"
apply (erule rev_mp)
apply (erule tls.induct,
frule_tac [7] CX_KB_is_pubKB, force, simp_all)
apply (blast intro: parts_insertI)
done
text‹C may be either A or B›
lemma Notes_master_imp_Crypt_PMS:
"⟦Notes C ⦃s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))⦄ ∈ set evs;
evs ∈ tls⟧
⟹ Crypt (pubK B) (Nonce PMS) ∈ parts (spies evs)"
apply (erule rev_mp)
apply (erule tls.induct, force, simp_all)
txt‹Fake›
apply (blast intro: parts_insertI)
txt‹Client, Server Accept›
apply (blast dest!: Notes_Crypt_parts_spies)+
done
text‹Compared with the theorem above, both premise and conclusion are stronger›
lemma Notes_master_imp_Notes_PMS:
"⟦Notes A ⦃s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))⦄ ∈ set evs;
evs ∈ tls⟧
⟹ Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs"
apply (erule rev_mp)
apply (erule tls.induct, force, simp_all)
txt‹ServerAccepts›
apply blast
done
subsubsection‹Protocol goal: if B receives CertVerify, then A sent it›
text‹B can check A's signature if he has received A's certificate.›
lemma TrustCertVerify_lemma:
"⟦X ∈ parts (spies evs);
X = Crypt (priK A) (Hash⦃nb, Agent B, pms⦄);
evs ∈ tls; A ∉ bad⟧
⟹ Says A B X ∈ set evs"
apply (erule rev_mp, erule ssubst)
apply (erule tls.induct, force, simp_all, blast)
done
text‹Final version: B checks X using the distributed KA instead of priK A›
lemma TrustCertVerify:
"⟦X ∈ parts (spies evs);
X = Crypt (invKey KA) (Hash⦃nb, Agent B, pms⦄);
certificate A KA ∈ parts (spies evs);
evs ∈ tls; A ∉ bad⟧
⟹ Says A B X ∈ set evs"
by (blast dest!: certificate_valid intro!: TrustCertVerify_lemma)
text‹If CertVerify is present then A has chosen PMS.›
lemma UseCertVerify_lemma:
"⟦Crypt (priK A) (Hash⦃nb, Agent B, Nonce PMS⦄) ∈ parts (spies evs);
evs ∈ tls; A ∉ bad⟧
⟹ Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs"
apply (erule rev_mp)
apply (erule tls.induct, force, simp_all, blast)
done
text‹Final version using the distributed KA instead of priK A›
lemma UseCertVerify:
"⟦Crypt (invKey KA) (Hash⦃nb, Agent B, Nonce PMS⦄)
∈ parts (spies evs);
certificate A KA ∈ parts (spies evs);
evs ∈ tls; A ∉ bad⟧
⟹ Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs"
by (blast dest!: certificate_valid intro!: UseCertVerify_lemma)
lemma no_Notes_A_PRF [simp]:
"evs ∈ tls ⟹ Notes A ⦃Agent B, Nonce (PRF x)⦄ ∉ set evs"
apply (erule tls.induct, force, simp_all)
txt‹ClientKeyExch: PMS is assumed to differ from any PRF.›
apply blast
done
lemma MS_imp_PMS [dest!]:
"⟦Nonce (PRF (PMS,NA,NB)) ∈ parts (spies evs); evs ∈ tls⟧
⟹ Nonce PMS ∈ parts (spies evs)"
apply (erule rev_mp)
apply (erule tls.induct, force, simp_all)
txt‹Fake›
apply (blast intro: parts_insertI)
txt‹Easy, e.g. by freshness›
apply (blast dest: Notes_Crypt_parts_spies)+
done
subsubsection‹Unicity results for PMS, the pre-master-secret›
text‹PMS determines B.›
lemma Crypt_unique_PMS:
"⟦Crypt(pubK B) (Nonce PMS) ∈ parts (spies evs);
Crypt(pubK B') (Nonce PMS) ∈ parts (spies evs);
Nonce PMS ∉ analz (spies evs);
evs ∈ tls⟧
⟹ B=B'"
apply (erule rev_mp, erule rev_mp, erule rev_mp)
apply (erule tls.induct, analz_mono_contra, force, simp_all (no_asm_simp))
txt‹Fake, ClientKeyExch›
apply blast+
done
text‹In A's internal Note, PMS determines A and B.›
lemma Notes_unique_PMS:
"⟦Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
Notes A' ⦃Agent B', Nonce PMS⦄ ∈ set evs;
evs ∈ tls⟧
⟹ A=A' ∧ B=B'"
apply (erule rev_mp, erule rev_mp)
apply (erule tls.induct, force, simp_all)
txt‹ClientKeyExch›
apply (blast dest!: Notes_Crypt_parts_spies)
done
subsection‹Secrecy Theorems›
text‹Key compromise lemma needed to prove \<^term>‹analz_image_keys›.
No collection of keys can help the spy get new private keys.›
lemma analz_image_priK [rule_format]:
"evs ∈ tls
⟹ ∀KK. (Key(priK B) ∈ analz (Key`KK ∪ (spies evs))) =
(priK B ∈ KK | B ∈ bad)"
apply (erule tls.induct)
apply (simp_all (no_asm_simp)
del: image_insert
add: image_Un [THEN sym]
insert_Key_image Un_assoc [THEN sym])
txt‹Fake›
apply spy_analz
done
text‹slightly speeds up the big simplification below›
lemma range_sessionkeys_not_priK:
"KK ⊆ range sessionK ⟹ priK B ∉ KK"
by blast
text‹Lemma for the trivial direction of the if-and-only-if›
lemma analz_image_keys_lemma:
"(X ∈ analz (G ∪ H)) ⟶ (X ∈ analz H) ⟹
(X ∈ analz (G ∪ H)) = (X ∈ analz H)"
by (blast intro: analz_mono [THEN subsetD])
lemma analz_image_keys [rule_format]:
"evs ∈ tls ⟹
∀KK. KK ⊆ range sessionK ⟶
(Nonce N ∈ analz (Key`KK ∪ (spies evs))) =
(Nonce N ∈ analz (spies evs))"
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (safe del: iffI)
apply (safe del: impI iffI intro!: analz_image_keys_lemma)
apply (simp_all (no_asm_simp)
del: image_insert imp_disjL
add: image_Un [THEN sym] Un_assoc [THEN sym]
insert_Key_singleton
range_sessionkeys_not_priK analz_image_priK)
apply (simp_all add: insert_absorb)
txt‹Fake›
apply spy_analz
done
text‹Knowing some session keys is no help in getting new nonces›
lemma analz_insert_key [simp]:
"evs ∈ tls ⟹
(Nonce N ∈ analz (insert (Key (sessionK z)) (spies evs))) =
(Nonce N ∈ analz (spies evs))"
by (simp del: image_insert
add: insert_Key_singleton analz_image_keys)
subsubsection‹Protocol goal: serverK(Na,Nb,M) and clientK(Na,Nb,M) remain secure›
text‹Lemma: session keys are never used if PMS is fresh.
Nonces don't have to agree, allowing session resumption.
Converse doesn't hold; revealing PMS doesn't force the keys to be sent.
THEY ARE NOT SUITABLE AS SAFE ELIM RULES.›
lemma PMS_lemma:
"⟦Nonce PMS ∉ parts (spies evs);
K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role);
evs ∈ tls⟧
⟹ Key K ∉ parts (spies evs) ∧ (∀Y. Crypt K Y ∉ parts (spies evs))"
apply (erule rev_mp, erule ssubst)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹Fake›
apply (blast intro: parts_insertI)
txt‹SpyKeys›
apply blast
txt‹Many others›
apply (force dest!: Notes_Crypt_parts_spies Notes_master_imp_Crypt_PMS)+
done
lemma PMS_sessionK_not_spied:
"⟦Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) ∈ parts (spies evs);
evs ∈ tls⟧
⟹ Nonce PMS ∈ parts (spies evs)"
by (blast dest: PMS_lemma)
lemma PMS_Crypt_sessionK_not_spied:
"⟦Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y
∈ parts (spies evs); evs ∈ tls⟧
⟹ Nonce PMS ∈ parts (spies evs)"
by (blast dest: PMS_lemma)
text‹Write keys are never sent if M (MASTER SECRET) is secure.
Converse fails; betraying M doesn't force the keys to be sent!
The strong Oops condition can be weakened later by unicity reasoning,
with some effort.
NO LONGER USED: see ‹clientK_not_spied› and ‹serverK_not_spied››
lemma sessionK_not_spied:
"⟦∀A. Says A Spy (Key (sessionK((NA,NB,M),role))) ∉ set evs;
Nonce M ∉ analz (spies evs); evs ∈ tls⟧
⟹ Key (sessionK((NA,NB,M),role)) ∉ parts (spies evs)"
apply (erule rev_mp, erule rev_mp)
apply (erule tls.induct, analz_mono_contra)
apply (force, simp_all (no_asm_simp))
txt‹Fake, SpyKeys›
apply blast+
done
text‹If A sends ClientKeyExch to an honest B, then the PMS will stay secret.›
lemma Spy_not_see_PMS:
"⟦Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
evs ∈ tls; A ∉ bad; B ∉ bad⟧
⟹ Nonce PMS ∉ analz (spies evs)"
apply (erule rev_mp, erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹Fake›
apply spy_analz
txt‹SpyKeys›
apply force
apply (simp_all add: insert_absorb)
txt‹ClientHello, ServerHello, ClientKeyExch: mostly freshness reasoning›
apply (blast dest: Notes_Crypt_parts_spies)
apply (blast dest: Notes_Crypt_parts_spies)
apply (blast dest: Notes_Crypt_parts_spies)
txt‹ClientAccepts and ServerAccepts: because \<^term>‹PMS ∉ range PRF››
apply force+
done
text‹If A sends ClientKeyExch to an honest B, then the MASTER SECRET
will stay secret.›
lemma Spy_not_see_MS:
"⟦Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
evs ∈ tls; A ∉ bad; B ∉ bad⟧
⟹ Nonce (PRF(PMS,NA,NB)) ∉ analz (spies evs)"
apply (erule rev_mp, erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹Fake›
apply spy_analz
txt‹SpyKeys: by secrecy of the PMS, Spy cannot make the MS›
apply (blast dest!: Spy_not_see_PMS)
apply (simp_all add: insert_absorb)
txt‹ClientAccepts and ServerAccepts: because PMS was already visible;
others, freshness etc.›
apply (blast dest: Notes_Crypt_parts_spies Spy_not_see_PMS
Notes_imp_knows_Spy [THEN analz.Inj])+
done
subsubsection‹Weakening the Oops conditions for leakage of clientK›
text‹If A created PMS then nobody else (except the Spy in replays)
would send a message using a clientK generated from that PMS.›
lemma Says_clientK_unique:
"⟦Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) ∈ set evs;
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
evs ∈ tls; A' ≠ Spy⟧
⟹ A = A'"
apply (erule rev_mp, erule rev_mp)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all)
txt‹ClientKeyExch›
apply (blast dest!: PMS_Crypt_sessionK_not_spied)
txt‹ClientFinished, ClientResume: by unicity of PMS›
apply (blast dest!: Notes_master_imp_Notes_PMS
intro: Notes_unique_PMS [THEN conjunct1])+
done
text‹If A created PMS and has not leaked her clientK to the Spy,
then it is completely secure: not even in parts!›
lemma clientK_not_spied:
"⟦Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) ∉ set evs;
A ∉ bad; B ∉ bad;
evs ∈ tls⟧
⟹ Key (clientK(Na,Nb,PRF(PMS,NA,NB))) ∉ parts (spies evs)"
apply (erule rev_mp, erule rev_mp)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹ClientKeyExch›
apply blast
txt‹SpyKeys›
apply (blast dest!: Spy_not_see_MS)
txt‹ClientKeyExch›
apply (blast dest!: PMS_sessionK_not_spied)
txt‹Oops›
apply (blast intro: Says_clientK_unique)
done
subsubsection‹Weakening the Oops conditions for leakage of serverK›
text‹If A created PMS for B, then nobody other than B or the Spy would
send a message using a serverK generated from that PMS.›
lemma Says_serverK_unique:
"⟦Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) ∈ set evs;
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
evs ∈ tls; A ∉ bad; B ∉ bad; B' ≠ Spy⟧
⟹ B = B'"
apply (erule rev_mp, erule rev_mp)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all)
txt‹ClientKeyExch›
apply (blast dest!: PMS_Crypt_sessionK_not_spied)
txt‹ServerResume, ServerFinished: by unicity of PMS›
apply (blast dest!: Notes_master_imp_Crypt_PMS
dest: Spy_not_see_PMS Notes_Crypt_parts_spies Crypt_unique_PMS)+
done
text‹If A created PMS for B, and B has not leaked his serverK to the Spy,
then it is completely secure: not even in parts!›
lemma serverK_not_spied:
"⟦Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs;
Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) ∉ set evs;
A ∉ bad; B ∉ bad; evs ∈ tls⟧
⟹ Key (serverK(Na,Nb,PRF(PMS,NA,NB))) ∉ parts (spies evs)"
apply (erule rev_mp, erule rev_mp)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹Fake›
apply blast
txt‹SpyKeys›
apply (blast dest!: Spy_not_see_MS)
txt‹ClientKeyExch›
apply (blast dest!: PMS_sessionK_not_spied)
txt‹Oops›
apply (blast intro: Says_serverK_unique)
done
subsubsection‹Protocol goals: if A receives ServerFinished, then B is present
and has used the quoted values PA, PB, etc. Note that it is up to A
to compare PA with what she originally sent.›
text‹The mention of her name (A) in X assures A that B knows who she is.›
lemma TrustServerFinished [rule_format]:
"⟦X = Crypt (serverK(Na,Nb,M))
(Hash⦃Number SID, Nonce M,
Nonce Na, Number PA, Agent A,
Nonce Nb, Number PB, Agent B⦄);
M = PRF(PMS,NA,NB);
evs ∈ tls; A ∉ bad; B ∉ bad⟧
⟹ Says B Spy (Key(serverK(Na,Nb,M))) ∉ set evs ⟶
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs ⟶
X ∈ parts (spies evs) ⟶ Says B A X ∈ set evs"
apply (erule ssubst)+
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹Fake: the Spy doesn't have the critical session key!›
apply (blast dest: serverK_not_spied)
txt‹ClientKeyExch›
apply (blast dest!: PMS_Crypt_sessionK_not_spied)
done
text‹This version refers not to ServerFinished but to any message from B.
We don't assume B has received CertVerify, and an intruder could
have changed A's identity in all other messages, so we can't be sure
that B sends his message to A. If CLIENT KEY EXCHANGE were augmented
to bind A's identity with PMS, then we could replace A' by A below.›
lemma TrustServerMsg [rule_format]:
"⟦M = PRF(PMS,NA,NB); evs ∈ tls; A ∉ bad; B ∉ bad⟧
⟹ Says B Spy (Key(serverK(Na,Nb,M))) ∉ set evs ⟶
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs ⟶
Crypt (serverK(Na,Nb,M)) Y ∈ parts (spies evs) ⟶
(∃A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) ∈ set evs)"
apply (erule ssubst)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp) add: ex_disj_distrib)
txt‹Fake: the Spy doesn't have the critical session key!›
apply (blast dest: serverK_not_spied)
txt‹ClientKeyExch›
apply (clarify, blast dest!: PMS_Crypt_sessionK_not_spied)
txt‹ServerResume, ServerFinished: by unicity of PMS›
apply (blast dest!: Notes_master_imp_Crypt_PMS
dest: Spy_not_see_PMS Notes_Crypt_parts_spies Crypt_unique_PMS)+
done
subsubsection‹Protocol goal: if B receives any message encrypted with clientK
then A has sent it›
text‹ASSUMING that A chose PMS. Authentication is
assumed here; B cannot verify it. But if the message is
ClientFinished, then B can then check the quoted values PA, PB, etc.›
lemma TrustClientMsg [rule_format]:
"⟦M = PRF(PMS,NA,NB); evs ∈ tls; A ∉ bad; B ∉ bad⟧
⟹ Says A Spy (Key(clientK(Na,Nb,M))) ∉ set evs ⟶
Notes A ⦃Agent B, Nonce PMS⦄ ∈ set evs ⟶
Crypt (clientK(Na,Nb,M)) Y ∈ parts (spies evs) ⟶
Says A B (Crypt (clientK(Na,Nb,M)) Y) ∈ set evs"
apply (erule ssubst)
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (force, simp_all (no_asm_simp))
txt‹Fake: the Spy doesn't have the critical session key!›
apply (blast dest: clientK_not_spied)
txt‹ClientKeyExch›
apply (blast dest!: PMS_Crypt_sessionK_not_spied)
txt‹ClientFinished, ClientResume: by unicity of PMS›
apply (blast dest!: Notes_master_imp_Notes_PMS dest: Notes_unique_PMS)+
done
subsubsection‹Protocol goal: if B receives ClientFinished, and if B is able to
check a CertVerify from A, then A has used the quoted
values PA, PB, etc. Even this one requires A to be uncompromised.›
lemma AuthClientFinished:
"⟦M = PRF(PMS,NA,NB);
Says A Spy (Key(clientK(Na,Nb,M))) ∉ set evs;
Says A' B (Crypt (clientK(Na,Nb,M)) Y) ∈ set evs;
certificate A KA ∈ parts (spies evs);
Says A'' B (Crypt (invKey KA) (Hash⦃nb, Agent B, Nonce PMS⦄))
∈ set evs;
evs ∈ tls; A ∉ bad; B ∉ bad⟧
⟹ Says A B (Crypt (clientK(Na,Nb,M)) Y) ∈ set evs"
by (blast intro!: TrustClientMsg UseCertVerify)
end