| 
10123
 | 
     1  | 
%
  | 
| 
 | 
     2  | 
\begin{isabellebody}%
 | 
| 
 | 
     3  | 
\def\isabellecontext{CTL}%
 | 
| 
17056
 | 
     4  | 
%
  | 
| 
 | 
     5  | 
\isadelimtheory
  | 
| 
 | 
     6  | 
%
  | 
| 
 | 
     7  | 
\endisadelimtheory
  | 
| 
 | 
     8  | 
%
  | 
| 
 | 
     9  | 
\isatagtheory
  | 
| 
 | 
    10  | 
%
  | 
| 
 | 
    11  | 
\endisatagtheory
  | 
| 
 | 
    12  | 
{\isafoldtheory}%
 | 
| 
 | 
    13  | 
%
  | 
| 
 | 
    14  | 
\isadelimtheory
  | 
| 
 | 
    15  | 
%
  | 
| 
 | 
    16  | 
\endisadelimtheory
  | 
| 
10133
 | 
    17  | 
%
  | 
| 
10971
 | 
    18  | 
\isamarkupsubsection{Computation Tree Logic --- CTL%
 | 
| 
10395
 | 
    19  | 
}
  | 
| 
11866
 | 
    20  | 
\isamarkuptrue%
  | 
| 
10149
 | 
    21  | 
%
  | 
| 
 | 
    22  | 
\begin{isamarkuptext}%
 | 
| 
10217
 | 
    23  | 
\label{sec:CTL}
 | 
| 
11494
 | 
    24  | 
