(* Title: HOL/MetisTest/Message.thy ID: $Id: Message.thy,v 1.6 2008/05/07 08:59:07 berghofe Exp $ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Testing the metis method *) theory Message imports Main begin (*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*) lemma strange_Un_eq [simp]: "A ∪ (B ∪ A) = B ∪ A" by blast types key = nat consts all_symmetric :: bool --{*true if all keys are symmetric*} invKey :: "key=>key" --{*inverse of a symmetric 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*} constdefs symKeys :: "key set" "symKeys == {K. invKey K = K}" datatype --{*We allow any number of friendly agents*} agent = Server | Friend nat | Spy datatype msg = Agent agent --{*Agent names*} | Number nat --{*Ordinary integers, timestamps, ...*} | Nonce nat --{*Unguessable nonces*} | Key key --{*Crypto keys*} | Hash msg --{*Hashing*} | MPair msg msg --{*Compound messages*} | Crypt key msg --{*Encryption, public- or shared-key*} text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*} syntax "@MTuple" :: "['a, args] => 'a * 'b" ("(2{|_,/ _|})") syntax (xsymbols) "@MTuple" :: "['a, args] => 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)") translations "{|x, y, z|}" == "{|x, {|y, z|}|}" "{|x, y|}" == "MPair x y" constdefs HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) --{*Message Y paired with a MAC computed with the help of X*} "Hash[X] Y == {| Hash{|X,Y|}, Y|}" keysFor :: "msg set => key set" --{*Keys useful to decrypt elements of a message set*} "keysFor H == invKey ` {K. ∃X. Crypt K X ∈ H}" subsubsection{*Inductive Definition of All Parts" of a Message*} inductive_set parts :: "msg set => msg set" for H :: "msg set" where Inj [intro]: "X ∈ H ==> X ∈ parts H" | Fst: "{|X,Y|} ∈ parts H ==> X ∈ parts H" | Snd: "{|X,Y|} ∈ parts H ==> Y ∈ parts H" | Body: "Crypt K X ∈ parts H ==> X ∈ parts H" ML{*ResAtp.problem_name := "Message__parts_mono"*} lemma parts_mono: "G ⊆ H ==> parts(G) ⊆ parts(H)" apply auto apply (erule parts.induct) apply (metis Inj set_mp) apply (metis Fst) apply (metis Snd) apply (metis Body) done text{*Equations hold because constructors are injective.*} lemma Friend_image_eq [simp]: "(Friend x ∈ Friend`A) = (x:A)" by auto lemma Key_image_eq [simp]: "(Key x ∈ Key`A) = (x∈A)" by auto lemma Nonce_Key_image_eq [simp]: "(Nonce x ∉ Key`A)" by auto subsubsection{*Inverse of keys *} ML{*ResAtp.problem_name := "Message__invKey_eq"*} lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')" by (metis invKey) subsection{*keysFor operator*} lemma keysFor_empty [simp]: "keysFor {} = {}" by (unfold keysFor_def, blast) lemma keysFor_Un [simp]: "keysFor (H ∪ H') = keysFor H ∪ keysFor H'" by (unfold keysFor_def, blast) lemma keysFor_UN [simp]: "keysFor (\<Union>i∈A. H i) = (\<Union>i∈A. keysFor (H i))" by (unfold keysFor_def, blast) text{*Monotonicity*} lemma keysFor_mono: "G ⊆ H ==> keysFor(G) ⊆ keysFor(H)" by (unfold keysFor_def, blast) lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_Crypt [simp]: "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)" by (unfold keysFor_def, auto) lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}" by (unfold keysFor_def, auto) lemma Crypt_imp_invKey_keysFor: "Crypt K X ∈ H ==> invKey K ∈ keysFor H" by (unfold keysFor_def, blast) subsection{*Inductive relation "parts"*} lemma MPair_parts: "[| {|X,Y|} ∈ parts H; [| X ∈ parts H; Y ∈ parts H |] ==> P |] ==> P" by (blast dest: parts.Fst parts.Snd) declare MPair_parts [elim!] parts.Body [dest!] text{*NB These two rules are UNSAFE in the formal sense, as they discard the compound message. They work well on THIS FILE. @{text MPair_parts} is left as SAFE because it speeds up proofs. The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*} lemma parts_increasing: "H ⊆ parts(H)" by blast lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard] lemma parts_empty [simp]: "parts{} = {}" apply safe apply (erule parts.induct) apply blast+ done lemma parts_emptyE [elim!]: "X∈ parts{} ==> P" by simp text{*WARNING: loops if H = {Y}, therefore must not be repeated!*} lemma parts_singleton: "X∈ parts H ==> ∃Y∈H. X∈ parts {Y}" apply (erule parts.induct) apply fast+ done subsubsection{*Unions *} lemma parts_Un_subset1: "parts(G) ∪ parts(H) ⊆ parts(G ∪ H)" by (intro Un_least parts_mono Un_upper1 Un_upper2) lemma parts_Un_subset2: "parts(G ∪ H) ⊆ parts(G) ∪ parts(H)" apply (rule subsetI) apply (erule parts.induct, blast+) done lemma parts_Un [simp]: "parts(G ∪ H) = parts(G) ∪ parts(H)" by (intro equalityI parts_Un_subset1 parts_Un_subset2) lemma parts_insert: "parts (insert X H) = parts {X} ∪ parts H" apply (subst insert_is_Un [of _ H]) apply (simp only: parts_Un) done ML{*ResAtp.problem_name := "Message__parts_insert_two"*} lemma parts_insert2: "parts (insert X (insert Y H)) = parts {X} ∪ parts {Y} ∪ parts H" by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un) lemma parts_UN_subset1: "(\<Union>x∈A. parts(H x)) ⊆ parts(\<Union>x∈A. H x)" by (intro UN_least parts_mono UN_upper) lemma parts_UN_subset2: "parts(\<Union>x∈A. H x) ⊆ (\<Union>x∈A. parts(H x))" apply (rule subsetI) apply (erule parts.induct, blast+) done lemma parts_UN [simp]: "parts(\<Union>x∈A. H x) = (\<Union>x∈A. parts(H x))" by (intro equalityI parts_UN_subset1 parts_UN_subset2) text{*Added to simplify arguments to parts, analz and synth. NOTE: the UN versions are no longer used!*} text{*This allows @{text blast} to simplify occurrences of @{term "parts(G∪H)"} in the assumption.*} lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] declare in_parts_UnE [elim!] lemma parts_insert_subset: "insert X (parts H) ⊆ parts(insert X H)" by (blast intro: parts_mono [THEN [2] rev_subsetD]) subsubsection{*Idempotence and transitivity *} lemma parts_partsD [dest!]: "X∈ parts (parts H) ==> X∈ parts H" by (erule parts.induct, blast+) lemma parts_idem [simp]: "parts (parts H) = parts H" by blast ML{*ResAtp.problem_name := "Message__parts_subset_iff"*} lemma parts_subset_iff [simp]: "(parts G ⊆ parts H) = (G ⊆ parts H)" apply (rule iffI) apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing) apply (metis parts_idem parts_mono) done lemma parts_trans: "[| X∈ parts G; G ⊆ parts H |] ==> X∈ parts H" by (blast dest: parts_mono); ML{*ResAtp.problem_name := "Message__parts_cut"*} lemma parts_cut: "[|Y∈ parts(insert X G); X∈ parts H|] ==> Y∈ parts(G ∪ H)" by (metis Un_subset_iff insert_subset parts_increasing parts_trans) subsubsection{*Rewrite rules for pulling out atomic messages *} lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset] lemma parts_insert_Agent [simp]: "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)" apply (rule parts_insert_eq_I) apply (erule parts.induct, auto) done lemma parts_insert_Nonce [simp]: "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)" apply (rule parts_insert_eq_I) apply (erule parts.induct, auto) done lemma parts_insert_Number [simp]: "parts (insert (Number N) H) = insert (Number N) (parts H)" apply (rule parts_insert_eq_I) apply (erule parts.induct, auto) done lemma parts_insert_Key [simp]: "parts (insert (Key K) H) = insert (Key K) (parts H)" apply (rule parts_insert_eq_I) apply (erule parts.induct, auto) done lemma parts_insert_Hash [simp]: "parts (insert (Hash X) H) = insert (Hash X) (parts H)" apply (rule parts_insert_eq_I) apply (erule parts.induct, auto) done lemma parts_insert_Crypt [simp]: "parts (insert (Crypt K X) H) = insert (Crypt K X) (parts (insert X H))" apply (rule equalityI) apply (rule subsetI) apply (erule parts.induct, auto) apply (blast intro: parts.Body) done lemma parts_insert_MPair [simp]: "parts (insert {|X,Y|} H) = insert {|X,Y|} (parts (insert X (insert Y H)))" apply (rule equalityI) apply (rule subsetI) apply (erule parts.induct, auto) apply (blast intro: parts.Fst parts.Snd)+ done lemma parts_image_Key [simp]: "parts (Key`N) = Key`N" apply auto apply (erule parts.induct, auto) done ML{*ResAtp.problem_name := "Message__msg_Nonce_supply"*} lemma msg_Nonce_supply: "∃N. ∀n. N≤n --> Nonce n ∉ parts {msg}" apply (induct_tac "msg") apply (simp_all add: parts_insert2) apply (metis Suc_n_not_le_n) apply (metis le_trans linorder_linear) done subsection{*Inductive relation "analz"*} text{*Inductive definition of "analz" -- what can be broken down from a set of messages, including keys. A form of downward closure. Pairs can be taken apart; messages decrypted with known keys. *} inductive_set analz :: "msg set => msg set" for H :: "msg set" where Inj [intro,simp] : "X ∈ H ==> X ∈ analz H" | Fst: "{|X,Y|} ∈ analz H ==> X ∈ analz H" | Snd: "{|X,Y|} ∈ analz H ==> Y ∈ analz H" | Decrypt [dest]: "[|Crypt K X ∈ analz H; Key(invKey K): analz H|] ==> X ∈ analz H" text{*Monotonicity; Lemma 1 of Lowe's paper*} lemma analz_mono: "G⊆H ==> analz(G) ⊆ analz(H)" apply auto apply (erule analz.induct) apply (auto dest: analz.Fst analz.Snd) done text{*Making it safe speeds up proofs*} lemma MPair_analz [elim!]: "[| {|X,Y|} ∈ analz H; [| X ∈ analz H; Y ∈ analz H |] ==> P |] ==> P" by (blast dest: analz.Fst analz.Snd) lemma analz_increasing: "H ⊆ analz(H)" by blast lemma analz_subset_parts: "analz H ⊆ parts H" apply (rule subsetI) apply (erule analz.induct, blast+) done lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard] lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard] ML{*ResAtp.problem_name := "Message__parts_analz"*} lemma parts_analz [simp]: "parts (analz H) = parts H" apply (rule equalityI) apply (metis analz_subset_parts parts_subset_iff) apply (metis analz_increasing parts_mono) done lemma analz_parts [simp]: "analz (parts H) = parts H" apply auto apply (erule analz.induct, auto) done lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard] subsubsection{*General equational properties *} lemma analz_empty [simp]: "analz{} = {}" apply safe apply (erule analz.induct, blast+) done text{*Converse fails: we can analz more from the union than from the separate parts, as a key in one might decrypt a message in the other*} lemma analz_Un: "analz(G) ∪ analz(H) ⊆ analz(G ∪ H)" by (intro Un_least analz_mono Un_upper1 Un_upper2) lemma analz_insert: "insert X (analz H) ⊆ analz(insert X H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) subsubsection{*Rewrite rules for pulling out atomic messages *} lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert] lemma analz_insert_Agent [simp]: "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)" apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done lemma analz_insert_Nonce [simp]: "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)" apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done lemma analz_insert_Number [simp]: "analz (insert (Number N) H) = insert (Number N) (analz H)" apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done lemma analz_insert_Hash [simp]: "analz (insert (Hash X) H) = insert (Hash X) (analz H)" apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done text{*Can only pull out Keys if they are not needed to decrypt the rest*} lemma analz_insert_Key [simp]: "K ∉ keysFor (analz H) ==> analz (insert (Key K) H) = insert (Key K) (analz H)" apply (unfold keysFor_def) apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done lemma analz_insert_MPair [simp]: "analz (insert {|X,Y|} H) = insert {|X,Y|} (analz (insert X (insert Y H)))" apply (rule equalityI) apply (rule subsetI) apply (erule analz.induct, auto) apply (erule analz.induct) apply (blast intro: analz.Fst analz.Snd)+ done text{*Can pull out enCrypted message if the Key is not known*} lemma analz_insert_Crypt: "Key (invKey K) ∉ analz H ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done lemma lemma1: "Key (invKey K) ∈ analz H ==> analz (insert (Crypt K X) H) ⊆ insert (Crypt K X) (analz (insert X H))" apply (rule subsetI) apply (erule_tac x = x in analz.induct, auto) done lemma lemma2: "Key (invKey K) ∈ analz H ==> insert (Crypt K X) (analz (insert X H)) ⊆ analz (insert (Crypt K X) H)" apply auto apply (erule_tac x = x in analz.induct, auto) apply (blast intro: analz_insertI analz.Decrypt) done lemma analz_insert_Decrypt: "Key (invKey K) ∈ analz H ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz (insert X H))" by (intro equalityI lemma1 lemma2) text{*Case analysis: either the message is secure, or it is not! Effective, but can cause subgoals to blow up! Use with @{text "split_if"}; apparently @{text "split_tac"} does not cope with patterns such as @{term"analz (insert (Crypt K X) H)"} *} lemma analz_Crypt_if [simp]: "analz (insert (Crypt K X) H) = (if (Key (invKey K) ∈ analz H) then insert (Crypt K X) (analz (insert X H)) else insert (Crypt K X) (analz H))" by (simp add: analz_insert_Crypt analz_insert_Decrypt) text{*This rule supposes "for the sake of argument" that we have the key.*} lemma analz_insert_Crypt_subset: "analz (insert (Crypt K X) H) ⊆ insert (Crypt K X) (analz (insert X H))" apply (rule subsetI) apply (erule analz.induct, auto) done lemma analz_image_Key [simp]: "analz (Key`N) = Key`N" apply auto apply (erule analz.induct, auto) done subsubsection{*Idempotence and transitivity *} lemma analz_analzD [dest!]: "X∈ analz (analz H) ==> X∈ analz H" by (erule analz.induct, blast+) lemma analz_idem [simp]: "analz (analz H) = analz H" by blast lemma analz_subset_iff [simp]: "(analz G ⊆ analz H) = (G ⊆ analz H)" apply (rule iffI) apply (iprover intro: subset_trans analz_increasing) apply (frule analz_mono, simp) done lemma analz_trans: "[| X∈ analz G; G ⊆ analz H |] ==> X∈ analz H" by (drule analz_mono, blast) ML{*ResAtp.problem_name := "Message__analz_cut"*} declare analz_trans[intro] lemma analz_cut: "[| Y∈ analz (insert X H); X∈ analz H |] ==> Y∈ analz H" (*TOO SLOW by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset) --{*317s*} ??*) by (erule analz_trans, blast) text{*This rewrite rule helps in the simplification of messages that involve the forwarding of unknown components (X). Without it, removing occurrences of X can be very complicated. *} lemma analz_insert_eq: "X∈ analz H ==> analz (insert X H) = analz H" by (blast intro: analz_cut analz_insertI) text{*A congruence rule for "analz" *} ML{*ResAtp.problem_name := "Message__analz_subset_cong"*} lemma analz_subset_cong: "[| analz G ⊆ analz G'; analz H ⊆ analz H' |] ==> analz (G ∪ H) ⊆ analz (G' ∪ H')" apply simp apply (metis Un_absorb2 Un_commute Un_subset_iff Un_upper1 Un_upper2 analz_mono) done lemma analz_cong: "[| analz G = analz G'; analz H = analz H' |] ==> analz (G ∪ H) = analz (G' ∪ H')" by (intro equalityI analz_subset_cong, simp_all) lemma analz_insert_cong: "analz H = analz H' ==> analz(insert X H) = analz(insert X H')" by (force simp only: insert_def intro!: analz_cong) text{*If there are no pairs or encryptions then analz does nothing*} lemma analz_trivial: "[| ∀X Y. {|X,Y|} ∉ H; ∀X K. Crypt K X ∉ H |] ==> analz H = H" apply safe apply (erule analz.induct, blast+) done text{*These two are obsolete (with a single Spy) but cost little to prove...*} lemma analz_UN_analz_lemma: "X∈ analz (\<Union>i∈A. analz (H i)) ==> X∈ analz (\<Union>i∈A. H i)" apply (erule analz.induct) apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ done lemma analz_UN_analz [simp]: "analz (\<Union>i∈A. analz (H i)) = analz (\<Union>i∈A. H i)" by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) subsection{*Inductive relation "synth"*} text{*Inductive definition of "synth" -- what can be built up from a set of messages. A form of upward closure. Pairs can be built, messages encrypted with known keys. Agent names are public domain. Numbers can be guessed, but Nonces cannot be. *} inductive_set synth :: "msg set => msg set" for H :: "msg set" where Inj [intro]: "X ∈ H ==> X ∈ synth H" | Agent [intro]: "Agent agt ∈ synth H" | Number [intro]: "Number n ∈ synth H" | Hash [intro]: "X ∈ synth H ==> Hash X ∈ synth H" | MPair [intro]: "[|X ∈ synth H; Y ∈ synth H|] ==> {|X,Y|} ∈ synth H" | Crypt [intro]: "[|X ∈ synth H; Key(K) ∈ H|] ==> Crypt K X ∈ synth H" text{*Monotonicity*} lemma synth_mono: "G⊆H ==> synth(G) ⊆ synth(H)" by (auto, erule synth.induct, auto) text{*NO @{text Agent_synth}, as any Agent name can be synthesized. The same holds for @{term Number}*} inductive_cases Nonce_synth [elim!]: "Nonce n ∈ synth H" inductive_cases Key_synth [elim!]: "Key K ∈ synth H" inductive_cases Hash_synth [elim!]: "Hash X ∈ synth H" inductive_cases MPair_synth [elim!]: "{|X,Y|} ∈ synth H" inductive_cases Crypt_synth [elim!]: "Crypt K X ∈ synth H" lemma synth_increasing: "H ⊆ synth(H)" by blast subsubsection{*Unions *} text{*Converse fails: we can synth more from the union than from the separate parts, building a compound message using elements of each.*} lemma synth_Un: "synth(G) ∪ synth(H) ⊆ synth(G ∪ H)" by (intro Un_least synth_mono Un_upper1 Un_upper2) ML{*ResAtp.problem_name := "Message__synth_insert"*} lemma synth_insert: "insert X (synth H) ⊆ synth(insert X H)" by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono) subsubsection{*Idempotence and transitivity *} lemma synth_synthD [dest!]: "X∈ synth (synth H) ==> X∈ synth H" by (erule synth.induct, blast+) lemma synth_idem: "synth (synth H) = synth H" by blast lemma synth_subset_iff [simp]: "(synth G ⊆ synth H) = (G ⊆ synth H)" apply (rule iffI) apply (iprover intro: subset_trans synth_increasing) apply (frule synth_mono, simp add: synth_idem) done lemma synth_trans: "[| X∈ synth G; G ⊆ synth H |] ==> X∈ synth H" by (drule synth_mono, blast) ML{*ResAtp.problem_name := "Message__synth_cut"*} lemma synth_cut: "[| Y∈ synth (insert X H); X∈ synth H |] ==> Y∈ synth H" (*TOO SLOW by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono) *) by (erule synth_trans, blast) lemma Agent_synth [simp]: "Agent A ∈ synth H" by blast lemma Number_synth [simp]: "Number n ∈ synth H" by blast lemma Nonce_synth_eq [simp]: "(Nonce N ∈ synth H) = (Nonce N ∈ H)" by blast lemma Key_synth_eq [simp]: "(Key K ∈ synth H) = (Key K ∈ H)" by blast lemma Crypt_synth_eq [simp]: "Key K ∉ H ==> (Crypt K X ∈ synth H) = (Crypt K X ∈ H)" by blast lemma keysFor_synth [simp]: "keysFor (synth H) = keysFor H ∪ invKey`{K. Key K ∈ H}" by (unfold keysFor_def, blast) subsubsection{*Combinations of parts, analz and synth *} ML{*ResAtp.problem_name := "Message__parts_synth"*} lemma parts_synth [simp]: "parts (synth H) = parts H ∪ synth H" apply (rule equalityI) apply (rule subsetI) apply (erule parts.induct) apply (metis UnCI) apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Fst parts_increasing) apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Snd parts_increasing) apply (metis Body Crypt_synth UnCI UnE insert_absorb insert_subset parts_increasing) apply (metis Un_subset_iff parts_increasing parts_mono synth_increasing) done ML{*ResAtp.problem_name := "Message__analz_analz_Un"*} lemma analz_analz_Un [simp]: "analz (analz G ∪ H) = analz (G ∪ H)" apply (rule equalityI); apply (metis analz_idem analz_subset_cong order_eq_refl) apply (metis analz_increasing analz_subset_cong order_eq_refl) done ML{*ResAtp.problem_name := "Message__analz_synth_Un"*} declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro] lemma analz_synth_Un [simp]: "analz (synth G ∪ H) = analz (G ∪ H) ∪ synth G" apply (rule equalityI) apply (rule subsetI) apply (erule analz.induct) apply (metis UnCI UnE Un_commute analz.Inj) apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Fst analz_increasing analz_mono insert_absorb insert_subset) apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Snd analz_increasing analz_mono insert_absorb insert_subset) apply (blast intro: analz.Decrypt) apply blast done ML{*ResAtp.problem_name := "Message__analz_synth"*} lemma analz_synth [simp]: "analz (synth H) = analz H ∪ synth H" proof (neg_clausify) assume 0: "analz (synth H) ≠ analz H ∪ synth H" have 1: "!!X1 X3. sup (analz (sup X3 X1)) (synth X3) = analz (sup (synth X3) X1)" by (metis analz_synth_Un sup_set_eq sup_set_eq sup_set_eq) have 2: "sup (analz H) (synth H) ≠ analz (synth H)" by (metis 0 sup_set_eq) have 3: "!!X1 X3. sup (synth X3) (analz (sup X3 X1)) = analz (sup (synth X3) X1)" by (metis 1 Un_commute sup_set_eq sup_set_eq) have 4: "!!X3. sup (synth X3) (analz X3) = analz (sup (synth X3) {})" by (metis 3 Un_empty_right sup_set_eq) have 5: "!!X3. sup (synth X3) (analz X3) = analz (synth X3)" by (metis 4 Un_empty_right sup_set_eq) have 6: "!!X3. sup (analz X3) (synth X3) = analz (synth X3)" by (metis 5 Un_commute sup_set_eq sup_set_eq) show "False" by (metis 2 6) qed subsubsection{*For reasoning about the Fake rule in traces *} ML{*ResAtp.problem_name := "Message__parts_insert_subset_Un"*} lemma parts_insert_subset_Un: "X∈ G ==> parts(insert X H) ⊆ parts G ∪ parts H" proof (neg_clausify) assume 0: "X ∈ G" assume 1: "¬ parts (insert X H) ⊆ parts G ∪ parts H" have 2: "¬ parts (insert X H) ⊆ parts (G ∪ H)" by (metis 1 parts_Un) have 3: "¬ insert X H ⊆ G ∪ H" by (metis 2 parts_mono) have 4: "X ∉ G ∪ H ∨ ¬ H ⊆ G ∪ H" by (metis 3 insert_subset) have 5: "X ∉ G ∪ H" by (metis 4 Un_upper2) have 6: "X ∉ G" by (metis 5 UnCI) show "False" by (metis 6 0) qed ML{*ResAtp.problem_name := "Message__Fake_parts_insert"*} lemma Fake_parts_insert: "X ∈ synth (analz H) ==> parts (insert X H) ⊆ synth (analz H) ∪ parts H" proof (neg_clausify) assume 0: "X ∈ synth (analz H)" assume 1: "¬ parts (insert X H) ⊆ synth (analz H) ∪ parts H" have 2: "!!X3. parts X3 ∪ synth (analz X3) = parts (synth (analz X3))" by (metis parts_synth parts_analz) have 3: "!!X3. analz X3 ∪ synth (analz X3) = analz (synth (analz X3))" by (metis analz_synth analz_idem) have 4: "!!X3. analz X3 ⊆ analz (synth X3)" by (metis Un_upper1 analz_synth) have 5: "¬ parts (insert X H) ⊆ parts H ∪ synth (analz H)" by (metis 1 Un_commute) have 6: "¬ parts (insert X H) ⊆ parts (synth (analz H))" by (metis 5 2) have 7: "¬ insert X H ⊆ synth (analz H)" by (metis 6 parts_mono) have 8: "X ∉ synth (analz H) ∨ ¬ H ⊆ synth (analz H)" by (metis 7 insert_subset) have 9: "¬ H ⊆ synth (analz H)" by (metis 8 0) have 10: "!!X3. X3 ⊆ analz (synth X3)" by (metis analz_subset_iff 4) have 11: "!!X3. X3 ⊆ analz (synth (analz X3))" by (metis analz_subset_iff 10) have 12: "!!X3. analz (synth (analz X3)) = synth (analz X3) ∨ ¬ analz X3 ⊆ synth (analz X3)" by (metis Un_absorb1 3) have 13: "!!X3. analz (synth (analz X3)) = synth (analz X3)" by (metis 12 synth_increasing) have 14: "!!X3. X3 ⊆ synth (analz X3)" by (metis 11 13) show "False" by (metis 9 14) qed lemma Fake_parts_insert_in_Un: "[|Z ∈ parts (insert X H); X: synth (analz H)|] ==> Z ∈ synth (analz H) ∪ parts H"; by (blast dest: Fake_parts_insert [THEN subsetD, dest]) ML{*ResAtp.problem_name := "Message__Fake_analz_insert"*} declare analz_mono [intro] synth_mono [intro] lemma Fake_analz_insert: "X∈ synth (analz G) ==> analz (insert X H) ⊆ synth (analz G) ∪ analz (G ∪ H)" by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un analz_mono analz_synth_Un equalityE insert_absorb order_le_less xt1(12)) ML{*ResAtp.problem_name := "Message__Fake_analz_insert_simpler"*} (*simpler problems? BUT METIS CAN'T PROVE lemma Fake_analz_insert_simpler: "X∈ synth (analz G) ==> analz (insert X H) ⊆ synth (analz G) ∪ analz (G ∪ H)" apply (rule subsetI) apply (subgoal_tac "x ∈ analz (synth (analz G) ∪ H) ") apply (metis Un_commute analz_analz_Un analz_synth_Un) apply (metis Un_commute Un_upper1 Un_upper2 analz_cut analz_increasing analz_mono insert_absorb insert_mono insert_subset) done *) end
lemma strange_Un_eq:
A ∪ (B ∪ A) = B ∪ A
lemma parts_mono:
G ⊆ H ==> parts G ⊆ parts H
lemma Friend_image_eq:
(Friend x ∈ Friend ` A) = (x ∈ A)
lemma Key_image_eq:
(Key x ∈ Key ` A) = (x ∈ A)
lemma Nonce_Key_image_eq:
Nonce x ∉ Key ` A
lemma invKey_eq:
(invKey K = invKey K') = (K = K')
lemma keysFor_empty:
keysFor {} = {}
lemma keysFor_Un:
keysFor (H ∪ H') = keysFor H ∪ keysFor H'
lemma keysFor_UN:
keysFor (UN i:A. H i) = (UN i:A. keysFor (H i))
lemma keysFor_mono:
G ⊆ H ==> keysFor G ⊆ keysFor H
lemma keysFor_insert_Agent:
keysFor (insert (Agent A) H) = keysFor H
lemma keysFor_insert_Nonce:
keysFor (insert (Nonce N) H) = keysFor H
lemma keysFor_insert_Number:
keysFor (insert (Number N) H) = keysFor H
lemma keysFor_insert_Key:
keysFor (insert (Key K) H) = keysFor H
lemma keysFor_insert_Hash:
keysFor (insert (Hash X) H) = keysFor H
lemma keysFor_insert_MPair:
keysFor (insert {|X, Y|} H) = keysFor H
lemma keysFor_insert_Crypt:
keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)
lemma keysFor_image_Key:
keysFor (Key ` E) = {}
lemma Crypt_imp_invKey_keysFor:
Crypt K X ∈ H ==> invKey K ∈ keysFor H
lemma MPair_parts:
[| {|X, Y|} ∈ parts H; [| X ∈ parts H; Y ∈ parts H |] ==> P |] ==> P
lemma parts_increasing:
H ⊆ parts H
lemma parts_insertI:
c ∈ parts G ==> c ∈ parts (insert a G)
lemma parts_empty:
parts {} = {}
lemma parts_emptyE:
X ∈ parts {} ==> P
lemma parts_singleton:
X ∈ parts H ==> ∃Y∈H. X ∈ parts {Y}
lemma parts_Un_subset1:
parts G ∪ parts H ⊆ parts (G ∪ H)
lemma parts_Un_subset2:
parts (G ∪ H) ⊆ parts G ∪ parts H
lemma parts_Un:
parts (G ∪ H) = parts G ∪ parts H
lemma parts_insert:
parts (insert X H) = parts {X} ∪ parts H
lemma parts_insert2:
parts (insert X (insert Y H)) = parts {X} ∪ parts {Y} ∪ parts H
lemma parts_UN_subset1:
(UN x:A. parts (H x)) ⊆ parts (UN x:A. H x)
lemma parts_UN_subset2:
parts (UN x:A. H x) ⊆ (UN x:A. parts (H x))
lemma parts_UN:
parts (UN x:A. H x) = (UN x:A. parts (H x))
lemma in_parts_UnE:
[| c ∈ parts (G3 ∪ H3); c ∈ parts G3 ==> P; c ∈ parts H3 ==> P |] ==> P
lemma parts_insert_subset:
insert X (parts H) ⊆ parts (insert X H)
lemma parts_partsD:
X ∈ parts (parts H) ==> X ∈ parts H
lemma parts_idem:
parts (parts H) = parts H
lemma parts_subset_iff:
(parts G ⊆ parts H) = (G ⊆ parts H)
lemma parts_trans:
[| X ∈ parts G; G ⊆ parts H |] ==> X ∈ parts H
lemma parts_cut:
[| Y ∈ parts (insert X G); X ∈ parts H |] ==> Y ∈ parts (G ∪ H)
lemma parts_insert_eq_I:
(!!x. x ∈ parts (insert X1 H1) ==> x ∈ insert X1 (parts H1))
==> parts (insert X1 H1) = insert X1 (parts H1)
lemma parts_insert_Agent:
parts (insert (Agent agt) H) = insert (Agent agt) (parts H)
lemma parts_insert_Nonce:
parts (insert (Nonce N) H) = insert (Nonce N) (parts H)
lemma parts_insert_Number:
parts (insert (Number N) H) = insert (Number N) (parts H)
lemma parts_insert_Key:
parts (insert (Key K) H) = insert (Key K) (parts H)
lemma parts_insert_Hash:
parts (insert (Hash X) H) = insert (Hash X) (parts H)
lemma parts_insert_Crypt:
parts (insert (Crypt K X) H) = insert (Crypt K X) (parts (insert X H))
lemma parts_insert_MPair:
parts (insert {|X, Y|} H) = insert {|X, Y|} (parts (insert X (insert Y H)))
lemma parts_image_Key:
parts (Key ` N) = Key ` N
lemma msg_Nonce_supply:
∃N. ∀n≥N. Nonce n ∉ parts {msg}
lemma analz_mono:
G ⊆ H ==> analz G ⊆ analz H
lemma MPair_analz:
[| {|X, Y|} ∈ analz H; [| X ∈ analz H; Y ∈ analz H |] ==> P |] ==> P
lemma analz_increasing:
H ⊆ analz H
lemma analz_subset_parts:
analz H ⊆ parts H
lemma analz_into_parts:
c ∈ analz H ==> c ∈ parts H
lemma not_parts_not_analz:
c ∉ parts H ==> c ∉ analz H
lemma parts_analz:
parts (analz H) = parts H
lemma analz_parts:
analz (parts H) = parts H
lemma analz_insertI:
c ∈ analz G ==> c ∈ analz (insert a G)
lemma analz_empty:
analz {} = {}
lemma analz_Un:
analz G ∪ analz H ⊆ analz (G ∪ H)
lemma analz_insert:
insert X (analz H) ⊆ analz (insert X H)
lemma analz_insert_eq_I:
(!!x. x ∈ analz (insert X1 H1) ==> x ∈ insert X1 (analz H1))
==> analz (insert X1 H1) = insert X1 (analz H1)
lemma analz_insert_Agent:
analz (insert (Agent agt) H) = insert (Agent agt) (analz H)
lemma analz_insert_Nonce:
analz (insert (Nonce N) H) = insert (Nonce N) (analz H)
lemma analz_insert_Number:
analz (insert (Number N) H) = insert (Number N) (analz H)
lemma analz_insert_Hash:
analz (insert (Hash X) H) = insert (Hash X) (analz H)
lemma analz_insert_Key:
K ∉ keysFor (analz H) ==> analz (insert (Key K) H) = insert (Key K) (analz H)
lemma analz_insert_MPair:
analz (insert {|X, Y|} H) = insert {|X, Y|} (analz (insert X (insert Y H)))
lemma analz_insert_Crypt:
Key (invKey K) ∉ analz H
==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)
lemma lemma1:
Key (invKey K) ∈ analz H
==> analz (insert (Crypt K X) H) ⊆ insert (Crypt K X) (analz (insert X H))
lemma lemma2:
Key (invKey K) ∈ analz H
==> insert (Crypt K X) (analz (insert X H)) ⊆ analz (insert (Crypt K X) H)
lemma analz_insert_Decrypt:
Key (invKey K) ∈ analz H
==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz (insert X H))
lemma analz_Crypt_if:
analz (insert (Crypt K X) H) =
(if Key (invKey K) ∈ analz H then insert (Crypt K X) (analz (insert X H))
else insert (Crypt K X) (analz H))
lemma analz_insert_Crypt_subset:
analz (insert (Crypt K X) H) ⊆ insert (Crypt K X) (analz (insert X H))
lemma analz_image_Key:
analz (Key ` N) = Key ` N
lemma analz_analzD:
X ∈ analz (analz H) ==> X ∈ analz H
lemma analz_idem:
analz (analz H) = analz H
lemma analz_subset_iff:
(analz G ⊆ analz H) = (G ⊆ analz H)
lemma analz_trans:
[| X ∈ analz G; G ⊆ analz H |] ==> X ∈ analz H
lemma analz_cut:
[| Y ∈ analz (insert X H); X ∈ analz H |] ==> Y ∈ analz H
lemma analz_insert_eq:
X ∈ analz H ==> analz (insert X H) = analz H
lemma analz_subset_cong:
[| analz G ⊆ analz G'; analz H ⊆ analz H' |] ==> analz (G ∪ H) ⊆ analz (G' ∪ H')
lemma analz_cong:
[| analz G = analz G'; analz H = analz H' |] ==> analz (G ∪ H) = analz (G' ∪ H')
lemma analz_insert_cong:
analz H = analz H' ==> analz (insert X H) = analz (insert X H')
lemma analz_trivial:
[| ∀X Y. {|X, Y|} ∉ H; ∀X K. Crypt K X ∉ H |] ==> analz H = H
lemma analz_UN_analz_lemma:
X ∈ analz (UN i:A. analz (H i)) ==> X ∈ analz (UN i:A. H i)
lemma analz_UN_analz:
analz (UN i:A. analz (H i)) = analz (UN i:A. H i)
lemma synth_mono:
G ⊆ H ==> synth G ⊆ synth H
lemma synth_increasing:
H ⊆ synth H
lemma synth_Un:
synth G ∪ synth H ⊆ synth (G ∪ H)
lemma synth_insert:
insert X (synth H) ⊆ synth (insert X H)
lemma synth_synthD:
X ∈ synth (synth H) ==> X ∈ synth H
lemma synth_idem:
synth (synth H) = synth H
lemma synth_subset_iff:
(synth G ⊆ synth H) = (G ⊆ synth H)
lemma synth_trans:
[| X ∈ synth G; G ⊆ synth H |] ==> X ∈ synth H
lemma synth_cut:
[| Y ∈ synth (insert X H); X ∈ synth H |] ==> Y ∈ synth H
lemma Agent_synth:
Agent A ∈ synth H
lemma Number_synth:
Number n ∈ synth H
lemma Nonce_synth_eq:
(Nonce N ∈ synth H) = (Nonce N ∈ H)
lemma Key_synth_eq:
(Key K ∈ synth H) = (Key K ∈ H)
lemma Crypt_synth_eq:
Key K ∉ H ==> (Crypt K X ∈ synth H) = (Crypt K X ∈ H)
lemma keysFor_synth:
keysFor (synth H) = keysFor H ∪ invKey ` {K. Key K ∈ H}
lemma parts_synth:
parts (synth H) = parts H ∪ synth H
lemma analz_analz_Un:
analz (analz G ∪ H) = analz (G ∪ H)
lemma analz_synth_Un:
analz (synth G ∪ H) = analz (G ∪ H) ∪ synth G
lemma analz_synth:
analz (synth H) = analz H ∪ synth H
lemma parts_insert_subset_Un:
X ∈ G ==> parts (insert X H) ⊆ parts G ∪ parts H
lemma Fake_parts_insert:
X ∈ synth (analz H) ==> parts (insert X H) ⊆ synth (analz H) ∪ parts H
lemma Fake_parts_insert_in_Un:
[| Z ∈ parts (insert X H); X ∈ synth (analz H) |]
==> Z ∈ synth (analz H) ∪ parts H
lemma Fake_analz_insert:
X ∈ synth (analz G) ==> analz (insert X H) ⊆ synth (analz G) ∪ analz (G ∪ H)