|
1 (*<*)theory Advanced imports Even begin |
|
2 ML_file "../../antiquote_setup.ML" |
|
3 setup Antiquote_Setup.setup |
|
4 (*>*) |
|
5 |
|
6 text {* |
|
7 The premises of introduction rules may contain universal quantifiers and |
|
8 monotone functions. A universal quantifier lets the rule |
|
9 refer to any number of instances of |
|
10 the inductively defined set. A monotone function lets the rule refer |
|
11 to existing constructions (such as ``list of'') over the inductively defined |
|
12 set. The examples below show how to use the additional expressiveness |
|
13 and how to reason from the resulting definitions. |
|
14 *} |
|
15 |
|
16 subsection{* Universal Quantifiers in Introduction Rules \label{sec:gterm-datatype} *} |
|
17 |
|
18 text {* |
|
19 \index{ground terms example|(}% |
|
20 \index{quantifiers!and inductive definitions|(}% |
|
21 As a running example, this section develops the theory of \textbf{ground |
|
22 terms}: terms constructed from constant and function |
|
23 symbols but not variables. To simplify matters further, we regard a |
|
24 constant as a function applied to the null argument list. Let us declare a |
|
25 datatype @{text gterm} for the type of ground terms. It is a type constructor |
|
26 whose argument is a type of function symbols. |
|
27 *} |
|
28 |
|
29 datatype 'f gterm = Apply 'f "'f gterm list" |
|
30 |
|
31 text {* |
|
32 To try it out, we declare a datatype of some integer operations: |
|
33 integer constants, the unary minus operator and the addition |
|
34 operator. |
|
35 *} |
|
36 |
|
37 datatype integer_op = Number int | UnaryMinus | Plus |
|
38 |
|
39 text {* |
|
40 Now the type @{typ "integer_op gterm"} denotes the ground |
|
41 terms built over those symbols. |
|
42 |
|
43 The type constructor @{text gterm} can be generalized to a function |
|
44 over sets. It returns |
|
45 the set of ground terms that can be formed over a set @{text F} of function symbols. For |
|
46 example, we could consider the set of ground terms formed from the finite |
|
47 set @{text "{Number 2, UnaryMinus, Plus}"}. |
|
48 |
|
49 This concept is inductive. If we have a list @{text args} of ground terms |
|
50 over~@{text F} and a function symbol @{text f} in @{text F}, then we |
|
51 can apply @{text f} to @{text args} to obtain another ground term. |
|
52 The only difficulty is that the argument list may be of any length. Hitherto, |
|
53 each rule in an inductive definition referred to the inductively |
|
54 defined set a fixed number of times, typically once or twice. |
|
55 A universal quantifier in the premise of the introduction rule |
|
56 expresses that every element of @{text args} belongs |
|
57 to our inductively defined set: is a ground term |
|
58 over~@{text F}. The function @{term set} denotes the set of elements in a given |
|
59 list. |
|
60 *} |
|
61 |
|
62 inductive_set |
|
63 gterms :: "'f set \<Rightarrow> 'f gterm set" |
|
64 for F :: "'f set" |
|
65 where |
|
66 step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> gterms F; f \<in> F\<rbrakk> |
|
67 \<Longrightarrow> (Apply f args) \<in> gterms F" |
|
68 |
|
69 text {* |
|
70 To demonstrate a proof from this definition, let us |
|
71 show that the function @{term gterms} |
|
72 is \textbf{monotone}. We shall need this concept shortly. |
|
73 *} |
|
74 |
|
75 lemma gterms_mono: "F\<subseteq>G \<Longrightarrow> gterms F \<subseteq> gterms G" |
|
76 apply clarify |
|
77 apply (erule gterms.induct) |
|
78 apply blast |
|
79 done |
|
80 (*<*) |
|
81 lemma gterms_mono: "F\<subseteq>G \<Longrightarrow> gterms F \<subseteq> gterms G" |
|
82 apply clarify |
|
83 apply (erule gterms.induct) |
|
84 (*>*) |
|
85 txt{* |
|
86 Intuitively, this theorem says that |
|
87 enlarging the set of function symbols enlarges the set of ground |
|
88 terms. The proof is a trivial rule induction. |
|
89 First we use the @{text clarify} method to assume the existence of an element of |
|
90 @{term "gterms F"}. (We could have used @{text "intro subsetI"}.) We then |
|
91 apply rule induction. Here is the resulting subgoal: |
|
92 @{subgoals[display,indent=0]} |
|
93 The assumptions state that @{text f} belongs |
|
94 to~@{text F}, which is included in~@{text G}, and that every element of the list @{text args} is |
|
95 a ground term over~@{text G}. The @{text blast} method finds this chain of reasoning easily. |
|
96 *} |
|
97 (*<*)oops(*>*) |
|
98 text {* |
|
99 \begin{warn} |
|
100 Why do we call this function @{text gterms} instead |
|
101 of @{text gterm}? A constant may have the same name as a type. However, |
|
102 name clashes could arise in the theorems that Isabelle generates. |
|
103 Our choice of names keeps @{text gterms.induct} separate from |
|
104 @{text gterm.induct}. |
|
105 \end{warn} |
|
106 |
|
107 Call a term \textbf{well-formed} if each symbol occurring in it is applied |
|
108 to the correct number of arguments. (This number is called the symbol's |
|
109 \textbf{arity}.) We can express well-formedness by |
|
110 generalizing the inductive definition of |
|
111 \isa{gterms}. |
|
112 Suppose we are given a function called @{text arity}, specifying the arities |
|
113 of all symbols. In the inductive step, we have a list @{text args} of such |
|
114 terms and a function symbol~@{text f}. If the length of the list matches the |
|
115 function's arity then applying @{text f} to @{text args} yields a well-formed |
|
116 term. |
|
117 *} |
|
118 |
|
119 inductive_set |
|
120 well_formed_gterm :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set" |
|
121 for arity :: "'f \<Rightarrow> nat" |
|
122 where |
|
123 step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> well_formed_gterm arity; |
|
124 length args = arity f\<rbrakk> |
|
125 \<Longrightarrow> (Apply f args) \<in> well_formed_gterm arity" |
|
126 |
|
127 text {* |
|
128 The inductive definition neatly captures the reasoning above. |
|
129 The universal quantification over the |
|
130 @{text set} of arguments expresses that all of them are well-formed.% |
|
131 \index{quantifiers!and inductive definitions|)} |
|
132 *} |
|
133 |
|
134 subsection{* Alternative Definition Using a Monotone Function *} |
|
135 |
|
136 text {* |
|
137 \index{monotone functions!and inductive definitions|(}% |
|
138 An inductive definition may refer to the |
|
139 inductively defined set through an arbitrary monotone function. To |
|
140 demonstrate this powerful feature, let us |
|
141 change the inductive definition above, replacing the |
|
142 quantifier by a use of the function @{term lists}. This |
|
143 function, from the Isabelle theory of lists, is analogous to the |
|
144 function @{term gterms} declared above: if @{text A} is a set then |
|
145 @{term "lists A"} is the set of lists whose elements belong to |
|
146 @{term A}. |
|
147 |
|
148 In the inductive definition of well-formed terms, examine the one |
|
149 introduction rule. The first premise states that @{text args} belongs to |
|
150 the @{text lists} of well-formed terms. This formulation is more |
|
151 direct, if more obscure, than using a universal quantifier. |
|
152 *} |
|
153 |
|
154 inductive_set |
|
155 well_formed_gterm' :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set" |
|
156 for arity :: "'f \<Rightarrow> nat" |
|
157 where |
|
158 step[intro!]: "\<lbrakk>args \<in> lists (well_formed_gterm' arity); |
|
159 length args = arity f\<rbrakk> |
|
160 \<Longrightarrow> (Apply f args) \<in> well_formed_gterm' arity" |
|
161 monos lists_mono |
|
162 |
|
163 text {* |
|
164 We cite the theorem @{text lists_mono} to justify |
|
165 using the function @{term lists}.% |
|
166 \footnote{This particular theorem is installed by default already, but we |
|
167 include the \isakeyword{monos} declaration in order to illustrate its syntax.} |
|
168 @{named_thms [display,indent=0] lists_mono [no_vars] (lists_mono)} |
|
169 Why must the function be monotone? An inductive definition describes |
|
170 an iterative construction: each element of the set is constructed by a |
|
171 finite number of introduction rule applications. For example, the |
|
172 elements of \isa{even} are constructed by finitely many applications of |
|
173 the rules |
|
174 @{thm [display,indent=0] even.intros [no_vars]} |
|
175 All references to a set in its |
|
176 inductive definition must be positive. Applications of an |
|
177 introduction rule cannot invalidate previous applications, allowing the |
|
178 construction process to converge. |
|
179 The following pair of rules do not constitute an inductive definition: |
|
180 \begin{trivlist} |
|
181 \item @{term "0 \<in> even"} |
|
182 \item @{term "n \<notin> even \<Longrightarrow> (Suc n) \<in> even"} |
|
183 \end{trivlist} |
|
184 Showing that 4 is even using these rules requires showing that 3 is not |
|
185 even. It is far from trivial to show that this set of rules |
|
186 characterizes the even numbers. |
|
187 |
|
188 Even with its use of the function \isa{lists}, the premise of our |
|
189 introduction rule is positive: |
|
190 @{thm [display,indent=0] (prem 1) step [no_vars]} |
|
191 To apply the rule we construct a list @{term args} of previously |
|
192 constructed well-formed terms. We obtain a |
|
193 new term, @{term "Apply f args"}. Because @{term lists} is monotone, |
|
194 applications of the rule remain valid as new terms are constructed. |
|
195 Further lists of well-formed |
|
196 terms become available and none are taken away.% |
|
197 \index{monotone functions!and inductive definitions|)} |
|
198 *} |
|
199 |
|
200 subsection{* A Proof of Equivalence *} |
|
201 |
|
202 text {* |
|
203 We naturally hope that these two inductive definitions of ``well-formed'' |
|
204 coincide. The equality can be proved by separate inclusions in |
|
205 each direction. Each is a trivial rule induction. |
|
206 *} |
|
207 |
|
208 lemma "well_formed_gterm arity \<subseteq> well_formed_gterm' arity" |
|
209 apply clarify |
|
210 apply (erule well_formed_gterm.induct) |
|
211 apply auto |
|
212 done |
|
213 (*<*) |
|
214 lemma "well_formed_gterm arity \<subseteq> well_formed_gterm' arity" |
|
215 apply clarify |
|
216 apply (erule well_formed_gterm.induct) |
|
217 (*>*) |
|
218 txt {* |
|
219 The @{text clarify} method gives |
|
220 us an element of @{term "well_formed_gterm arity"} on which to perform |
|
221 induction. The resulting subgoal can be proved automatically: |
|
222 @{subgoals[display,indent=0]} |
|
223 This proof resembles the one given in |
|
224 {\S}\ref{sec:gterm-datatype} above, especially in the form of the |
|
225 induction hypothesis. Next, we consider the opposite inclusion: |
|
226 *} |
|
227 (*<*)oops(*>*) |
|
228 lemma "well_formed_gterm' arity \<subseteq> well_formed_gterm arity" |
|
229 apply clarify |
|
230 apply (erule well_formed_gterm'.induct) |
|
231 apply auto |
|
232 done |
|
233 (*<*) |
|
234 lemma "well_formed_gterm' arity \<subseteq> well_formed_gterm arity" |
|
235 apply clarify |
|
236 apply (erule well_formed_gterm'.induct) |
|
237 (*>*) |
|
238 txt {* |
|
239 The proof script is virtually identical, |
|
240 but the subgoal after applying induction may be surprising: |
|
241 @{subgoals[display,indent=0,margin=65]} |
|
242 The induction hypothesis contains an application of @{term lists}. Using a |
|
243 monotone function in the inductive definition always has this effect. The |
|
244 subgoal may look uninviting, but fortunately |
|
245 @{term lists} distributes over intersection: |
|
246 @{named_thms [display,indent=0] lists_Int_eq [no_vars] (lists_Int_eq)} |
|
247 Thanks to this default simplification rule, the induction hypothesis |
|
248 is quickly replaced by its two parts: |
|
249 \begin{trivlist} |
|
250 \item @{term "args \<in> lists (well_formed_gterm' arity)"} |
|
251 \item @{term "args \<in> lists (well_formed_gterm arity)"} |
|
252 \end{trivlist} |
|
253 Invoking the rule @{text well_formed_gterm.step} completes the proof. The |
|
254 call to @{text auto} does all this work. |
|
255 |
|
256 This example is typical of how monotone functions |
|
257 \index{monotone functions} can be used. In particular, many of them |
|
258 distribute over intersection. Monotonicity implies one direction of |
|
259 this set equality; we have this theorem: |
|
260 @{named_thms [display,indent=0] mono_Int [no_vars] (mono_Int)} |
|
261 *} |
|
262 (*<*)oops(*>*) |
|
263 |
|
264 |
|
265 subsection{* Another Example of Rule Inversion *} |
|
266 |
|
267 text {* |
|
268 \index{rule inversion|(}% |
|
269 Does @{term gterms} distribute over intersection? We have proved that this |
|
270 function is monotone, so @{text mono_Int} gives one of the inclusions. The |
|
271 opposite inclusion asserts that if @{term t} is a ground term over both of the |
|
272 sets |
|
273 @{term F} and~@{term G} then it is also a ground term over their intersection, |
|
274 @{term "F \<inter> G"}. |
|
275 *} |
|
276 |
|
277 lemma gterms_IntI: |
|
278 "t \<in> gterms F \<Longrightarrow> t \<in> gterms G \<longrightarrow> t \<in> gterms (F\<inter>G)" |
|
279 (*<*)oops(*>*) |
|
280 text {* |
|
281 Attempting this proof, we get the assumption |
|
282 @{term "Apply f args \<in> gterms G"}, which cannot be broken down. |
|
283 It looks like a job for rule inversion:\cmmdx{inductive\protect\_cases} |
|
284 *} |
|
285 |
|
286 inductive_cases gterm_Apply_elim [elim!]: "Apply f args \<in> gterms F" |
|
287 |
|
288 text {* |
|
289 Here is the result. |
|
290 @{named_thms [display,indent=0,margin=50] gterm_Apply_elim [no_vars] (gterm_Apply_elim)} |
|
291 This rule replaces an assumption about @{term "Apply f args"} by |
|
292 assumptions about @{term f} and~@{term args}. |
|
293 No cases are discarded (there was only one to begin |
|
294 with) but the rule applies specifically to the pattern @{term "Apply f args"}. |
|
295 It can be applied repeatedly as an elimination rule without looping, so we |
|
296 have given the @{text "elim!"} attribute. |
|
297 |
|
298 Now we can prove the other half of that distributive law. |
|
299 *} |
|
300 |
|
301 lemma gterms_IntI [rule_format, intro!]: |
|
302 "t \<in> gterms F \<Longrightarrow> t \<in> gterms G \<longrightarrow> t \<in> gterms (F\<inter>G)" |
|
303 apply (erule gterms.induct) |
|
304 apply blast |
|
305 done |
|
306 (*<*) |
|
307 lemma "t \<in> gterms F \<Longrightarrow> t \<in> gterms G \<longrightarrow> t \<in> gterms (F\<inter>G)" |
|
308 apply (erule gterms.induct) |
|
309 (*>*) |
|
310 txt {* |
|
311 The proof begins with rule induction over the definition of |
|
312 @{term gterms}, which leaves a single subgoal: |
|
313 @{subgoals[display,indent=0,margin=65]} |
|
314 To prove this, we assume @{term "Apply f args \<in> gterms G"}. Rule inversion, |
|
315 in the form of @{text gterm_Apply_elim}, infers |
|
316 that every element of @{term args} belongs to |
|
317 @{term "gterms G"}; hence (by the induction hypothesis) it belongs |
|
318 to @{term "gterms (F \<inter> G)"}. Rule inversion also yields |
|
319 @{term "f \<in> G"} and hence @{term "f \<in> F \<inter> G"}. |
|
320 All of this reasoning is done by @{text blast}. |
|
321 |
|
322 \smallskip |
|
323 Our distributive law is a trivial consequence of previously-proved results: |
|
324 *} |
|
325 (*<*)oops(*>*) |
|
326 lemma gterms_Int_eq [simp]: |
|
327 "gterms (F \<inter> G) = gterms F \<inter> gterms G" |
|
328 by (blast intro!: mono_Int monoI gterms_mono) |
|
329 |
|
330 text_raw {* |
|
331 \index{rule inversion|)}% |
|
332 \index{ground terms example|)} |
|
333 |
|
334 |
|
335 \begin{isamarkuptext} |
|
336 \begin{exercise} |
|
337 A function mapping function symbols to their |
|
338 types is called a \textbf{signature}. Given a type |
|
339 ranging over type symbols, we can represent a function's type by a |
|
340 list of argument types paired with the result type. |
|
341 Complete this inductive definition: |
|
342 \begin{isabelle} |
|
343 *} |
|
344 |
|
345 inductive_set |
|
346 well_typed_gterm :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set" |
|
347 for sig :: "'f \<Rightarrow> 't list * 't" |
|
348 (*<*) |
|
349 where |
|
350 step[intro!]: |
|
351 "\<lbrakk>\<forall>pair \<in> set args. pair \<in> well_typed_gterm sig; |
|
352 sig f = (map snd args, rtype)\<rbrakk> |
|
353 \<Longrightarrow> (Apply f (map fst args), rtype) |
|
354 \<in> well_typed_gterm sig" |
|
355 (*>*) |
|
356 text_raw {* |
|
357 \end{isabelle} |
|
358 \end{exercise} |
|
359 \end{isamarkuptext} |
|
360 *} |
|
361 |
|
362 (*<*) |
|
363 |
|
364 text{*the following declaration isn't actually used*} |
|
365 primrec |
|
366 integer_arity :: "integer_op \<Rightarrow> nat" |
|
367 where |
|
368 "integer_arity (Number n) = 0" |
|
369 | "integer_arity UnaryMinus = 1" |
|
370 | "integer_arity Plus = 2" |
|
371 |
|
372 text{* the rest isn't used: too complicated. OK for an exercise though.*} |
|
373 |
|
374 inductive_set |
|
375 integer_signature :: "(integer_op * (unit list * unit)) set" |
|
376 where |
|
377 Number: "(Number n, ([], ())) \<in> integer_signature" |
|
378 | UnaryMinus: "(UnaryMinus, ([()], ())) \<in> integer_signature" |
|
379 | Plus: "(Plus, ([(),()], ())) \<in> integer_signature" |
|
380 |
|
381 inductive_set |
|
382 well_typed_gterm' :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set" |
|
383 for sig :: "'f \<Rightarrow> 't list * 't" |
|
384 where |
|
385 step[intro!]: |
|
386 "\<lbrakk>args \<in> lists(well_typed_gterm' sig); |
|
387 sig f = (map snd args, rtype)\<rbrakk> |
|
388 \<Longrightarrow> (Apply f (map fst args), rtype) |
|
389 \<in> well_typed_gterm' sig" |
|
390 monos lists_mono |
|
391 |
|
392 |
|
393 lemma "well_typed_gterm sig \<subseteq> well_typed_gterm' sig" |
|
394 apply clarify |
|
395 apply (erule well_typed_gterm.induct) |
|
396 apply auto |
|
397 done |
|
398 |
|
399 lemma "well_typed_gterm' sig \<subseteq> well_typed_gterm sig" |
|
400 apply clarify |
|
401 apply (erule well_typed_gterm'.induct) |
|
402 apply auto |
|
403 done |
|
404 |
|
405 |
|
406 end |
|
407 (*>*) |