\index{CTL|(}%
 | 
| 
10867
 | 
    25  | 
The semantics of PDL only needs reflexive transitive closure.
  | 
| 
 | 
    26  | 
Let us be adventurous and introduce a more expressive temporal operator.
  | 
| 
 | 
    27  | 
We extend the datatype
  | 
| 
10149
 | 
    28  | 
\isa{formula} by a new constructor%
 | 
| 
 | 
    29  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    30  | 
\isamarkuptrue%
  | 
| 
 | 
    31  | 
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AF\ formula%
 | 
| 
10149
 | 
    32  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    33  | 
\noindent
  | 
| 
10983
 | 
    34  | 
which stands for ``\emph{A}lways in the \emph{F}uture'':
 | 
| 
 | 
    35  | 
on all infinite paths, at some point the formula holds.
  | 
| 
 | 
    36  | 
Formalizing the notion of an infinite path is easy
  | 
| 
10159
 | 
    37  | 
in HOL: it is simply a function from \isa{nat} to \isa{state}.%
 | 
| 
10149
 | 
    38  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    39  | 
\isamarkuptrue%
  | 
| 
 | 
    40  | 
\isacommand{constdefs}\isamarkupfalse%
 | 
| 
 | 
    41  | 
\ Paths\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}state\ {\isasymRightarrow}\ {\isacharparenleft}nat\ {\isasymRightarrow}\ state{\isacharparenright}set{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
    42  | 
\ \ \ \ \ \ \ \ \ {\isachardoublequoteopen}Paths\ s\ {\isasymequiv}\ {\isacharbraceleft}p{\isachardot}\ s\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p{\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
10149
 | 
    43  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    44  | 
\noindent
  | 
| 
11494
 | 
    45  | 
This definition allows a succinct statement of the semantics of \isa{AF}:
 | 
| 
10867
 | 
    46  | 
\footnote{Do not be misled: neither datatypes nor recursive functions can be
 | 
| 
10149
 | 
    47  | 
extended by new constructors or equations. This is just a trick of the
  | 
| 
12815
 | 
    48  | 
presentation (see \S\ref{sec:doc-prep-suppress}). In reality one has to define
 | 
| 
 | 
    49  | 
a new datatype and a new function.}%
  | 
| 
10149
 | 
    50  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    51  | 
\isamarkuptrue%
  | 
| 
 | 
    52  | 
{\isachardoublequoteopen}s\ {\isasymTurnstile}\ AF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
10149
 | 
    53  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    54  | 
\noindent
  | 
| 
 | 
    55  | 
Model checking \isa{AF} involves a function which
 | 
| 
10159
 | 
    56  | 
is just complicated enough to warrant a separate definition:%
  | 
| 
10149
 | 
    57  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    58  | 
\isamarkuptrue%
  | 
| 
 | 
    59  | 
\isacommand{constdefs}\isamarkupfalse%
 | 
| 
 | 
    60  | 
\ af\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}state\ set\ {\isasymRightarrow}\ state\ set\ {\isasymRightarrow}\ state\ set{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
    61  | 
\ \ \ \ \ \ \ \ \ {\isachardoublequoteopen}af\ A\ T\ {\isasymequiv}\ A\ {\isasymunion}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymlongrightarrow}\ t\ {\isasymin}\ T{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
10149
 | 
    62  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    63  | 
\noindent
  | 
| 
10867
 | 
    64  | 
Now we define \isa{mc\ {\isacharparenleft}AF\ f{\isacharparenright}} as the least set \isa{T} that includes
 | 
| 
10159
 | 
    65  | 
\isa{mc\ f} and all states all of whose direct successors are in \isa{T}:%
 | 
| 
 | 
    66  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    67  | 
\isamarkuptrue%
  | 
| 
 | 
    68  | 
{\isachardoublequoteopen}mc{\isacharparenleft}AF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}af{\isacharparenleft}mc\ f{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
10159
 | 
    69  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    70  | 
\noindent
  | 
| 
 | 
    71  | 
Because \isa{af} is monotone in its second argument (and also its first, but
 | 
| 
10983
 | 
    72  | 
that is irrelevant), \isa{af\ A} has a least fixed point:%
 | 
| 
10149
 | 
    73  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    74  | 
\isamarkuptrue%
  | 
| 
 | 
    75  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
    76  | 
\ mono{\isacharunderscore}af{\isacharcolon}\ {\isachardoublequoteopen}mono{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
    77  | 
%
  | 
| 
 | 
    78  | 
\isadelimproof
  | 
| 
 | 
    79  | 
%
  | 
| 
 | 
    80  | 
\endisadelimproof
  | 
| 
 | 
    81  | 
%
  | 
| 
 | 
    82  | 
\isatagproof
  | 
| 
 | 
    83  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
    84  | 
{\isacharparenleft}simp\ add{\isacharcolon}\ mono{\isacharunderscore}def\ af{\isacharunderscore}def{\isacharparenright}\isanewline
 | 
| 
 | 
    85  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
    86  | 
\ blast\isanewline
  | 
| 
 | 
    87  | 
\isacommand{done}\isamarkupfalse%
 | 
| 
 | 
    88  | 
%
  | 
| 
 | 
    89  | 
\endisatagproof
  | 
| 
 | 
    90  | 
{\isafoldproof}%
 | 
| 
 | 
    91  | 
%
  | 
| 
 | 
    92  | 
\isadelimproof
  | 
| 
 | 
    93  | 
%
  | 
| 
 | 
    94  | 
\endisadelimproof
  | 
| 
17056
 | 
    95  | 
%
  | 
| 
 | 
    96  | 
\isadelimproof
  | 
| 
 | 
    97  | 
%
  | 
| 
 | 
    98  | 
\endisadelimproof
  | 
| 
 | 
    99  | 
%
  | 
| 
 | 
   100  | 
\isatagproof
  | 
| 
17175
 | 
   101  | 
%
  | 
| 
17056
 | 
   102  | 
\endisatagproof
  | 
| 
 | 
   103  | 
{\isafoldproof}%
 | 
| 
 | 
   104  | 
%
  | 
| 
 | 
   105  | 
\isadelimproof
  | 
| 
 | 
   106  | 
%
  | 
| 
 | 
   107  | 
\endisadelimproof
  | 
| 
 | 
   108  | 
%
  | 
| 
 | 
   109  | 
\isadelimproof
  | 
| 
 | 
   110  | 
%
  | 
| 
 | 
   111  | 
\endisadelimproof
  | 
| 
 | 
   112  | 
%
  | 
| 
 | 
   113  | 
\isatagproof
  | 
| 
 | 
   114  | 
%
  | 
| 
 | 
   115  | 
\endisatagproof
  | 
| 
 | 
   116  | 
{\isafoldproof}%
 | 
| 
 | 
   117  | 
%
  | 
| 
 | 
   118  | 
\isadelimproof
  | 
| 
 | 
   119  | 
%
  | 
| 
 | 
   120  | 
\endisadelimproof
  | 
| 
 | 
   121  | 
%
  | 
| 
10149
 | 
   122  | 
\begin{isamarkuptext}%
 | 
| 
10867
 | 
   123  | 
All we need to prove now is  \isa{mc\ {\isacharparenleft}AF\ f{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ AF\ f{\isacharbraceright}}, which states
 | 
| 
 | 
   124  | 
that \isa{mc} and \isa{{\isasymTurnstile}} agree for \isa{AF}\@.
 | 
| 
 | 
   125  | 
This time we prove the two inclusions separately, starting
  | 
| 
10159
 | 
   126  | 
with the easy one:%
  | 
| 
 | 
   127  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   128  | 
\isamarkuptrue%
  | 
| 
 | 
   129  | 
\isacommand{theorem}\isamarkupfalse%
 | 
| 
 | 
   130  | 
\ AF{\isacharunderscore}lemma{\isadigit{1}}{\isacharcolon}\isanewline
 | 
| 
 | 
   131  | 
\ \ {\isachardoublequoteopen}lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
17056
 | 
   132  | 
\isadelimproof
  | 
| 
 | 
   133  | 
%
  | 
| 
 | 
   134  | 
\endisadelimproof
  | 
| 
 | 
   135  | 
%
  | 
| 
 | 
   136  | 
\isatagproof
  | 
| 
16069
 | 
   137  | 
%
  | 
| 
 | 
   138  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   139  | 
\noindent
  | 
| 
 | 
   140  | 
In contrast to the analogous proof for \isa{EF}, and just
 | 
| 
 | 
   141  | 
for a change, we do not use fixed point induction.  Park-induction,
  | 
| 
 | 
   142  | 
named after David Park, is weaker but sufficient for this proof:
  | 
| 
 | 
   143  | 
\begin{center}
 | 
| 
21261
 | 
   144  | 
\isa{f\ S\ {\isasymle}\ S\ {\isasymLongrightarrow}\ lfp\ f\ {\isasymle}\ S} \hfill (\isa{lfp{\isacharunderscore}lowerbound})
 | 
| 
16069
 | 
   145  | 
\end{center}
 | 
| 
19654
 | 
   146  | 
The instance of the premise \isa{f\ S\ {\isasymsubseteq}\ S} is proved pointwise,
 | 
| 
16069
 | 
   147  | 
a decision that \isa{auto} takes for us:%
 | 
| 
 | 
   148  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   149  | 
\isamarkuptrue%
  | 
| 
 | 
   150  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   151  | 
{\isacharparenleft}rule\ lfp{\isacharunderscore}lowerbound{\isacharparenright}\isanewline
 | 
| 
 | 
   152  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   153  | 
{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ af{\isacharunderscore}def\ Paths{\isacharunderscore}def{\isacharparenright}%
 | 
| 
16069
 | 
   154  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   155  | 
\begin{isabelle}%
 | 
| 
 | 
   156  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isasymlbrakk}{\isasymforall}t{\isachardot}\ {\isacharparenleft}p\ {\isadigit{0}}{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymlongrightarrow}\isanewline
 | 
| 
 | 
   157  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isasymlbrakk}{\isasymforall}t{\isachardot}\ }{\isacharparenleft}{\isasymforall}p{\isachardot}\ t\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p\ {\isacharparenleft}Suc\ i{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M{\isacharparenright}\ {\isasymlongrightarrow}\isanewline
 | 
| 
 | 
   158  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isasymlbrakk}{\isasymforall}t{\isachardot}\ {\isacharparenleft}{\isasymforall}p{\isachardot}\ }{\isacharparenleft}{\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline
 | 
| 
 | 
   159  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ \ }{\isasymforall}i{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p\ {\isacharparenleft}Suc\ i{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M{\isasymrbrakk}\isanewline
 | 
| 
 | 
   160  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ }{\isasymLongrightarrow}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A%
 | 
| 
 | 
   161  | 
\end{isabelle}
 | 
| 
 | 
   162  | 
In this remaining case, we set \isa{t} to \isa{p\ {\isadigit{1}}}.
 | 
| 
 | 
   163  | 
The rest is automatic, which is surprising because it involves
  | 
| 
 | 
   164  | 
finding the instantiation \isa{{\isasymlambda}i{\isachardot}\ p\ {\isacharparenleft}i\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}}
 | 
| 
 | 
   165  | 
for \isa{{\isasymforall}p}.%
 | 
| 
 | 
   166  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   167  | 
\isamarkuptrue%
  | 
| 
 | 
   168  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   169  | 
{\isacharparenleft}erule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}p\ {\isadigit{1}}{\isachardoublequoteclose}\ \isakeyword{in}\ allE{\isacharparenright}\isanewline
 | 
