| 10225 |      1 | %
 | 
|  |      2 | \begin{isabellebody}%
 | 
|  |      3 | \def\isabellecontext{Star}%
 | 
| 11866 |      4 | \isamarkupfalse%
 | 
| 10225 |      5 | %
 | 
| 10878 |      6 | \isamarkupsection{The Reflexive Transitive Closure%
 | 
| 10395 |      7 | }
 | 
| 11866 |      8 | \isamarkuptrue%
 | 
| 10225 |      9 | %
 | 
|  |     10 | \begin{isamarkuptext}%
 | 
| 10242 |     11 | \label{sec:rtc}
 | 
| 11494 |     12 | \index{reflexive transitive closure!defining inductively|(}%
 | 
| 10878 |     13 | An inductive definition may accept parameters, so it can express 
 | 
|  |     14 | functions that yield sets.
 | 
|  |     15 | Relations too can be defined inductively, since they are just sets of pairs.
 | 
|  |     16 | A perfect example is the function that maps a relation to its
 | 
|  |     17 | reflexive transitive closure.  This concept was already
 | 
| 11147 |     18 | introduced in \S\ref{sec:Relations}, where the operator \isa{\isactrlsup {\isacharasterisk}} was
 | 
| 10520 |     19 | defined as a least fixed point because inductive definitions were not yet
 | 
|  |     20 | available. But now they are:%
 | 
| 10225 |     21 | \end{isamarkuptext}%
 | 
| 11866 |     22 | \isamarkuptrue%
 | 
| 10242 |     23 | \isacommand{consts}\ rtc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set{\isachardoublequote}\ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharunderscore}{\isacharasterisk}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline
 | 
| 11866 |     24 | \isamarkupfalse%
 | 
| 10225 |     25 | \isacommand{inductive}\ {\isachardoublequote}r{\isacharasterisk}{\isachardoublequote}\isanewline
 | 
|  |     26 | \isakeyword{intros}\isanewline
 | 
| 10242 |     27 | rtc{\isacharunderscore}refl{\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline
 | 
| 11866 |     28 | rtc{\isacharunderscore}step{\isacharcolon}\ \ \ \ \ \ \ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkupfalse%
 | 
|  |     29 | %
 | 
| 10242 |     30 | \begin{isamarkuptext}%
 | 
|  |     31 | \noindent
 | 
|  |     32 | The function \isa{rtc} is annotated with concrete syntax: instead of
 | 
| 11494 |     33 | \isa{rtc\ r} we can write \isa{r{\isacharasterisk}}. The actual definition
 | 
| 10520 |     34 | consists of two rules. Reflexivity is obvious and is immediately given the
 | 
|  |     35 | \isa{iff} attribute to increase automation. The
 | 
| 10363 |     36 | second rule, \isa{rtc{\isacharunderscore}step}, says that we can always add one more
 | 
|  |     37 | \isa{r}-step to the left. Although we could make \isa{rtc{\isacharunderscore}step} an
 | 
| 10520 |     38 | introduction rule, this is dangerous: the recursion in the second premise
 | 
|  |     39 | slows down and may even kill the automatic tactics.
 | 
| 10242 |     40 | 
 | 
|  |     41 | The above definition of the concept of reflexive transitive closure may
 | 
|  |     42 | be sufficiently intuitive but it is certainly not the only possible one:
 | 
| 10878 |     43 | for a start, it does not even mention transitivity.
 | 
| 10242 |     44 | The rest of this section is devoted to proving that it is equivalent to
 | 
| 10878 |     45 | the standard definition. We start with a simple lemma:%
 | 
| 10242 |     46 | \end{isamarkuptext}%
 | 
| 11866 |     47 | \isamarkuptrue%
 | 
| 11308 |     48 | \isacommand{lemma}\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline
 | 
| 11866 |     49 | \isamarkupfalse%
 | 
|  |     50 | \isacommand{by}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}\isamarkupfalse%
 | 
|  |     51 | %
 | 
| 10242 |     52 | \begin{isamarkuptext}%
 | 
|  |     53 | \noindent
 | 
|  |     54 | Although the lemma itself is an unremarkable consequence of the basic rules,
 | 
|  |     55 | it has the advantage that it can be declared an introduction rule without the
 | 
|  |     56 | danger of killing the automatic tactics because \isa{r{\isacharasterisk}} occurs only in
 | 
