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