| 
 | 
   170  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   171  | 
{\isacharparenleft}auto{\isacharparenright}\isanewline
 | 
| 
 | 
   172  | 
\isacommand{done}\isamarkupfalse%
 | 
| 
 | 
   173  | 
%
  | 
| 
17056
 | 
   174  | 
\endisatagproof
  | 
| 
 | 
   175  | 
{\isafoldproof}%
 | 
| 
 | 
   176  | 
%
  | 
| 
 | 
   177  | 
\isadelimproof
  | 
| 
 | 
   178  | 
%
  | 
| 
 | 
   179  | 
\endisadelimproof
  | 
| 
11866
 | 
   180  | 
%
  | 
| 
10123
 | 
   181  | 
\begin{isamarkuptext}%
 | 
| 
10867
 | 
   182  | 
The opposite inclusion is proved by contradiction: if some state
  | 
| 
10159
 | 
   183  | 
\isa{s} is not in \isa{lfp\ {\isacharparenleft}af\ A{\isacharparenright}}, then we can construct an
 | 
| 
11494
 | 
   184  | 
infinite \isa{A}-avoiding path starting from~\isa{s}. The reason is
 | 
| 
10123
 | 
   185  | 
that by unfolding \isa{lfp} we find that if \isa{s} is not in
 | 
| 
 | 
   186  | 
\isa{lfp\ {\isacharparenleft}af\ A{\isacharparenright}}, then \isa{s} is not in \isa{A} and there is a
 | 
| 
10983
 | 
   187  | 
direct successor of \isa{s} that is again not in \mbox{\isa{lfp\ {\isacharparenleft}af\ A{\isacharparenright}}}. Iterating this argument yields the promised infinite
 | 
| 
10123
 | 
   188  | 
\isa{A}-avoiding path. Let us formalize this sketch.
 | 
| 
 | 
   189  | 
  | 
| 
10867
 | 
   190  | 
The one-step argument in the sketch above
  | 
| 
 | 
   191  | 
is proved by a variant of contraposition:%
  | 
| 
10123
 | 
   192  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   193  | 
\isamarkuptrue%
  | 
| 
 | 
   194  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
   195  | 
\ not{\isacharunderscore}in{\isacharunderscore}lfp{\isacharunderscore}afD{\isacharcolon}\isanewline
 | 
| 
 | 
   196  | 
\ {\isachardoublequoteopen}s\ {\isasymnotin}\ lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isasymLongrightarrow}\ s\ {\isasymnotin}\ A\ {\isasymand}\ {\isacharparenleft}{\isasymexists}\ t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ t\ {\isasymnotin}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
17056
 | 
   197  | 
%
  | 
| 
 | 
   198  | 
\isadelimproof
  | 
| 
 | 
   199  | 
%
  | 
| 
 | 
   200  | 
\endisadelimproof
  | 
| 
 | 
   201  | 
%
  | 
| 
 | 
   202  | 
\isatagproof
  | 
| 
17175
 | 
   203  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   204  | 
{\isacharparenleft}erule\ contrapos{\isacharunderscore}np{\isacharparenright}\isanewline
 | 
| 
 | 
   205  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   206  | 
{\isacharparenleft}subst\ lfp{\isacharunderscore}unfold{\isacharbrackleft}OF\ mono{\isacharunderscore}af{\isacharbrackright}{\isacharparenright}\isanewline
 | 
| 
 | 
   207  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   208  | 
{\isacharparenleft}simp\ add{\isacharcolon}\ af{\isacharunderscore}def{\isacharparenright}\isanewline
 | 
| 
 | 
   209  | 
\isacommand{done}\isamarkupfalse%
 | 
| 
 | 
   210  | 
%
  | 
| 
17056
 | 
   211  | 
\endisatagproof
  | 
| 
 | 
   212  | 
{\isafoldproof}%
 | 
| 
 | 
   213  | 
%
  | 
| 
 | 
   214  | 
\isadelimproof
  | 
| 
 | 
   215  | 
%
  | 
| 
 | 
   216  | 
\endisadelimproof
  | 
| 
11866
 | 
   217  | 
%
  | 
| 
10123
 | 
   218  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   219  | 
\noindent
  | 
| 
10867
 | 
   220  | 
We assume the negation of the conclusion and prove \isa{s\ {\isasymin}\ lfp\ {\isacharparenleft}af\ A{\isacharparenright}}.
 | 
| 
10237
 | 
   221  | 
Unfolding \isa{lfp} once and
 | 
| 
10123
 | 
   222  | 
simplifying with the definition of \isa{af} finishes the proof.
 | 
| 
 | 
   223  | 
  | 
| 
 | 
   224  | 
