| 17914 |      1 | (*<*)theory Star imports Main begin(*>*)
 | 
| 10225 |      2 | 
 | 
| 10898 |      3 | section{*The Reflexive Transitive Closure*}
 | 
| 10225 |      4 | 
 | 
| 10242 |      5 | text{*\label{sec:rtc}
 | 
| 11494 |      6 | \index{reflexive transitive closure!defining inductively|(}%
 | 
| 10898 |      7 | An inductive definition may accept parameters, so it can express 
 | 
|  |      8 | functions that yield sets.
 | 
|  |      9 | Relations too can be defined inductively, since they are just sets of pairs.
 | 
|  |     10 | A perfect example is the function that maps a relation to its
 | 
|  |     11 | reflexive transitive closure.  This concept was already
 | 
| 11147 |     12 | introduced in \S\ref{sec:Relations}, where the operator @{text"\<^sup>*"} was
 | 
| 10520 |     13 | defined as a least fixed point because inductive definitions were not yet
 | 
|  |     14 | available. But now they are:
 | 
| 10225 |     15 | *}
 | 
|  |     16 | 
 | 
| 10242 |     17 | consts rtc :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set"   ("_*" [1000] 999)
 | 
| 10225 |     18 | inductive "r*"
 | 
|  |     19 | intros
 | 
| 10242 |     20 | rtc_refl[iff]:  "(x,x) \<in> r*"
 | 
|  |     21 | rtc_step:       "\<lbrakk> (x,y) \<in> r; (y,z) \<in> r* \<rbrakk> \<Longrightarrow> (x,z) \<in> r*"
 | 
|  |     22 | 
 | 
|  |     23 | text{*\noindent
 | 
|  |     24 | The function @{term rtc} is annotated with concrete syntax: instead of
 | 
| 11494 |     25 | @{text"rtc r"} we can write @{term"r*"}. The actual definition
 | 
| 10520 |     26 | consists of two rules. Reflexivity is obvious and is immediately given the
 | 
|  |     27 | @{text iff} attribute to increase automation. The
 | 
| 10363 |     28 | second rule, @{thm[source]rtc_step}, says that we can always add one more
 | 
|  |     29 | @{term r}-step to the left. Although we could make @{thm[source]rtc_step} an
 | 
| 10520 |     30 | introduction rule, this is dangerous: the recursion in the second premise
 | 
|  |     31 | slows down and may even kill the automatic tactics.
 | 
| 10242 |     32 | 
 | 
|  |     33 | The above definition of the concept of reflexive transitive closure may
 | 
|  |     34 | be sufficiently intuitive but it is certainly not the only possible one:
 | 
| 10898 |     35 | for a start, it does not even mention transitivity.
 | 
| 10242 |     36 | The rest of this section is devoted to proving that it is equivalent to
 | 
| 10898 |     37 | the standard definition. We start with a simple lemma:
 | 
| 10242 |     38 | *}
 | 
| 10225 |     39 | 
 | 
| 11308 |     40 | lemma [intro]: "(x,y) \<in> r \<Longrightarrow> (x,y) \<in> r*"
 | 
| 10237 |     41 | by(blast intro: rtc_step);
 | 
| 10225 |     42 | 
 | 
| 10242 |     43 | text{*\noindent
 | 
|  |     44 | Although the lemma itself is an unremarkable consequence of the basic rules,
 | 
|  |     45 | it has the advantage that it can be declared an introduction rule without the
 | 
|  |     46 | danger of killing the automatic tactics because @{term"r*"} occurs only in
 | 
|  |     47 | the conclusion and not in the premise. Thus some proofs that would otherwise
 | 
|  |     48 | need @{thm[source]rtc_step} can now be found automatically. The proof also
 | 
| 10898 |     49 | shows that @{text blast} is able to handle @{thm[source]rtc_step}. But
 | 
| 10242 |     50 | some of the other automatic tactics are more sensitive, and even @{text
 | 
|  |     51 | blast} can be lead astray in the presence of large numbers of rules.
 | 
|  |     52 | 
 | 
| 10520 |     53 | To prove transitivity, we need rule induction, i.e.\ theorem
 | 
|  |     54 | @{thm[source]rtc.induct}:
 | 
|  |     55 | @{thm[display]rtc.induct}
 | 
|  |     56 | It says that @{text"?P"} holds for an arbitrary pair @{text"(?xb,?xa) \<in>
 | 
|  |     57 | ?r*"} if @{text"?P"} is preserved by all rules of the inductive definition,
 | 
|  |     58 | i.e.\ if @{text"?P"} holds for the conclusion provided it holds for the
 | 
|  |     59 | premises. In general, rule induction for an $n$-ary inductive relation $R$
 | 
|  |     60 | expects a premise of the form $(x@1,\dots,x@n) \in R$.
 | 
|  |     61 | 
 | 
|  |     62 | Now we turn to the inductive proof of transitivity:
 | 
| 10242 |     63 | *}
 | 
|  |     64 | 
 | 
| 10520 |     65 | lemma rtc_trans: "\<lbrakk> (x,y) \<in> r*; (y,z) \<in> r* \<rbrakk> \<Longrightarrow> (x,z) \<in> r*"
 | 
| 10363 |     66 | apply(erule rtc.induct)
 | 
|  |     67 | 
 | 
|  |     68 | txt{*\noindent
 | 
| 11494 |     69 | Unfortunately, even the base case is a problem:
 | 
| 10363 |     70 | @{subgoals[display,indent=0,goals_limit=1]}
 | 
| 11494 |     71 | We have to abandon this proof attempt.
 | 
| 10520 |     72 | To understand what is going on, let us look again at @{thm[source]rtc.induct}.
 | 
|  |     73 | In the above application of @{text erule}, the first premise of
 | 
|  |     74 | @{thm[source]rtc.induct} is unified with the first suitable assumption, which
 | 
|  |     75 | is @{term"(x,y) \<in> r*"} rather than @{term"(y,z) \<in> r*"}. Although that
 | 
|  |     76 | is what we want, it is merely due to the order in which the assumptions occur
 | 
|  |     77 | in the subgoal, which it is not good practice to rely on. As a result,
 | 
|  |     78 | @{text"?xb"} becomes @{term x}, @{text"?xa"} becomes
 | 
|  |     79 | @{term y} and @{text"?P"} becomes @{term"%u v. (u,z) : r*"}, thus
 | 
| 10242 |     80 | yielding the above subgoal. So what went wrong?
 | 
|  |     81 | 
 | 
| 10520 |     82 | When looking at the instantiation of @{text"?P"} we see that it does not
 | 
|  |     83 | depend on its second parameter at all. The reason is that in our original
 | 
|  |     84 | goal, of the pair @{term"(x,y)"} only @{term x} appears also in the
 | 
|  |     85 | conclusion, but not @{term y}. Thus our induction statement is too
 | 
|  |     86 | weak. Fortunately, it can easily be strengthened:
 | 
| 10363 |     87 | transfer the additional premise @{prop"(y,z):r*"} into the conclusion:*}
 | 
|  |     88 | (*<*)oops(*>*)
 | 
| 10242 |     89 | lemma rtc_trans[rule_format]:
 | 
|  |     90 |   "(x,y) \<in> r* \<Longrightarrow> (y,z) \<in> r* \<longrightarrow> (x,z) \<in> r*"
 | 
|  |     91 | 
 | 
|  |     92 | txt{*\noindent
 | 
|  |     93 | This is not an obscure trick but a generally applicable heuristic:
 | 
|  |     94 | \begin{quote}\em
 | 
| 11257 |     95 | When proving a statement by rule induction on $(x@1,\dots,x@n) \in R$,
 | 
| 10242 |     96 | pull all other premises containing any of the $x@i$ into the conclusion
 | 
|  |     97 | using $\longrightarrow$.
 | 
|  |     98 | \end{quote}
 | 
|  |     99 | A similar heuristic for other kinds of inductions is formulated in
 | 
|  |    100 | \S\ref{sec:ind-var-in-prems}. The @{text rule_format} directive turns
 | 
| 11147 |    101 | @{text"\<longrightarrow>"} back into @{text"\<Longrightarrow>"}: in the end we obtain the original
 | 
| 10242 |    102 | statement of our lemma.
 | 
|  |    103 | *}
 | 
|  |    104 | 
 | 
| 10363 |    105 | apply(erule rtc.induct)
 | 
|  |    106 | 
 | 
|  |    107 | txt{*\noindent
 | 
|  |    108 | Now induction produces two subgoals which are both proved automatically:
 | 
|  |    109 | @{subgoals[display,indent=0]}
 | 
|  |    110 | *}
 | 
|  |    111 | 
 | 
| 10225 |    112 |  apply(blast);
 | 
| 10237 |    113 | apply(blast intro: rtc_step);
 | 
| 10225 |    114 | done
 | 
|  |    115 | 
 | 
| 10242 |    116 | text{*
 | 
|  |    117 | Let us now prove that @{term"r*"} is really the reflexive transitive closure
 | 
|  |    118 | of @{term r}, i.e.\ the least reflexive and transitive
 | 
|  |    119 | relation containing @{term r}. The latter is easily formalized
 | 
|  |    120 | *}
 | 
| 10225 |    121 | 
 | 
| 10237 |    122 | consts rtc2 :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set"
 | 
|  |    123 | inductive "rtc2 r"
 | 
| 10225 |    124 | intros
 | 
| 10237 |    125 | "(x,y) \<in> r \<Longrightarrow> (x,y) \<in> rtc2 r"
 | 
|  |    126 | "(x,x) \<in> rtc2 r"
 | 
|  |    127 | "\<lbrakk> (x,y) \<in> rtc2 r; (y,z) \<in> rtc2 r \<rbrakk> \<Longrightarrow> (x,z) \<in> rtc2 r"
 | 
| 10225 |    128 | 
 | 
| 10242 |    129 | text{*\noindent
 | 
|  |    130 | and the equivalence of the two definitions is easily shown by the obvious rule
 | 
| 10237 |    131 | inductions:
 | 
|  |    132 | *}
 | 
| 10225 |    133 | 
 | 
| 10237 |    134 | lemma "(x,y) \<in> rtc2 r \<Longrightarrow> (x,y) \<in> r*"
 | 
|  |    135 | apply(erule rtc2.induct);
 | 
| 10225 |    136 |   apply(blast);
 | 
| 10237 |    137 |  apply(blast);
 | 
|  |    138 | apply(blast intro: rtc_trans);
 | 
|  |    139 | done
 | 
|  |    140 | 
 | 
|  |    141 | lemma "(x,y) \<in> r* \<Longrightarrow> (x,y) \<in> rtc2 r"
 | 
|  |    142 | apply(erule rtc.induct);
 | 
|  |    143 |  apply(blast intro: rtc2.intros);
 | 
|  |    144 | apply(blast intro: rtc2.intros);
 | 
| 10225 |    145 | done
 | 
|  |    146 | 
 | 
| 10242 |    147 | text{*
 | 
|  |    148 | So why did we start with the first definition? Because it is simpler. It
 | 
|  |    149 | contains only two rules, and the single step rule is simpler than
 | 
|  |    150 | transitivity.  As a consequence, @{thm[source]rtc.induct} is simpler than
 | 
| 10898 |    151 | @{thm[source]rtc2.induct}. Since inductive proofs are hard enough
 | 
| 11147 |    152 | anyway, we should always pick the simplest induction schema available.
 | 
| 10242 |    153 | Hence @{term rtc} is the definition of choice.
 | 
| 11494 |    154 | \index{reflexive transitive closure!defining inductively|)}
 | 
| 10242 |    155 | 
 | 
| 10520 |    156 | \begin{exercise}\label{ex:converse-rtc-step}
 | 
| 10242 |    157 | Show that the converse of @{thm[source]rtc_step} also holds:
 | 
|  |    158 | @{prop[display]"[| (x,y) : r*; (y,z) : r |] ==> (x,z) : r*"}
 | 
|  |    159 | \end{exercise}
 | 
| 10520 |    160 | \begin{exercise}
 | 
|  |    161 | Repeat the development of this section, but starting with a definition of
 | 
|  |    162 | @{term rtc} where @{thm[source]rtc_step} is replaced by its converse as shown
 | 
|  |    163 | in exercise~\ref{ex:converse-rtc-step}.
 | 
|  |    164 | \end{exercise}
 | 
| 10242 |    165 | *}
 | 
|  |    166 | (*<*)
 | 
|  |    167 | lemma rtc_step2[rule_format]: "(x,y) : r* \<Longrightarrow> (y,z) : r --> (x,z) : r*"
 | 
|  |    168 | apply(erule rtc.induct);
 | 
|  |    169 |  apply blast;
 | 
| 12815 |    170 | apply(blast intro: rtc_step)
 | 
| 10242 |    171 | done
 | 
|  |    172 | 
 | 
|  |    173 | end
 | 
|  |    174 | (*>*)
 |