diff -r c8a2755bf220 -r ff784d5a5bfb src/Doc/Tutorial/CTL/PDL.thy --- a/src/Doc/Tutorial/CTL/PDL.thy Sat Jan 05 17:00:43 2019 +0100 +++ b/src/Doc/Tutorial/CTL/PDL.thy Sat Jan 05 17:24:33 2019 +0100 @@ -36,8 +36,8 @@ text\\noindent The first three equations should be self-explanatory. The temporal formula -@{term"AX f"} means that @{term f} is true in \emph{A}ll ne\emph{X}t states whereas -@{term"EF f"} means that there \emph{E}xists some \emph{F}uture state in which @{term f} is +\<^term>\AX f\ means that \<^term>\f\ is true in \emph{A}ll ne\emph{X}t states whereas +\<^term>\EF f\ means that there \emph{E}xists some \emph{F}uture state in which \<^term>\f\ is true. The future is expressed via \\<^sup>*\, the reflexive transitive closure. Because of reflexivity, the future includes the present. @@ -53,17 +53,17 @@ "mc(EF f) = lfp(\T. mc f \ (M\ `` T))" text\\noindent -Only the equation for @{term EF} deserves some comments. Remember that the +Only the equation for \<^term>\EF\ deserves some comments. Remember that the postfix \\\ and the infix \``\ are predefined and denote the converse of a relation and the image of a set under a relation. Thus -@{term "M\ `` T"} is the set of all predecessors of @{term T} and the least -fixed point (@{term lfp}) of @{term"\T. mc f \ M\ `` T"} is the least set -@{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you -find it hard to see that @{term"mc(EF f)"} contains exactly those states from -which there is a path to a state where @{term f} is true, do not worry --- this +\<^term>\M\ `` T\ is the set of all predecessors of \<^term>\T\ and the least +fixed point (\<^term>\lfp\) of \<^term>\\T. mc f \ M\ `` T\ is the least set +\<^term>\T\ containing \<^term>\mc f\ and all predecessors of \<^term>\T\. If you +find it hard to see that \<^term>\mc(EF f)\ contains exactly those states from +which there is a path to a state where \<^term>\f\ is true, do not worry --- this will be proved in a moment. -First we prove monotonicity of the function inside @{term lfp} +First we prove monotonicity of the function inside \<^term>\lfp\ in order to make sure it really has a least fixed point. \ @@ -92,7 +92,7 @@ txt\\noindent Simplification leaves us with the following first subgoal @{subgoals[display,indent=0,goals_limit=1]} -which is proved by @{term lfp}-induction: +which is proved by \<^term>\lfp\-induction: \ apply(erule lfp_induct_set) @@ -123,15 +123,15 @@ txt\\noindent After simplification and clarification we are left with @{subgoals[display,indent=0,goals_limit=1]} -This goal is proved by induction on @{term"(s,t)\M\<^sup>*"}. But since the model -checker works backwards (from @{term t} to @{term s}), we cannot use the +This goal is proved by induction on \<^term>\(s,t)\M\<^sup>*\. But since the model +checker works backwards (from \<^term>\t\ to \<^term>\s\), we cannot use the induction theorem @{thm[source]rtrancl_induct}: it works in the forward direction. Fortunately the converse induction theorem @{thm[source]converse_rtrancl_induct} already exists: @{thm[display,margin=60]converse_rtrancl_induct[no_vars]} -It says that if @{prop"(a,b)\r\<^sup>*"} and we know @{prop"P b"} then we can infer -@{prop"P a"} provided each step backwards from a predecessor @{term z} of -@{term b} preserves @{term P}. +It says that if \<^prop>\(a,b)\r\<^sup>*\ and we know \<^prop>\P b\ then we can infer +\<^prop>\P a\ provided each step backwards from a predecessor \<^term>\z\ of +\<^term>\b\ preserves \<^term>\P\. \ apply(erule converse_rtrancl_induct) @@ -139,7 +139,7 @@ txt\\noindent The base case @{subgoals[display,indent=0,goals_limit=1]} -is solved by unrolling @{term lfp} once +is solved by unrolling \<^term>\lfp\ once \ apply(subst lfp_unfold[OF mono_ef]) @@ -171,15 +171,15 @@ text\ \begin{exercise} -@{term AX} has a dual operator @{term EN} +\<^term>\AX\ has a dual operator \<^term>\EN\ (``there exists a next state such that'')% \footnote{We cannot use the customary \EX\: it is reserved as the \textsc{ascii}-equivalent of \\\.} with the intended semantics @{prop[display]"(s \ EN f) = (\t. (s,t) \ M \ t \ f)"} -Fortunately, @{term"EN f"} can already be expressed as a PDL formula. How? +Fortunately, \<^term>\EN f\ can already be expressed as a PDL formula. How? -Show that the semantics for @{term EF} satisfies the following recursion equation: +Show that the semantics for \<^term>\EF\ satisfies the following recursion equation: @{prop[display]"(s \ EF f) = (s \ f | s \ EN(EF f))"} \end{exercise} \index{PDL|)}