Now we iterate this process. The following construction of the desired
  | 
| 
10895
 | 
   225  | 
path is parameterized by a predicate \isa{Q} that should hold along the path:%
 | 
| 
10123
 | 
   226  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   227  | 
\isamarkuptrue%
  | 
| 
 | 
   228  | 
\isacommand{consts}\isamarkupfalse%
 | 
| 
 | 
   229  | 
\ path\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}state\ {\isasymRightarrow}\ {\isacharparenleft}state\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}nat\ {\isasymRightarrow}\ state{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
   230  | 
\isacommand{primrec}\isamarkupfalse%
 | 
| 
 | 
   231  | 
\isanewline
  | 
| 
 | 
   232  | 
{\isachardoublequoteopen}path\ s\ Q\ {\isadigit{0}}\ {\isacharequal}\ s{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
   233  | 
{\isachardoublequoteopen}path\ s\ Q\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}SOME\ t{\isachardot}\ {\isacharparenleft}path\ s\ Q\ n{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
10123
 | 
   234  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   235  | 
\noindent
  | 
| 
12699
 | 
   236  | 
Element \isa{n\ {\isacharplus}\ {\isadigit{1}}} on this path is some arbitrary successor
 | 
| 
10895
 | 
   237  | 
\isa{t} of element \isa{n} such that \isa{Q\ t} holds.  Remember that \isa{SOME\ t{\isachardot}\ R\ t}
 | 
| 
10654
 | 
   238  | 
is some arbitrary but fixed \isa{t} such that \isa{R\ t} holds (see \S\ref{sec:SOME}). Of
 | 
| 
10867
 | 
   239  | 
course, such a \isa{t} need not exist, but that is of no
 | 
| 
 | 
   240  | 
concern to us since we will only use \isa{path} when a
 | 
| 
10123
 | 
   241  | 
suitable \isa{t} does exist.
 | 
| 
 | 
   242  | 
  | 
| 
10895
 | 
   243  | 
Let us show that if each state \isa{s} that satisfies \isa{Q}
 | 
| 
 | 
   244  | 
has a successor that again satisfies \isa{Q}, then there exists an infinite \isa{Q}-path:%
 | 
| 
10123
 | 
   245  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   246  | 
\isamarkuptrue%
  | 
| 
 | 
   247  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
   248  | 
\ infinity{\isacharunderscore}lemma{\isacharcolon}\isanewline
 | 
| 
 | 
   249  | 
\ \ {\isachardoublequoteopen}{\isasymlbrakk}\ Q\ s{\isacharsemicolon}\ {\isasymforall}s{\isachardot}\ Q\ s\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}\ t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\isanewline
 | 
| 
 | 
   250  | 
\ \ \ {\isasymexists}p{\isasymin}Paths\ s{\isachardot}\ {\isasymforall}i{\isachardot}\ Q{\isacharparenleft}p\ i{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
17056
 | 
   251  | 
\isadelimproof
  | 
| 
 | 
   252  | 
%
  | 
| 
 | 
   253  | 
\endisadelimproof
  | 
| 
 | 
   254  | 
%
  | 
| 
 | 
   255  | 
\isatagproof
  | 
| 
16069
 | 
   256  | 
%
  | 
| 
 | 
   257  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   258  | 
\noindent
  | 
| 
 | 
   259  | 
First we rephrase the conclusion slightly because we need to prove simultaneously
  | 
| 
 | 
   260  | 
both the path property and the fact that \isa{Q} holds:%
 | 
| 
 | 
   261  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   262  | 
\isamarkuptrue%
  | 
| 
 | 
   263  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   264  | 
{\isacharparenleft}subgoal{\isacharunderscore}tac\isanewline
 | 
| 
 | 
   265  | 
\ \ {\isachardoublequoteopen}{\isasymexists}p{\isachardot}\ s\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p{\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q{\isacharparenleft}p\ i{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}%
 | 
| 
16069
 | 
   266  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   267  | 
\noindent
  | 
| 
 | 
   268  | 
From this proposition the original goal follows easily:%
  | 
| 
 | 
   269  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   270  | 
\isamarkuptrue%
  | 
| 
 | 
   271  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   272  | 
{\isacharparenleft}simp\ add{\isacharcolon}\ Paths{\isacharunderscore}def{\isacharcomma}\ blast{\isacharparenright}%
 | 
| 
16069
 | 
   273  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   274  | 
\noindent
  | 
| 
 | 
   275  | 
The new subgoal is proved by providing the witness \isa{path\ s\ Q} for \isa{p}:%
 | 
| 
 | 
   276  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   277  | 
\isamarkuptrue%
  | 
| 
 | 
   278  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   279  | 
{\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}path\ s\ Q{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline
 | 
| 
 | 
   280  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   281  | 
{\isacharparenleft}clarsimp{\isacharparenright}%
 | 
| 
16069
 | 
   282  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   283  | 
\noindent
  | 
| 
 | 
   284  | 
After simplification and clarification, the subgoal has the following form:
  | 
| 
 | 
   285  | 
\begin{isabelle}%
 | 
| 
 | 
   286  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}i{\isachardot}\ {\isasymlbrakk}Q\ s{\isacharsemicolon}\ {\isasymforall}s{\isachardot}\ Q\ s\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}{\isasymrbrakk}\isanewline
 | 
| 
 | 
   287  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}i{\isachardot}\ }{\isasymLongrightarrow}\ {\isacharparenleft}path\ s\ Q\ i{\isacharcomma}\ SOME\ t{\isachardot}\ {\isacharparenleft}path\ s\ Q\ i{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\isanewline
 | 
| 
 | 
   288  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}i{\isachardot}\ {\isasymLongrightarrow}\ }Q\ {\isacharparenleft}path\ s\ Q\ i{\isacharparenright}%
 | 
| 
 | 
   289  | 
\end{isabelle}
 | 
| 
 | 
   290  | 
It invites a proof by induction on \isa{i}:%
 | 
