diff -r 4939494ed791 -r ce654b0e6d69 src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Tue Feb 13 14:24:50 2018 +0100 +++ b/src/Doc/Tutorial/Protocol/Message.thy Thu Feb 15 12:11:00 2018 +0100 @@ -39,7 +39,7 @@ specification (invKey) invKey [simp]: "invKey (invKey K) = K" - invKey_symmetric: "all_symmetric --> invKey = id" + invKey_symmetric: "all_symmetric \ invKey = id" by (rule exI [of _ id], auto) @@ -81,13 +81,13 @@ (*<*) text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ syntax - "_MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") + "_MTuple" :: "['a, args] \ 'a * 'b" ("(2\_,/ _\)") translations "\x, y, z\" == "\x, \y, z\\" "\x, y\" == "CONST MPair x y" -definition keysFor :: "msg set => key set" where +definition keysFor :: "msg set \ key set" where \ \Keys useful to decrypt elements of a message set\ "keysFor H == invKey ` {K. \X. Crypt K X \ H}" @@ -95,17 +95,17 @@ subsubsection\Inductive Definition of All Parts" of a Message\ inductive_set - parts :: "msg set => msg 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" + 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" text\Monotonicity\ -lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" +lemma parts_mono: "G \ H \ parts(G) \ parts(H)" apply auto apply (erule parts.induct) apply (blast dest: parts.Fst parts.Snd parts.Body)+ @@ -113,7 +113,7 @@ text\Equations hold because constructors are injective.\ -lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" +lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x\A)" by auto lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x\A)" @@ -143,7 +143,7 @@ by (unfold keysFor_def, blast) text\Monotonicity\ -lemma keysFor_mono: "G \ H ==> keysFor(G) \ keysFor(H)" +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" @@ -165,7 +165,7 @@ 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" +lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H \ invKey K \ keysFor H" by (unfold keysFor_def, blast) @@ -192,11 +192,11 @@ apply (erule parts.induct, blast+) done -lemma parts_emptyE [elim!]: "X\ parts{} ==> P" +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}" +lemma parts_singleton: "X\ parts H \ \Y\H. X\ parts {Y}" by (erule parts.induct, fast+) @@ -252,7 +252,7 @@ subsubsection\Idempotence and transitivity\ -lemma parts_partsD [dest!]: "X\ parts (parts H) ==> X\ parts H" +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" @@ -324,7 +324,7 @@ 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}" +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!\ @@ -363,7 +363,7 @@ \ X \ analz H" (*<*) text\Monotonicity; Lemma 1 of Lowe's paper\ -lemma analz_mono: "G\H ==> analz(G) \ analz(H)" +lemma analz_mono: "G\H \ analz(G) \ analz(H)" apply auto apply (erule analz.induct) apply (auto dest: analz.Fst analz.Snd) @@ -435,7 +435,7 @@ 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) @@ -455,20 +455,20 @@ 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)" + \ 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 ==> +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 ==> +lemma lemma2: "Key (invKey K) \ analz H \ insert (Crypt K X) (analz (insert X H)) \ analz (insert (Crypt K X) H)" apply auto @@ -477,7 +477,7 @@ 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) @@ -511,7 +511,7 @@ subsubsection\Idempotence and transitivity\ -lemma analz_analzD [dest!]: "X\ analz (analz H) ==> X\ analz H" +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" @@ -531,13 +531,13 @@ by (erule analz_trans, blast) (*Cut can be proved easily by induction on - "Y: analz (insert X H) ==> X: analz H --> Y: analz H" + "Y \ analz (insert X H) \ X \ analz H \ Y \ analz H" *) 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" +lemma analz_insert_eq: "X\ analz H \ analz (insert X H) = analz H" by (blast intro: analz_cut analz_insertI) @@ -556,7 +556,7 @@ by (intro equalityI analz_subset_cong, simp_all) lemma analz_insert_cong: - "analz H = analz H' ==> analz(insert X H) = analz(insert X H')" + "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\ @@ -568,7 +568,7 @@ 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)" + "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 @@ -598,7 +598,7 @@ | Crypt [intro]: "\X \ synth H; Key K \ H\ \ Crypt K X \ synth H" (*<*) -lemma synth_mono: "G\H ==> synth(G) \ synth(H)" +lemma synth_mono: "G\H \ synth(G) \ synth(H)" by (auto, erule synth.induct, auto) inductive_cases Key_synth [elim!]: "Key K \ synth H" @@ -668,7 +668,7 @@ subsubsection\Idempotence and transitivity\ -lemma synth_synthD [dest!]: "X\ synth (synth H) ==> X\ synth H" +lemma synth_synthD [dest!]: "X\ synth (synth H) \ X\ synth H" by (erule synth.induct, blast+) lemma synth_idem: "synth (synth H) = synth H" @@ -697,7 +697,7 @@ by blast lemma Crypt_synth_eq [simp]: - "Key K \ H ==> (Crypt K X \ synth H) = (Crypt K X \ H)" + "Key K \ H \ (Crypt K X \ synth H) = (Crypt K X \ H)" by blast @@ -724,13 +724,13 @@ subsubsection\For reasoning about the Fake rule in traces\ -lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" +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"}\ lemma Fake_parts_insert: - "X \ synth (analz H) ==> + "X \ synth (analz H) \ parts (insert X H) \ synth (analz H) \ parts H" apply (drule parts_insert_subset_Un) apply (simp (no_asm_use)) @@ -738,14 +738,14 @@ done 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 (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"}.\ 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) ") @@ -869,11 +869,11 @@ lemma Crypt_notin_image_Key [simp]: "Crypt K X \ Key ` A" by auto -lemma synth_analz_mono: "G\H ==> synth (analz(G)) \ synth (analz(H))" +lemma synth_analz_mono: "G\H \ synth (analz(G)) \ synth (analz(H))" by (iprover intro: synth_mono analz_mono) lemma Fake_analz_eq [simp]: - "X \ synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)" + "X \ synth(analz H) \ synth (analz (insert X H)) = synth (analz H)" apply (drule Fake_analz_insert[of _ _ "H"]) apply (simp add: synth_increasing[THEN Un_absorb2]) apply (drule synth_mono) @@ -885,18 +885,18 @@ 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" + "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) - ==> ALL G. H \ G --> (Key K \ analz (insert X G)) = (Key K \ analz G)" + \ \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 lemma Fake_parts_sing: - "X \ synth (analz H) ==> parts{X} \ synth (analz H) \ parts H" + "X \ synth (analz H) \ parts{X} \ synth (analz H) \ parts H" apply (rule subset_trans) apply (erule_tac [2] Fake_parts_insert) apply (rule parts_mono, blast)