|  |     57 | the conclusion and not in the premise. Thus some proofs that would otherwise
 | 
|  |     58 | need \isa{rtc{\isacharunderscore}step} can now be found automatically. The proof also
 | 
| 10878 |     59 | shows that \isa{blast} is able to handle \isa{rtc{\isacharunderscore}step}. But
 | 
| 10242 |     60 | some of the other automatic tactics are more sensitive, and even \isa{blast} can be lead astray in the presence of large numbers of rules.
 | 
|  |     61 | 
 | 
| 10520 |     62 | To prove transitivity, we need rule induction, i.e.\ theorem
 | 
|  |     63 | \isa{rtc{\isachardot}induct}:
 | 
|  |     64 | \begin{isabelle}%
 | 
| 10696 |     65 | \ \ \ \ \ {\isasymlbrakk}{\isacharparenleft}{\isacharquery}xb{\isacharcomma}\ {\isacharquery}xa{\isacharparenright}\ {\isasymin}\ {\isacharquery}r{\isacharasterisk}{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ {\isacharquery}P\ x\ x{\isacharsemicolon}\isanewline
 | 
| 10950 |     66 | \isaindent{\ \ \ \ \ \ \ \ }{\isasymAnd}x\ y\ z{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ {\isacharquery}r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ {\isacharquery}r{\isacharasterisk}{\isacharsemicolon}\ {\isacharquery}P\ y\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ x\ z{\isasymrbrakk}\isanewline
 | 
|  |     67 | \isaindent{\ \ \ \ \ }{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}xb\ {\isacharquery}xa%
 | 
| 10520 |     68 | \end{isabelle}
 | 
|  |     69 | It says that \isa{{\isacharquery}P} holds for an arbitrary pair \isa{{\isacharparenleft}{\isacharquery}xb{\isacharcomma}{\isacharquery}xa{\isacharparenright}\ {\isasymin}\ {\isacharquery}r{\isacharasterisk}} if \isa{{\isacharquery}P} is preserved by all rules of the inductive definition,
 | 
|  |     70 | i.e.\ if \isa{{\isacharquery}P} holds for the conclusion provided it holds for the
 | 
|  |     71 | premises. In general, rule induction for an $n$-ary inductive relation $R$
 | 
|  |     72 | expects a premise of the form $(x@1,\dots,x@n) \in R$.
 | 
|  |     73 | 
 | 
|  |     74 | Now we turn to the inductive proof of transitivity:%
 | 
| 10242 |     75 | \end{isamarkuptext}%
 | 
| 11866 |     76 | \isamarkuptrue%
 | 
| 10520 |     77 | \isacommand{lemma}\ rtc{\isacharunderscore}trans{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline
 | 
| 11866 |     78 | \isamarkupfalse%
 | 
|  |     79 | \isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkupfalse%
 | 
|  |     80 | %
 | 
| 10363 |     81 | \begin{isamarkuptxt}%
 | 
| 10242 |     82 | \noindent
 | 
| 11494 |     83 | Unfortunately, even the base case is a problem:
 | 
| 10363 |     84 | \begin{isabelle}%
 | 
|  |     85 | \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}%
 | 
| 10242 |     86 | \end{isabelle}
 | 
| 11494 |     87 | We have to abandon this proof attempt.
 | 
| 10520 |     88 | To understand what is going on, let us look again at \isa{rtc{\isachardot}induct}.
 | 
|  |     89 | In the above application of \isa{erule}, the first premise of
 | 
|  |     90 | \isa{rtc{\isachardot}induct} is unified with the first suitable assumption, which
 | 
|  |     91 | is \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}} rather than \isa{{\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}}. Although that
 | 
|  |     92 | is what we want, it is merely due to the order in which the assumptions occur
 | 
|  |     93 | in the subgoal, which it is not good practice to rely on. As a result,
 | 
|  |     94 | \isa{{\isacharquery}xb} becomes \isa{x}, \isa{{\isacharquery}xa} becomes
 | 
|  |     95 | \isa{y} and \isa{{\isacharquery}P} becomes \isa{{\isasymlambda}u\ v{\isachardot}\ {\isacharparenleft}u{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}}, thus
 | 
| 10242 |     96 | yielding the above subgoal. So what went wrong?
 | 
|  |     97 | 
 | 
