10365
|
1 |
%
|
|
2 |
\begin{isabellebody}%
|
|
3 |
\def\isabellecontext{Even}%
|
17056
|
4 |
%
|
|
5 |
\isadelimtheory
|
|
6 |
%
|
|
7 |
\endisadelimtheory
|
|
8 |
%
|
|
9 |
\isatagtheory
|
23848
|
10 |
%
|
17056
|
11 |
\endisatagtheory
|
|
12 |
{\isafoldtheory}%
|
|
13 |
%
|
|
14 |
\isadelimtheory
|
|
15 |
%
|
|
16 |
\endisadelimtheory
|
23848
|
17 |
%
|
|
18 |
\isamarkupsection{The Set of Even Numbers%
|
|
19 |
}
|
|
20 |
\isamarkuptrue%
|
|
21 |
%
|
|
22 |
\begin{isamarkuptext}%
|
|
23 |
\index{even numbers!defining inductively|(}%
|
|
24 |
The set of even numbers can be inductively defined as the least set
|
|
25 |
containing 0 and closed under the operation $+2$. Obviously,
|
|
26 |
\emph{even} can also be expressed using the divides relation (\isa{dvd}).
|
|
27 |
We shall prove below that the two formulations coincide. On the way we
|
|
28 |
shall examine the primary means of reasoning about inductively defined
|
|
29 |
sets: rule induction.%
|
|
30 |
\end{isamarkuptext}%
|
|
31 |
\isamarkuptrue%
|
|
32 |
%
|
|
33 |
\isamarkupsubsection{Making an Inductive Definition%
|
|
34 |
}
|
|
35 |
\isamarkuptrue%
|
|
36 |
%
|
|
37 |
\begin{isamarkuptext}%
|
23928
|
38 |
Using \commdx{inductive\protect\_set}, we declare the constant \isa{even} to be
|
23848
|
39 |
a set of natural numbers with the desired properties.%
|
|
40 |
\end{isamarkuptext}%
|
|
41 |
\isamarkuptrue%
|
40406
|
42 |
\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse%
|
|
43 |
\ even\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ set{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
|
|
44 |
zero{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{0}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
|
|
45 |
step{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}%
|
10365
|
46 |
\begin{isamarkuptext}%
|
23848
|
47 |
An inductive definition consists of introduction rules. The first one
|
|
48 |
above states that 0 is even; the second states that if $n$ is even, then so
|
|
49 |
is~$n+2$. Given this declaration, Isabelle generates a fixed point
|
|
50 |
definition for \isa{even} and proves theorems about it,
|
|
51 |
thus following the definitional approach (see {\S}\ref{sec:definitional}).
|
|
52 |
These theorems
|
|
53 |
include the introduction rules specified in the declaration, an elimination
|
|
54 |
rule for case analysis and an induction rule. We can refer to these
|
|
55 |
theorems by automatically-generated names. Here are two examples:
|
10365
|
56 |
\begin{isabelle}%
|
40406
|
57 |
{\isadigit{0}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\rulename{even{\isaliteral{2E}{\isachardot}}zero}\par\smallskip%
|
|
58 |
n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\rulename{even{\isaliteral{2E}{\isachardot}}step}%
|
10365
|
59 |
\end{isabelle}
|
|
60 |
|
23848
|
61 |
The introduction rules can be given attributes. Here
|
|
62 |
both rules are specified as \isa{intro!},%
|
|
63 |
\index{intro"!@\isa {intro"!} (attribute)}
|
|
64 |
directing the classical reasoner to
|
|
65 |
apply them aggressively. Obviously, regarding 0 as even is safe. The
|
|
66 |
\isa{step} rule is also safe because $n+2$ is even if and only if $n$ is
|
|
67 |
even. We prove this equivalence later.%
|
|
68 |
\end{isamarkuptext}%
|
|
69 |
\isamarkuptrue%
|
|
70 |
%
|
|
71 |
\isamarkupsubsection{Using Introduction Rules%
|
|
72 |
}
|
|
73 |
\isamarkuptrue%
|
|
74 |
%
|
|
75 |
\begin{isamarkuptext}%
|
|
76 |
Our first lemma states that numbers of the form $2\times k$ are even.
|
|
77 |
Introduction rules are used to show that specific values belong to the
|
|
78 |
inductive set. Such proofs typically involve
|
|
79 |
induction, perhaps over some other inductive set.%
|
10365
|
80 |
\end{isamarkuptext}%
|
17175
|
81 |
\isamarkuptrue%
|
|
82 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
83 |
\ two{\isaliteral{5F}{\isacharunderscore}}times{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{2}}{\isaliteral{2A}{\isacharasterisk}}k\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17056
|
84 |
%
|
|
85 |
\isadelimproof
|
|
86 |
%
|
|
87 |
\endisadelimproof
|
|
88 |
%
|
|
89 |
\isatagproof
|
17175
|
90 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
91 |
\ {\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ k{\isaliteral{29}{\isacharparenright}}\isanewline
|
17175
|
92 |
\ \isacommand{apply}\isamarkupfalse%
|
|
93 |
\ auto\isanewline
|
|
94 |
\isacommand{done}\isamarkupfalse%
|
|
95 |
%
|
17056
|
96 |
\endisatagproof
|
|
97 |
{\isafoldproof}%
|
|
98 |
%
|
|
99 |
\isadelimproof
|
|
100 |
%
|
|
101 |
\endisadelimproof
|
11866
|
102 |
%
|
23848
|
103 |
\isadelimproof
|
|
104 |
%
|
|
105 |
\endisadelimproof
|
|
106 |
%
|
|
107 |
\isatagproof
|
|
108 |
%
|
|
109 |
\begin{isamarkuptxt}%
|
|
110 |
\noindent
|
|
111 |
The first step is induction on the natural number \isa{k}, which leaves
|
|
112 |
two subgoals:
|
|
113 |
\begin{isabelle}%
|
40406
|
114 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ {\isadigit{0}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\isanewline
|
|
115 |
\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ Suc\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even%
|
23848
|
116 |
\end{isabelle}
|
|
117 |
Here \isa{auto} simplifies both subgoals so that they match the introduction
|
|
118 |
rules, which are then applied automatically.
|
|
119 |
|
|
120 |
Our ultimate goal is to prove the equivalence between the traditional
|
|
121 |
definition of \isa{even} (using the divides relation) and our inductive
|
|
122 |
definition. One direction of this equivalence is immediate by the lemma
|
40406
|
123 |
just proved, whose \isa{intro{\isaliteral{21}{\isacharbang}}} attribute ensures it is applied automatically.%
|
23848
|
124 |
\end{isamarkuptxt}%
|
17175
|
125 |
\isamarkuptrue%
|
23848
|
126 |
%
|
|
127 |
\endisatagproof
|
|
128 |
{\isafoldproof}%
|
|
129 |
%
|
|
130 |
\isadelimproof
|
|
131 |
%
|
|
132 |
\endisadelimproof
|
17175
|
133 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
134 |
\ dvd{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{2}}\ dvd\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17056
|
135 |
%
|
|
136 |
\isadelimproof
|
|
137 |
%
|
|
138 |
\endisadelimproof
|
|
139 |
%
|
|
140 |
\isatagproof
|
17175
|
141 |
\isacommand{by}\isamarkupfalse%
|
40406
|
142 |
\ {\isaliteral{28}{\isacharparenleft}}auto\ simp\ add{\isaliteral{3A}{\isacharcolon}}\ dvd{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
|
17056
|
143 |
\endisatagproof
|
|
144 |
{\isafoldproof}%
|
|
145 |
%
|
|
146 |
\isadelimproof
|
|
147 |
%
|
|
148 |
\endisadelimproof
|
11866
|
149 |
%
|
23848
|
150 |
\isamarkupsubsection{Rule Induction \label{sec:rule-induction}%
|
|
151 |
}
|
|
152 |
\isamarkuptrue%
|
|
153 |
%
|
10365
|
154 |
\begin{isamarkuptext}%
|
23848
|
155 |
\index{rule induction|(}%
|
|
156 |
From the definition of the set
|
|
157 |
\isa{even}, Isabelle has
|
|
158 |
generated an induction rule:
|
|
159 |
\begin{isabelle}%
|
40406
|
160 |
{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ P\ {\isadigit{0}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
|
|
161 |
\isaindent{\ }{\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ P\ n{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
|
|
162 |
{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ x\rulename{even{\isaliteral{2E}{\isachardot}}induct}%
|
23848
|
163 |
\end{isabelle}
|
|
164 |
A property \isa{P} holds for every even number provided it
|
|
165 |
holds for~\isa{{\isadigit{0}}} and is closed under the operation
|
|
166 |
\isa{Suc(Suc \(\cdot\))}. Then \isa{P} is closed under the introduction
|
|
167 |
rules for \isa{even}, which is the least set closed under those rules.
|
|
168 |
This type of inductive argument is called \textbf{rule induction}.
|
|
169 |
|
|
170 |
Apart from the double application of \isa{Suc}, the induction rule above
|
|
171 |
resembles the familiar mathematical induction, which indeed is an instance
|
|
172 |
of rule induction; the natural numbers can be defined inductively to be
|
|
173 |
the least set containing \isa{{\isadigit{0}}} and closed under~\isa{Suc}.
|
|
174 |
|
|
175 |
Induction is the usual way of proving a property of the elements of an
|
|
176 |
inductively defined set. Let us prove that all members of the set
|
|
177 |
\isa{even} are multiples of two.%
|
10365
|
178 |
\end{isamarkuptext}%
|
17175
|
179 |
\isamarkuptrue%
|
|
180 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
181 |
\ even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}dvd{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isadigit{2}}\ dvd\ n{\isaliteral{22}{\isachardoublequoteclose}}%
|
17056
|
182 |
\isadelimproof
|
|
183 |
%
|
|
184 |
\endisadelimproof
|
|
185 |
%
|
|
186 |
\isatagproof
|
23848
|
187 |
%
|
|
188 |
\begin{isamarkuptxt}%
|
40406
|
189 |
We begin by applying induction. Note that \isa{even{\isaliteral{2E}{\isachardot}}induct} has the form
|
23848
|
190 |
of an elimination rule, so we use the method \isa{erule}. We get two
|
|
191 |
subgoals:%
|
|
192 |
\end{isamarkuptxt}%
|
|
193 |
\isamarkuptrue%
|
17175
|
194 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
195 |
\ {\isaliteral{28}{\isacharparenleft}}erule\ even{\isaliteral{2E}{\isachardot}}induct{\isaliteral{29}{\isacharparenright}}%
|
16069
|
196 |
\begin{isamarkuptxt}%
|
|
197 |
\begin{isabelle}%
|
40406
|
198 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isadigit{2}}\ dvd\ {\isadigit{0}}\isanewline
|
|
199 |
\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ {\isadigit{2}}\ dvd\ n{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isadigit{2}}\ dvd\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}%
|
23848
|
200 |
\end{isabelle}
|
|
201 |
We unfold the definition of \isa{dvd} in both subgoals, proving the first
|
|
202 |
one and simplifying the second:%
|
16069
|
203 |
\end{isamarkuptxt}%
|
17175
|
204 |
\isamarkuptrue%
|
|
205 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
206 |
\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ dvd{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
|
16069
|
207 |
\begin{isamarkuptxt}%
|
|
208 |
\begin{isabelle}%
|
40406
|
209 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}k{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ k{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}k{\isaliteral{2E}{\isachardot}}\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ k%
|
23848
|
210 |
\end{isabelle}
|
|
211 |
The next command eliminates the existential quantifier from the assumption
|
40406
|
212 |
and replaces \isa{n} by \isa{{\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ k}.%
|
16069
|
213 |
\end{isamarkuptxt}%
|
17175
|
214 |
\isamarkuptrue%
|
|
215 |
\isacommand{apply}\isamarkupfalse%
|
|
216 |
\ clarify%
|
16069
|
217 |
\begin{isamarkuptxt}%
|
|
218 |
\begin{isabelle}%
|
40406
|
219 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n\ k{\isaliteral{2E}{\isachardot}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ k\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}ka{\isaliteral{2E}{\isachardot}}\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ k{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ ka%
|
23848
|
220 |
\end{isabelle}
|
|
221 |
To conclude, we tell Isabelle that the desired value is
|
|
222 |
\isa{Suc\ k}. With this hint, the subgoal falls to \isa{simp}.%
|
16069
|
223 |
\end{isamarkuptxt}%
|
17175
|
224 |
\isamarkuptrue%
|
|
225 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
226 |
\ {\isaliteral{28}{\isacharparenleft}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ x\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}Suc\ k{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{in}\ exI{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{29}{\isacharparenright}}%
|
17056
|
227 |
\endisatagproof
|
|
228 |
{\isafoldproof}%
|
|
229 |
%
|
|
230 |
\isadelimproof
|
|
231 |
%
|
|
232 |
\endisadelimproof
|
11866
|
233 |
%
|
10365
|
234 |
\begin{isamarkuptext}%
|
23848
|
235 |
Combining the previous two results yields our objective, the
|
|
236 |
equivalence relating \isa{even} and \isa{dvd}.
|
|
237 |
%
|
|
238 |
%we don't want [iff]: discuss?%
|
10365
|
239 |
\end{isamarkuptext}%
|
17056
|
240 |
\isamarkuptrue%
|
17175
|
241 |
\isacommand{theorem}\isamarkupfalse%
|
40406
|
242 |
\ even{\isaliteral{5F}{\isacharunderscore}}iff{\isaliteral{5F}{\isacharunderscore}}dvd{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}\ dvd\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17056
|
243 |
%
|
|
244 |
\isadelimproof
|
|
245 |
%
|
|
246 |
\endisadelimproof
|
|
247 |
%
|
|
248 |
\isatagproof
|
17175
|
249 |
\isacommand{by}\isamarkupfalse%
|
40406
|
250 |
\ {\isaliteral{28}{\isacharparenleft}}blast\ intro{\isaliteral{3A}{\isacharcolon}}\ dvd{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even\ even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}dvd{\isaliteral{29}{\isacharparenright}}%
|
17056
|
251 |
\endisatagproof
|
|
252 |
{\isafoldproof}%
|
|
253 |
%
|
|
254 |
\isadelimproof
|
|
255 |
%
|
|
256 |
\endisadelimproof
|
11866
|
257 |
%
|
23848
|
258 |
\isamarkupsubsection{Generalization and Rule Induction \label{sec:gen-rule-induction}%
|
|
259 |
}
|
|
260 |
\isamarkuptrue%
|
|
261 |
%
|
10365
|
262 |
\begin{isamarkuptext}%
|
23848
|
263 |
\index{generalizing for induction}%
|
|
264 |
Before applying induction, we typically must generalize
|
|
265 |
the induction formula. With rule induction, the required generalization
|
|
266 |
can be hard to find and sometimes requires a complete reformulation of the
|
|
267 |
problem. In this example, our first attempt uses the obvious statement of
|
|
268 |
the result. It fails:%
|
10365
|
269 |
\end{isamarkuptext}%
|
17175
|
270 |
\isamarkuptrue%
|
|
271 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
272 |
\ {\isaliteral{22}{\isachardoublequoteopen}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17056
|
273 |
%
|
|
274 |
\isadelimproof
|
|
275 |
%
|
|
276 |
\endisadelimproof
|
|
277 |
%
|
|
278 |
\isatagproof
|
17175
|
279 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
280 |
\ {\isaliteral{28}{\isacharparenleft}}erule\ even{\isaliteral{2E}{\isachardot}}induct{\isaliteral{29}{\isacharparenright}}\isanewline
|
17175
|
281 |
\isacommand{oops}\isamarkupfalse%
|
|
282 |
%
|
|
283 |
\endisatagproof
|
|
284 |
{\isafoldproof}%
|
|
285 |
%
|
|
286 |
\isadelimproof
|
16069
|
287 |
%
|
17175
|
288 |
\endisadelimproof
|
|
289 |
%
|
23848
|
290 |
\isadelimproof
|
|
291 |
%
|
|
292 |
\endisadelimproof
|
|
293 |
%
|
|
294 |
\isatagproof
|
|
295 |
%
|
|
296 |
\begin{isamarkuptxt}%
|
40406
|
297 |
Rule induction finds no occurrences of \isa{Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}} in the
|
23848
|
298 |
conclusion, which it therefore leaves unchanged. (Look at
|
40406
|
299 |
\isa{even{\isaliteral{2E}{\isachardot}}induct} to see why this happens.) We have these subgoals:
|
23848
|
300 |
\begin{isabelle}%
|
40406
|
301 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\isanewline
|
|
302 |
\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}na{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}na\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even%
|
23848
|
303 |
\end{isabelle}
|
|
304 |
The first one is hopeless. Rule induction on
|
|
305 |
a non-variable term discards information, and usually fails.
|
|
306 |
How to deal with such situations
|
|
307 |
in general is described in {\S}\ref{sec:ind-var-in-prems} below.
|
|
308 |
In the current case the solution is easy because
|
|
309 |
we have the necessary inverse, subtraction:%
|
|
310 |
\end{isamarkuptxt}%
|
17175
|
311 |
\isamarkuptrue%
|
23848
|
312 |
%
|
|
313 |
\endisatagproof
|
|
314 |
{\isafoldproof}%
|
|
315 |
%
|
|
316 |
\isadelimproof
|
|
317 |
%
|
|
318 |
\endisadelimproof
|
17175
|
319 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
320 |
\ even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{5F}{\isacharunderscore}}minus{\isaliteral{5F}{\isacharunderscore}}{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{2}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17175
|
321 |
%
|
|
322 |
\isadelimproof
|
|
323 |
%
|
|
324 |
\endisadelimproof
|
|
325 |
%
|
|
326 |
\isatagproof
|
|
327 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
328 |
\ {\isaliteral{28}{\isacharparenleft}}erule\ even{\isaliteral{2E}{\isachardot}}induct{\isaliteral{29}{\isacharparenright}}\isanewline
|
23848
|
329 |
\ \isacommand{apply}\isamarkupfalse%
|
17175
|
330 |
\ auto\isanewline
|
|
331 |
\isacommand{done}\isamarkupfalse%
|
|
332 |
%
|
17056
|
333 |
\endisatagproof
|
|
334 |
{\isafoldproof}%
|
|
335 |
%
|
|
336 |
\isadelimproof
|
|
337 |
%
|
|
338 |
\endisadelimproof
|
11866
|
339 |
%
|
23848
|
340 |
\isadelimproof
|
|
341 |
%
|
|
342 |
\endisadelimproof
|
|
343 |
%
|
|
344 |
\isatagproof
|
|
345 |
%
|
|
346 |
\begin{isamarkuptxt}%
|
|
347 |
This lemma is trivially inductive. Here are the subgoals:
|
|
348 |
\begin{isabelle}%
|
40406
|
349 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isadigit{0}}\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{2}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\isanewline
|
|
350 |
\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ n\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{2}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{2}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even%
|
23848
|
351 |
\end{isabelle}
|
40406
|
352 |
The first is trivial because \isa{{\isadigit{0}}\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{2}}} simplifies to \isa{{\isadigit{0}}}, which is
|
|
353 |
even. The second is trivial too: \isa{Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{2}}} simplifies to
|
23848
|
354 |
\isa{n}, matching the assumption.%
|
|
355 |
\index{rule induction|)} %the sequel isn't really about induction
|
|
356 |
|
|
357 |
\medskip
|
|
358 |
Using our lemma, we can easily prove the result we originally wanted:%
|
|
359 |
\end{isamarkuptxt}%
|
17175
|
360 |
\isamarkuptrue%
|
23848
|
361 |
%
|
|
362 |
\endisatagproof
|
|
363 |
{\isafoldproof}%
|
|
364 |
%
|
|
365 |
\isadelimproof
|
|
366 |
%
|
|
367 |
\endisadelimproof
|
17175
|
368 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
369 |
\ Suc{\isaliteral{5F}{\isacharunderscore}}Suc{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17056
|
370 |
%
|
|
371 |
\isadelimproof
|
|
372 |
%
|
|
373 |
\endisadelimproof
|
|
374 |
%
|
|
375 |
\isatagproof
|
17175
|
376 |
\isacommand{by}\isamarkupfalse%
|
40406
|
377 |
\ {\isaliteral{28}{\isacharparenleft}}drule\ even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{5F}{\isacharunderscore}}minus{\isaliteral{5F}{\isacharunderscore}}{\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{29}{\isacharparenright}}%
|
17056
|
378 |
\endisatagproof
|
|
379 |
{\isafoldproof}%
|
|
380 |
%
|
|
381 |
\isadelimproof
|
|
382 |
%
|
|
383 |
\endisadelimproof
|
23848
|
384 |
%
|
|
385 |
\begin{isamarkuptext}%
|
40406
|
386 |
We have just proved the converse of the introduction rule \isa{even{\isaliteral{2E}{\isachardot}}step}.
|
23848
|
387 |
This suggests proving the following equivalence. We give it the
|
|
388 |
\attrdx{iff} attribute because of its obvious value for simplification.%
|
|
389 |
\end{isamarkuptext}%
|
|
390 |
\isamarkuptrue%
|
17175
|
391 |
\isacommand{lemma}\isamarkupfalse%
|
40406
|
392 |
\ {\isaliteral{5B}{\isacharbrackleft}}iff{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
|
17056
|
393 |
%
|
|
394 |
\isadelimproof
|
|
395 |
%
|
|
396 |
\endisadelimproof
|
|
397 |
%
|
|
398 |
\isatagproof
|
17175
|
399 |
\isacommand{by}\isamarkupfalse%
|
40406
|
400 |
\ {\isaliteral{28}{\isacharparenleft}}blast\ dest{\isaliteral{3A}{\isacharcolon}}\ Suc{\isaliteral{5F}{\isacharunderscore}}Suc{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{29}{\isacharparenright}}%
|
17056
|
401 |
\endisatagproof
|
|
402 |
{\isafoldproof}%
|
|
403 |
%
|
|
404 |
\isadelimproof
|
|
405 |
%
|
|
406 |
\endisadelimproof
|
|
407 |
%
|
23848
|
408 |
\isamarkupsubsection{Rule Inversion \label{sec:rule-inversion}%
|
|
409 |
}
|
|
410 |
\isamarkuptrue%
|
|
411 |
%
|
|
412 |
\begin{isamarkuptext}%
|
|
413 |
\index{rule inversion|(}%
|
|
414 |
Case analysis on an inductive definition is called \textbf{rule
|
|
415 |
inversion}. It is frequently used in proofs about operational
|
|
416 |
semantics. It can be highly effective when it is applied
|
|
417 |
automatically. Let us look at how rule inversion is done in
|
|
418 |
Isabelle/HOL\@.
|
|
419 |
|
|
420 |
Recall that \isa{even} is the minimal set closed under these two rules:
|
|
421 |
\begin{isabelle}%
|
40406
|
422 |
{\isadigit{0}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\isasep\isanewline%
|
|
423 |
n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even%
|
23848
|
424 |
\end{isabelle}
|
|
425 |
Minimality means that \isa{even} contains only the elements that these
|
|
426 |
rules force it to contain. If we are told that \isa{a}
|
|
427 |
belongs to
|
|
428 |
\isa{even} then there are only two possibilities. Either \isa{a} is \isa{{\isadigit{0}}}
|
40406
|
429 |
or else \isa{a} has the form \isa{Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}}, for some suitable \isa{n}
|
23848
|
430 |
that belongs to
|
|
431 |
\isa{even}. That is the gist of the \isa{cases} rule, which Isabelle proves
|
|
432 |
for us when it accepts an inductive definition:
|
|
433 |
\begin{isabelle}%
|
40406
|
434 |
{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ a\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{3B}{\isacharsemicolon}}\isanewline
|
|
435 |
\isaindent{\ }{\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}a\ {\isaliteral{3D}{\isacharequal}}\ Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
|
|
436 |
{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\rulename{even{\isaliteral{2E}{\isachardot}}cases}%
|
23848
|
437 |
\end{isabelle}
|
|
438 |
This general rule is less useful than instances of it for
|
|
439 |
specific patterns. For example, if \isa{a} has the form
|
40406
|
440 |
\isa{Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}} then the first case becomes irrelevant, while the second
|
23848
|
441 |
case tells us that \isa{n} belongs to \isa{even}. Isabelle will generate
|
|
442 |
this instance for us:%
|
|
443 |
\end{isamarkuptext}%
|
|
444 |
\isamarkuptrue%
|
40406
|
445 |
\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}\isamarkupfalse%
|
|
446 |
\ Suc{\isaliteral{5F}{\isacharunderscore}}Suc{\isaliteral{5F}{\isacharunderscore}}cases\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}Suc{\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}%
|
23848
|
447 |
\begin{isamarkuptext}%
|
|
448 |
The \commdx{inductive\protect\_cases} command generates an instance of
|
|
449 |
the \isa{cases} rule for the supplied pattern and gives it the supplied name:
|
|
450 |
\begin{isabelle}%
|
40406
|
451 |
{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{3B}{\isacharsemicolon}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\rulename{Suc{\isaliteral{5F}{\isacharunderscore}}Suc{\isaliteral{5F}{\isacharunderscore}}cases}%
|
23848
|
452 |
\end{isabelle}
|
40406
|
453 |
Applying this as an elimination rule yields one case where \isa{even{\isaliteral{2E}{\isachardot}}cases}
|
23848
|
454 |
would yield two. Rule inversion works well when the conclusions of the
|
40406
|
455 |
introduction rules involve datatype constructors like \isa{Suc} and \isa{{\isaliteral{23}{\isacharhash}}}
|
23848
|
456 |
(list ``cons''); freeness reasoning discards all but one or two cases.
|
|
457 |
|
|
458 |
In the \isacommand{inductive\_cases} command we supplied an
|
40406
|
459 |
attribute, \isa{elim{\isaliteral{21}{\isacharbang}}},
|
23848
|
460 |
\index{elim"!@\isa {elim"!} (attribute)}%
|
|
461 |
indicating that this elimination rule can be
|
|
462 |
applied aggressively. The original
|
|
463 |
\isa{cases} rule would loop if used in that manner because the
|
|
464 |
pattern~\isa{a} matches everything.
|
|
465 |
|
40406
|
466 |
The rule \isa{Suc{\isaliteral{5F}{\isacharunderscore}}Suc{\isaliteral{5F}{\isacharunderscore}}cases} is equivalent to the following implication:
|
23848
|
467 |
\begin{isabelle}%
|
40406
|
468 |
Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even%
|
23848
|
469 |
\end{isabelle}
|
|
470 |
Just above we devoted some effort to reaching precisely
|
|
471 |
this result. Yet we could have obtained it by a one-line declaration,
|
40406
|
472 |
dispensing with the lemma \isa{even{\isaliteral{5F}{\isacharunderscore}}imp{\isaliteral{5F}{\isacharunderscore}}even{\isaliteral{5F}{\isacharunderscore}}minus{\isaliteral{5F}{\isacharunderscore}}{\isadigit{2}}}.
|
23848
|
473 |
This example also justifies the terminology
|
|
474 |
\textbf{rule inversion}: the new rule inverts the introduction rule
|
40406
|
475 |
\isa{even{\isaliteral{2E}{\isachardot}}step}. In general, a rule can be inverted when the set of elements
|
23848
|
476 |
it introduces is disjoint from those of the other introduction rules.
|
|
477 |
|
|
478 |
For one-off applications of rule inversion, use the \methdx{ind_cases} method.
|
|
479 |
Here is an example:%
|
|
480 |
\end{isamarkuptext}%
|
|
481 |
\isamarkuptrue%
|
|
482 |
%
|
|
483 |
\isadelimproof
|
|
484 |
%
|
|
485 |
\endisadelimproof
|
|
486 |
%
|
|
487 |
\isatagproof
|
|
488 |
\isacommand{apply}\isamarkupfalse%
|
40406
|
489 |
\ {\isaliteral{28}{\isacharparenleft}}ind{\isaliteral{5F}{\isacharunderscore}}cases\ {\isaliteral{22}{\isachardoublequoteopen}}Suc{\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ even{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}%
|
23848
|
490 |
\endisatagproof
|
|
491 |
{\isafoldproof}%
|
|
492 |
%
|
|
493 |
\isadelimproof
|
|
494 |
%
|
|
495 |
\endisadelimproof
|
|
496 |
%
|
|
497 |
\begin{isamarkuptext}%
|
|
498 |
The specified instance of the \isa{cases} rule is generated, then applied
|
|
499 |
as an elimination rule.
|
|
500 |
|
|
501 |
To summarize, every inductive definition produces a \isa{cases} rule. The
|
|
502 |
\commdx{inductive\protect\_cases} command stores an instance of the
|
|
503 |
\isa{cases} rule for a given pattern. Within a proof, the
|
40406
|
504 |
\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases} method applies an instance of the \isa{cases}
|
23848
|
505 |
rule.
|
|
506 |
|
|
507 |
The even numbers example has shown how inductive definitions can be
|
|
508 |
used. Later examples will show that they are actually worth using.%
|
|
509 |
\index{rule inversion|)}%
|
|
510 |
\index{even numbers!defining inductively|)}%
|
|
511 |
\end{isamarkuptext}%
|
|
512 |
\isamarkuptrue%
|
|
513 |
%
|
17056
|
514 |
\isadelimtheory
|
|
515 |
%
|
|
516 |
\endisadelimtheory
|
|
517 |
%
|
|
518 |
\isatagtheory
|
17175
|
519 |
%
|
17056
|
520 |
\endisatagtheory
|
|
521 |
{\isafoldtheory}%
|
|
522 |
%
|
|
523 |
\isadelimtheory
|
|
524 |
%
|
|
525 |
\endisadelimtheory
|
10365
|
526 |
\end{isabellebody}%
|
|
527 |
%%% Local Variables:
|
|
528 |
%%% mode: latex
|
|
529 |
%%% TeX-master: "root"
|
|
530 |
%%% End:
|