doc-src/TutorialI/CTL/Base.thy
changeset 11458 09a6c44a48ea
parent 10983 59961d32b1ae
child 17914 99ead7a7eb42
--- a/doc-src/TutorialI/CTL/Base.thy	Thu Jul 26 18:23:38 2001 +0200
+++ b/doc-src/TutorialI/CTL/Base.thy	Fri Aug 03 18:04:55 2001 +0200
@@ -10,7 +10,7 @@
 (specifications) \cite{ClarkeGP-book,Huth-Ryan-book}. Its foundations are set theoretic
 and this section will explore them in HOL\@. This is done in two steps.  First
 we consider a simple modal logic called propositional dynamic
-logic (PDL), which we then extend to the temporal logic CTL, which is
+logic (PDL)\@.  We then proceed to the temporal logic CTL, which is
 used in many real
 model checkers. In each case we give both a traditional semantics (@{text \<Turnstile>}) and a
 recursive function @{term mc} that maps a formula into the set of all states of
@@ -21,8 +21,9 @@
 
 \underscoreon
 
-Our models are \emph{transition systems}, i.e.\ sets of \emph{states} with
-transitions between them, as shown in this simple example:
+Our models are \emph{transition systems}:\index{transition systems}
+sets of \emph{states} with
+transitions between them.  Here is a simple example:
 \begin{center}
 \unitlength.5mm
 \thicklines
@@ -45,21 +46,20 @@
 \put(91,21){\makebox(0,0)[bl]{$s_2$}}
 \end{picture}
 \end{center}
-Each state has a unique name or number ($s_0,s_1,s_2$), and in each
-state certain \emph{atomic propositions} ($p,q,r$) are true.
-The aim of temporal logic is to formalize statements such as ``there is no
-path starting from $s_2$ leading to a state where $p$ or $q$
-are true'', which is true, and ``on all paths starting from $s_0$ $q$ is always true'',
-which is false.
+Each state has a unique name or number ($s_0,s_1,s_2$), and in each state
+certain \emph{atomic propositions} ($p,q,r$) hold.  The aim of temporal logic
+is to formalize statements such as ``there is no path starting from $s_2$
+leading to a state where $p$ or $q$ holds,'' which is true, and ``on all paths
+starting from $s_0$, $q$ always holds,'' which is false.
 
-Abstracting from this concrete example, we assume there is some type of
+Abstracting from this concrete example, we assume there is a type of
 states:
 *}
 
 typedecl state
 
 text{*\noindent
-Command \isacommand{typedecl} merely declares a new type but without
+Command \commdx{typedecl} merely declares a new type but without
 defining it (see \S\ref{sec:typedecl}). Thus we know nothing
 about the type other than its existence. That is exactly what we need
 because @{typ state} really is an implicit parameter of our model.  Of