diff -r 863b4f9f6bd7 -r 378ae9e46175 src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Sat Jan 25 18:34:05 2014 +0100 +++ b/src/Doc/Tutorial/Protocol/Message.thy Sat Jan 25 21:52:04 2014 +0100 @@ -190,7 +190,7 @@ lemma parts_increasing: "H \ parts(H)" by blast -lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard] +lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD] lemma parts_empty [simp]: "parts{} = {}" apply safe @@ -388,9 +388,9 @@ apply (erule analz.induct, blast+) done -lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard] +lemmas analz_into_parts = analz_subset_parts [THEN subsetD] -lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard] +lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD] lemma parts_analz [simp]: "parts (analz H) = parts H" @@ -404,7 +404,7 @@ apply (erule analz.induct, auto) done -lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard] +lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD] subsubsection{*General equational properties *} @@ -786,21 +786,23 @@ 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 [standard] = +lemmas pushKeys = insert_commute [of "Key K" "Agent C"] insert_commute [of "Key K" "Nonce N"] insert_commute [of "Key K" "Number N"] insert_commute [of "Key K" "Hash X"] insert_commute [of "Key K" "MPair X Y"] insert_commute [of "Key K" "Crypt X K'"] + for K C N X Y K' -lemmas pushCrypts [standard] = +lemmas pushCrypts = insert_commute [of "Crypt X K" "Agent C"] insert_commute [of "Crypt X K" "Agent C"] insert_commute [of "Crypt X K" "Nonce N"] insert_commute [of "Crypt X K" "Number N"] insert_commute [of "Crypt X K" "Hash X'"] 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. *}