Theory Message_SET
section‹The Message Theory, Modified for SET›
theory Message_SET
imports Main "HOL-Library.Nat_Bijection"
begin
subsection‹General Lemmas›
text‹Needed occasionally with ‹spy_analz_tac›, e.g. in
‹analz_insert_Key_newK››
lemma Un_absorb3 [simp] : "A ∪ (B ∪ A) = B ∪ A"
by blast
text‹Collapses redundant cases in the huge protocol proofs›
lemmas disj_simps = disj_comms disj_left_absorb disj_assoc
text‹Effective with assumptions like \<^term>‹K ∉ range pubK› and
\<^term>‹K ∉ invKey`range pubK››
lemma notin_image_iff: "(y ∉ f`I) = (∀i∈I. f i ≠ y)"
by blast
text‹Effective with the assumption \<^term>‹KK ⊆ - (range(invKey o pubK))››
lemma disjoint_image_iff: "(A ⊆ - (f`I)) = (∀i∈I. f i ∉ A)"
by blast
type_synonym key = nat
consts
all_symmetric :: bool
invKey :: "key⇒key"
specification (invKey)
invKey [simp]: "invKey (invKey K) = K"
invKey_symmetric: "all_symmetric ⟶ invKey = id"
by (rule exI [of _ id], auto)
text‹The inverse of a symmetric key is itself; that of a public key
is the private key and vice versa›
definition symKeys :: "key set" where
"symKeys == {K. invKey K = K}"
text‹Agents. We allow any number of certification authorities, cardholders
merchants, and payment gateways.›