header{*Theory of smartcards*}
theory Smartcard imports EventSC begin
text{*
As smartcards handle long-term (symmetric) keys, this theoy extends and
supersedes theory Private.thy
An agent is bad if she reveals her PIN to the spy, not the shared key that
is embedded in her card. An agent's being bad implies nothing about her
smartcard, which independently may be stolen or cloned.
*}
consts
shrK :: "agent => key"
crdK :: "card => key"
pin :: "agent => key"
Pairkey :: "agent * agent => nat"
pairK :: "agent * agent => key"
axioms
inj_shrK: "inj shrK" --{*No two smartcards store the same key*}
inj_crdK: "inj crdK" --{*Nor do two cards*}
inj_pin : "inj pin" --{*Nor do two agents have the same pin*}
inj_pairK [iff]: "(pairK(A,B) = pairK(A',B')) = (A = A' & B = B')"
comm_Pairkey [iff]: "Pairkey(A,B) = Pairkey(B,A)"
pairK_disj_crdK [iff]: "pairK(A,B) ≠ crdK C"
pairK_disj_shrK [iff]: "pairK(A,B) ≠ shrK P"
pairK_disj_pin [iff]: "pairK(A,B) ≠ pin P"
shrK_disj_crdK [iff]: "shrK P ≠ crdK C"
shrK_disj_pin [iff]: "shrK P ≠ pin Q"
crdK_disj_pin [iff]: "crdK C ≠ pin P"
text{*All keys are symmetric*}
defs all_symmetric_def: "all_symmetric == True"
lemma isSym_keys: "K ∈ symKeys"
by (simp add: symKeys_def all_symmetric_def invKey_symmetric)
constdefs
legalUse :: "card => bool" ("legalUse (_)")
"legalUse C == C ∉ stolen"
consts
illegalUse :: "card => bool"
primrec
illegalUse_def:
"illegalUse (Card A) = ( (Card A ∈ stolen ∧ A ∈ bad) ∨ Card A ∈ cloned )"
text{*initState must be defined with care*}
primrec
initState_Server: "initState Server =
(Key`(range shrK ∪ range crdK ∪ range pin ∪ range pairK)) ∪
(Nonce`(range Pairkey))"
initState_Friend: "initState (Friend i) = {Key (pin (Friend i))}"
initState_Spy: "initState Spy =
(Key`((pin`bad) ∪ (pin `{A. Card A ∈ cloned}) ∪
(shrK`{A. Card A ∈ cloned}) ∪
(crdK`cloned) ∪
(pairK`{(X,A). Card A ∈ cloned})))
∪ (Nonce`(Pairkey`{(A,B). Card A ∈ cloned & Card B ∈ cloned}))"
text{*Still relying on axioms*}
axioms
Key_supply_ax: "finite KK ==> ∃ K. K ∉ KK & Key K ∉ used evs"
Nonce_supply_ax: "finite NN ==> ∃ N. N ∉ NN & Nonce N ∉ used evs"
subsection{*Basic properties of shrK*}
declare inj_shrK [THEN inj_eq, iff]
declare inj_crdK [THEN inj_eq, iff]
declare inj_pin [THEN inj_eq, iff]
lemma invKey_K [simp]: "invKey K = K"
apply (insert isSym_keys)
apply (simp add: symKeys_def)
done
lemma analz_Decrypt' [dest]:
"[| Crypt K X ∈ analz H; Key K ∈ analz H |] ==> X ∈ analz H"
by auto
text{*Now cancel the @{text dest} attribute given to
@{text analz.Decrypt} in its declaration.*}
declare analz.Decrypt [rule del]
text{*Rewrites should not refer to @{term "initState(Friend i)"} because
that expression is not in normal form.*}
text{*Added to extend initstate with set of nonces*}
lemma parts_image_Nonce [simp]: "parts (Nonce`N) = Nonce`N"
apply auto
apply (erule parts.induct)
apply auto
done
lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
apply (unfold keysFor_def)
apply (induct_tac "C", auto)
done
lemma keysFor_parts_insert:
"[| K ∈ keysFor (parts (insert X G)); X ∈ synth (analz H) |]
==> K ∈ keysFor (parts (G ∪ H)) | Key K ∈ parts H";
by (force dest: EventSC.keysFor_parts_insert)
lemma Crypt_imp_keysFor: "Crypt K X ∈ H ==> K ∈ keysFor H"
by (drule Crypt_imp_invKey_keysFor, simp)
subsection{*Function "knows"*}
lemma Spy_knows_bad [intro!]: "A ∈ bad ==> Key (pin A) ∈ knows Spy evs"
apply (induct_tac "evs")
apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
done
lemma Spy_knows_cloned [intro!]:
"Card A ∈ cloned ==> Key (crdK (Card A)) ∈ knows Spy evs &
Key (shrK A) ∈ knows Spy evs &
Key (pin A) ∈ knows Spy evs &
(∀ B. Key (pairK(B,A)) ∈ knows Spy evs)"
apply (induct_tac "evs")
apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
done
lemma Spy_knows_cloned1 [intro!]: "C ∈ cloned ==> Key (crdK C) ∈ knows Spy evs"
apply (induct_tac "evs")
apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
done
lemma Spy_knows_cloned2 [intro!]: "[| Card A ∈ cloned; Card B ∈ cloned |]
==> Nonce (Pairkey(A,B))∈ knows Spy evs"
apply (induct_tac "evs")
apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
done
lemma Spy_knows_Spy_bad [intro!]: "A∈ bad ==> Key (pin A) ∈ knows Spy evs"
apply (induct_tac "evs")
apply (simp_all (no_asm_simp) add: imageI knows_Cons split add: event.split)
done
lemma Crypt_Spy_analz_bad:
"[| Crypt (pin A) X ∈ analz (knows Spy evs); A∈bad |]
==> X ∈ analz (knows Spy evs)"
apply (force dest!: analz.Decrypt)
done
lemma shrK_in_initState [iff]: "Key (shrK A) ∈ initState Server"
apply (induct_tac "A")
apply auto
done
lemma shrK_in_used [iff]: "Key (shrK A) ∈ used evs"
apply (rule initState_into_used)
apply blast
done
lemma crdK_in_initState [iff]: "Key (crdK A) ∈ initState Server"
apply (induct_tac "A")
apply auto
done
lemma crdK_in_used [iff]: "Key (crdK A) ∈ used evs"
apply (rule initState_into_used)
apply blast
done
lemma pin_in_initState [iff]: "Key (pin A) ∈ initState A"
apply (induct_tac "A")
apply auto
done
lemma pin_in_used [iff]: "Key (pin A) ∈ used evs"
apply (rule initState_into_used)
apply blast
done
lemma pairK_in_initState [iff]: "Key (pairK X) ∈ initState Server"
apply (induct_tac "X")
apply auto
done
lemma pairK_in_used [iff]: "Key (pairK X) ∈ used evs"
apply (rule initState_into_used)
apply blast
done
lemma Key_not_used [simp]: "Key K ∉ used evs ==> K ∉ range shrK"
by blast
lemma shrK_neq [simp]: "Key K ∉ used evs ==> shrK B ≠ K"
by blast
lemma crdK_not_used [simp]: "Key K ∉ used evs ==> K ∉ range crdK"
apply clarify
done
lemma crdK_neq [simp]: "Key K ∉ used evs ==> crdK C ≠ K"
apply clarify
done
lemma pin_not_used [simp]: "Key K ∉ used evs ==> K ∉ range pin"
apply clarify
done
lemma pin_neq [simp]: "Key K ∉ used evs ==> pin A ≠ K"
apply clarify
done
lemma pairK_not_used [simp]: "Key K ∉ used evs ==> K ∉ range pairK"
apply clarify
done
lemma pairK_neq [simp]: "Key K ∉ used evs ==> pairK(A,B) ≠ K"
apply clarify
done
declare shrK_neq [THEN not_sym, simp]
declare crdK_neq [THEN not_sym, simp]
declare pin_neq [THEN not_sym, simp]
declare pairK_neq [THEN not_sym, simp]
subsection{*Fresh nonces*}
lemma Nonce_notin_initState [iff]: "Nonce N ∉ parts (initState (Friend i))"
by auto
subsection{*Supply fresh nonces for possibility theorems.*}
lemma Nonce_supply1: "∃N. Nonce N ∉ used evs"
apply (rule finite.emptyI [THEN Nonce_supply_ax, THEN exE], blast)
done
lemma Nonce_supply2:
"∃N N'. Nonce N ∉ used evs & Nonce N' ∉ used evs' & N ≠ N'"
apply (cut_tac evs = evs in finite.emptyI [THEN Nonce_supply_ax])
apply (erule exE)
apply (cut_tac evs = evs' in finite.emptyI [THEN finite.insertI, THEN Nonce_supply_ax])
apply auto
done
lemma Nonce_supply3: "∃N N' N''. Nonce N ∉ used evs & Nonce N' ∉ used evs' &
Nonce N'' ∉ used evs'' & N ≠ N' & N' ≠ N'' & N ≠ N''"
apply (cut_tac evs = evs in finite.emptyI [THEN Nonce_supply_ax])
apply (erule exE)
apply (cut_tac evs = evs' and a1 = N in finite.emptyI [THEN finite.insertI, THEN Nonce_supply_ax])
apply (erule exE)
apply (cut_tac evs = evs'' and a1 = Na and a2 = N in finite.emptyI [THEN finite.insertI, THEN finite.insertI, THEN Nonce_supply_ax])
apply blast
done
lemma Nonce_supply: "Nonce (@ N. Nonce N ∉ used evs) ∉ used evs"
apply (rule finite.emptyI [THEN Nonce_supply_ax, THEN exE])
apply (rule someI, blast)
done
text{*Unlike the corresponding property of nonces, we cannot prove
@{term "finite KK ==> ∃K. K ∉ KK & Key K ∉ used evs"}.
We have infinitely many agents and there is nothing to stop their
long-term keys from exhausting all the natural numbers. Instead,
possibility theorems must assume the existence of a few keys.*}
subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
lemma subset_Compl_range_shrK: "A ⊆ - (range shrK) ==> shrK x ∉ A"
by blast
lemma subset_Compl_range_crdK: "A ⊆ - (range crdK) ==> crdK x ∉ A"
apply blast
done
lemma subset_Compl_range_pin: "A ⊆ - (range pin) ==> pin x ∉ A"
apply blast
done
lemma subset_Compl_range_pairK: "A ⊆ - (range pairK) ==> pairK x ∉ A"
apply blast
done
lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} ∪ H"
by blast
lemma insert_Key_image: "insert (Key K) (Key`KK ∪ C) = Key`(insert K KK) ∪ C"
by blast
lemmas analz_image_freshK_simps =
simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
disj_comms
image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset
analz_insert_eq Un_upper2 [THEN analz_mono, THEN [2] rev_subsetD]
insert_Key_singleton subset_Compl_range_shrK subset_Compl_range_crdK
subset_Compl_range_pin subset_Compl_range_pairK
Key_not_used insert_Key_image Un_assoc [THEN sym]
lemma analz_image_freshK_lemma:
"(Key K ∈ analz (Key`nE ∪ H)) --> (K ∈ nE | Key K ∈ analz H) ==>
(Key K ∈ analz (Key`nE ∪ H)) = (K ∈ nE | Key K ∈ analz H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
subsection{*Tactics for possibility theorems*}
ML
{*
structure Smartcard =
struct
(*Omitting used_Says makes the tactic much faster: it leaves expressions
such as Nonce ?N ∉ used evs that match Nonce_supply*)
fun possibility_tac ctxt =
(REPEAT
(ALLGOALS (simp_tac (local_simpset_of ctxt
delsimps [@{thm used_Says}, @{thm used_Notes}, @{thm used_Gets},
@{thm used_Inputs}, @{thm used_C_Gets}, @{thm used_Outpts}, @{thm used_A_Gets}]
setSolver safe_solver))
THEN
REPEAT_FIRST (eq_assume_tac ORELSE'
resolve_tac [refl, conjI, @{thm Nonce_supply}])))
(*For harder protocols (such as Recur) where we have to set up some
nonces and keys initially*)
fun basic_possibility_tac ctxt =
REPEAT
(ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver))
THEN
REPEAT_FIRST (resolve_tac [refl, conjI]))
val analz_image_freshK_ss =
@{simpset} delsimps [image_insert, image_Un]
delsimps [@{thm imp_disjL}] (*reduces blow-up*)
addsimps @{thms analz_image_freshK_simps}
end
*}
lemma invKey_shrK_iff [iff]:
"(Key (invKey K) ∈ X) = (Key K ∈ X)"
by auto
method_setup analz_freshK = {*
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
(EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
ALLGOALS (asm_simp_tac (Simplifier.context ctxt Smartcard.analz_image_freshK_ss))]))) *}
"for proving the Session Key Compromise theorem"
method_setup possibility = {*
Scan.succeed (fn ctxt =>
SIMPLE_METHOD (Smartcard.possibility_tac ctxt)) *}
"for proving possibility theorems"
method_setup basic_possibility = {*
Scan.succeed (fn ctxt =>
SIMPLE_METHOD (Smartcard.basic_possibility_tac ctxt)) *}
"for proving possibility theorems"
lemma knows_subset_knows_Cons: "knows A evs ⊆ knows A (e # evs)"
by (induct e) (auto simp: knows_Cons)
declare shrK_disj_crdK[THEN not_sym, iff]
declare shrK_disj_pin[THEN not_sym, iff]
declare pairK_disj_shrK[THEN not_sym, iff]
declare pairK_disj_crdK[THEN not_sym, iff]
declare pairK_disj_pin[THEN not_sym, iff]
declare crdK_disj_pin[THEN not_sym, iff]
declare legalUse_def [iff] illegalUse_def [iff]
end