10225
|
1 |
(*<*)theory Star = Main:(*>*)
|
|
2 |
|
10237
|
3 |
section{*The reflexive transitive closure*}
|
10225
|
4 |
|
10242
|
5 |
text{*\label{sec:rtc}
|
|
6 |
{\bf Say something about inductive relations as opposed to sets? Or has that
|
|
7 |
been said already? If not, explain induction!}
|
|
8 |
|
10237
|
9 |
A perfect example of an inductive definition is the reflexive transitive
|
10225
|
10 |
closure of a relation. This concept was already introduced in
|
|
11 |
\S\ref{sec:rtrancl}, but it was not shown how it is defined. In fact,
|
|
12 |
the operator @{text"^*"} is not defined inductively but via a least
|
10242
|
13 |
fixed point because at that point in the theory hierarchy
|
10225
|
14 |
inductive definitions are not yet available. But now they are:
|
|
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
|
|
25 |
@{text"rtc r"} we can read and write {term"r*"}. The actual definition
|
|
26 |
consists of two rules. Reflexivity is obvious and is immediately declared an
|
|
27 |
equivalence rule. Thus the automatic tools will apply it automatically. The
|
|
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
|
|
30 |
introduction rule, this is dangerous: the recursion slows down and may
|
|
31 |
even kill the automatic tactics.
|
|
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:
|
|
35 |
for a start, it does not even mention transitivity explicitly.
|
|
36 |
The rest of this section is devoted to proving that it is equivalent to
|
|
37 |
the ``standard'' definition. We start with a simple lemma:
|
|
38 |
*}
|
10225
|
39 |
|
|
40 |
lemma [intro]: "(x,y) : 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
|
|
49 |
shows that @{term blast} is quite able to handle @{thm[source]rtc_step}. But
|
|
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 |
|
|
53 |
Let us now turn to transitivity. It should be a consequence of the definition.
|
|
54 |
*}
|
|
55 |
|
|
56 |
lemma rtc_trans:
|
|
57 |
"\<lbrakk> (x,y) \<in> r*; (y,z) \<in> r* \<rbrakk> \<Longrightarrow> (x,z) \<in> r*"
|
|
58 |
|
|
59 |
txt{*\noindent
|
|
60 |
The proof starts canonically by rule induction:
|
|
61 |
*}
|
|
62 |
|
|
63 |
apply(erule rtc.induct)(*pr(latex xsymbols symbols);*)(*<*)oops(*>*)
|
|
64 |
text{*\noindent
|
|
65 |
However, even the resulting base case is a problem
|
|
66 |
\begin{isabelle}
|
|
67 |
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}
|
|
68 |
\end{isabelle}
|
|
69 |
and maybe not what you had expected. We have to abandon this proof attempt.
|
|
70 |
To understand what is going on,
|
|
71 |
let us look at the induction rule @{thm[source]rtc.induct}:
|
|
72 |
\[ \frac{(x,y) \in r^* \qquad \bigwedge x.~P~x~x \quad \dots}{P~x~y} \]
|
|
73 |
When applying this rule, $x$ becomes @{term x}, $y$ becomes
|
|
74 |
@{term y} and $P~x~y$ becomes @{prop"(x,z) : r*"}, thus
|
|
75 |
yielding the above subgoal. So what went wrong?
|
|
76 |
|
|
77 |
When looking at the instantiation of $P~x~y$ we see
|
|
78 |
that $P$ does not depend on its second parameter at
|
|
79 |
all. The reason is that in our original goal, of the pair @{term"(x,y)"} only
|
|
80 |
@{term x} appears also in the conclusion, but not @{term y}. Thus our
|
|
81 |
induction statement is too weak. Fortunately, it can easily be strengthened:
|
|
82 |
transfer the additional premise @{prop"(y,z):r*"} into the conclusion:
|
|
83 |
*}
|
|
84 |
|
|
85 |
lemma rtc_trans[rule_format]:
|
|
86 |
"(x,y) \<in> r* \<Longrightarrow> (y,z) \<in> r* \<longrightarrow> (x,z) \<in> r*"
|
|
87 |
|
|
88 |
txt{*\noindent
|
|
89 |
This is not an obscure trick but a generally applicable heuristic:
|
|
90 |
\begin{quote}\em
|
|
91 |
Whe proving a statement by rule induction on $(x@1,\dots,x@n) \in R$,
|
|
92 |
pull all other premises containing any of the $x@i$ into the conclusion
|
|
93 |
using $\longrightarrow$.
|
|
94 |
\end{quote}
|
|
95 |
A similar heuristic for other kinds of inductions is formulated in
|
|
96 |
\S\ref{sec:ind-var-in-prems}. The @{text rule_format} directive turns
|
|
97 |
@{text"\<longrightarrow>"} back into @{text"\<Longrightarrow>"}. Thus in the end we obtain the original
|
|
98 |
statement of our lemma.
|
|
99 |
|
|
100 |
Now induction produces two subgoals which are both proved automatically:
|
|
101 |
\begin{isabelle}
|
|
102 |
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\isanewline
|
|
103 |
\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y\ za{\isachardot}\isanewline
|
|
104 |
\ \ \ \ \ \ \ {\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
|
|
105 |
\ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isacharparenleft}za{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}
|
|
106 |
\end{isabelle}
|
|
107 |
*}
|
|
108 |
|
|
109 |
apply(erule rtc.induct)(*pr(latex xsymbols symbols);*)
|
10225
|
110 |
apply(blast);
|
10237
|
111 |
apply(blast intro: rtc_step);
|
10225
|
112 |
done
|
|
113 |
|
10242
|
114 |
text{*
|
|
115 |
Let us now prove that @{term"r*"} is really the reflexive transitive closure
|
|
116 |
of @{term r}, i.e.\ the least reflexive and transitive
|
|
117 |
relation containing @{term r}. The latter is easily formalized
|
|
118 |
*}
|
10225
|
119 |
|
10237
|
120 |
consts rtc2 :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set"
|
|
121 |
inductive "rtc2 r"
|
10225
|
122 |
intros
|
10237
|
123 |
"(x,y) \<in> r \<Longrightarrow> (x,y) \<in> rtc2 r"
|
|
124 |
"(x,x) \<in> rtc2 r"
|
|
125 |
"\<lbrakk> (x,y) \<in> rtc2 r; (y,z) \<in> rtc2 r \<rbrakk> \<Longrightarrow> (x,z) \<in> rtc2 r"
|
10225
|
126 |
|
10242
|
127 |
text{*\noindent
|
|
128 |
and the equivalence of the two definitions is easily shown by the obvious rule
|
10237
|
129 |
inductions:
|
|
130 |
*}
|
10225
|
131 |
|
10237
|
132 |
lemma "(x,y) \<in> rtc2 r \<Longrightarrow> (x,y) \<in> r*"
|
|
133 |
apply(erule rtc2.induct);
|
10225
|
134 |
apply(blast);
|
10237
|
135 |
apply(blast);
|
|
136 |
apply(blast intro: rtc_trans);
|
|
137 |
done
|
|
138 |
|
|
139 |
lemma "(x,y) \<in> r* \<Longrightarrow> (x,y) \<in> rtc2 r"
|
|
140 |
apply(erule rtc.induct);
|
|
141 |
apply(blast intro: rtc2.intros);
|
|
142 |
apply(blast intro: rtc2.intros);
|
10225
|
143 |
done
|
|
144 |
|
10242
|
145 |
text{*
|
|
146 |
So why did we start with the first definition? Because it is simpler. It
|
|
147 |
contains only two rules, and the single step rule is simpler than
|
|
148 |
transitivity. As a consequence, @{thm[source]rtc.induct} is simpler than
|
|
149 |
@{thm[source]rtc2.induct}. Since inductive proofs are hard enough, we should
|
|
150 |
certainly pick the simplest induction schema available for any concept.
|
|
151 |
Hence @{term rtc} is the definition of choice.
|
|
152 |
|
|
153 |
\begin{exercise}
|
|
154 |
Show that the converse of @{thm[source]rtc_step} also holds:
|
|
155 |
@{prop[display]"[| (x,y) : r*; (y,z) : r |] ==> (x,z) : r*"}
|
|
156 |
\end{exercise}
|
|
157 |
*}
|
|
158 |
(*<*)
|
|
159 |
lemma rtc_step2[rule_format]: "(x,y) : r* \<Longrightarrow> (y,z) : r --> (x,z) : r*"
|
|
160 |
apply(erule rtc.induct);
|
|
161 |
apply blast;
|
|
162 |
apply(blast intro:rtc_step)
|
|
163 |
done
|
|
164 |
|
|
165 |
end
|
|
166 |
(*>*)
|