| 
 | 
   291  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   292  | 
\isamarkuptrue%
  | 
| 
 | 
   293  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   294  | 
{\isacharparenleft}induct{\isacharunderscore}tac\ i{\isacharparenright}\isanewline
 | 
| 
 | 
   295  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   296  | 
{\isacharparenleft}simp{\isacharparenright}%
 | 
| 
16069
 | 
   297  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   298  | 
\noindent
  | 
| 
 | 
   299  | 
After simplification, the base case boils down to
  | 
| 
 | 
   300  | 
\begin{isabelle}%
 | 
| 
 | 
   301  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}Q\ s{\isacharsemicolon}\ {\isasymforall}s{\isachardot}\ Q\ s\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}{\isasymrbrakk}\isanewline
 | 
| 
 | 
   302  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ }{\isasymLongrightarrow}\ {\isacharparenleft}s{\isacharcomma}\ SOME\ t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}\ {\isasymin}\ M%
 | 
| 
 | 
   303  | 
\end{isabelle}
 | 
| 
 | 
   304  | 
The conclusion looks exceedingly trivial: after all, \isa{t} is chosen such that \isa{{\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M}
 | 
| 
 | 
   305  | 
holds. However, we first have to show that such a \isa{t} actually exists! This reasoning
 | 
| 
 | 
   306  | 
is embodied in the theorem \isa{someI{\isadigit{2}}{\isacharunderscore}ex}:
 | 
| 
 | 
   307  | 
\begin{isabelle}%
 | 
| 
 | 
   308  | 
\ \ \ \ \ {\isasymlbrakk}{\isasymexists}a{\isachardot}\ {\isacharquery}P\ a{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ {\isacharquery}P\ x\ {\isasymLongrightarrow}\ {\isacharquery}Q\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}Q\ {\isacharparenleft}SOME\ x{\isachardot}\ {\isacharquery}P\ x{\isacharparenright}%
 | 
| 
 | 
   309  | 
\end{isabelle}
 | 
| 
 | 
   310  | 
When we apply this theorem as an introduction rule, \isa{{\isacharquery}P\ x} becomes
 | 
| 
 | 
   311  | 
\isa{{\isacharparenleft}s{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ x} and \isa{{\isacharquery}Q\ x} becomes \isa{{\isacharparenleft}s{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ M} and we have to prove
 | 
| 
 | 
   312  | 
two subgoals: \isa{{\isasymexists}a{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ a{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ a}, which follows from the assumptions, and
 | 
| 
 | 
   313  | 
\isa{{\isacharparenleft}s{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ x\ {\isasymLongrightarrow}\ {\isacharparenleft}s{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ M}, which is trivial. Thus it is not surprising that
 | 
| 
 | 
   314  | 
\isa{fast} can prove the base case quickly:%
 | 
| 
 | 
   315  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   316  | 
\isamarkuptrue%
  | 
| 
 | 
   317  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   318  | 
{\isacharparenleft}fast\ intro{\isacharcolon}\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharparenright}%
 | 
| 
16069
 | 
   319  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   320  | 
\noindent
  | 
| 
 | 
   321  | 
What is worth noting here is that we have used \methdx{fast} rather than
 | 
| 
 | 
   322  | 
\isa{blast}.  The reason is that \isa{blast} would fail because it cannot
 | 
| 
 | 
   323  | 
cope with \isa{someI{\isadigit{2}}{\isacharunderscore}ex}: unifying its conclusion with the current
 | 
| 
 | 
   324  | 
subgoal is non-trivial because of the nested schematic variables. For
  | 
| 
 | 
   325  | 
efficiency reasons \isa{blast} does not even attempt such unifications.
 | 
| 
 | 
   326  | 
Although \isa{fast} can in principle cope with complicated unification
 | 
| 
 | 
   327  | 
problems, in practice the number of unifiers arising is often prohibitive and
  | 
| 
 | 
   328  | 
the offending rule may need to be applied explicitly rather than
  | 
| 
 | 
   329  | 
automatically. This is what happens in the step case.
  | 
| 
 | 
   330  | 
  | 
| 
 | 
   331  | 
The induction step is similar, but more involved, because now we face nested
  | 
| 
 | 
   332  | 
occurrences of \isa{SOME}. As a result, \isa{fast} is no longer able to
 | 
| 
 | 
   333  | 
solve the subgoal and we apply \isa{someI{\isadigit{2}}{\isacharunderscore}ex} by hand.  We merely
 | 
| 
 | 
   334  | 
show the proof commands but do not describe the details:%
  | 
| 
 | 
   335  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   336  | 
\isamarkuptrue%
  | 
| 
 | 
   337  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   338  | 
{\isacharparenleft}simp{\isacharparenright}\isanewline
 | 
| 
 | 
   339  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   340  | 
{\isacharparenleft}rule\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharparenright}\isanewline
 | 
| 
 | 
   341  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   342  | 
{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
| 
 | 
   343  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   344  | 
{\isacharparenleft}rule\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharparenright}\isanewline
 | 
| 
 | 
   345  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   346  | 
{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
| 
 | 
   347  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   348  | 
{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
| 
 | 
   349  | 
\isacommand{done}\isamarkupfalse%
 | 
| 
 | 
   350  | 
%
  | 
| 
17056
 | 
   351  | 
\endisatagproof
  | 
| 
 | 
   352  | 
{\isafoldproof}%
 | 
| 
 | 
   353  | 
%
  | 
| 
 | 
   354  | 
\isadelimproof
  | 
| 
 | 
   355  | 
%
  | 
| 
 | 
   356  | 
\endisadelimproof
  | 
| 
11866
 | 
   357  | 
%
  | 
| 
10159
 | 
   358  | 
\begin{isamarkuptext}%
 | 
| 
10867
 | 
   359  | 
Function \isa{path} has fulfilled its purpose now and can be forgotten.
 | 
| 
 | 
   360  | 
It was merely defined to provide the witness in the proof of the
  | 
| 
10171
 | 
   361  | 
\isa{infinity{\isacharunderscore}lemma}. Aficionados of minimal proofs might like to know
 | 
| 
10159
 | 
   362  | 
that we could have given the witness without having to define a new function:
  | 
| 
 | 
   363  | 
the term
  | 
| 
 | 
   364  | 
\begin{isabelle}%
 | 
| 
10895
 | 
   365  | 
\ \ \ \ \ nat{\isacharunderscore}rec\ s\ {\isacharparenleft}{\isasymlambda}n\ t{\isachardot}\ SOME\ u{\isachardot}\ {\isacharparenleft}t{\isacharcomma}\ u{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ u{\isacharparenright}%
 | 
| 
10159
 | 
   366  | 
\end{isabelle}
 | 
| 
10895
 | 
   367  | 
is extensionally equal to \isa{path\ s\ Q},
 | 
| 
10867
 | 
   368  | 
where \isa{nat{\isacharunderscore}rec} is the predefined primitive recursor on \isa{nat}.%
 | 
| 
10159
 | 
   369  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   370  | 
\isamarkuptrue%
  | 
| 
17056
 | 
   371  | 
%
  | 
| 
 | 
   372  | 
\isadelimproof
  | 
| 
 | 
   373  | 
%
  | 
| 
 | 
   374  | 
\endisadelimproof
  | 
| 
 | 
   375  | 
%
  | 
| 
 | 
   376  | 
\isatagproof
  | 
| 
 | 
   377  | 
%
  | 
| 
 | 
   378  | 
\endisatagproof
  | 
| 
 | 
   379  | 
{\isafoldproof}%
 | 
| 
 | 
   380  | 
%
  | 
| 
 | 
   381  | 
\isadelimproof
  | 
| 
 | 
   382  | 
%
  | 
| 
 | 
   383  | 
\endisadelimproof
  | 
| 
10159
 | 
   384  | 
%
  | 
| 
 | 
   385  | 
\begin{isamarkuptext}%
 | 
| 
10187
 | 
   386  | 
At last we can prove the opposite direction of \isa{AF{\isacharunderscore}lemma{\isadigit{1}}}:%
 | 
| 
10159
 | 
   387  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   388  | 
\isamarkuptrue%
  | 
| 
 | 
   389  | 
\isacommand{theorem}\isamarkupfalse%
 | 
| 
 | 
   390  | 
\ AF{\isacharunderscore}lemma{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}\ {\isasymsubseteq}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
17056
 | 
   391  | 
\isadelimproof
  | 
| 
 | 
   392  | 
%
  | 
| 
 | 
   393  | 
\endisadelimproof
  | 
| 
 | 
   394  | 
%
  | 
| 
 | 
   395  | 
\isatagproof
  | 
| 
16069
 | 
   396  | 
%
  | 
| 
 | 
   397  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   398  | 
\noindent
  | 
| 
 | 
   399  | 
The proof is again pointwise and then by contraposition:%
  | 
| 
 | 
   400  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   401  | 
\isamarkuptrue%
  | 
| 
 | 
   402  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   403  | 
{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline
 | 
| 
 | 
   404  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   405  | 
{\isacharparenleft}erule\ contrapos{\isacharunderscore}pp{\isacharparenright}\isanewline
 | 
| 
 | 
   406  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   407  | 
\ simp%
  | 
| 
16069
 | 
   408  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   409  | 
\begin{isabelle}%
 | 
| 
 | 
   410  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymnotin}\ lfp\ {\isacharparenleft}af\ A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymexists}p{\isasymin}Paths\ x{\isachardot}\ {\isasymforall}i{\isachardot}\ p\ i\ {\isasymnotin}\ A%
 | 
| 
 | 
   411  | 
\end{isabelle}
 | 
| 
 | 
   412  | 
Applying the \isa{infinity{\isacharunderscore}lemma} as a destruction rule leaves two subgoals, the second
 | 
| 
 | 
   413  | 
premise of \isa{infinity{\isacharunderscore}lemma} and the original subgoal:%
 | 
| 
 | 
   414  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   415  | 
\isamarkuptrue%
  | 
| 
 | 
   416  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   417  | 
{\isacharparenleft}drule\ infinity{\isacharunderscore}lemma{\isacharparenright}%
 | 
| 
16069
 | 
   418  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   419  | 
\begin{isabelle}%
 | 
| 
 | 
   420  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isasymforall}s{\isachardot}\ s\ {\isasymnotin}\ lfp\ {\isacharparenleft}af\ A{\isacharparenright}\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ t\ {\isasymnotin}\ lfp\ {\isacharparenleft}af\ A{\isacharparenright}{\isacharparenright}\isanewline
 | 
| 
 | 
   421  | 
\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isasymexists}p{\isasymin}Paths\ x{\isachardot}\ {\isasymforall}i{\isachardot}\ p\ i\ {\isasymnotin}\ lfp\ {\isacharparenleft}af\ A{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
 | 
| 
 | 
   422  | 
\isaindent{\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ }{\isasymexists}p{\isasymin}Paths\ x{\isachardot}\ {\isasymforall}i{\isachardot}\ p\ i\ {\isasymnotin}\ A%
 | 
| 
 | 
   423  | 
\end{isabelle}
 | 
| 
 | 
   424  | 
Both are solved automatically:%
  | 
| 
 | 
   425  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   426  | 
\isamarkuptrue%
  | 
| 
 | 
   427  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   428  | 
{\isacharparenleft}auto\ dest{\isacharcolon}\ not{\isacharunderscore}in{\isacharunderscore}lfp{\isacharunderscore}afD{\isacharparenright}\isanewline
 | 
| 
 | 
   429  | 
\isacommand{done}\isamarkupfalse%
 | 
| 
 | 
   430  | 
%
  | 
| 
17056
 | 
   431  | 
\endisatagproof
  | 
| 
 | 
   432  | 
{\isafoldproof}%
 | 
| 
 | 
   433  | 
%
  | 
| 
 | 
   434  | 
\isadelimproof
  | 
| 
 | 
   435  | 
%
  | 
| 
 | 
   436  | 
\endisadelimproof
  | 
| 
11866
 | 
   437  | 
%
  | 
| 
10159
 | 
   438  | 
\begin{isamarkuptext}%
 | 
| 
10867
 | 
   439  | 
If you find these proofs too complicated, we recommend that you read
  | 
| 
 | 
   440  | 
\S\ref{sec:CTL-revisited}, where we show how inductive definitions lead to
 | 
| 
10217
 | 
   441  | 
simpler arguments.
  | 
| 
 | 
   442  | 
  | 
| 
 | 
   443  | 
The main theorem is proved as for PDL, except that we also derive the
  | 
| 
 | 
   444  | 
necessary equality \isa{lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isacharequal}\ {\isachardot}{\isachardot}{\isachardot}} by combining
 | 
| 
 | 
   445  | 
\isa{AF{\isacharunderscore}lemma{\isadigit{1}}} and \isa{AF{\isacharunderscore}lemma{\isadigit{2}}} on the spot:%
 | 
| 
10159
 | 
   446  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   447  | 
\isamarkuptrue%
  | 
| 
 | 
   448  | 
\isacommand{theorem}\isamarkupfalse%
 | 
| 
 | 
   449  | 
\ {\isachardoublequoteopen}mc\ f\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ f{\isacharbraceright}{\isachardoublequoteclose}\isanewline
 | 
| 
17056
 | 
   450  | 
%
  | 
| 
 | 
   451  | 
\isadelimproof
  | 
| 
 | 
   452  | 
%
  | 
| 
 | 
   453  | 
\endisadelimproof
  | 
| 
 | 
   454  | 
%
  | 
| 
 | 
   455  | 
\isatagproof
  | 
| 
17175
 | 
   456  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   457  | 
{\isacharparenleft}induct{\isacharunderscore}tac\ f{\isacharparenright}\isanewline
 | 
| 
 | 
   458  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   459  | 
{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ EF{\isacharunderscore}lemma\ equalityI{\isacharbrackleft}OF\ AF{\isacharunderscore}lemma{\isadigit{1}}\ AF{\isacharunderscore}lemma{\isadigit{2}}{\isacharbrackright}{\isacharparenright}\isanewline
 | 
| 
 | 
   460  | 
\isacommand{done}\isamarkupfalse%
 | 
| 
 | 
   461  | 
%
  | 
| 
17056
 | 
   462  | 
\endisatagproof
  | 
| 
 | 
   463  | 
{\isafoldproof}%
 | 
| 
 | 
   464  | 
%
  | 
| 
 | 
   465  | 
\isadelimproof
  | 
| 
 | 
   466  | 
%
  | 
| 
 | 
   467  | 
\endisadelimproof
  | 
| 
11866
 | 
   468  | 
%
  | 
| 
10159
 | 
   469  | 
\begin{isamarkuptext}%
 | 
| 
10867
 | 
   470  | 
The language defined above is not quite CTL\@. The latter also includes an
  | 
| 
10983
 | 
   471  | 
until-operator \isa{EU\ f\ g} with semantics ``there \emph{E}xists a path
 | 
| 
11494
 | 
   472  | 
where \isa{f} is true \emph{U}ntil \isa{g} becomes true''.  We need
 | 
| 
 | 
   473  | 
an auxiliary function:%
  | 
| 
10281
 | 
   474  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   475  | 
\isamarkuptrue%
  | 
| 
 | 
   476  | 
\isacommand{consts}\isamarkupfalse%
 | 
| 
 | 
   477  | 
\ until{\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}state\ set\ {\isasymRightarrow}\ state\ set\ {\isasymRightarrow}\ state\ {\isasymRightarrow}\ state\ list\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
   478  | 
\isacommand{primrec}\isamarkupfalse%
 | 
| 
 | 
   479  | 
\isanewline
  | 
| 
 | 
   480  | 
{\isachardoublequoteopen}until\ A\ B\ s\ {\isacharbrackleft}{\isacharbrackright}\ \ \ \ {\isacharequal}\ {\isacharparenleft}s\ {\isasymin}\ B{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
17181
 | 
   481  | 
{\isachardoublequoteopen}until\ A\ B\ s\ {\isacharparenleft}t{\isacharhash}p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}s\ {\isasymin}\ A\ {\isasymand}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ until\ A\ B\ t\ p{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
10281
 | 
   482  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   483  | 
\noindent
  | 
| 
11494
 | 
   484  | 
Expressing the semantics of \isa{EU} is now straightforward:
 | 
| 
10171
 | 
   485  | 
\begin{isabelle}%
 | 
| 
10983
 | 
   486  | 
\ \ \ \ \ s\ {\isasymTurnstile}\ EU\ f\ g\ {\isacharequal}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ until\ {\isacharbraceleft}t{\isachardot}\ t\ {\isasymTurnstile}\ f{\isacharbraceright}\ {\isacharbraceleft}t{\isachardot}\ t\ {\isasymTurnstile}\ g{\isacharbraceright}\ s\ p{\isacharparenright}%
 | 
| 
10171
 | 
   487  | 
\end{isabelle}
 | 
| 
10281
 | 
   488  | 
Note that \isa{EU} is not definable in terms of the other operators!
 | 
| 
 | 
   489  | 
  | 
| 
 | 
   490  | 
Model checking \isa{EU} is again a least fixed point construction:
 | 
| 
10171
 | 
   491  | 
\begin{isabelle}%
 | 
| 
10839
 | 
   492  | 
\ \ \ \ \ mc{\isacharparenleft}EU\ f\ g{\isacharparenright}\ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ g\ {\isasymunion}\ mc\ f\ {\isasyminter}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}%
 | 
| 
10171
 | 
   493  | 
\end{isabelle}
 | 
| 
10281
 | 
   494  | 
  | 
| 
 | 
   495  | 
\begin{exercise}
 | 
| 
 | 
   496  | 
Extend the datatype of formulae by the above until operator
  | 
| 
 | 
   497  | 
and prove the equivalence between semantics and model checking, i.e.\ that
  | 
| 
10186
 | 
   498  | 
\begin{isabelle}%
 | 
| 
 | 
   499  | 
\ \ \ \ \ mc\ {\isacharparenleft}EU\ f\ g{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ EU\ f\ g{\isacharbraceright}%
 | 
| 
 | 
   500  | 
\end{isabelle}
 | 
| 
 | 
   501  | 
%For readability you may want to annotate {term EU} with its customary syntax
 | 
| 
 | 
   502  | 
%{text[display]"| EU formula formula    E[_ U _]"}
 | 
| 
 | 
   503  | 
%which enables you to read and write {text"E[f U g]"} instead of {term"EU f g"}.
 | 
| 
 | 
   504  | 
\end{exercise}
 | 
| 
10867
 | 
   505  | 
For more CTL exercises see, for example, Huth and Ryan \cite{Huth-Ryan-book}.%
 | 
| 
10281
 | 
   506  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   507  | 
\isamarkuptrue%
  | 
| 
17056
 | 
   508  | 
%
  | 
| 
 | 
   509  | 
\isadelimproof
  | 
| 
 | 
   510  | 
%
  | 
| 
 | 
   511  | 
\endisadelimproof
  | 
| 
 | 
   512  | 
%
  | 
| 
 | 
   513  | 
\isatagproof
  | 
| 
 | 
   514  | 
%
  | 
| 
 | 
   515  | 
\endisatagproof
  | 
| 
 | 
   516  | 
{\isafoldproof}%
 | 
| 
 | 
   517  | 
%
  | 
| 
 | 
   518  | 
\isadelimproof
  | 
| 
 | 
   519  | 
%
  | 
| 
 | 
   520  | 
\endisadelimproof
  | 
| 
 | 
   521  | 
%
  | 
| 
 | 
   522  | 
\isadelimproof
  | 
| 
 | 
   523  | 
%
  | 
| 
 | 
   524  | 
\endisadelimproof
  | 
| 
 | 
   525  | 
%
  | 
| 
 | 
   526  | 
\isatagproof
  | 
| 
 | 
   527  | 
%
  | 
| 
 | 
   528  | 
\endisatagproof
  | 
| 
 | 
   529  | 
{\isafoldproof}%
 | 
| 
 | 
   530  | 
%
  | 
| 
 | 
   531  | 
\isadelimproof
  | 
| 
 | 
   532  | 
%
  | 
| 
 | 
   533  | 
\endisadelimproof
  | 
| 
 | 
   534  | 
%
  | 
| 
 | 
   535  | 
\isadelimproof
  | 
| 
 | 
   536  | 
%
  | 
| 
 | 
   537  | 
\endisadelimproof
  | 
| 
 | 
   538  | 
%
  | 
| 
 | 
   539  | 
\isatagproof
  | 
| 
 | 
   540  | 
%
  | 
| 
 | 
   541  | 
\endisatagproof
  | 
| 
 | 
   542  | 
{\isafoldproof}%
 | 
| 
 | 
   543  | 
%
  | 
| 
 | 
   544  | 
\isadelimproof
  | 
| 
 | 
   545  | 
%
  | 
| 
 | 
   546  | 
\endisadelimproof
  | 
| 
10281
 | 
   547  | 
%
  | 
| 
 | 
   548  | 
\begin{isamarkuptext}%
 | 
| 
12334
 | 
   549  | 
Let us close this section with a few words about the executability of
  | 
| 
 | 
   550  | 
our model checkers.  It is clear that if all sets are finite, they can be
  | 
| 
 | 
   551  | 
represented as lists and the usual set operations are easily
  | 
| 
 | 
   552  | 
implemented. Only \isa{lfp} requires a little thought.  Fortunately, theory
 | 
| 
12473
 | 
   553  | 
\isa{While{\isacharunderscore}Combinator} in the Library~\cite{HOL-Library} provides a
 | 
| 
12334
 | 
   554  | 
theorem stating that in the case of finite sets and a monotone
  | 
| 
 | 
   555  | 
function~\isa{F}, the value of \mbox{\isa{lfp\ F}} can be computed by
 | 
| 
 | 
   556  | 
iterated application of \isa{F} to~\isa{{\isacharbraceleft}{\isacharbraceright}} until a fixed point is
 | 
| 
 | 
   557  | 
reached. It is actually possible to generate executable functional programs
  | 
| 
10159
 | 
   558  | 
from HOL definitions, but that is beyond the scope of the tutorial.%
  | 
| 
11494
 | 
   559  | 
\index{CTL|)}%
 | 
| 
10159
 | 
   560  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   561  | 
\isamarkuptrue%
  | 
| 
17056
 | 
   562  | 
%
  | 
| 
 | 
   563  | 
\isadelimtheory
  | 
| 
 | 
   564  | 
%
  | 
| 
 | 
   565  | 
\endisadelimtheory
  | 
| 
 | 
   566  | 
%
  | 
| 
 | 
   567  | 
\isatagtheory
  | 
| 
 | 
   568  | 
%
  | 
| 
 | 
   569  | 
\endisatagtheory
  | 
| 
 | 
   570  | 
{\isafoldtheory}%
 | 
| 
 | 
   571  | 
%
  | 
| 
 | 
   572  | 
\isadelimtheory
  | 
| 
 | 
   573  | 
%
  | 
| 
 | 
   574  | 
\endisadelimtheory
  | 
| 
10123
 | 
   575  | 
\end{isabellebody}%
 | 
| 
 | 
   576  | 
%%% Local Variables:
  | 
| 
 | 
   577  | 
%%% mode: latex
  | 
| 
 | 
   578  | 
%%% TeX-master: "root"
  | 
| 
 | 
   579  | 
%%% End:
  |