| 10520 |     98 | When looking at the instantiation of \isa{{\isacharquery}P} we see that it does not
 | 
|  |     99 | depend on its second parameter at all. The reason is that in our original
 | 
|  |    100 | goal, of the pair \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}} only \isa{x} appears also in the
 | 
|  |    101 | conclusion, but not \isa{y}. Thus our induction statement is too
 | 
|  |    102 | weak. Fortunately, it can easily be strengthened:
 | 
| 10242 |    103 | transfer the additional premise \isa{{\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}} into the conclusion:%
 | 
| 10363 |    104 | \end{isamarkuptxt}%
 | 
| 11866 |    105 | \isamarkuptrue%
 | 
|  |    106 | \isamarkupfalse%
 | 
| 10242 |    107 | \isacommand{lemma}\ rtc{\isacharunderscore}trans{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 11866 |    108 | \ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkupfalse%
 | 
|  |    109 | %
 | 
| 10242 |    110 | \begin{isamarkuptxt}%
 | 
|  |    111 | \noindent
 | 
|  |    112 | This is not an obscure trick but a generally applicable heuristic:
 | 
|  |    113 | \begin{quote}\em
 | 
| 11257 |    114 | When proving a statement by rule induction on $(x@1,\dots,x@n) \in R$,
 | 
| 10242 |    115 | pull all other premises containing any of the $x@i$ into the conclusion
 | 
|  |    116 | using $\longrightarrow$.
 | 
|  |    117 | \end{quote}
 | 
|  |    118 | A similar heuristic for other kinds of inductions is formulated in
 | 
|  |    119 | \S\ref{sec:ind-var-in-prems}. The \isa{rule{\isacharunderscore}format} directive turns
 | 
| 11147 |    120 | \isa{{\isasymlongrightarrow}} back into \isa{{\isasymLongrightarrow}}: in the end we obtain the original
 | 
| 10363 |    121 | statement of our lemma.%
 | 
|  |    122 | \end{isamarkuptxt}%
 | 
| 11866 |    123 | \isamarkuptrue%
 | 
|  |    124 | \isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkupfalse%
 | 
|  |    125 | %
 | 
| 10363 |    126 | \begin{isamarkuptxt}%
 | 
|  |    127 | \noindent
 | 
| 10242 |    128 | Now induction produces two subgoals which are both proved automatically:
 | 
| 10363 |    129 | \begin{isabelle}%
 | 
| 10242 |    130 | \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\isanewline
 | 
|  |    131 | \ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y\ za{\isachardot}\isanewline
 | 
| 10950 |    132 | \isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}\ za{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isacharsemicolon}\ {\isacharparenleft}za{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isasymrbrakk}\isanewline
 | 
|  |    133 | \isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ {\isacharparenleft}za{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}%
 | 
| 10242 |    134 | \end{isabelle}%
 | 
|  |    135 | \end{isamarkuptxt}%
 | 
| 11866 |    136 | \ \isamarkuptrue%
 | 
|  |    137 | \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
|  |    138 | \isamarkupfalse%
 | 
| 10237 |    139 | \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}\isanewline
 | 
| 11866 |    140 | \isamarkupfalse%
 | 
|  |    141 | \isacommand{done}\isamarkupfalse%
 | 
|  |    142 | %
 | 
| 10242 |    143 | \begin{isamarkuptext}%
 | 
|  |    144 | Let us now prove that \isa{r{\isacharasterisk}} is really the reflexive transitive closure
 | 
|  |    145 | of \isa{r}, i.e.\ the least reflexive and transitive
 | 
|  |    146 | relation containing \isa{r}. The latter is easily formalized%
 | 
|  |    147 | \end{isamarkuptext}%
 | 
| 11866 |    148 | \isamarkuptrue%
 | 
