# HG changeset patch # User paulson # Date 1665673222 -3600 # Node ID a6cc15ec45b2f7a70bdef49e16bfb93bdf4f1586 # Parent b82ac7ef65ecfa97c2b73727cf73a2ea5376dd4c Mostly trivial simplifications diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/Guard/Extensions.thy --- a/src/HOL/Auth/Guard/Extensions.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/Guard/Extensions.thy Thu Oct 13 16:00:22 2022 +0100 @@ -306,7 +306,7 @@ "one_step p == \evs ev. ev#evs \ p \ evs \ p" lemma one_step_Cons [dest]: "\one_step p; ev#evs \ p\ \ evs \ p" - unfolding one_step_def by (blast) + unfolding one_step_def by blast lemma one_step_app: "\evs@evs' \ p; one_step p; [] \ p\ \ evs' \ p" by (induct evs, auto) @@ -320,7 +320,7 @@ lemma has_only_SaysD: "\ev \ set evs; evs \ p; has_only_Says p\ \ \A B X. ev = Says A B X" - unfolding has_only_Says_def by (blast) + unfolding has_only_Says_def by blast lemma in_has_only_Says [dest]: "\has_only_Says p; evs \ p; ev \ set evs\ \ \A B X. ev = Says A B X" diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/Guard/Proto.thy --- a/src/HOL/Auth/Guard/Proto.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/Guard/Proto.thy Thu Oct 13 16:00:22 2022 +0100 @@ -164,12 +164,12 @@ lemma freshD: "fresh p R s n Ks evs \ (\evs1 evs2. evs = evs2 @ ap' s R # evs1 \ Nonce n \ used evs1 \ R \ p \ ok evs1 R s \ Nonce n \ parts {apm' s R} \ apm' s R \ guard n Ks)" - unfolding fresh_def by (blast) + unfolding fresh_def by blast lemma freshI [intro]: "\Nonce n \ used evs1; R \ p; Nonce n \ parts {apm' s R}; ok evs1 R s; apm' s R \ guard n Ks\ \ fresh p R s n Ks (list @ ap' s R # evs1)" - unfolding fresh_def by (blast) + unfolding fresh_def by blast lemma freshI': "\Nonce n \ used evs1; (l,r) \ p; Nonce n \ parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r) \ guard n Ks\ @@ -220,10 +220,10 @@ "safe Ks G \ \K. K \ Ks \ Key K \ analz G" lemma safeD [dest]: "\safe Ks G; K \ Ks\ \ Key K \ analz G" - unfolding safe_def by (blast) + unfolding safe_def by blast lemma safe_insert: "safe Ks (insert X G) \ safe Ks G" - unfolding safe_def by (blast) + unfolding safe_def by blast lemma Guard_safe: "\Guard n Ks G; safe Ks G\ \ Nonce n \ analz G" by (blast dest: Guard_invKey) @@ -238,7 +238,7 @@ lemma preservD: "\preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); safe Ks (spies evs); fresh p R' s' n Ks evs; R \ p; ok evs R s; keys R' s' n evs \ Ks\ \ apm' s R \ guard n Ks" - unfolding preserv_def by (blast) + unfolding preserv_def by blast lemma preservD': "\preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); safe Ks (spies evs); fresh p R' s' n Ks evs; (l,Says A B X) \ p; @@ -253,7 +253,7 @@ lemma monotonD [dest]: "\keys R' s' n (ev # evs) \ Ks; monoton p keys; ev # evs \ tr p\ \ keys R' s' n evs \ Ks" - unfolding monoton_def by (blast) + unfolding monoton_def by blast subsection\guardedness theorem\ @@ -328,13 +328,13 @@ secret R n s Ks \ parts (spies evs); secret R' n' s' Ks \ parts (spies evs); apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks\ \ secret R n s Ks = secret R' n' s' Ks" - unfolding uniq_def by (blast) + unfolding uniq_def by blast definition ord :: "proto \ (rule \ rule \ bool) \ bool" where "ord p inff \ \R R'. R \ p \ R' \ p \ \ inff R R' \ inff R' R" lemma ordD: "\ord p inff; \ inff R R'; R \ p; R' \ p\ \ inff R' R" - unfolding ord_def by (blast) + unfolding ord_def by blast definition uniq' :: "proto \ (rule \ rule \ bool) \ secfun \ bool" where "uniq' p inff secret \ \evs R R' n n' Ks s s'. R \ p \ R' \ p \ diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/KerberosIV.thy --- a/src/HOL/Auth/KerberosIV.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/KerberosIV.thy Thu Oct 13 16:00:22 2022 +0100 @@ -304,27 +304,27 @@ ev \ Says Kas A (Crypt (shrK A) \akey, Agent Peer, Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, akey, Ta\)\)) \ authKeys (ev # evs) = authKeys evs" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeys_insert: "authKeys (Says Kas A (Crypt (shrK A) \Key K, Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K, Number Ta\)\) # evs) = insert K (authKeys evs)" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeys_simp: "K \ authKeys (Says Kas A (Crypt (shrK A) \Key K', Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K', Number Ta\)\) # evs) \ K = K' | K \ authKeys evs" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeysI: "Says Kas A (Crypt (shrK A) \Key K, Agent Tgs, Number Ta, (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key K, Number Ta\)\) \ set evs \ K \ authKeys evs" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeys_used: "K \ authKeys evs \ Key K \ used evs" by (simp add: authKeys_def, blast) @@ -1044,7 +1044,7 @@ (with respect to a given trace). *) lemma Serv_fresh_not_AKcryptSK: "Key servK \ used evs \ \ AKcryptSK authK servK evs" - unfolding AKcryptSK_def by (blast) + unfolding AKcryptSK_def by blast lemma authK_not_AKcryptSK: "\ Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, tk\ diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/KerberosIV_Gets.thy --- a/src/HOL/Auth/KerberosIV_Gets.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/KerberosIV_Gets.thy Thu Oct 13 16:00:22 2022 +0100 @@ -270,27 +270,27 @@ ev \ Says Kas A (Crypt (shrK A) \akey, Agent Peer, Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, akey, Ta\)\)) \ authKeys (ev # evs) = authKeys evs" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeys_insert: "authKeys (Says Kas A (Crypt (shrK A) \Key K, Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K, Number Ta\)\) # evs) = insert K (authKeys evs)" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeys_simp: "K \ authKeys (Says Kas A (Crypt (shrK A) \Key K', Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key K', Number Ta\)\) # evs) \ K = K' | K \ authKeys evs" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeysI: "Says Kas A (Crypt (shrK A) \Key K, Agent Tgs, Number Ta, (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key K, Number Ta\)\) \ set evs \ K \ authKeys evs" - unfolding authKeys_def by (auto) + unfolding authKeys_def by auto lemma authKeys_used: "K \ authKeys evs \ Key K \ used evs" by (simp add: authKeys_def, blast) @@ -905,7 +905,7 @@ (with respect to a given trace). *) lemma Serv_fresh_not_AKcryptSK: "Key servK \ used evs \ \ AKcryptSK authK servK evs" - unfolding AKcryptSK_def by (blast) + unfolding AKcryptSK_def by blast lemma authK_not_AKcryptSK: "\ Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, tk\ diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/Message.thy Thu Oct 13 16:00:22 2022 +0100 @@ -78,7 +78,7 @@ text\Monotonicity\ -lemma parts_mono_aux: "\G \ H; x \ parts G\ \ x \ parts H" +lemma parts_mono_aux: "\G \ H; X \ parts G\ \ X \ parts H" by (erule parts.induct) (auto dest: parts.Fst parts.Snd parts.Body) lemma parts_mono: "G \ H \ parts(G) \ parts(H)" @@ -105,53 +105,53 @@ subsection\keysFor operator\ lemma keysFor_empty [simp]: "keysFor {} = {}" - unfolding keysFor_def by (blast) + unfolding keysFor_def by blast lemma keysFor_Un [simp]: "keysFor (H \ H') = keysFor H \ keysFor H'" - unfolding keysFor_def by (blast) + unfolding keysFor_def by blast lemma keysFor_UN [simp]: "keysFor (\i\A. H i) = (\i\A. keysFor (H i))" - unfolding keysFor_def by (blast) + unfolding keysFor_def by blast text\Monotonicity\ lemma keysFor_mono: "G \ H \ keysFor(G) \ keysFor(H)" - unfolding keysFor_def by (blast) + unfolding keysFor_def by blast lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_insert_MPair [simp]: "keysFor (insert \X,Y\ H) = keysFor H" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_insert_Crypt [simp]: "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}" - unfolding keysFor_def by (auto) + unfolding keysFor_def by auto lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H \ invKey K \ keysFor H" - unfolding keysFor_def by (blast) + unfolding keysFor_def by blast subsection\Inductive relation "parts"\ lemma MPair_parts: - "\\X,Y\ \ parts H; + "\\X,Y\ \ parts H; \X \ parts H; Y \ parts H\ \ P\ \ P" -by (blast dest: parts.Fst parts.Snd) + 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 @@ -160,52 +160,53 @@ The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\ lemma parts_increasing: "H \ parts(H)" -by blast + by blast lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD] +lemma parts_empty_aux: "X \ parts{} \ False" + by (induction rule: parts.induct) (blast+) + lemma parts_empty [simp]: "parts{} = {}" -apply safe -apply (erule parts.induct, blast+) -done + using parts_empty_aux by blast lemma parts_emptyE [elim!]: "X\ parts{} \ P" -by simp + by simp text\WARNING: loops if H = {Y}, therefore must not be repeated!\ -lemma parts_singleton: "X\ parts H \ \Y\H. X\ parts {Y}" -by (erule parts.induct, fast+) +lemma parts_singleton: "X \ parts H \ \Y\H. X \ parts {Y}" + by (erule parts.induct, fast+) subsubsection\Unions\ lemma parts_Un_subset1: "parts(G) \ parts(H) \ parts(G \ H)" -by (intro Un_least parts_mono Un_upper1 Un_upper2) + 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 + 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) + by (intro equalityI parts_Un_subset1 parts_Un_subset2) lemma parts_insert: "parts (insert X H) = parts {X} \ parts H" -by (metis insert_is_Un parts_Un) + by (metis insert_is_Un parts_Un) text\TWO inserts to avoid looping. This rewrite is better than nothing. But its behaviour can be strange.\ lemma parts_insert2: - "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" -by (metis Un_commute Un_empty_right Un_insert_right insert_is_Un parts_Un) + "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" + by (metis Un_commute Un_empty_right Un_insert_right insert_is_Un parts_Un) lemma parts_UN_subset1: "(\x\A. parts(H x)) \ parts(\x\A. H x)" -by (intro UN_least parts_mono UN_upper) + by (intro UN_least parts_mono UN_upper) lemma parts_UN_subset2: "parts(\x\A. H x) \ (\x\A. parts(H x))" -apply (rule subsetI) -apply (erule parts.induct, blast+) -done + apply (rule subsetI) + apply (erule parts.induct, blast+) + done lemma parts_UN [simp]: "parts (\x\A. H x) = (\x\A. parts (H x))" @@ -214,7 +215,7 @@ lemma parts_image [simp]: "parts (f ` A) = (\x\A. parts {f x})" apply auto - apply (metis (mono_tags, opaque_lifting) image_iff parts_singleton) + apply (metis (mono_tags, opaque_lifting) image_iff parts_singleton) apply (metis empty_subsetI image_eqI insert_absorb insert_subset parts_mono) done @@ -229,29 +230,29 @@ lemma parts_insert_subset: "insert X (parts H) \ parts(insert X H)" -by (blast intro: parts_mono [THEN [2] rev_subsetD]) + 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+) + by (erule parts.induct, blast+) lemma parts_idem [simp]: "parts (parts H) = parts H" -by blast + by blast lemma parts_subset_iff [simp]: "(parts G \ parts H) = (G \ parts H)" -by (metis parts_idem parts_increasing parts_mono subset_trans) + by (metis parts_idem parts_increasing parts_mono subset_trans) lemma parts_trans: "\X\ parts G; G \ parts H\ \ X\ parts H" -by (metis parts_subset_iff subsetD) + by (metis parts_subset_iff subsetD) text\Cut\ lemma parts_cut: - "\Y\ parts (insert X G); X\ parts H\ \ Y\ parts (G \ H)" -by (blast intro: parts_trans) + "\Y\ parts (insert X G); X\ parts H\ \ Y\ parts (G \ H)" + by (blast intro: parts_trans) lemma parts_cut_eq [simp]: "X\ parts H \ parts (insert X H) = parts H" -by (metis insert_absorb parts_idem parts_insert) + by (metis insert_absorb parts_idem parts_insert) subsubsection\Rewrite rules for pulling out atomic messages\ @@ -260,65 +261,65 @@ 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 + "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 + "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 + "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 + "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 + "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 + "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) = + "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 + 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" -by auto + by auto text\In any message, there is an upper bound N on its greatest nonce.\ lemma msg_Nonce_supply: "\N. \n. N\n \ Nonce n \ parts {msg}" proof (induct msg) case (Nonce n) - show ?case - by simp (metis Suc_n_not_le_n) + show ?case + by simp (metis Suc_n_not_le_n) next case (MPair X Y) - then show ?case \ \metis works out the necessary sum itself!\ - by (simp add: parts_insert2) (metis le_trans nat_le_linear) + then show ?case \ \metis works out the necessary sum itself!\ + by (simp add: parts_insert2) (metis le_trans nat_le_linear) qed auto subsection\Inductive relation "analz"\ @@ -335,30 +336,30 @@ | 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" + "\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 + 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,Y\ \ analz H; \X \ analz H; Y \ analz H\ \ P \ \ P" -by (blast dest: analz.Fst analz.Snd) + by (blast dest: analz.Fst analz.Snd) lemma analz_increasing: "H \ analz(H)" -by blast + by blast lemma analz_subset_parts: "analz H \ parts H" -apply (rule subsetI) -apply (erule analz.induct, blast+) -done + apply (rule subsetI) + apply (erule analz.induct, blast+) + done lemmas analz_into_parts = analz_subset_parts [THEN subsetD] @@ -366,151 +367,151 @@ lemma parts_analz [simp]: "parts (analz H) = parts H" -by (metis analz_increasing analz_subset_parts equalityI parts_mono parts_subset_iff) + by (metis analz_increasing analz_subset_parts equalityI parts_mono parts_subset_iff) lemma analz_parts [simp]: "analz (parts H) = parts H" -apply auto -apply (erule analz.induct, auto) -done + apply auto + apply (erule analz.induct, auto) + done lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD] subsubsection\General equational properties\ lemma analz_empty [simp]: "analz{} = {}" -apply safe -apply (erule analz.induct, blast+) -done + 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) + 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]) + 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 + "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 + "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 + "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 + "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) \ + "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 + 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) = + "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 + 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 + "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) + apply (rule analz_insert_eq_I) + apply (erule analz.induct, auto) -done + 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 + 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 + 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 \ + "Key (invKey K) \ analz H \ analz (insert (Crypt K X) H) = insert (Crypt K X) (analz (insert X H))" -by (intro equalityI lemma1 lemma2) + 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 \if_split\; apparently \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) = + "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) + 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) \ + "analz (insert (Crypt K X) H) \ insert (Crypt K X) (analz (insert X H))" -apply (rule subsetI) -apply (erule analz.induct, auto) -done + 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 + 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+) + by (erule analz.induct, blast+) lemma analz_idem [simp]: "analz (analz H) = analz H" -by blast + by blast lemma analz_subset_iff [simp]: "(analz G \ analz H) = (G \ analz H)" -by (metis analz_idem analz_increasing analz_mono subset_trans) + by (metis analz_idem analz_increasing analz_mono subset_trans) lemma analz_trans: "\X\ analz G; G \ analz H\ \ X\ analz H" -by (drule analz_mono, blast) + by (drule analz_mono, blast) text\Cut; Lemma 2 of Lowe\ lemma analz_cut: "\Y\ analz (insert X H); X\ analz H\ \ Y\ analz H" -by (erule analz_trans, blast) + by (erule analz_trans, blast) (*Cut can be proved easily by induction on "Y: analz (insert X H) \ X: analz H \ Y: analz H" @@ -520,41 +521,41 @@ 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 (metis analz_cut analz_insert_eq_I insert_absorb) + by (metis analz_cut analz_insert_eq_I insert_absorb) text\A congruence rule for "analz"\ lemma analz_subset_cong: - "\analz G \ analz G'; analz H \ analz H'\ + "\analz G \ analz G'; analz H \ analz H'\ \ analz (G \ H) \ analz (G' \ H')" -by (metis Un_mono analz_Un analz_subset_iff subset_trans) + by (metis Un_mono analz_Un analz_subset_iff subset_trans) lemma analz_cong: - "\analz G = analz G'; analz H = analz H'\ + "\analz G = analz G'; analz H = analz H'\ \ analz (G \ H) = analz (G' \ H')" -by (intro equalityI analz_subset_cong, simp_all) + 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) + "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 + "\\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 (\i\A. analz (H i)) \ X\ analz (\i\A. H i)" -apply (erule analz.induct) -apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ -done + "X\ analz (\i\A. analz (H i)) \ X\ analz (\i\A. H i)" + apply (erule analz.induct) + apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ + done lemma analz_UN_analz [simp]: "analz (\i\A. analz (H i)) = analz (\i\A. H i)" -by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) + by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) subsection\Inductive relation "synth"\ @@ -583,132 +584,127 @@ The same holds for \<^term>\Number\\ inductive_simps synth_simps [iff]: - "Nonce n \ synth H" - "Key K \ synth H" - "Hash X \ synth H" - "\X,Y\ \ synth H" - "Crypt K X \ synth H" + "Nonce n \ synth H" + "Key K \ synth H" + "Hash X \ synth H" + "\X,Y\ \ synth H" + "Crypt K X \ synth H" lemma synth_increasing: "H \ synth(H)" -by blast + 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) + by (intro Un_least synth_mono Un_upper1 Un_upper2) lemma synth_insert: "insert X (synth H) \ synth(insert X H)" -by (blast intro: synth_mono [THEN [2] rev_subsetD]) + by (blast intro: synth_mono [THEN [2] rev_subsetD]) subsubsection\Idempotence and transitivity\ lemma synth_synthD [dest!]: "X\ synth (synth H) \ X\ synth H" -by (erule synth.induct, auto) + by (erule synth.induct, auto) lemma synth_idem: "synth (synth H) = synth H" -by blast + by blast lemma synth_subset_iff [simp]: "(synth G \ synth H) = (G \ synth H)" -by (metis subset_trans synth_idem synth_increasing synth_mono) + by (metis subset_trans synth_idem synth_increasing synth_mono) lemma synth_trans: "\X\ synth G; G \ synth H\ \ X\ synth H" -by (drule synth_mono, blast) + by (drule synth_mono, blast) text\Cut; Lemma 2 of Lowe\ lemma synth_cut: "\Y\ synth (insert X H); X\ synth H\ \ Y\ synth H" -by (erule synth_trans, blast) + by (erule synth_trans, blast) lemma Crypt_synth_eq [simp]: - "Key K \ H \ (Crypt K X \ synth H) = (Crypt K X \ H)" -by blast + "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}" - unfolding keysFor_def by (blast) + "keysFor (synth H) = keysFor H \ invKey`{K. Key K \ H}" + unfolding keysFor_def by blast subsubsection\Combinations of parts, analz and synth\ lemma parts_synth [simp]: "parts (synth H) = parts H \ synth H" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct) -apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD] - parts.Fst parts.Snd parts.Body)+ -done + apply (rule equalityI) + apply (rule subsetI) + apply (erule parts.induct) + apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD] + parts.Fst parts.Snd parts.Body)+ + done lemma analz_analz_Un [simp]: "analz (analz G \ H) = analz (G \ H)" -apply (intro equalityI analz_subset_cong)+ -apply simp_all -done + apply (intro equalityI analz_subset_cong)+ + apply simp_all + done lemma analz_synth_Un [simp]: "analz (synth G \ H) = analz (G \ H) \ synth G" -apply (rule equalityI) -apply (rule subsetI) -apply (erule analz.induct) -prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD]) -apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+ -done + apply (rule equalityI) + apply (rule subsetI) + apply (erule analz.induct) + prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD]) + apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+ + done lemma analz_synth [simp]: "analz (synth H) = analz H \ synth H" -by (metis Un_empty_right analz_synth_Un) + by (metis Un_empty_right analz_synth_Un) subsubsection\For reasoning about the Fake rule in traces\ lemma parts_insert_subset_Un: "X\ G \ parts(insert X H) \ parts G \ parts H" -by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono) + by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono) text\More specifically for Fake. See also \Fake_parts_sing\ below\ lemma Fake_parts_insert: - "X \ synth (analz H) \ + "X \ synth (analz H) \ parts (insert X H) \ synth (analz H) \ parts H" -by (metis Un_commute analz_increasing insert_subset parts_analz parts_mono - parts_synth synth_mono synth_subset_iff) + by (metis Un_commute analz_increasing insert_subset parts_analz parts_mono + parts_synth synth_mono synth_subset_iff) lemma Fake_parts_insert_in_Un: - "\Z \ parts (insert X H); X \ synth (analz H)\ + "\Z \ parts (insert X H); X \ synth (analz H)\ \ Z \ synth (analz H) \ parts H" -by (metis Fake_parts_insert subsetD) + by (metis Fake_parts_insert subsetD) text\\<^term>\H\ is sometimes \<^term>\Key ` KK \ spies evs\, so can't put \<^term>\G=H\.\ lemma Fake_analz_insert: - "X\ synth (analz G) \ + "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)", force) -apply (blast intro: analz_mono [THEN [2] rev_subsetD] analz_mono [THEN synth_mono, THEN [2] rev_subsetD]) -done + by (metis UnCI Un_commute Un_upper1 analz_analz_Un analz_mono analz_synth_Un insert_subset) lemma analz_conj_parts [simp]: - "(X \ analz H \ X \ parts H) = (X \ analz H)" -by (blast intro: analz_subset_parts [THEN subsetD]) + "(X \ analz H \ X \ parts H) = (X \ analz H)" + by (blast intro: analz_subset_parts [THEN subsetD]) lemma analz_disj_parts [simp]: - "(X \ analz H | X \ parts H) = (X \ parts H)" -by (blast intro: analz_subset_parts [THEN subsetD]) + "(X \ analz H | X \ parts H) = (X \ parts H)" + by (blast intro: analz_subset_parts [THEN subsetD]) text\Without this equation, other rules for synth and analz would yield redundant cases\ lemma MPair_synth_analz [iff]: - "(\X,Y\ \ synth (analz H)) = - (X \ synth (analz H) \ Y \ synth (analz H))" -by blast + "\X,Y\ \ synth (analz H) \ X \ synth (analz H) \ Y \ synth (analz H)" + by blast lemma Crypt_synth_analz: - "\Key K \ analz H; Key (invKey K) \ analz H\ + "\Key K \ analz H; Key (invKey K) \ analz H\ \ (Crypt K X \ synth (analz H)) = (X \ synth (analz H))" -by blast - + by blast lemma Hash_synth_analz [simp]: - "X \ synth (analz H) + "X \ synth (analz H) \ (Hash\X,Y\ \ synth (analz H)) = (Hash\X,Y\ \ analz H)" -by blast + by blast subsection\HPair: a combination of Hash and MPair\ @@ -734,43 +730,43 @@ unfolding HPair_def by simp lemmas HPair_neqs = Agent_neq_HPair Nonce_neq_HPair Number_neq_HPair - Key_neq_HPair Hash_neq_HPair Crypt_neq_HPair + Key_neq_HPair Hash_neq_HPair Crypt_neq_HPair declare HPair_neqs [iff] declare HPair_neqs [symmetric, iff] lemma HPair_eq [iff]: "(Hash[X'] Y' = Hash[X] Y) = (X' = X \ Y'=Y)" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma MPair_eq_HPair [iff]: - "(\X',Y'\ = Hash[X] Y) = (X' = Hash\X,Y\ \ Y'=Y)" -by (simp add: HPair_def) + "(\X',Y'\ = Hash[X] Y) = (X' = Hash\X,Y\ \ Y'=Y)" + by (simp add: HPair_def) lemma HPair_eq_MPair [iff]: - "(Hash[X] Y = \X',Y'\) = (X' = Hash\X,Y\ \ Y'=Y)" -by (auto simp add: HPair_def) + "(Hash[X] Y = \X',Y'\) = (X' = Hash\X,Y\ \ Y'=Y)" + by (auto simp add: HPair_def) subsubsection\Specialized laws, proved in terms of those for Hash and MPair\ lemma keysFor_insert_HPair [simp]: "keysFor (insert (Hash[X] Y) H) = keysFor H" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma parts_insert_HPair [simp]: - "parts (insert (Hash[X] Y) H) = + "parts (insert (Hash[X] Y) H) = insert (Hash[X] Y) (insert (Hash\X,Y\) (parts (insert Y H)))" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma analz_insert_HPair [simp]: - "analz (insert (Hash[X] Y) H) = + "analz (insert (Hash[X] Y) H) = insert (Hash[X] Y) (insert (Hash\X,Y\) (analz (insert Y H)))" -by (simp add: HPair_def) + by (simp add: HPair_def) lemma HPair_synth_analz [simp]: - "X \ synth (analz H) + "X \ synth (analz H) \ (Hash[X] Y \ synth (analz H)) = (Hash \X, Y\ \ analz H \ Y \ synth (analz H))" -by (auto simp add: HPair_def) + by (auto simp add: HPair_def) text\We do NOT want Crypt... messages broken up in protocols!!\ @@ -830,13 +826,13 @@ (*The key-free part of a set of messages can be removed from the scope of the analz operator.*) lemma analz_keyfree_into_Un: "\X \ analz (G \ H); G \ keyfree\ \ X \ parts G \ analz H" -apply (erule analz.induct, auto dest: parts.Body) -apply (metis Un_absorb2 keyfree_KeyE parts_Un parts_keyfree UnI2) -done + apply (erule analz.induct, auto dest: parts.Body) + apply (metis Un_absorb2 keyfree_KeyE parts_Un parts_keyfree UnI2) + done subsection\Tactics useful for many protocol proofs\ ML -\ + \ (*Analysis of Fake cases. Also works for messages that forward unknown parts, but this application is no longer necessary if analz_insert_eq is used. DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) @@ -882,47 +878,47 @@ lemma Crypt_notin_image_Key [simp]: "Crypt K X \ Key ` A" -by auto + by auto lemma Hash_notin_image_Key [simp] :"Hash X \ Key ` A" -by auto + by auto lemma synth_analz_mono: "G\H \ synth (analz(G)) \ synth (analz(H))" -by (iprover intro: synth_mono analz_mono) + by (iprover intro: synth_mono analz_mono) lemma Fake_analz_eq [simp]: - "X \ synth(analz H) \ synth (analz (insert X H)) = synth (analz H)" -by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute - subset_insertI synth_analz_mono synth_increasing synth_subset_iff) + "X \ synth(analz H) \ synth (analz (insert X H)) = synth (analz H)" + by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute + subset_insertI synth_analz_mono synth_increasing synth_subset_iff) text\Two generalizations of \analz_insert_eq\\ lemma gen_analz_insert_eq [rule_format]: - "X \ analz H \ \G. H \ G \ analz (insert X G) = analz G" -by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) + "X \ analz H \ \G. H \ G \ analz (insert X G) = analz G" + by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) lemma synth_analz_insert_eq [rule_format]: - "X \ synth (analz H) + "X \ synth (analz H) \ \G. H \ G \ (Key K \ analz (insert X G)) = (Key K \ analz G)" -apply (erule synth.induct) -apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) -done + apply (erule synth.induct) + apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) + done lemma Fake_parts_sing: - "X \ synth (analz H) \ parts{X} \ synth (analz H) \ parts H" -by (metis Fake_parts_insert empty_subsetI insert_mono parts_mono subset_trans) + "X \ synth (analz H) \ parts{X} \ synth (analz H) \ parts H" + by (metis Fake_parts_insert empty_subsetI insert_mono parts_mono subset_trans) lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD] method_setup spy_analz = \ Scan.succeed (SIMPLE_METHOD' o spy_analz_tac)\ - "for proving the Fake case when analz is involved" + "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = \ Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac)\ - "for debugging spy_analz" + "for debugging spy_analz" method_setup Fake_insert_simp = \ Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac)\ - "for debugging spy_analz" + "for debugging spy_analz" end diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/OtwayReesBella.thy --- a/src/HOL/Auth/OtwayReesBella.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/OtwayReesBella.thy Thu Oct 13 16:00:22 2022 +0100 @@ -185,7 +185,7 @@ A \ bad; evs \ orb\ \ Says A B \Nonce M, Agent A, Agent B, Crypt (shrK A) \Nonce Na, Nonce M, Agent A, Agent B\\ \ set evs" apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all) -apply (blast) +apply blast done @@ -312,7 +312,7 @@ txt\Oops\ prefer 4 apply (blast dest: analz_insert_freshCryptK) txt\OR4 - ii\ -prefer 3 apply (blast) +prefer 3 apply blast txt\OR3\ (*adding Gets_imp_ and Says_imp_ for efficiency*) apply (blast dest: diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/Public.thy Thu Oct 13 16:00:22 2022 +0100 @@ -95,7 +95,7 @@ by blast lemma symKeys_invKey_iff [iff]: "(invKey K \ symKeys) = (K \ symKeys)" - unfolding symKeys_def by (auto) + unfolding symKeys_def by auto lemma analz_symKeys_Decrypt: "\Crypt K X \ analz H; K \ symKeys; Key K \ analz H\ diff -r b82ac7ef65ec -r a6cc15ec45b2 src/HOL/Auth/Yahalom.thy --- a/src/HOL/Auth/Yahalom.thy Thu Oct 13 15:49:09 2022 +0100 +++ b/src/HOL/Auth/Yahalom.thy Thu Oct 13 16:00:22 2022 +0100 @@ -332,7 +332,7 @@ "Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB\, X\ \ set evs \ KeyWithNonce K NB evs" - unfolding KeyWithNonce_def by (blast) + unfolding KeyWithNonce_def by blast lemma KeyWithNonce_Says [simp]: "KeyWithNonce K NB (Says S A X # evs) = @@ -354,7 +354,7 @@ (with respect to a given trace).\ lemma fresh_not_KeyWithNonce: "Key K \ used evs \ \ KeyWithNonce K NB evs" - unfolding KeyWithNonce_def by (blast) + unfolding KeyWithNonce_def by blast text\The Server message associates K with NB' and therefore not with any other nonce NB.\