diff -r e9ab4ad7bd15 -r 23307fd33906 src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Thu Jan 11 13:48:17 2018 +0100 +++ b/src/Doc/Tutorial/Protocol/Message.thy Fri Jan 12 14:08:53 2018 +0100 @@ -5,7 +5,7 @@ Inductive relations "parts", "analz" and "synth" *)(*<*) -section{*Theory of Agents and Messages for Security Protocols*} +section\Theory of Agents and Messages for Security Protocols\ theory Message imports Main begin ML_file "../../antiquote_setup.ML" @@ -15,27 +15,27 @@ by blast (*>*) -section{* Agents and Messages *} +section\Agents and Messages\ -text {* +text \ All protocol specifications refer to a syntactic theory of messages. Datatype @{text agent} introduces the constant @{text Server} (a trusted central machine, needed for some protocols), an infinite population of friendly agents, and the~@{text Spy}: -*} +\ datatype agent = Server | Friend nat | Spy -text {* +text \ Keys are just natural numbers. Function @{text invKey} maps a public key to the matching private key, and vice versa: -*} +\ type_synonym key = nat consts invKey :: "key \ key" (*<*) -consts all_symmetric :: bool --{*true if all keys are symmetric*} +consts all_symmetric :: bool \\true if all keys are symmetric\ specification (invKey) invKey [simp]: "invKey (invKey K) = K" @@ -43,18 +43,18 @@ 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*} +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 {* +text \ Datatype @{text msg} introduces the message forms, which include agent names, nonces, keys, compound messages, and encryptions. -*} +\ datatype msg = Agent agent @@ -63,7 +63,7 @@ | MPair msg msg | Crypt key msg -text {* +text \ \noindent The notation $\comp{X\sb 1,\ldots X\sb{n-1},X\sb n}$ abbreviates @@ -76,10 +76,10 @@ wrong key succeeds but yields garbage. Our model of encryption is realistic if encryption adds some redundancy to the plaintext, such as a checksum, so that garbage can be detected. -*} +\ (*<*) -text{*Concrete syntax: messages appear as \\A,B,NA\\, etc...*} +text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ syntax "_MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") translations @@ -88,11 +88,11 @@ definition keysFor :: "msg set => key set" where - --{*Keys useful to decrypt elements of a message 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*} +subsubsection\Inductive Definition of All Parts" of a Message\ inductive_set parts :: "msg set => msg set" @@ -104,7 +104,7 @@ | Body: "Crypt K X \ parts H ==> X \ parts H" -text{*Monotonicity*} +text\Monotonicity\ lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" apply auto apply (erule parts.induct) @@ -112,7 +112,7 @@ done -text{*Equations hold because constructors are injective.*} +text\Equations hold because constructors are injective.\ lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" by auto @@ -123,7 +123,7 @@ by auto -subsubsection{*Inverse of keys *} +subsubsection\Inverse of keys\ lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')" apply safe @@ -131,7 +131,7 @@ done -subsection{*keysFor operator*} +subsection\keysFor operator\ lemma keysFor_empty [simp]: "keysFor {} = {}" by (unfold keysFor_def, blast) @@ -142,7 +142,7 @@ lemma keysFor_UN [simp]: "keysFor (\i\A. H i) = (\i\A. keysFor (H i))" by (unfold keysFor_def, blast) -text{*Monotonicity*} +text\Monotonicity\ lemma keysFor_mono: "G \ H ==> keysFor(G) \ keysFor(H)" by (unfold keysFor_def, blast) @@ -169,7 +169,7 @@ by (unfold keysFor_def, blast) -subsection{*Inductive relation "parts"*} +subsection\Inductive relation "parts"\ lemma MPair_parts: "[| \X,Y\ \ parts H; @@ -177,10 +177,10 @@ 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 +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.*} + The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\ lemma parts_increasing: "H \ parts(H)" by blast @@ -195,12 +195,12 @@ lemma parts_emptyE [elim!]: "X\ parts{} ==> P" by simp -text{*WARNING: loops if H = {Y}, therefore must not be repeated!*} +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+) -subsubsection{*Unions *} +subsubsection\Unions\ lemma parts_Un_subset1: "parts(G) \ parts(H) \ parts(G \ H)" by (intro Un_least parts_mono Un_upper1 Un_upper2) @@ -218,8 +218,8 @@ apply (simp only: parts_Un) done -text{*TWO inserts to avoid looping. This rewrite is better than nothing. - Not suitable for Addsimps: its behaviour can be strange.*} +text\TWO inserts to avoid looping. This rewrite is better than nothing. + Not suitable for Addsimps: its behaviour can be strange.\ lemma parts_insert2: "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" apply (simp add: Un_assoc) @@ -237,12 +237,12 @@ lemma parts_UN [simp]: "parts(\x\A. H x) = (\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\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.*} +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!] @@ -250,7 +250,7 @@ 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 *} +subsubsection\Idempotence and transitivity\ lemma parts_partsD [dest!]: "X\ parts (parts H) ==> X\ parts H" by (erule parts.induct, blast+) @@ -267,7 +267,7 @@ lemma parts_trans: "[| X\ parts G; G \ parts H |] ==> X\ parts H" by (drule parts_mono, blast) -text{*Cut*} +text\Cut\ lemma parts_cut: "[| Y\ parts (insert X G); X\ parts H |] ==> Y\ parts (G \ H)" by (blast intro: parts_trans) @@ -277,7 +277,7 @@ by (force dest!: parts_cut intro: parts_insertI) -subsubsection{*Rewrite rules for pulling out atomic messages *} +subsubsection\Rewrite rules for pulling out atomic messages\ lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset] @@ -323,21 +323,21 @@ done -text{*In any message, there is an upper bound N on its greatest nonce.*} +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}" apply (induct_tac "msg") apply (simp_all (no_asm_simp) add: exI parts_insert2) - txt{*MPair case: blast works out the necessary sum itself!*} + txt\MPair case: blast works out the necessary sum itself!\ prefer 2 apply auto apply (blast elim!: add_leE) -txt{*Nonce case*} +txt\Nonce case\ apply (rename_tac nat) apply (rule_tac x = "N + Suc nat" in exI, auto) done (*>*) -section{* Modelling the Adversary *} +section\Modelling the Adversary\ -text {* +text \ The spy is part of the system and must be built into the model. He is a malicious user who does not have to follow the protocol. He watches the network and uses any keys he knows to decrypt messages. @@ -349,7 +349,7 @@ messages. The set @{text "analz H"} formalizes what the adversary can learn from the set of messages~$H$. The closure properties of this set are defined inductively. -*} +\ inductive_set analz :: "msg set \ msg set" @@ -362,14 +362,14 @@ "\Crypt K X \ analz H; Key(invKey K) \ analz H\ \ X \ analz H" (*<*) -text{*Monotonicity; Lemma 1 of Lowe's paper*} +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*} +text\Making it safe speeds up proofs\ lemma MPair_analz [elim!]: "[| \X,Y\ \ analz H; [| X \ analz H; Y \ analz H |] ==> P @@ -402,22 +402,22 @@ lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD] -subsubsection{*General equational properties *} +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*} +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 *} +subsubsection\Rewrite rules for pulling out atomic messages\ lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert] @@ -433,7 +433,7 @@ apply (erule analz.induct, auto) done -text{*Can only pull out Keys if they are not needed to decrypt the rest*} +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)" @@ -452,7 +452,7 @@ apply (blast intro: analz.Fst analz.Snd)+ done -text{*Can pull out enCrypted message if the Key is not known*} +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)" @@ -482,10 +482,10 @@ 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, +text\Case analysis: either the message is secure, or it is not! Effective, but can cause subgoals to blow up! Use with @{text "if_split"}; apparently @{text "split_tac"} does not cope with patterns such as @{term"analz (insert -(Crypt K X) H)"} *} +(Crypt K X) H)"}\ lemma analz_Crypt_if [simp]: "analz (insert (Crypt K X) H) = (if (Key (invKey K) \ analz H) @@ -494,7 +494,7 @@ by (simp add: analz_insert_Crypt analz_insert_Decrypt) -text{*This rule supposes "for the sake of argument" that we have the key.*} +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))" @@ -509,7 +509,7 @@ done -subsubsection{*Idempotence and transitivity *} +subsubsection\Idempotence and transitivity\ lemma analz_analzD [dest!]: "X\ analz (analz H) ==> X\ analz H" by (erule analz.induct, blast+) @@ -526,7 +526,7 @@ lemma analz_trans: "[| X\ analz G; G \ analz H |] ==> X\ analz H" by (drule analz_mono, blast) -text{*Cut; Lemma 2 of Lowe*} +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) @@ -534,14 +534,14 @@ "Y: analz (insert X H) ==> X: analz H --> Y: analz H" *) -text{*This rewrite rule helps in the simplification of messages that involve +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. *} + 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" *} +text\A congruence rule for "analz"\ lemma analz_subset_cong: "[| analz G \ analz G'; analz H \ analz H' |] @@ -559,14 +559,14 @@ "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*} +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...*} +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) @@ -576,7 +576,7 @@ 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]) (*>*) -text {* +text \ Note the @{text Decrypt} rule: the spy can decrypt a message encrypted with key~$K$ if he has the matching key,~$K^{-1}$. Properties proved by rule induction include the following: @@ -585,7 +585,7 @@ The set of fake messages that an intruder could invent starting from~@{text H} is @{text "synth(analz H)"}, where @{text "synth H"} formalizes what the adversary can build from the set of messages~$H$. -*} +\ inductive_set synth :: "msg set \ msg set" @@ -618,7 +618,7 @@ apply (simp (no_asm_use)) done (*>*) -text {* +text \ The set includes all agent names. Nonces and keys are assumed to be unguessable, so none are included beyond those already in~$H$. Two elements of @{term "synth H"} can be combined, and an element can be encrypted @@ -629,11 +629,11 @@ @{named_thms [display,indent=0] analz_synth [no_vars] (analz_synth)} Rule inversion plays a major role in reasoning about @{text synth}, through declarations such as this one: -*} +\ inductive_cases Nonce_synth [elim!]: "Nonce n \ synth H" -text {* +text \ \noindent The resulting elimination rule replaces every assumption of the form @{term "Nonce n \ synth H"} by @{term "Nonce n \ H"}, @@ -651,22 +651,22 @@ use @{text parts} to express general well-formedness properties of a protocol, for example, that an uncompromised agent's private key will never be included as a component of any message. -*} +\ (*<*) lemma synth_increasing: "H \ synth(H)" by blast -subsubsection{*Unions *} +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.*} +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) lemma synth_insert: "insert X (synth H) \ synth(insert X H)" by (blast intro: synth_mono [THEN [2] rev_subsetD]) -subsubsection{*Idempotence and transitivity *} +subsubsection\Idempotence and transitivity\ lemma synth_synthD [dest!]: "X\ synth (synth H) ==> X\ synth H" by (erule synth.induct, blast+) @@ -683,7 +683,7 @@ lemma synth_trans: "[| X\ synth G; G \ synth H |] ==> X\ synth H" by (drule synth_mono, blast) -text{*Cut; Lemma 2 of Lowe*} +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) @@ -706,7 +706,7 @@ by (unfold keysFor_def, blast) -subsubsection{*Combinations of parts, analz and synth *} +subsubsection\Combinations of parts, analz and synth\ lemma parts_synth [simp]: "parts (synth H) = parts H \ synth H" apply (rule equalityI) @@ -722,13 +722,13 @@ done -subsubsection{*For reasoning about the Fake rule in traces *} +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 (rule subset_trans [OF parts_mono parts_Un_subset2], blast) -text{*More specifically for Fake. Very occasionally we could do with a version - of the form @{term"parts{X} \ synth (analz H) \ parts H"} *} +text\More specifically for Fake. Very occasionally we could do with a version + of the form @{term"parts{X} \ synth (analz H) \ parts H"}\ lemma Fake_parts_insert: "X \ synth (analz H) ==> parts (insert X H) \ synth (analz H) \ parts H" @@ -742,8 +742,8 @@ ==> Z \ synth (analz H) \ parts H" by (blast dest: Fake_parts_insert [THEN subsetD, dest]) -text{*@{term H} is sometimes @{term"Key ` KK \ spies evs"}, so can't put - @{term "G=H"}.*} +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) ==> analz (insert X H) \ synth (analz G) \ analz (G \ H)" @@ -762,8 +762,8 @@ "(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*} +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))" @@ -775,12 +775,12 @@ by blast -text{*We do NOT want Crypt... messages broken up in protocols!!*} +text\We do NOT want Crypt... messages broken up in protocols!!\ declare parts.Body [rule del] -text{*Rewrites to push in Key and Crypt messages, so that other messages can - be pulled out using the @{text analz_insert} rules*} +text\Rewrites to push in Key and Crypt messages, so that other messages can + be pulled out using the @{text analz_insert} rules\ lemmas pushKeys = insert_commute [of "Key K" "Agent C"] @@ -800,14 +800,14 @@ insert_commute [of "Crypt X K" "MPair X' Y"] for X K C N X' Y -text{*Cannot be added with @{text "[simp]"} -- messages should not always be - re-ordered. *} +text\Cannot be added with @{text "[simp]"} -- messages should not always be + re-ordered.\ lemmas pushes = pushKeys pushCrypts -subsection{*Tactics useful for many protocol proofs*} +subsection\Tactics useful for many protocol proofs\ ML -{* +\ val invKey = @{thm invKey}; val keysFor_def = @{thm keysFor_def}; val symKeys_def = @{thm symKeys_def}; @@ -858,11 +858,11 @@ simp_tac ctxt 1, REPEAT (FIRSTGOAL (resolve_tac ctxt [allI,impI,notI,conjI,iffI])), DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i); -*} +\ -text{*By default only @{text o_apply} is built-in. But in the presence of +text\By default only @{text o_apply} is built-in. But in the presence of eta-expansion this means that some terms displayed as @{term "f o g"} will be -rewritten, and others will not!*} +rewritten, and others will not!\ declare o_def [simp] @@ -883,7 +883,7 @@ apply (rule synth_analz_mono, blast) done -text{*Two generalizations of @{text analz_insert_eq}*} +text\Two generalizations of @{text analz_insert_eq}\ lemma gen_analz_insert_eq [rule_format]: "X \ analz H ==> ALL G. H \ G --> analz (insert X G) = analz G" by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) @@ -904,16 +904,16 @@ 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) *} +method_setup spy_analz = \ + Scan.succeed (SIMPLE_METHOD' o spy_analz_tac)\ "for proving the Fake case when analz is involved" -method_setup atomic_spy_analz = {* - Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac) *} +method_setup atomic_spy_analz = \ + Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac)\ "for debugging spy_analz" -method_setup Fake_insert_simp = {* - Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac) *} +method_setup Fake_insert_simp = \ + Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac)\ "for debugging spy_analz"