| 10237 |    149 | \isacommand{consts}\ rtc{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set{\isachardoublequote}\isanewline
 | 
| 11866 |    150 | \isamarkupfalse%
 | 
| 10237 |    151 | \isacommand{inductive}\ {\isachardoublequote}rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline
 | 
| 10225 |    152 | \isakeyword{intros}\isanewline
 | 
| 10237 |    153 | {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline
 | 
|  |    154 | {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline
 | 
| 11866 |    155 | {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isamarkupfalse%
 | 
|  |    156 | %
 | 
| 10237 |    157 | \begin{isamarkuptext}%
 | 
| 10242 |    158 | \noindent
 | 
|  |    159 | and the equivalence of the two definitions is easily shown by the obvious rule
 | 
| 10237 |    160 | inductions:%
 | 
|  |    161 | \end{isamarkuptext}%
 | 
| 11866 |    162 | \isamarkuptrue%
 | 
| 10237 |    163 | \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline
 | 
| 11866 |    164 | \isamarkupfalse%
 | 
| 10237 |    165 | \isacommand{apply}{\isacharparenleft}erule\ rtc{\isadigit{2}}{\isachardot}induct{\isacharparenright}\isanewline
 | 
| 11866 |    166 | \ \ \isamarkupfalse%
 | 
|  |    167 | \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
|  |    168 | \ \isamarkupfalse%
 | 
|  |    169 | \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
|  |    170 | \isamarkupfalse%
 | 
| 10237 |    171 | \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}trans{\isacharparenright}\isanewline
 | 
| 11866 |    172 | \isamarkupfalse%
 | 
| 10237 |    173 | \isacommand{done}\isanewline
 | 
| 10225 |    174 | \isanewline
 | 
| 11866 |    175 | \isamarkupfalse%
 | 
| 10237 |    176 | \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline
 | 
| 11866 |    177 | \isamarkupfalse%
 | 
| 10237 |    178 | \isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isanewline
 | 
| 11866 |    179 | \ \isamarkupfalse%
 | 
|  |    180 | \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isadigit{2}}{\isachardot}intros{\isacharparenright}\isanewline
 | 
|  |    181 | \isamarkupfalse%
 | 
| 10237 |    182 | \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isadigit{2}}{\isachardot}intros{\isacharparenright}\isanewline
 | 
| 11866 |    183 | \isamarkupfalse%
 | 
|  |    184 | \isacommand{done}\isamarkupfalse%
 | 
|  |    185 | %
 | 
| 10242 |    186 | \begin{isamarkuptext}%
 | 
|  |    187 | So why did we start with the first definition? Because it is simpler. It
 | 
|  |    188 | contains only two rules, and the single step rule is simpler than
 | 
|  |    189 | transitivity.  As a consequence, \isa{rtc{\isachardot}induct} is simpler than
 | 
| 10878 |    190 | \isa{rtc{\isadigit{2}}{\isachardot}induct}. Since inductive proofs are hard enough
 | 
| 11147 |    191 | anyway, we should always pick the simplest induction schema available.
 | 
| 10242 |    192 | Hence \isa{rtc} is the definition of choice.
 | 
| 11494 |    193 | \index{reflexive transitive closure!defining inductively|)}
 | 
| 10242 |    194 | 
 | 
| 10520 |    195 | \begin{exercise}\label{ex:converse-rtc-step}
 | 
| 10242 |    196 | Show that the converse of \isa{rtc{\isacharunderscore}step} also holds:
 | 
|  |    197 | \begin{isabelle}%
 | 
| 10696 |    198 | \ \ \ \ \ {\isasymlbrakk}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}%
 | 
| 10242 |    199 | \end{isabelle}
 | 
| 10520 |    200 | \end{exercise}
 | 
|  |    201 | \begin{exercise}
 | 
|  |    202 | Repeat the development of this section, but starting with a definition of
 | 
|  |    203 | \isa{rtc} where \isa{rtc{\isacharunderscore}step} is replaced by its converse as shown
 | 
|  |    204 | in exercise~\ref{ex:converse-rtc-step}.
 | 
| 10242 |    205 | \end{exercise}%
 | 
|  |    206 | \end{isamarkuptext}%
 | 
| 11866 |    207 | \isamarkuptrue%
 | 
|  |    208 | \isamarkupfalse%
 | 
|  |    209 | \isamarkupfalse%
 | 
|  |    210 | \isamarkupfalse%
 | 
|  |    211 | \isamarkupfalse%
 | 
|  |    212 | \isamarkupfalse%
 | 
|  |    213 | \isamarkupfalse%
 | 
| 10225 |    214 | \end{isabellebody}%
 | 
|  |    215 | %%% Local Variables:
 | 
|  |    216 | %%% mode: latex
 | 
|  |    217 | %%% TeX-master: "root"
 | 
|  |    218 | %%% End:
 |