diff -r c8a2755bf220 -r ff784d5a5bfb src/Doc/Tutorial/Protocol/Event.thy --- a/src/Doc/Tutorial/Protocol/Event.thy Sat Jan 05 17:00:43 2019 +0100 +++ b/src/Doc/Tutorial/Protocol/Event.thy Sat Jan 05 17:24:33 2019 +0100 @@ -73,9 +73,9 @@ Says A B X \ parts {X} \ used evs | Gets A X \ used evs | Notes A X \ parts {X} \ used evs)" - \ \The case for @{term Gets} seems anomalous, but @{term Gets} always - follows @{term Says} in real protocols. Seems difficult to change. - See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}.\ + \ \The case for \<^term>\Gets\ seems anomalous, but \<^term>\Gets\ always + follows \<^term>\Says\ in real protocols. Seems difficult to change. + See \<^text>\Gets_correct\ in theory \<^text>\Guard/Extensions.thy\.\ lemma Notes_imp_used [rule_format]: "Notes A X \ set evs \ X \ used evs" apply (induct_tac evs) @@ -88,7 +88,7 @@ done -subsection\Function @{term knows}\ +subsection\Function \<^term>\knows\\ (*Simplifying parts(insert X (knows Spy evs)) = parts{X} \ parts(knows Spy evs). @@ -100,7 +100,7 @@ by simp text\Letting the Spy see "bad" agents' notes avoids redundant case-splits - on whether @{term "A=Spy"} and whether @{term "A\bad"}\ + on whether \<^term>\A=Spy\ and whether \<^term>\A\bad\\ lemma knows_Spy_Notes [simp]: "knows Spy (Notes A X # evs) = (if A\bad then insert X (knows Spy evs) else knows Spy evs)" @@ -246,10 +246,10 @@ used_Nil [simp del] used_Cons [simp del] -text\For proving theorems of the form @{term "X \ analz (knows Spy evs) \ P"} +text\For proving theorems of the form \<^term>\X \ analz (knows Spy evs) \ P\ New events added by induction to "evs" are discarded. Provided this information isn't needed, the proof will be much shorter, since - it will omit complicated reasoning about @{term analz}.\ + it will omit complicated reasoning about \<^term>\analz\.\ lemmas analz_mono_contra = knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD] @@ -358,14 +358,14 @@ Sometimes the protocol requires an agent to generate a new nonce. The probability that a 20-byte random number has appeared before is effectively -zero. To formalize this important property, the set @{term "used evs"} +zero. To formalize this important property, the set \<^term>\used evs\ denotes the set of all items mentioned in the trace~\evs\. The function \used\ has a straightforward recursive definition. Here is the case for \Says\ event: @{thm [display,indent=5] used_Says [no_vars]} The function \knows\ formalizes an agent's knowledge. Mostly we only -care about the spy's knowledge, and @{term "knows Spy evs"} is the set of items +care about the spy's knowledge, and \<^term>\knows Spy evs\ is the set of items available to the spy in the trace~\evs\. Already in the empty trace, the spy starts with some secrets at his disposal, such as the private keys of compromised users. After each \Says\ event, the spy learns the @@ -374,9 +374,9 @@ Combinations of functions express other important sets of messages derived from~\evs\: \begin{itemize} -\item @{term "analz (knows Spy evs)"} is everything that the spy could +\item \<^term>\analz (knows Spy evs)\ is everything that the spy could learn by decryption -\item @{term "synth (analz (knows Spy evs))"} is everything that the spy +\item \<^term>\synth (analz (knows Spy evs))\ is everything that the spy could generate \end{itemize} \