adam@398
|
1 (* Copyright (c) 2008-2012, Adam Chlipala
|
adamc@26
|
2 *
|
adamc@26
|
3 * This work is licensed under a
|
adamc@26
|
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
|
adamc@26
|
5 * Unported License.
|
adamc@26
|
6 * The license text is available at:
|
adamc@26
|
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
|
adamc@26
|
8 *)
|
adamc@26
|
9
|
adamc@26
|
10 (* begin hide *)
|
adamc@26
|
11 Require Import List.
|
adamc@26
|
12
|
adam@314
|
13 Require Import CpdtTactics.
|
adamc@26
|
14
|
adamc@26
|
15 Set Implicit Arguments.
|
adamc@26
|
16 (* end hide *)
|
adamc@26
|
17
|
adamc@26
|
18
|
adamc@74
|
19 (** %\part{Basic Programming and Proving}
|
adamc@74
|
20
|
adamc@74
|
21 \chapter{Introducing Inductive Types}% *)
|
adamc@26
|
22
|
adam@315
|
23 (** In a sense, CIC is built from just two relatively straightforward features: function types and inductive types. From this modest foundation, we can prove effectively all of the theorems of math and carry out effectively all program verifications, with enough effort expended. This chapter introduces induction and recursion for functional programming in Coq. Most of our examples reproduce functionality from the Coq standard library, and we have tried to copy the standard library's choices of identifiers, where possible, so many of the definitions here are already available in the default Coq environment.
|
adam@315
|
24
|
adam@315
|
25 The last chapter took a deep dive into some of the more advanced Coq features, to highlight the unusual approach that I advocate in this book. However, from this point on, we will rewind and go back to basics, presenting the relevant features of Coq in a more bottom-up manner. A useful first step is a discussion of the differences and relationships between proofs and programs in Coq. *)
|
adam@315
|
26
|
adam@315
|
27
|
adam@315
|
28 (** * Proof Terms *)
|
adam@315
|
29
|
adam@465
|
30 (** Mainstream presentations of mathematics treat proofs as objects that exist outside of the universe of mathematical objects. However, for a variety of reasoning tasks, it is convenient to encode proofs, traditional mathematical objects, and programs within a single formal language. Validity checks on mathematical objects are useful in any setting, to catch typos and other uninteresting errors. The benefits of static typing for programs are widely recognized, and Coq brings those benefits to both mathematical objects and programs via a uniform mechanism. In fact, from this point on, we will not bother to distinguish between programs and mathematical objects. Many mathematical formalisms are most easily encoded in terms of programs.
|
adam@315
|
31
|
adam@400
|
32 Proofs are fundamentally different from programs, because any two proofs of a theorem are considered equivalent, from a formal standpoint if not from an engineering standpoint. However, we can use the same type-checking technology to check proofs as we use to validate our programs. This is the%\index{Curry-Howard correspondence}% _Curry-Howard correspondence_ %\cite{Curry,Howard}%, an approach for relating proofs and programs. We represent mathematical theorems as types, such that a theorem's proofs are exactly those programs that type-check at the corresponding type.
|
adam@315
|
33
|
adam@315
|
34 The last chapter's example already snuck in an instance of Curry-Howard. We used the token [->] to stand for both function types and logical implications. One reasonable conclusion upon seeing this might be that some fancy overloading of notations is at work. In fact, functions and implications are precisely identical according to Curry-Howard! That is, they are just two ways of describing the same computational phenomenon.
|
adam@315
|
35
|
adam@315
|
36 A short demonstration should explain how this can be. The identity function over the natural numbers is certainly not a controversial program. *)
|
adam@315
|
37
|
adam@315
|
38 Check (fun x : nat => x).
|
adam@315
|
39 (** [: nat -> nat] *)
|
adam@315
|
40
|
adam@439
|
41 (** %\smallskip{}%Consider this alternate program, which is almost identical to the last one. *)
|
adam@315
|
42
|
adam@315
|
43 Check (fun x : True => x).
|
adam@315
|
44 (** [: True -> True] *)
|
adam@315
|
45
|
adam@439
|
46 (** %\smallskip{}%The identity program is interpreted as a proof that %\index{Gallina terms!True}%[True], the always-true proposition, implies itself! What we see is that Curry-Howard interprets implications as functions, where an input is a proposition being assumed and an output is a proposition being deduced. This intuition is not too far from a common one for informal theorem proving, where we might already think of an implication proof as a process for transforming a hypothesis into a conclusion.
|
adam@315
|
47
|
adam@315
|
48 There are also more primitive proof forms available. For instance, the term %\index{Gallina terms!I}%[I] is the single proof of [True], applicable in any context. *)
|
adam@315
|
49
|
adam@315
|
50 Check I.
|
adam@315
|
51 (** [: True] *)
|
adam@315
|
52
|
adam@439
|
53 (** %\smallskip{}%With [I], we can prove another simple propositional theorem. *)
|
adam@315
|
54
|
adam@315
|
55 Check (fun _ : False => I).
|
adam@315
|
56 (** [: False -> True] *)
|
adam@315
|
57
|
adam@439
|
58 (** %\smallskip{}%No proofs of %\index{Gallina terms!False}%[False] exist in the top-level context, but the implication-as-function analogy gives us an easy way to, for example, show that [False] implies itself. *)
|
adam@315
|
59
|
adam@315
|
60 Check (fun x : False => x).
|
adam@315
|
61 (** [: False -> False] *)
|
adam@315
|
62
|
adam@439
|
63 (** %\smallskip{}%Every one of these example programs whose type looks like a logical formula is a%\index{proof term}% _proof term_. We use that name for any Gallina term of a logical type, and we will elaborate shortly on what makes a type logical.
|
adam@315
|
64
|
adam@469
|
65 In the rest of this chapter, we will introduce different ways of defining types. Every example type can be interpreted alternatively as a type of programs or proofs.
|
adam@317
|
66
|
adam@469
|
67 One of the first types we introduce will be [bool], with constructors [true] and [false]. Newcomers to Coq often wonder about the distinction between [True] and [true] and the distinction between [False] and [false]. One glib answer is that [True] and [False] are types, but [true] and [false] are not. A more useful answer is that Coq's metatheory guarantees that any term of type [bool] _evaluates_ to either [true] or [false]. This means that we have an _algorithm_ for answering any question phrased as an expression of type [bool]. Conversely, most propositions do not evaluate to [True] or [False]; the language of inductively defined propositions is much richer than that. We ought to be glad that we have no algorithm for deciding mathematical truth, since otherwise it would be clear that we could not formalize undecidable properties, like almost any interesting property of general-purpose programs. *)
|
adamc@26
|
68
|
adamc@26
|
69
|
adamc@26
|
70 (** * Enumerations *)
|
adamc@26
|
71
|
adam@419
|
72 (** Coq inductive types generalize the %\index{algebraic datatypes}%algebraic datatypes found in %\index{Haskell}%Haskell and %\index{ML}%ML. Confusingly enough, inductive types also generalize %\index{generalized algebraic datatypes}%generalized algebraic datatypes (GADTs), by adding the possibility for type dependency. Even so, it is worth backing up from the examples of the last chapter and going over basic, algebraic datatype uses of inductive datatypes, because the chance to prove things about the values of these types adds new wrinkles beyond usual practice in Haskell and ML.
|
adamc@26
|
73
|
adam@315
|
74 The singleton type [unit] is an inductive type:%\index{Gallina terms!unit}\index{Gallina terms!tt}% *)
|
adamc@26
|
75
|
adamc@26
|
76 Inductive unit : Set :=
|
adamc@26
|
77 | tt.
|
adamc@26
|
78
|
adamc@26
|
79 (** This vernacular command defines a new inductive type [unit] whose only value is [tt], as we can see by checking the types of the two identifiers: *)
|
adamc@26
|
80
|
adamc@26
|
81 Check unit.
|
adamc@208
|
82 (** [unit : Set] *)
|
adamc@26
|
83
|
adamc@26
|
84 Check tt.
|
adamc@208
|
85 (** [tt : unit] *)
|
adamc@26
|
86
|
adam@439
|
87 (** %\smallskip{}%We can prove that [unit] is a genuine singleton type. *)
|
adamc@26
|
88
|
adamc@26
|
89 Theorem unit_singleton : forall x : unit, x = tt.
|
adamc@208
|
90
|
adam@315
|
91 (** The important thing about an inductive type is, unsurprisingly, that you can do induction over its values, and induction is the key to proving this theorem. We ask to proceed by induction on the variable [x].%\index{tactics!induction}% *)
|
adamc@208
|
92
|
adamc@41
|
93 (* begin thide *)
|
adamc@26
|
94 induction x.
|
adamc@26
|
95
|
adamc@208
|
96 (** The goal changes to:
|
adamc@208
|
97 [[
|
adamc@26
|
98 tt = tt
|
adam@302
|
99 ]]
|
adam@302
|
100 *)
|
adamc@208
|
101
|
adam@448
|
102 (** %\noindent{}%...which we can discharge trivially. *)
|
adamc@208
|
103
|
adamc@26
|
104 reflexivity.
|
adamc@26
|
105 Qed.
|
adamc@41
|
106 (* end thide *)
|
adamc@26
|
107
|
adam@315
|
108 (** It seems kind of odd to write a proof by induction with no inductive hypotheses. We could have arrived at the same result by beginning the proof with:%\index{tactics!destruct}% [[
|
adamc@26
|
109 destruct x.
|
adamc@205
|
110 ]]
|
adamc@205
|
111
|
adam@420
|
112 %\noindent%...which corresponds to "proof by case analysis" in classical math. For non-recursive inductive types, the two tactics will always have identical behavior. Often case analysis is sufficient, even in proofs about recursive types, and it is nice to avoid introducing unneeded induction hypotheses.
|
adamc@26
|
113
|
adam@398
|
114 What exactly _is_ the %\index{induction principles}%induction principle for [unit]? We can ask Coq: *)
|
adamc@26
|
115
|
adamc@26
|
116 Check unit_ind.
|
adamc@208
|
117 (** [unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u] *)
|
adamc@26
|
118
|
adam@439
|
119 (** %\smallskip{}%Every [Inductive] command defining a type [T] also defines an induction principle named [T_ind]. Recall from the last section that our type, operations over it, and principles for reasoning about it all live in the same language and are described by the same type system. The key to telling what is a program and what is a proof lies in the distinction between the type %\index{Gallina terms!Prop}%[Prop], which appears in our induction principle; and the type %\index{Gallina terms!Set}%[Set], which we have seen a few times already.
|
adamc@26
|
120
|
adam@315
|
121 The convention goes like this: [Set] is the type of normal types used in programming, and the values of such types are programs. [Prop] is the type of logical propositions, and the values of such types are proofs. Thus, an induction principle has a type that shows us that it is a function for building proofs.
|
adamc@26
|
122
|
adam@400
|
123 Specifically, [unit_ind] quantifies over a predicate [P] over [unit] values. If we can present a proof that [P] holds of [tt], then we are rewarded with a proof that [P] holds for any value [u] of type [unit]. In our last proof, the predicate was [(fun u : unit => u = tt)].
|
adam@315
|
124
|
adam@315
|
125 The definition of [unit] places the type in [Set]. By replacing [Set] with [Prop], [unit] with [True], and [tt] with [I], we arrive at precisely the definition of [True] that the Coq standard library employs! The program type [unit] is the Curry-Howard equivalent of the proposition [True]. We might make the tongue-in-cheek claim that, while philosophers have expended much ink on the nature of truth, we have now determined that truth is the [unit] type of functional programming.
|
adamc@26
|
126
|
adamc@26
|
127 %\medskip%
|
adamc@26
|
128
|
adam@315
|
129 We can define an inductive type even simpler than [unit]:%\index{Gallina terms!Empty\_set}% *)
|
adamc@26
|
130
|
adamc@26
|
131 Inductive Empty_set : Set := .
|
adamc@26
|
132
|
adamc@26
|
133 (** [Empty_set] has no elements. We can prove fun theorems about it: *)
|
adamc@26
|
134
|
adamc@26
|
135 Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5.
|
adamc@41
|
136 (* begin thide *)
|
adamc@26
|
137 destruct 1.
|
adamc@26
|
138 Qed.
|
adamc@41
|
139 (* end thide *)
|
adamc@26
|
140
|
adam@317
|
141 (** Because [Empty_set] has no elements, the fact of having an element of this type implies anything. We use [destruct 1] instead of [destruct x] in the proof because unused quantified variables are relegated to being referred to by number. (There is a good reason for this, related to the unity of quantifiers and implication. At least within Coq's logical foundation of %\index{constructive logic}%constructive logic, which we elaborate on more in the next chapter, an implication is just a quantification over a proof, where the quantified variable is never used. It generally makes more sense to refer to implication hypotheses by number than by name, and Coq treats our quantifier over an unused variable as an implication in determining the proper behavior.)
|
adamc@26
|
142
|
adamc@26
|
143 We can see the induction principle that made this proof so easy: *)
|
adamc@26
|
144
|
adamc@26
|
145 Check Empty_set_ind.
|
adam@400
|
146 (** [Empty_set_ind : forall (P : Empty_set -> Prop) (e : Empty_set), P e] *)
|
adamc@26
|
147
|
adam@439
|
148 (** %\smallskip{}%In other words, any predicate over values from the empty set holds vacuously of every such element. In the last proof, we chose the predicate [(fun _ : Empty_set => 2 + 2 = 5)].
|
adamc@26
|
149
|
adamc@26
|
150 We can also apply this get-out-of-jail-free card programmatically. Here is a lazy way of converting values of [Empty_set] to values of [unit]: *)
|
adamc@26
|
151
|
adamc@26
|
152 Definition e2u (e : Empty_set) : unit := match e with end.
|
adamc@26
|
153
|
adam@315
|
154 (** We employ [match] pattern matching as in the last chapter. Since we match on a value whose type has no constructors, there is no need to provide any branches. This idiom may look familiar; we employed it with proofs of [False] in the last section. In fact, [Empty_set] is the Curry-Howard equivalent of [False]. As for why [Empty_set] starts with a capital letter and not a lowercase letter like [unit] does, we must refer the reader to the authors of the Coq standard library, to which we try to be faithful.
|
adamc@26
|
155
|
adamc@26
|
156 %\medskip%
|
adamc@26
|
157
|
adam@448
|
158 Moving up the ladder of complexity, we can define the Booleans:%\index{Gallina terms!bool}\index{Gallina terms!true}\index{Gallina terms!false}% *)
|
adamc@26
|
159
|
adamc@26
|
160 Inductive bool : Set :=
|
adamc@26
|
161 | true
|
adamc@26
|
162 | false.
|
adamc@26
|
163
|
adam@448
|
164 (** We can use less vacuous pattern matching to define Boolean negation.%\index{Gallina terms!negb}% *)
|
adamc@26
|
165
|
adam@279
|
166 Definition negb (b : bool) : bool :=
|
adamc@26
|
167 match b with
|
adamc@26
|
168 | true => false
|
adamc@26
|
169 | false => true
|
adamc@26
|
170 end.
|
adamc@26
|
171
|
adam@317
|
172 (** An alternative definition desugars to the above, thanks to an %\index{Gallina terms!if}%[if] notation overloaded to work with any inductive type that has exactly two constructors: *)
|
adamc@27
|
173
|
adam@279
|
174 Definition negb' (b : bool) : bool :=
|
adamc@27
|
175 if b then false else true.
|
adamc@27
|
176
|
adam@279
|
177 (** We might want to prove that [negb] is its own inverse operation. *)
|
adamc@26
|
178
|
adam@279
|
179 Theorem negb_inverse : forall b : bool, negb (negb b) = b.
|
adamc@41
|
180 (* begin thide *)
|
adamc@26
|
181 destruct b.
|
adamc@26
|
182
|
adamc@208
|
183 (** After we case-analyze on [b], we are left with one subgoal for each constructor of [bool].
|
adam@439
|
184 [[
|
adam@439
|
185 2 subgoals
|
adamc@26
|
186
|
adamc@26
|
187 ============================
|
adam@279
|
188 negb (negb true) = true
|
adam@439
|
189
|
adam@439
|
190 subgoal 2 is
|
adam@439
|
191
|
adam@279
|
192 negb (negb false) = false
|
adamc@26
|
193 ]]
|
adamc@26
|
194
|
adamc@26
|
195 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
|
adamc@26
|
196
|
adamc@26
|
197 reflexivity.
|
adamc@26
|
198
|
adam@315
|
199 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification.%\index{Vernacular commands!Restart}% *)
|
adamc@26
|
200
|
adamc@26
|
201 Restart.
|
adam@315
|
202
|
adamc@26
|
203 destruct b; reflexivity.
|
adamc@26
|
204 Qed.
|
adamc@41
|
205 (* end thide *)
|
adamc@27
|
206
|
adam@448
|
207 (** Another theorem about Booleans illustrates another useful tactic.%\index{tactics!discriminate}% *)
|
adamc@27
|
208
|
adam@279
|
209 Theorem negb_ineq : forall b : bool, negb b <> b.
|
adamc@41
|
210 (* begin thide *)
|
adamc@27
|
211 destruct b; discriminate.
|
adamc@27
|
212 Qed.
|
adamc@41
|
213 (* end thide *)
|
adamc@27
|
214
|
adam@448
|
215 (** The [discriminate] tactic is used to prove that two values of an inductive type are not equal, whenever the values are formed with different constructors. In this case, the different constructors are [true] and [false].
|
adamc@27
|
216
|
adamc@27
|
217 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *)
|
adamc@27
|
218
|
adamc@27
|
219 Check bool_ind.
|
adamc@208
|
220 (** [bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b] *)
|
adamc@28
|
221
|
adam@439
|
222 (** %\smallskip{}%That is, to prove that a property describes all [bool]s, prove that it describes both [true] and [false].
|
adam@315
|
223
|
adam@392
|
224 There is no interesting Curry-Howard analogue of [bool]. Of course, we can define such a type by replacing [Set] by [Prop] above, but the proposition we arrive at is not very useful. It is logically equivalent to [True], but it provides two indistinguishable primitive proofs, [true] and [false]. In the rest of the chapter, we will skip commenting on Curry-Howard versions of inductive definitions where such versions are not interesting. *)
|
adam@315
|
225
|
adamc@28
|
226
|
adamc@28
|
227 (** * Simple Recursive Types *)
|
adamc@28
|
228
|
adam@315
|
229 (** The natural numbers are the simplest common example of an inductive type that actually deserves the name.%\index{Gallina terms!nat}\index{Gallina terms!O}\index{Gallina terms!S}% *)
|
adamc@28
|
230
|
adamc@28
|
231 Inductive nat : Set :=
|
adamc@28
|
232 | O : nat
|
adamc@28
|
233 | S : nat -> nat.
|
adamc@28
|
234
|
adam@439
|
235 (** The constructor [O] is zero, and [S] is the successor function, so that [0] is syntactic sugar for [O], [1] for [S O], [2] for [S (S O)], and so on.
|
adamc@28
|
236
|
adam@316
|
237 Pattern matching works as we demonstrated in the last chapter:%\index{Gallina terms!pred}% *)
|
adamc@28
|
238
|
adamc@28
|
239 Definition isZero (n : nat) : bool :=
|
adamc@28
|
240 match n with
|
adamc@28
|
241 | O => true
|
adamc@28
|
242 | S _ => false
|
adamc@28
|
243 end.
|
adamc@28
|
244
|
adamc@28
|
245 Definition pred (n : nat) : nat :=
|
adamc@28
|
246 match n with
|
adamc@28
|
247 | O => O
|
adamc@28
|
248 | S n' => n'
|
adamc@28
|
249 end.
|
adamc@28
|
250
|
adam@469
|
251 (** We can prove theorems by case analysis with [destruct] as for simpler inductive types, but we can also now get into genuine inductive theorems. First, we will need a recursive function, to make things interesting.%\index{Gallina terms!plus}% *)
|
adamc@28
|
252
|
adamc@208
|
253 Fixpoint plus (n m : nat) : nat :=
|
adamc@28
|
254 match n with
|
adamc@28
|
255 | O => m
|
adamc@28
|
256 | S n' => S (plus n' m)
|
adamc@28
|
257 end.
|
adamc@28
|
258
|
adamc@208
|
259 (** Recall that [Fixpoint] is Coq's mechanism for recursive function definitions. Some theorems about [plus] can be proved without induction. *)
|
adamc@28
|
260
|
adamc@28
|
261 Theorem O_plus_n : forall n : nat, plus O n = n.
|
adamc@41
|
262 (* begin thide *)
|
adamc@28
|
263 intro; reflexivity.
|
adamc@28
|
264 Qed.
|
adamc@41
|
265 (* end thide *)
|
adamc@28
|
266
|
adamc@208
|
267 (** Coq's computation rules automatically simplify the application of [plus], because unfolding the definition of [plus] gives us a [match] expression where the branch to be taken is obvious from syntax alone. If we just reverse the order of the arguments, though, this no longer works, and we need induction. *)
|
adamc@28
|
268
|
adamc@28
|
269 Theorem n_plus_O : forall n : nat, plus n O = n.
|
adamc@41
|
270 (* begin thide *)
|
adamc@28
|
271 induction n.
|
adamc@28
|
272
|
adam@398
|
273 (** Our first subgoal is [plus O O = O], which _is_ trivial by computation. *)
|
adamc@28
|
274
|
adamc@28
|
275 reflexivity.
|
adamc@28
|
276
|
adam@469
|
277 (** Our second subgoal requires more work and also demonstrates our first inductive hypothesis.
|
adamc@28
|
278
|
adamc@28
|
279 [[
|
adamc@28
|
280 n : nat
|
adamc@28
|
281 IHn : plus n O = n
|
adamc@28
|
282 ============================
|
adamc@28
|
283 plus (S n) O = S n
|
adamc@208
|
284
|
adamc@28
|
285 ]]
|
adamc@28
|
286
|
adam@315
|
287 We can start out by using computation to simplify the goal as far as we can.%\index{tactics!simpl}% *)
|
adamc@28
|
288
|
adamc@28
|
289 simpl.
|
adamc@28
|
290
|
adam@400
|
291 (** Now the conclusion is [S (plus n O) = S n]. Using our inductive hypothesis: *)
|
adamc@28
|
292
|
adamc@28
|
293 rewrite IHn.
|
adamc@28
|
294
|
adam@448
|
295 (** %\noindent{}%...we get a trivial conclusion [S n = S n]. *)
|
adamc@28
|
296
|
adamc@28
|
297 reflexivity.
|
adamc@28
|
298
|
adam@315
|
299 (** Not much really went on in this proof, so the [crush] tactic from the [CpdtTactics] module can prove this theorem automatically. *)
|
adamc@28
|
300
|
adamc@28
|
301 Restart.
|
adam@315
|
302
|
adamc@28
|
303 induction n; crush.
|
adamc@28
|
304 Qed.
|
adamc@41
|
305 (* end thide *)
|
adamc@28
|
306
|
adamc@28
|
307 (** We can check out the induction principle at work here: *)
|
adamc@28
|
308
|
adamc@28
|
309 Check nat_ind.
|
adamc@208
|
310 (** %\vspace{-.15in}% [[
|
adamc@208
|
311 nat_ind : forall P : nat -> Prop,
|
adamc@208
|
312 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
|
adamc@208
|
313 ]]
|
adamc@28
|
314
|
adam@442
|
315 Each of the two cases of our last proof came from the type of one of the arguments to [nat_ind]. We chose [P] to be [(fun n : nat => plus n O = n)]. The first proof case corresponded to [P O] and the second case to [(forall n : nat, P n -> P (S n))]. The free variable [n] and inductive hypothesis [IHn] came from the argument types given here.
|
adamc@28
|
316
|
adam@315
|
317 Since [nat] has a constructor that takes an argument, we may sometimes need to know that that constructor is injective.%\index{tactics!injection}\index{tactics!trivial}% *)
|
adamc@28
|
318
|
adamc@28
|
319 Theorem S_inj : forall n m : nat, S n = S m -> n = m.
|
adamc@41
|
320 (* begin thide *)
|
adamc@28
|
321 injection 1; trivial.
|
adamc@28
|
322 Qed.
|
adamc@41
|
323 (* end thide *)
|
adamc@28
|
324
|
adam@400
|
325 (** The [injection] tactic refers to a premise by number, adding new equalities between the corresponding arguments of equated terms that are formed with the same constructor. We end up needing to prove [n = m -> n = m], so it is unsurprising that a tactic named [trivial] is able to finish the proof.
|
adamc@28
|
326
|
adam@417
|
327 There is also a very useful tactic called %\index{tactics!congruence}%[congruence] that can prove this theorem immediately. The [congruence] tactic generalizes [discriminate] and [injection], and it also adds reasoning about the general properties of equality, such as that a function returns equal results on equal arguments. That is, [congruence] is a%\index{theory of equality and uninterpreted functions}% _complete decision procedure for the theory of equality and uninterpreted functions_, plus some smarts about inductive types.
|
adamc@29
|
328
|
adamc@29
|
329 %\medskip%
|
adamc@29
|
330
|
adamc@29
|
331 We can define a type of lists of natural numbers. *)
|
adamc@29
|
332
|
adamc@29
|
333 Inductive nat_list : Set :=
|
adamc@29
|
334 | NNil : nat_list
|
adamc@29
|
335 | NCons : nat -> nat_list -> nat_list.
|
adamc@29
|
336
|
adam@470
|
337 (** Recursive definitions over [nat_list] are straightforward extensions of what we have seen before. *)
|
adamc@29
|
338
|
adamc@29
|
339 Fixpoint nlength (ls : nat_list) : nat :=
|
adamc@29
|
340 match ls with
|
adamc@29
|
341 | NNil => O
|
adamc@29
|
342 | NCons _ ls' => S (nlength ls')
|
adamc@29
|
343 end.
|
adamc@29
|
344
|
adamc@208
|
345 Fixpoint napp (ls1 ls2 : nat_list) : nat_list :=
|
adamc@29
|
346 match ls1 with
|
adamc@29
|
347 | NNil => ls2
|
adamc@29
|
348 | NCons n ls1' => NCons n (napp ls1' ls2)
|
adamc@29
|
349 end.
|
adamc@29
|
350
|
adamc@29
|
351 (** Inductive theorem proving can again be automated quite effectively. *)
|
adamc@29
|
352
|
adamc@29
|
353 Theorem nlength_napp : forall ls1 ls2 : nat_list, nlength (napp ls1 ls2)
|
adamc@29
|
354 = plus (nlength ls1) (nlength ls2).
|
adamc@41
|
355 (* begin thide *)
|
adamc@29
|
356 induction ls1; crush.
|
adamc@29
|
357 Qed.
|
adamc@41
|
358 (* end thide *)
|
adamc@29
|
359
|
adamc@29
|
360 Check nat_list_ind.
|
adamc@208
|
361 (** %\vspace{-.15in}% [[
|
adamc@208
|
362 nat_list_ind
|
adamc@29
|
363 : forall P : nat_list -> Prop,
|
adamc@29
|
364 P NNil ->
|
adamc@29
|
365 (forall (n : nat) (n0 : nat_list), P n0 -> P (NCons n n0)) ->
|
adamc@29
|
366 forall n : nat_list, P n
|
adamc@29
|
367 ]]
|
adamc@29
|
368
|
adamc@29
|
369 %\medskip%
|
adamc@29
|
370
|
adam@420
|
371 In general, we can implement any "tree" types as inductive types. For example, here are binary trees of naturals. *)
|
adamc@29
|
372
|
adamc@29
|
373 Inductive nat_btree : Set :=
|
adamc@29
|
374 | NLeaf : nat_btree
|
adamc@29
|
375 | NNode : nat_btree -> nat -> nat_btree -> nat_btree.
|
adamc@29
|
376
|
adamc@29
|
377 Fixpoint nsize (tr : nat_btree) : nat :=
|
adamc@29
|
378 match tr with
|
adamc@35
|
379 | NLeaf => S O
|
adamc@29
|
380 | NNode tr1 _ tr2 => plus (nsize tr1) (nsize tr2)
|
adamc@29
|
381 end.
|
adamc@29
|
382
|
adamc@208
|
383 Fixpoint nsplice (tr1 tr2 : nat_btree) : nat_btree :=
|
adamc@29
|
384 match tr1 with
|
adamc@35
|
385 | NLeaf => NNode tr2 O NLeaf
|
adamc@29
|
386 | NNode tr1' n tr2' => NNode (nsplice tr1' tr2) n tr2'
|
adamc@29
|
387 end.
|
adamc@29
|
388
|
adamc@29
|
389 Theorem plus_assoc : forall n1 n2 n3 : nat, plus (plus n1 n2) n3 = plus n1 (plus n2 n3).
|
adamc@41
|
390 (* begin thide *)
|
adamc@29
|
391 induction n1; crush.
|
adamc@29
|
392 Qed.
|
adamc@41
|
393 (* end thide *)
|
adamc@29
|
394
|
adamc@29
|
395 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2)
|
adamc@29
|
396 = plus (nsize tr2) (nsize tr1).
|
adamc@41
|
397 (* begin thide *)
|
adam@375
|
398 Hint Rewrite n_plus_O plus_assoc.
|
adamc@29
|
399
|
adamc@29
|
400 induction tr1; crush.
|
adamc@29
|
401 Qed.
|
adamc@41
|
402 (* end thide *)
|
adamc@29
|
403
|
adam@315
|
404 (** It is convenient that these proofs go through so easily, but it is useful to check that the tree induction principle works as usual. *)
|
adam@315
|
405
|
adamc@29
|
406 Check nat_btree_ind.
|
adamc@208
|
407 (** %\vspace{-.15in}% [[
|
adamc@208
|
408 nat_btree_ind
|
adamc@29
|
409 : forall P : nat_btree -> Prop,
|
adamc@29
|
410 P NLeaf ->
|
adamc@29
|
411 (forall n : nat_btree,
|
adamc@29
|
412 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) ->
|
adamc@29
|
413 forall n : nat_btree, P n
|
adam@302
|
414 ]]
|
adam@315
|
415
|
adam@315
|
416 We have the usual two cases, one for each constructor of [nat_btree]. *)
|
adamc@30
|
417
|
adamc@30
|
418
|
adamc@30
|
419 (** * Parameterized Types *)
|
adamc@30
|
420
|
adam@316
|
421 (** We can also define %\index{polymorphism}%polymorphic inductive types, as with algebraic datatypes in Haskell and ML.%\index{Gallina terms!list}\index{Gallina terms!Nil}\index{Gallina terms!Cons}\index{Gallina terms!length}\index{Gallina terms!app}% *)
|
adamc@30
|
422
|
adamc@30
|
423 Inductive list (T : Set) : Set :=
|
adamc@30
|
424 | Nil : list T
|
adamc@30
|
425 | Cons : T -> list T -> list T.
|
adamc@30
|
426
|
adamc@30
|
427 Fixpoint length T (ls : list T) : nat :=
|
adamc@30
|
428 match ls with
|
adamc@30
|
429 | Nil => O
|
adamc@30
|
430 | Cons _ ls' => S (length ls')
|
adamc@30
|
431 end.
|
adamc@30
|
432
|
adamc@208
|
433 Fixpoint app T (ls1 ls2 : list T) : list T :=
|
adamc@30
|
434 match ls1 with
|
adamc@30
|
435 | Nil => ls2
|
adamc@30
|
436 | Cons x ls1' => Cons x (app ls1' ls2)
|
adamc@30
|
437 end.
|
adamc@30
|
438
|
adamc@30
|
439 Theorem length_app : forall T (ls1 ls2 : list T), length (app ls1 ls2)
|
adamc@30
|
440 = plus (length ls1) (length ls2).
|
adamc@41
|
441 (* begin thide *)
|
adamc@30
|
442 induction ls1; crush.
|
adamc@30
|
443 Qed.
|
adamc@41
|
444 (* end thide *)
|
adamc@30
|
445
|
adam@420
|
446 (** There is a useful shorthand for writing many definitions that share the same parameter, based on Coq's%\index{sections}\index{Vernacular commands!Section}\index{Vernacular commands!Variable}% _section_ mechanism. The following block of code is equivalent to the above: *)
|
adamc@30
|
447
|
adamc@30
|
448 (* begin hide *)
|
adamc@30
|
449 Reset list.
|
adamc@30
|
450 (* end hide *)
|
adamc@30
|
451
|
adamc@30
|
452 Section list.
|
adamc@30
|
453 Variable T : Set.
|
adamc@30
|
454
|
adamc@30
|
455 Inductive list : Set :=
|
adamc@30
|
456 | Nil : list
|
adamc@30
|
457 | Cons : T -> list -> list.
|
adamc@30
|
458
|
adamc@30
|
459 Fixpoint length (ls : list) : nat :=
|
adamc@30
|
460 match ls with
|
adamc@30
|
461 | Nil => O
|
adamc@30
|
462 | Cons _ ls' => S (length ls')
|
adamc@30
|
463 end.
|
adamc@30
|
464
|
adamc@208
|
465 Fixpoint app (ls1 ls2 : list) : list :=
|
adamc@30
|
466 match ls1 with
|
adamc@30
|
467 | Nil => ls2
|
adamc@30
|
468 | Cons x ls1' => Cons x (app ls1' ls2)
|
adamc@30
|
469 end.
|
adamc@30
|
470
|
adamc@30
|
471 Theorem length_app : forall ls1 ls2 : list, length (app ls1 ls2)
|
adamc@30
|
472 = plus (length ls1) (length ls2).
|
adamc@41
|
473 (* begin thide *)
|
adamc@30
|
474 induction ls1; crush.
|
adamc@30
|
475 Qed.
|
adamc@41
|
476 (* end thide *)
|
adamc@30
|
477 End list.
|
adamc@30
|
478
|
adamc@35
|
479 Implicit Arguments Nil [T].
|
adamc@35
|
480
|
adam@469
|
481 (** After we end the section, the [Variable]s we used are added as extra function parameters for each defined identifier, as needed. With an [Implicit Arguments]%~\index{Vernacular commands!Implicit Arguments}% command, we ask that [T] be inferred when we use [Nil]; Coq's heuristics already decided to apply a similar policy to [Cons], because of the [Set Implicit Arguments]%~\index{Vernacular commands!Set Implicit Arguments}% command elided at the beginning of this chapter. We verify that our definitions have been saved properly using the [Print] command, a cousin of [Check] which shows the definition of a symbol, rather than just its type. *)
|
adamc@30
|
482
|
adamc@202
|
483 Print list.
|
adamc@208
|
484 (** %\vspace{-.15in}% [[
|
adamc@208
|
485 Inductive list (T : Set) : Set :=
|
adam@316
|
486 Nil : list T | Cons : T -> list T -> list T
|
adamc@202
|
487 ]]
|
adamc@30
|
488
|
adam@442
|
489 The final definition is the same as what we wrote manually before. The other elements of the section are altered similarly, turning out exactly as they were before, though we managed to write their definitions more succinctly. *)
|
adamc@30
|
490
|
adamc@30
|
491 Check length.
|
adamc@208
|
492 (** %\vspace{-.15in}% [[
|
adamc@208
|
493 length
|
adamc@30
|
494 : forall T : Set, list T -> nat
|
adamc@30
|
495 ]]
|
adamc@30
|
496
|
adam@442
|
497 The parameter [T] is treated as a new argument to the induction principle, too. *)
|
adamc@30
|
498
|
adamc@30
|
499 Check list_ind.
|
adamc@208
|
500 (** %\vspace{-.15in}% [[
|
adamc@208
|
501 list_ind
|
adamc@30
|
502 : forall (T : Set) (P : list T -> Prop),
|
adamc@30
|
503 P (Nil T) ->
|
adamc@30
|
504 (forall (t : T) (l : list T), P l -> P (Cons t l)) ->
|
adamc@30
|
505 forall l : list T, P l
|
adamc@30
|
506 ]]
|
adamc@30
|
507
|
adam@442
|
508 Thus, even though we just saw that [T] is added as an extra argument to the constructor [Cons], there is no quantifier for [T] in the type of the inductive case like there is for each of the other arguments. *)
|
adamc@31
|
509
|
adamc@31
|
510
|
adamc@31
|
511 (** * Mutually Inductive Types *)
|
adamc@31
|
512
|
adamc@31
|
513 (** We can define inductive types that refer to each other: *)
|
adamc@31
|
514
|
adamc@31
|
515 Inductive even_list : Set :=
|
adamc@31
|
516 | ENil : even_list
|
adamc@31
|
517 | ECons : nat -> odd_list -> even_list
|
adamc@31
|
518
|
adamc@31
|
519 with odd_list : Set :=
|
adamc@31
|
520 | OCons : nat -> even_list -> odd_list.
|
adamc@31
|
521
|
adamc@31
|
522 Fixpoint elength (el : even_list) : nat :=
|
adamc@31
|
523 match el with
|
adamc@31
|
524 | ENil => O
|
adamc@31
|
525 | ECons _ ol => S (olength ol)
|
adamc@31
|
526 end
|
adamc@31
|
527
|
adamc@31
|
528 with olength (ol : odd_list) : nat :=
|
adamc@31
|
529 match ol with
|
adamc@31
|
530 | OCons _ el => S (elength el)
|
adamc@31
|
531 end.
|
adamc@31
|
532
|
adamc@208
|
533 Fixpoint eapp (el1 el2 : even_list) : even_list :=
|
adamc@31
|
534 match el1 with
|
adamc@31
|
535 | ENil => el2
|
adamc@31
|
536 | ECons n ol => ECons n (oapp ol el2)
|
adamc@31
|
537 end
|
adamc@31
|
538
|
adamc@208
|
539 with oapp (ol : odd_list) (el : even_list) : odd_list :=
|
adamc@31
|
540 match ol with
|
adamc@31
|
541 | OCons n el' => OCons n (eapp el' el)
|
adamc@31
|
542 end.
|
adamc@31
|
543
|
adamc@31
|
544 (** Everything is going roughly the same as in past examples, until we try to prove a theorem similar to those that came before. *)
|
adamc@31
|
545
|
adamc@31
|
546 Theorem elength_eapp : forall el1 el2 : even_list,
|
adamc@31
|
547 elength (eapp el1 el2) = plus (elength el1) (elength el2).
|
adamc@41
|
548 (* begin thide *)
|
adamc@31
|
549 induction el1; crush.
|
adamc@31
|
550
|
adamc@31
|
551 (** One goal remains: [[
|
adamc@31
|
552
|
adamc@31
|
553 n : nat
|
adamc@31
|
554 o : odd_list
|
adamc@31
|
555 el2 : even_list
|
adamc@31
|
556 ============================
|
adamc@31
|
557 S (olength (oapp o el2)) = S (plus (olength o) (elength el2))
|
adamc@31
|
558 ]]
|
adamc@31
|
559
|
adamc@31
|
560 We have no induction hypothesis, so we cannot prove this goal without starting another induction, which would reach a similar point, sending us into a futile infinite chain of inductions. The problem is that Coq's generation of [T_ind] principles is incomplete. We only get non-mutual induction principles generated by default. *)
|
adamc@31
|
561
|
adamc@31
|
562 Abort.
|
adamc@31
|
563 Check even_list_ind.
|
adamc@208
|
564 (** %\vspace{-.15in}% [[
|
adamc@208
|
565 even_list_ind
|
adamc@31
|
566 : forall P : even_list -> Prop,
|
adamc@31
|
567 P ENil ->
|
adamc@31
|
568 (forall (n : nat) (o : odd_list), P (ECons n o)) ->
|
adamc@31
|
569 forall e : even_list, P e
|
adamc@31
|
570 ]]
|
adamc@31
|
571
|
adam@442
|
572 We see that no inductive hypotheses are included anywhere in the type. To get them, we must ask for mutual principles as we need them, using the %\index{Vernacular commands!Scheme}%[Scheme] command. *)
|
adamc@31
|
573
|
adamc@31
|
574 Scheme even_list_mut := Induction for even_list Sort Prop
|
adamc@31
|
575 with odd_list_mut := Induction for odd_list Sort Prop.
|
adamc@31
|
576
|
adam@316
|
577 (** This invocation of [Scheme] asks for the creation of induction principles [even_list_mut] for the type [even_list] and [odd_list_mut] for the type [odd_list]. The [Induction] keyword says we want standard induction schemes, since [Scheme] supports more exotic choices. Finally, [Sort Prop] establishes that we really want induction schemes, not recursion schemes, which are the same according to Curry-Howard, save for the [Prop]/[Set] distinction. *)
|
adam@316
|
578
|
adamc@31
|
579 Check even_list_mut.
|
adamc@208
|
580 (** %\vspace{-.15in}% [[
|
adamc@208
|
581 even_list_mut
|
adamc@31
|
582 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
|
adamc@31
|
583 P ENil ->
|
adamc@31
|
584 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
|
adamc@31
|
585 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
|
adamc@31
|
586 forall e : even_list, P e
|
adamc@31
|
587 ]]
|
adamc@31
|
588
|
adam@316
|
589 This is the principle we wanted in the first place. There is one more wrinkle left in using it: the [induction] tactic will not apply it for us automatically. It will be helpful to look at how to prove one of our past examples without using [induction], so that we can then generalize the technique to mutual inductive types.%\index{tactics!apply}% *)
|
adamc@31
|
590
|
adamc@31
|
591 Theorem n_plus_O' : forall n : nat, plus n O = n.
|
adamc@31
|
592 apply (nat_ind (fun n => plus n O = n)); crush.
|
adamc@31
|
593 Qed.
|
adamc@31
|
594
|
adamc@31
|
595 (** From this example, we can see that [induction] is not magic. It only does some bookkeeping for us to make it easy to apply a theorem, which we can do directly with the [apply] tactic. We apply not just an identifier but a partial application of it, specifying the predicate we mean to prove holds for all naturals.
|
adamc@31
|
596
|
adamc@31
|
597 This technique generalizes to our mutual example: *)
|
adamc@31
|
598
|
adamc@31
|
599 Theorem elength_eapp : forall el1 el2 : even_list,
|
adamc@31
|
600 elength (eapp el1 el2) = plus (elength el1) (elength el2).
|
adamc@41
|
601
|
adamc@31
|
602 apply (even_list_mut
|
adamc@31
|
603 (fun el1 : even_list => forall el2 : even_list,
|
adamc@31
|
604 elength (eapp el1 el2) = plus (elength el1) (elength el2))
|
adamc@31
|
605 (fun ol : odd_list => forall el : even_list,
|
adamc@31
|
606 olength (oapp ol el) = plus (olength ol) (elength el))); crush.
|
adamc@31
|
607 Qed.
|
adamc@41
|
608 (* end thide *)
|
adamc@31
|
609
|
adamc@31
|
610 (** We simply need to specify two predicates, one for each of the mutually inductive types. In general, it would not be a good idea to assume that a proof assistant could infer extra predicates, so this way of applying mutual induction is about as straightforward as we could hope for. *)
|
adamc@33
|
611
|
adamc@33
|
612
|
adamc@33
|
613 (** * Reflexive Types *)
|
adamc@33
|
614
|
adam@469
|
615 (** A kind of inductive type called a _reflexive type_ includes at least one constructor that takes as an argument _a function returning the same type we are defining_. One very useful class of examples is in modeling variable binders. Our example will be an encoding of the syntax of first-order logic. Since the idea of syntactic encodings of logic may require a bit of acclimation, let us first consider a simpler formula type for a subset of propositional logic. We are not yet using a reflexive type, but later we will extend the example reflexively. *)
|
adam@316
|
616
|
adam@316
|
617 Inductive pformula : Set :=
|
adam@316
|
618 | Truth : pformula
|
adam@316
|
619 | Falsehood : pformula
|
adam@316
|
620 | Conjunction : pformula -> pformula -> pformula.
|
adam@316
|
621
|
adam@448
|
622 (* begin hide *)
|
adam@448
|
623 (* begin thide *)
|
adam@448
|
624 Definition prod' := prod.
|
adam@448
|
625 (* end thide *)
|
adam@448
|
626 (* end hide *)
|
adam@448
|
627
|
adam@470
|
628 (** A key distinction here is between, for instance, the _syntax_ [Truth] and its _semantics_ [True]. We can make the semantics explicit with a recursive function. This function uses the infix operator %\index{Gallina operators!/\textbackslash}%[/\], which desugars to uses of the type family %\index{Gallina terms!and}%[and] from the standard library. The family [and] implements conjunction, the [Prop] Curry-Howard analogue of the usual pair type from functional programming (which is the type family %\index{Gallina terms!prod}%[prod] in Coq's standard library). *)
|
adam@316
|
629
|
adam@316
|
630 Fixpoint pformulaDenote (f : pformula) : Prop :=
|
adam@316
|
631 match f with
|
adam@316
|
632 | Truth => True
|
adam@316
|
633 | Falsehood => False
|
adam@316
|
634 | Conjunction f1 f2 => pformulaDenote f1 /\ pformulaDenote f2
|
adam@316
|
635 end.
|
adam@316
|
636
|
adam@392
|
637 (** This is just a warm-up that does not use reflexive types, the new feature we mean to introduce. When we set our sights on first-order logic instead, it becomes very handy to give constructors recursive arguments that are functions. *)
|
adamc@33
|
638
|
adamc@33
|
639 Inductive formula : Set :=
|
adamc@33
|
640 | Eq : nat -> nat -> formula
|
adamc@33
|
641 | And : formula -> formula -> formula
|
adamc@33
|
642 | Forall : (nat -> formula) -> formula.
|
adamc@33
|
643
|
adam@470
|
644 (** Our kinds of formulas are equalities between naturals, conjunction, and universal quantification over natural numbers. We avoid needing to include a notion of "variables" in our type, by using Coq functions to encode the syntax of quantification. For instance, here is the encoding of [forall x : nat, x = x]:%\index{Vernacular commands!Example}% *)
|
adamc@33
|
645
|
adamc@33
|
646 Example forall_refl : formula := Forall (fun x => Eq x x).
|
adamc@33
|
647
|
adamc@33
|
648 (** We can write recursive functions over reflexive types quite naturally. Here is one translating our formulas into native Coq propositions. *)
|
adamc@33
|
649
|
adamc@33
|
650 Fixpoint formulaDenote (f : formula) : Prop :=
|
adamc@33
|
651 match f with
|
adamc@33
|
652 | Eq n1 n2 => n1 = n2
|
adamc@33
|
653 | And f1 f2 => formulaDenote f1 /\ formulaDenote f2
|
adamc@33
|
654 | Forall f' => forall n : nat, formulaDenote (f' n)
|
adamc@33
|
655 end.
|
adamc@33
|
656
|
adamc@33
|
657 (** We can also encode a trivial formula transformation that swaps the order of equality and conjunction operands. *)
|
adamc@33
|
658
|
adamc@33
|
659 Fixpoint swapper (f : formula) : formula :=
|
adamc@33
|
660 match f with
|
adamc@33
|
661 | Eq n1 n2 => Eq n2 n1
|
adamc@33
|
662 | And f1 f2 => And (swapper f2) (swapper f1)
|
adamc@33
|
663 | Forall f' => Forall (fun n => swapper (f' n))
|
adamc@33
|
664 end.
|
adamc@33
|
665
|
adamc@33
|
666 (** It is helpful to prove that this transformation does not make true formulas false. *)
|
adamc@33
|
667
|
adamc@33
|
668 Theorem swapper_preserves_truth : forall f, formulaDenote f -> formulaDenote (swapper f).
|
adamc@41
|
669 (* begin thide *)
|
adamc@33
|
670 induction f; crush.
|
adamc@33
|
671 Qed.
|
adamc@41
|
672 (* end thide *)
|
adamc@33
|
673
|
adamc@33
|
674 (** We can take a look at the induction principle behind this proof. *)
|
adamc@33
|
675
|
adamc@33
|
676 Check formula_ind.
|
adamc@208
|
677 (** %\vspace{-.15in}% [[
|
adamc@208
|
678 formula_ind
|
adamc@33
|
679 : forall P : formula -> Prop,
|
adamc@33
|
680 (forall n n0 : nat, P (Eq n n0)) ->
|
adamc@33
|
681 (forall f0 : formula,
|
adamc@33
|
682 P f0 -> forall f1 : formula, P f1 -> P (And f0 f1)) ->
|
adamc@33
|
683 (forall f1 : nat -> formula,
|
adamc@33
|
684 (forall n : nat, P (f1 n)) -> P (Forall f1)) ->
|
adamc@33
|
685 forall f2 : formula, P f2
|
adamc@208
|
686 ]]
|
adamc@33
|
687
|
adam@442
|
688 Focusing on the [Forall] case, which comes third, we see that we are allowed to assume that the theorem holds _for any application of the argument function [f1]_. That is, Coq induction principles do not follow a simple rule that the textual representations of induction variables must get shorter in appeals to induction hypotheses. Luckily for us, the people behind the metatheory of Coq have verified that this flexibility does not introduce unsoundness.
|
adamc@33
|
689
|
adamc@33
|
690 %\medskip%
|
adamc@33
|
691
|
adam@469
|
692 Up to this point, we have seen how to encode in Coq more and more of what is possible with algebraic datatypes in %\index{Haskell}%Haskell and %\index{ML}%ML. This may have given the inaccurate impression that inductive types are a strict extension of algebraic datatypes. In fact, Coq must rule out some types allowed by Haskell and ML, for reasons of soundness. Reflexive types provide our first good example of such a case; only some of them are legal.
|
adamc@33
|
693
|
adam@400
|
694 Given our last example of an inductive type, many readers are probably eager to try encoding the syntax of %\index{lambda calculus}%lambda calculus. Indeed, the function-based representation technique that we just used, called%\index{higher-order abstract syntax}\index{HOAS|see{higher-order abstract syntax}}% _higher-order abstract syntax_ (HOAS)%~\cite{HOAS}%, is the representation of choice for lambda calculi in %\index{Twelf}%Twelf and in many applications implemented in Haskell and ML. Let us try to import that choice to Coq: *)
|
adam@400
|
695 (* begin hide *)
|
adam@437
|
696 (* begin thide *)
|
adam@400
|
697 Inductive term : Set := App | Abs.
|
adam@400
|
698 Reset term.
|
adam@420
|
699 Definition uhoh := O.
|
adam@437
|
700 (* end thide *)
|
adam@400
|
701 (* end hide *)
|
adamc@33
|
702 (** [[
|
adamc@33
|
703 Inductive term : Set :=
|
adamc@33
|
704 | App : term -> term -> term
|
adamc@33
|
705 | Abs : (term -> term) -> term.
|
adamc@33
|
706 ]]
|
adamc@33
|
707
|
adam@316
|
708 <<
|
adam@316
|
709 Error: Non strictly positive occurrence of "term" in "(term -> term) -> term"
|
adam@316
|
710 >>
|
adam@316
|
711
|
adam@400
|
712 We have run afoul of the%\index{strict positivity requirement}\index{positivity requirement}% _strict positivity requirement_ for inductive definitions, which says that the type being defined may not occur to the left of an arrow in the type of a constructor argument. It is important that the type of a constructor is viewed in terms of a series of arguments and a result, since obviously we need recursive occurrences to the lefts of the outermost arrows if we are to have recursive occurrences at all. Our candidate definition above violates the positivity requirement because it involves an argument of type [term -> term], where the type [term] that we are defining appears to the left of an arrow. The candidate type of [App] is fine, however, since every occurrence of [term] is either a constructor argument or the final result type.
|
adamc@33
|
713
|
adamc@33
|
714 Why must Coq enforce this restriction? Imagine that our last definition had been accepted, allowing us to write this function:
|
adamc@33
|
715
|
adam@439
|
716 %\vspace{-.15in}%[[
|
adamc@33
|
717 Definition uhoh (t : term) : term :=
|
adamc@33
|
718 match t with
|
adamc@33
|
719 | Abs f => f t
|
adamc@33
|
720 | _ => t
|
adamc@33
|
721 end.
|
adamc@205
|
722 ]]
|
adamc@205
|
723
|
adamc@33
|
724 Using an informal idea of Coq's semantics, it is easy to verify that the application [uhoh (Abs uhoh)] will run forever. This would be a mere curiosity in OCaml and Haskell, where non-termination is commonplace, though the fact that we have a non-terminating program without explicit recursive function definitions is unusual.
|
adamc@33
|
725
|
adam@316
|
726 %\index{termination checking}%For Coq, however, this would be a disaster. The possibility of writing such a function would destroy all our confidence that proving a theorem means anything. Since Coq combines programs and proofs in one language, we would be able to prove every theorem with an infinite loop.
|
adamc@33
|
727
|
adam@439
|
728 Nonetheless, the basic insight of HOAS is a very useful one, and there are ways to realize most benefits of HOAS in Coq. We will study a particular technique of this kind in the final chapter, on programming language syntax and semantics. *)
|
adamc@34
|
729
|
adamc@34
|
730
|
adam@317
|
731 (** * An Interlude on Induction Principles *)
|
adamc@34
|
732
|
adam@317
|
733 (** As we have emphasized a few times already, Coq proofs are actually programs, written in the same language we have been using in our examples all along. We can get a first sense of what this means by taking a look at the definitions of some of the %\index{induction principles}%induction principles we have used. A close look at the details here will help us construct induction principles manually, which we will see is necessary for some more advanced inductive definitions. *)
|
adamc@34
|
734
|
adam@470
|
735 Print nat_ind.
|
adam@437
|
736 (** %\vspace{-.15in}%[[
|
adam@470
|
737 nat_ind =
|
adam@470
|
738 fun P : nat -> Prop => nat_rect P
|
adam@470
|
739 : forall P : nat -> Prop,
|
adam@470
|
740 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
|
adamc@34
|
741 ]]
|
adamc@34
|
742
|
adam@470
|
743 We see that this induction principle is defined in terms of a more general principle, [nat_rect]. The <<rec>> stands for "recursion principle," and the <<t>> at the end stands for [Type]. *)
|
adamc@34
|
744
|
adam@470
|
745 Check nat_rect.
|
adamc@208
|
746 (** %\vspace{-.15in}% [[
|
adam@470
|
747 nat_rect
|
adam@470
|
748 : forall P : nat -> Type,
|
adam@470
|
749 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
|
adamc@34
|
750 ]]
|
adamc@34
|
751
|
adam@470
|
752 The principle [nat_rect] gives [P] type [nat -> Type] instead of [nat -> Prop]. This [Type] is another universe, like [Set] and [Prop]. In fact, it is a common supertype of both. Later on, we will discuss exactly what the significances of the different universes are. For now, it is just important that we can use [Type] as a sort of meta-universe that may turn out to be either [Set] or [Prop]. We can see the symmetry inherent in the subtyping relationship by printing the definition of another principle that was generated for [nat] automatically: *)
|
adamc@34
|
753
|
adam@470
|
754 Print nat_rec.
|
adam@437
|
755 (** %\vspace{-.15in}%[[
|
adam@470
|
756 nat_rec =
|
adam@470
|
757 fun P : nat -> Set => nat_rect P
|
adam@470
|
758 : forall P : nat -> Set,
|
adam@470
|
759 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
|
adamc@34
|
760 ]]
|
adamc@34
|
761
|
adam@470
|
762 This is identical to the definition for [nat_ind], except that we have substituted [Set] for [Prop]. For most inductive types [T], then, we get not just induction principles [T_ind], but also %\index{recursion principles}%recursion principles [T_rec]. We can use [T_rec] to write recursive definitions without explicit [Fixpoint] recursion. For instance, the following two definitions are equivalent: *)
|
adamc@34
|
763
|
adam@470
|
764 Fixpoint plus_recursive (n : nat) : nat -> nat :=
|
adam@470
|
765 match n with
|
adam@470
|
766 | O => fun m => m
|
adam@470
|
767 | S n' => fun m => S (plus_recursive n' m)
|
adamc@34
|
768 end.
|
adamc@34
|
769
|
adam@470
|
770 Definition plus_rec : nat -> nat -> nat :=
|
adam@470
|
771 nat_rec (fun _ : nat => nat -> nat) (fun _ => O) (fun _ r m => S (r m)).
|
adamc@34
|
772
|
adam@470
|
773 (** Going even further down the rabbit hole, [nat_rect] itself is not even a primitive. It is a functional program that we can write manually. *)
|
adamc@34
|
774
|
adam@470
|
775 Print nat_rect.
|
adam@437
|
776 (** %\vspace{-.15in}%[[
|
adam@470
|
777 nat_rect =
|
adam@470
|
778 fun (P : nat -> Type) (f : P O) (f0 : forall n : nat, P n -> P (S n)) =>
|
adam@470
|
779 fix F (n : nat) : P n :=
|
adam@470
|
780 match n as n0 return (P n0) with
|
adam@470
|
781 | O => f
|
adam@470
|
782 | S n0 => f0 n0 (F n0)
|
adamc@208
|
783 end
|
adam@470
|
784 : forall P : nat -> Type,
|
adam@470
|
785 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
|
adamc@34
|
786 ]]
|
adamc@34
|
787
|
adam@470
|
788 The only new wrinkle heres are, first, an anonymous recursive function definition, using the %\index{Gallina terms!fix}%[fix] keyword of Gallina (which is like [fun] with recursion supported); and, second, the annotations on the [match] expression. This is a%\index{dependent pattern matching}% _dependently typed_ pattern match, because the _type_ of the expression depends on the _value_ being matched on. We will meet more involved examples later, especially in Part II of the book.
|
adam@317
|
789
|
adam@470
|
790 %\index{type inference}%Type inference for dependent pattern matching is undecidable, which can be proved by reduction from %\index{higher-order unification}%higher-order unification%~\cite{HOU}%. Thus, we often find ourselves needing to annotate our programs in a way that explains dependencies to the type checker. In the example of [nat_rect], we have an %\index{Gallina terms!as}%[as] clause, which binds a name for the discriminee; and a %\index{Gallina terms!return}%[return] clause, which gives a way to compute the [match] result type as a function of the discriminee.
|
adamc@34
|
791
|
adam@470
|
792 To prove that [nat_rect] is nothing special, we can reimplement it manually. *)
|
adamc@34
|
793
|
adam@470
|
794 Fixpoint nat_rect' (P : nat -> Type)
|
adam@470
|
795 (HO : P O)
|
adam@470
|
796 (HS : forall n, P n -> P (S n)) (n : nat) :=
|
adam@470
|
797 match n return P n with
|
adam@470
|
798 | O => HO
|
adam@470
|
799 | S n' => HS n' (nat_rect' P HO HS n')
|
adamc@34
|
800 end.
|
adamc@34
|
801
|
adam@470
|
802 (** We can understand the definition of [nat_rect] better by reimplementing [nat_ind] using sections. *)
|
adamc@34
|
803
|
adam@317
|
804 Section nat_ind'.
|
adamc@208
|
805 (** First, we have the property of natural numbers that we aim to prove. *)
|
adamc@34
|
806
|
adam@317
|
807 Variable P : nat -> Prop.
|
adamc@34
|
808
|
adam@317
|
809 (** Then we require a proof of the [O] case, which we declare with the command %\index{Vernacular commands!Hypothesis}%[Hypothesis], which is a synonym for [Variable] that, by convention, is used for variables whose types are propositions. *)
|
adamc@34
|
810
|
adam@317
|
811 Hypothesis O_case : P O.
|
adamc@34
|
812
|
adamc@208
|
813 (** Next is a proof of the [S] case, which may assume an inductive hypothesis. *)
|
adamc@34
|
814
|
adam@317
|
815 Hypothesis S_case : forall n : nat, P n -> P (S n).
|
adamc@34
|
816
|
adamc@208
|
817 (** Finally, we define a recursive function to tie the pieces together. *)
|
adamc@34
|
818
|
adam@317
|
819 Fixpoint nat_ind' (n : nat) : P n :=
|
adam@317
|
820 match n with
|
adam@317
|
821 | O => O_case
|
adam@317
|
822 | S n' => S_case (nat_ind' n')
|
adam@317
|
823 end.
|
adam@317
|
824 End nat_ind'.
|
adamc@34
|
825
|
adam@400
|
826 (** Closing the section adds the [Variable]s and [Hypothesis]es as new [fun]-bound arguments to [nat_ind'], and, modulo the use of [Prop] instead of [Type], we end up with the exact same definition that was generated automatically for [nat_rect].
|
adamc@34
|
827
|
adam@317
|
828 %\medskip%
|
adamc@34
|
829
|
adam@317
|
830 We can also examine the definition of [even_list_mut], which we generated with [Scheme] for a mutually recursive type. *)
|
adamc@34
|
831
|
adam@317
|
832 Print even_list_mut.
|
adam@439
|
833 (** %\vspace{-.15in}%[[
|
adam@317
|
834 even_list_mut =
|
adam@317
|
835 fun (P : even_list -> Prop) (P0 : odd_list -> Prop)
|
adam@317
|
836 (f : P ENil) (f0 : forall (n : nat) (o : odd_list), P0 o -> P (ECons n o))
|
adam@317
|
837 (f1 : forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) =>
|
adam@317
|
838 fix F (e : even_list) : P e :=
|
adam@317
|
839 match e as e0 return (P e0) with
|
adam@317
|
840 | ENil => f
|
adam@317
|
841 | ECons n o => f0 n o (F0 o)
|
adam@317
|
842 end
|
adam@317
|
843 with F0 (o : odd_list) : P0 o :=
|
adam@317
|
844 match o as o0 return (P0 o0) with
|
adam@317
|
845 | OCons n e => f1 n e (F e)
|
adam@317
|
846 end
|
adam@317
|
847 for F
|
adam@317
|
848 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
|
adam@317
|
849 P ENil ->
|
adam@317
|
850 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
|
adam@317
|
851 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
|
adam@317
|
852 forall e : even_list, P e
|
adam@317
|
853 ]]
|
adamc@34
|
854
|
adam@442
|
855 We see a mutually recursive [fix], with the different functions separated by %\index{Gallina terms!with}%[with] in the same way that they would be separated by <<and>> in ML. A final %\index{Gallina terms!for}%[for] clause identifies which of the mutually recursive functions should be the final value of the [fix] expression. Using this definition as a template, we can reimplement [even_list_mut] directly. *)
|
adamc@208
|
856
|
adam@317
|
857 Section even_list_mut'.
|
adam@317
|
858 (** First, we need the properties that we are proving. *)
|
adamc@208
|
859
|
adam@317
|
860 Variable Peven : even_list -> Prop.
|
adam@317
|
861 Variable Podd : odd_list -> Prop.
|
adamc@208
|
862
|
adam@317
|
863 (** Next, we need proofs of the three cases. *)
|
adamc@208
|
864
|
adam@317
|
865 Hypothesis ENil_case : Peven ENil.
|
adam@317
|
866 Hypothesis ECons_case : forall (n : nat) (o : odd_list), Podd o -> Peven (ECons n o).
|
adam@317
|
867 Hypothesis OCons_case : forall (n : nat) (e : even_list), Peven e -> Podd (OCons n e).
|
adamc@208
|
868
|
adam@317
|
869 (** Finally, we define the recursive functions. *)
|
adamc@208
|
870
|
adam@317
|
871 Fixpoint even_list_mut' (e : even_list) : Peven e :=
|
adam@317
|
872 match e with
|
adam@317
|
873 | ENil => ENil_case
|
adam@317
|
874 | ECons n o => ECons_case n (odd_list_mut' o)
|
adam@317
|
875 end
|
adam@317
|
876 with odd_list_mut' (o : odd_list) : Podd o :=
|
adam@317
|
877 match o with
|
adam@317
|
878 | OCons n e => OCons_case n (even_list_mut' e)
|
adam@317
|
879 end.
|
adamc@34
|
880 End even_list_mut'.
|
adamc@34
|
881
|
adamc@34
|
882 (** Even induction principles for reflexive types are easy to implement directly. For our [formula] type, we can use a recursive definition much like those we wrote above. *)
|
adamc@34
|
883
|
adamc@34
|
884 Section formula_ind'.
|
adamc@34
|
885 Variable P : formula -> Prop.
|
adamc@38
|
886 Hypothesis Eq_case : forall n1 n2 : nat, P (Eq n1 n2).
|
adamc@38
|
887 Hypothesis And_case : forall f1 f2 : formula,
|
adamc@34
|
888 P f1 -> P f2 -> P (And f1 f2).
|
adamc@38
|
889 Hypothesis Forall_case : forall f : nat -> formula,
|
adamc@34
|
890 (forall n : nat, P (f n)) -> P (Forall f).
|
adamc@34
|
891
|
adamc@34
|
892 Fixpoint formula_ind' (f : formula) : P f :=
|
adamc@208
|
893 match f with
|
adamc@34
|
894 | Eq n1 n2 => Eq_case n1 n2
|
adamc@34
|
895 | And f1 f2 => And_case (formula_ind' f1) (formula_ind' f2)
|
adamc@34
|
896 | Forall f' => Forall_case f' (fun n => formula_ind' (f' n))
|
adamc@34
|
897 end.
|
adamc@34
|
898 End formula_ind'.
|
adamc@34
|
899
|
adam@317
|
900 (** It is apparent that induction principle implementations involve some tedium but not terribly much creativity. *)
|
adam@317
|
901
|
adamc@35
|
902
|
adamc@35
|
903 (** * Nested Inductive Types *)
|
adamc@35
|
904
|
adamc@35
|
905 (** Suppose we want to extend our earlier type of binary trees to trees with arbitrary finite branching. We can use lists to give a simple definition. *)
|
adamc@35
|
906
|
adamc@35
|
907 Inductive nat_tree : Set :=
|
adamc@35
|
908 | NLeaf' : nat_tree
|
adamc@35
|
909 | NNode' : nat -> list nat_tree -> nat_tree.
|
adamc@35
|
910
|
adam@465
|
911 (** This is an example of a%\index{nested inductive type}% _nested_ inductive type definition, because we use the type we are defining as an argument to a parameterized type family. Coq will not allow all such definitions; it effectively pretends that we are defining [nat_tree] mutually with a version of [list] specialized to [nat_tree], checking that the resulting expanded definition satisfies the usual rules. For instance, if we replaced [list] with a type family that used its parameter as a function argument, then the definition would be rejected as violating the positivity restriction.
|
adamc@35
|
912
|
adam@317
|
913 Like we encountered for mutual inductive types, we find that the automatically generated induction principle for [nat_tree] is too weak. *)
|
adamc@35
|
914
|
adamc@35
|
915 Check nat_tree_ind.
|
adamc@208
|
916 (** %\vspace{-.15in}% [[
|
adamc@208
|
917 nat_tree_ind
|
adamc@35
|
918 : forall P : nat_tree -> Prop,
|
adamc@35
|
919 P NLeaf' ->
|
adamc@35
|
920 (forall (n : nat) (l : list nat_tree), P (NNode' n l)) ->
|
adamc@35
|
921 forall n : nat_tree, P n
|
adamc@35
|
922 ]]
|
adamc@35
|
923
|
adam@442
|
924 There is no command like [Scheme] that will implement an improved principle for us. In general, it takes creativity to figure out how to incorporate nested uses of different type families. This is roughly the same creativity employed in the traditional task of strengthening an induction hypothesis. Now that we know how to implement induction principles manually, we are in a position to apply just such creativity to this problem.
|
adamc@35
|
925
|
adamc@35
|
926 First, we will need an auxiliary definition, characterizing what it means for a property to hold of every element of a list. *)
|
adamc@35
|
927
|
adamc@35
|
928 Section All.
|
adamc@35
|
929 Variable T : Set.
|
adamc@35
|
930 Variable P : T -> Prop.
|
adamc@35
|
931
|
adamc@35
|
932 Fixpoint All (ls : list T) : Prop :=
|
adamc@35
|
933 match ls with
|
adamc@35
|
934 | Nil => True
|
adamc@35
|
935 | Cons h t => P h /\ All t
|
adamc@35
|
936 end.
|
adamc@35
|
937 End All.
|
adamc@35
|
938
|
adam@317
|
939 (** It will be useful to review the definitions of [True] and [/\], since we will want to write manual proofs of them below. *)
|
adamc@35
|
940
|
adamc@35
|
941 Print True.
|
adam@439
|
942 (** %\vspace{-.15in}%[[
|
adamc@208
|
943 Inductive True : Prop := I : True
|
adamc@208
|
944 ]]
|
adamc@35
|
945
|
adam@442
|
946 That is, [True] is a proposition with exactly one proof, [I], which we may always supply trivially.
|
adamc@35
|
947
|
adam@400
|
948 Finding the definition of [/\] takes a little more work. Coq supports user registration of arbitrary parsing rules, and it is such a rule that is letting us write [/\] instead of an application of some inductive type family. We can find the underlying inductive type with the %\index{Vernacular commands!Locate}%[Locate] command, whose argument may be a parsing token.%\index{Gallina terms!and}% *)
|
adamc@35
|
949
|
adamc@35
|
950 Locate "/\".
|
adam@439
|
951 (** %\vspace{-.15in}%[[
|
adam@317
|
952 "A /\ B" := and A B : type_scope (default interpretation)
|
adam@302
|
953 ]]
|
adam@302
|
954 *)
|
adamc@35
|
955
|
adamc@35
|
956 Print and.
|
adam@439
|
957 (** %\vspace{-.15in}%[[
|
adamc@208
|
958 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
|
adam@317
|
959 ]]
|
adam@317
|
960 %\vspace{-.1in}%
|
adam@317
|
961 <<
|
adamc@208
|
962 For conj: Arguments A, B are implicit
|
adam@317
|
963 >>
|
adamc@35
|
964
|
adam@400
|
965 In addition to the definition of [and] itself, we get information on %\index{implicit arguments}%implicit arguments (and some other information that we omit here). The implicit argument information tells us that we build a proof of a conjunction by calling the constructor [conj] on proofs of the conjuncts, with no need to include the types of those proofs as explicit arguments.
|
adamc@35
|
966
|
adamc@35
|
967 %\medskip%
|
adamc@35
|
968
|
adam@448
|
969 Now we create a section for our induction principle, following the same basic plan as in the previous section of this chapter. *)
|
adamc@35
|
970
|
adamc@35
|
971 Section nat_tree_ind'.
|
adamc@35
|
972 Variable P : nat_tree -> Prop.
|
adamc@35
|
973
|
adamc@38
|
974 Hypothesis NLeaf'_case : P NLeaf'.
|
adamc@38
|
975 Hypothesis NNode'_case : forall (n : nat) (ls : list nat_tree),
|
adamc@35
|
976 All P ls -> P (NNode' n ls).
|
adamc@35
|
977
|
adam@420
|
978 (* begin hide *)
|
adam@437
|
979 (* begin thide *)
|
adam@420
|
980 Definition list_nat_tree_ind := O.
|
adam@437
|
981 (* end thide *)
|
adam@420
|
982 (* end hide *)
|
adam@420
|
983
|
adamc@35
|
984 (** A first attempt at writing the induction principle itself follows the intuition that nested inductive type definitions are expanded into mutual inductive definitions.
|
adamc@35
|
985
|
adam@439
|
986 %\vspace{-.15in}%[[
|
adamc@35
|
987 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
|
adamc@208
|
988 match tr with
|
adamc@35
|
989 | NLeaf' => NLeaf'_case
|
adamc@35
|
990 | NNode' n ls => NNode'_case n ls (list_nat_tree_ind ls)
|
adamc@35
|
991 end
|
adamc@35
|
992
|
adamc@35
|
993 with list_nat_tree_ind (ls : list nat_tree) : All P ls :=
|
adamc@208
|
994 match ls with
|
adamc@35
|
995 | Nil => I
|
adamc@35
|
996 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
|
adamc@35
|
997 end.
|
adamc@205
|
998 ]]
|
adamc@205
|
999
|
adam@442
|
1000 Coq rejects this definition, saying
|
adam@317
|
1001 <<
|
adam@317
|
1002 Recursive call to nat_tree_ind' has principal argument equal to "tr"
|
adam@317
|
1003 instead of rest.
|
adam@317
|
1004 >>
|
adam@317
|
1005
|
adam@420
|
1006 There is no deep theoretical reason why this program should be rejected; Coq applies incomplete termination-checking heuristics, and it is necessary to learn a few of the most important rules. The term "nested inductive type" hints at the solution to this particular problem. Just like true mutually inductive types require mutually recursive induction principles, nested types require nested recursion. *)
|
adamc@35
|
1007
|
adamc@35
|
1008 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
|
adamc@208
|
1009 match tr with
|
adamc@35
|
1010 | NLeaf' => NLeaf'_case
|
adamc@35
|
1011 | NNode' n ls => NNode'_case n ls
|
adamc@35
|
1012 ((fix list_nat_tree_ind (ls : list nat_tree) : All P ls :=
|
adamc@208
|
1013 match ls with
|
adamc@35
|
1014 | Nil => I
|
adamc@35
|
1015 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
|
adamc@35
|
1016 end) ls)
|
adamc@35
|
1017 end.
|
adamc@35
|
1018
|
adam@398
|
1019 (** We include an anonymous [fix] version of [list_nat_tree_ind] that is literally _nested_ inside the definition of the recursive function corresponding to the inductive definition that had the nested use of [list]. *)
|
adamc@35
|
1020
|
adamc@35
|
1021 End nat_tree_ind'.
|
adamc@35
|
1022
|
adamc@35
|
1023 (** We can try our induction principle out by defining some recursive functions on [nat_tree]s and proving a theorem about them. First, we define some helper functions that operate on lists. *)
|
adamc@35
|
1024
|
adamc@35
|
1025 Section map.
|
adamc@35
|
1026 Variables T T' : Set.
|
adam@317
|
1027 Variable F : T -> T'.
|
adamc@35
|
1028
|
adamc@35
|
1029 Fixpoint map (ls : list T) : list T' :=
|
adamc@35
|
1030 match ls with
|
adamc@35
|
1031 | Nil => Nil
|
adam@317
|
1032 | Cons h t => Cons (F h) (map t)
|
adamc@35
|
1033 end.
|
adamc@35
|
1034 End map.
|
adamc@35
|
1035
|
adamc@35
|
1036 Fixpoint sum (ls : list nat) : nat :=
|
adamc@35
|
1037 match ls with
|
adamc@35
|
1038 | Nil => O
|
adamc@35
|
1039 | Cons h t => plus h (sum t)
|
adamc@35
|
1040 end.
|
adamc@35
|
1041
|
adamc@35
|
1042 (** Now we can define a size function over our trees. *)
|
adamc@35
|
1043
|
adamc@35
|
1044 Fixpoint ntsize (tr : nat_tree) : nat :=
|
adamc@35
|
1045 match tr with
|
adamc@35
|
1046 | NLeaf' => S O
|
adamc@35
|
1047 | NNode' _ trs => S (sum (map ntsize trs))
|
adamc@35
|
1048 end.
|
adamc@35
|
1049
|
adamc@35
|
1050 (** Notice that Coq was smart enough to expand the definition of [map] to verify that we are using proper nested recursion, even through a use of a higher-order function. *)
|
adamc@35
|
1051
|
adamc@208
|
1052 Fixpoint ntsplice (tr1 tr2 : nat_tree) : nat_tree :=
|
adamc@35
|
1053 match tr1 with
|
adamc@35
|
1054 | NLeaf' => NNode' O (Cons tr2 Nil)
|
adamc@35
|
1055 | NNode' n Nil => NNode' n (Cons tr2 Nil)
|
adamc@35
|
1056 | NNode' n (Cons tr trs) => NNode' n (Cons (ntsplice tr tr2) trs)
|
adamc@35
|
1057 end.
|
adamc@35
|
1058
|
adamc@35
|
1059 (** We have defined another arbitrary notion of tree splicing, similar to before, and we can prove an analogous theorem about its relationship with tree size. We start with a useful lemma about addition. *)
|
adamc@35
|
1060
|
adamc@41
|
1061 (* begin thide *)
|
adamc@35
|
1062 Lemma plus_S : forall n1 n2 : nat,
|
adamc@35
|
1063 plus n1 (S n2) = S (plus n1 n2).
|
adamc@35
|
1064 induction n1; crush.
|
adamc@35
|
1065 Qed.
|
adamc@41
|
1066 (* end thide *)
|
adamc@35
|
1067
|
adamc@35
|
1068 (** Now we begin the proof of the theorem, adding the lemma [plus_S] as a hint. *)
|
adamc@35
|
1069
|
adamc@35
|
1070 Theorem ntsize_ntsplice : forall tr1 tr2 : nat_tree, ntsize (ntsplice tr1 tr2)
|
adamc@35
|
1071 = plus (ntsize tr2) (ntsize tr1).
|
adamc@41
|
1072 (* begin thide *)
|
adam@375
|
1073 Hint Rewrite plus_S.
|
adamc@35
|
1074
|
adam@317
|
1075 (** We know that the standard induction principle is insufficient for the task, so we need to provide a %\index{tactics!using}%[using] clause for the [induction] tactic to specify our alternate principle. *)
|
adamc@208
|
1076
|
adamc@35
|
1077 induction tr1 using nat_tree_ind'; crush.
|
adamc@35
|
1078
|
adamc@35
|
1079 (** One subgoal remains: [[
|
adamc@35
|
1080 n : nat
|
adamc@35
|
1081 ls : list nat_tree
|
adamc@35
|
1082 H : All
|
adamc@35
|
1083 (fun tr1 : nat_tree =>
|
adamc@35
|
1084 forall tr2 : nat_tree,
|
adamc@35
|
1085 ntsize (ntsplice tr1 tr2) = plus (ntsize tr2) (ntsize tr1)) ls
|
adamc@35
|
1086 tr2 : nat_tree
|
adamc@35
|
1087 ============================
|
adamc@35
|
1088 ntsize
|
adamc@35
|
1089 match ls with
|
adamc@35
|
1090 | Nil => NNode' n (Cons tr2 Nil)
|
adamc@35
|
1091 | Cons tr trs => NNode' n (Cons (ntsplice tr tr2) trs)
|
adamc@35
|
1092 end = S (plus (ntsize tr2) (sum (map ntsize ls)))
|
adamc@208
|
1093
|
adamc@35
|
1094 ]]
|
adamc@35
|
1095
|
adamc@35
|
1096 After a few moments of squinting at this goal, it becomes apparent that we need to do a case analysis on the structure of [ls]. The rest is routine. *)
|
adamc@35
|
1097
|
adamc@35
|
1098 destruct ls; crush.
|
adamc@35
|
1099
|
adam@317
|
1100 (** We can go further in automating the proof by exploiting the hint mechanism.%\index{Vernacular commands!Hint Extern}% *)
|
adamc@35
|
1101
|
adamc@35
|
1102 Restart.
|
adam@317
|
1103
|
adamc@35
|
1104 Hint Extern 1 (ntsize (match ?LS with Nil => _ | Cons _ _ => _ end) = _) =>
|
adamc@35
|
1105 destruct LS; crush.
|
adamc@35
|
1106 induction tr1 using nat_tree_ind'; crush.
|
adamc@35
|
1107 Qed.
|
adamc@41
|
1108 (* end thide *)
|
adamc@35
|
1109
|
adamc@35
|
1110 (** We will go into great detail on hints in a later chapter, but the only important thing to note here is that we register a pattern that describes a conclusion we expect to encounter during the proof. The pattern may contain unification variables, whose names are prefixed with question marks, and we may refer to those bound variables in a tactic that we ask to have run whenever the pattern matches.
|
adamc@35
|
1111
|
adam@317
|
1112 The advantage of using the hint is not very clear here, because the original proof was so short. However, the hint has fundamentally improved the readability of our proof. Before, the proof referred to the local variable [ls], which has an automatically generated name. To a human reading the proof script without stepping through it interactively, it was not clear where [ls] came from. The hint explains to the reader the process for choosing which variables to case analyze, and the hint can continue working even if the rest of the proof structure changes significantly. *)
|
adamc@36
|
1113
|
adamc@36
|
1114
|
adamc@36
|
1115 (** * Manual Proofs About Constructors *)
|
adamc@36
|
1116
|
adam@317
|
1117 (** It can be useful to understand how tactics like %\index{tactics!discriminate}%[discriminate] and %\index{tactics!injection}%[injection] work, so it is worth stepping through a manual proof of each kind. We will start with a proof fit for [discriminate]. *)
|
adamc@36
|
1118
|
adamc@36
|
1119 Theorem true_neq_false : true <> false.
|
adamc@208
|
1120
|
adamc@41
|
1121 (* begin thide *)
|
adam@420
|
1122 (** We begin with the tactic %\index{tactics!red}%[red], which is short for "one step of reduction," to unfold the definition of logical negation. *)
|
adamc@36
|
1123
|
adamc@36
|
1124 red.
|
adam@439
|
1125 (** %\vspace{-.15in}%[[
|
adamc@36
|
1126 ============================
|
adamc@36
|
1127 true = false -> False
|
adamc@36
|
1128 ]]
|
adamc@36
|
1129
|
adam@442
|
1130 The negation is replaced with an implication of falsehood. We use the tactic %\index{tactics!intro}%[intro H] to change the assumption of the implication into a hypothesis named [H]. *)
|
adamc@36
|
1131
|
adamc@36
|
1132 intro H.
|
adam@439
|
1133 (** %\vspace{-.15in}%[[
|
adamc@36
|
1134 H : true = false
|
adamc@36
|
1135 ============================
|
adamc@36
|
1136 False
|
adamc@36
|
1137 ]]
|
adamc@36
|
1138
|
adam@442
|
1139 This is the point in the proof where we apply some creativity. We define a function whose utility will become clear soon. *)
|
adamc@36
|
1140
|
adam@317
|
1141 Definition toProp (b : bool) := if b then True else False.
|
adamc@36
|
1142
|
adam@448
|
1143 (** It is worth recalling the difference between the lowercase and uppercase versions of truth and falsehood: [True] and [False] are logical propositions, while [true] and [false] are Boolean values that we can case-analyze. We have defined [toProp] such that our conclusion of [False] is computationally equivalent to [toProp false]. Thus, the %\index{tactics!change}%[change] tactic will let us change the conclusion to [toProp false]. The general form [change e] replaces the conclusion with [e], whenever Coq's built-in computation rules suffice to establish the equivalence of [e] with the original conclusion. *)
|
adamc@36
|
1144
|
adam@317
|
1145 change (toProp false).
|
adam@439
|
1146 (** %\vspace{-.15in}%[[
|
adamc@36
|
1147 H : true = false
|
adamc@36
|
1148 ============================
|
adam@317
|
1149 toProp false
|
adamc@36
|
1150 ]]
|
adamc@36
|
1151
|
adam@448
|
1152 Now the righthand side of [H]'s equality appears in the conclusion, so we can rewrite, using the notation [<-] to request to replace the righthand side of the equality with the lefthand side.%\index{tactics!rewrite}% *)
|
adamc@36
|
1153
|
adamc@36
|
1154 rewrite <- H.
|
adam@439
|
1155 (** %\vspace{-.15in}%[[
|
adamc@36
|
1156 H : true = false
|
adamc@36
|
1157 ============================
|
adam@317
|
1158 toProp true
|
adamc@36
|
1159 ]]
|
adamc@36
|
1160
|
adam@442
|
1161 We are almost done. Just how close we are to done is revealed by computational simplification. *)
|
adamc@36
|
1162
|
adamc@36
|
1163 simpl.
|
adam@439
|
1164 (** %\vspace{-.15in}%[[
|
adamc@36
|
1165 H : true = false
|
adamc@36
|
1166 ============================
|
adamc@36
|
1167 True
|
adam@302
|
1168 ]]
|
adam@302
|
1169 *)
|
adamc@36
|
1170
|
adamc@36
|
1171 trivial.
|
adamc@36
|
1172 Qed.
|
adamc@41
|
1173 (* end thide *)
|
adamc@36
|
1174
|
adamc@36
|
1175 (** I have no trivial automated version of this proof to suggest, beyond using [discriminate] or [congruence] in the first place.
|
adamc@36
|
1176
|
adamc@36
|
1177 %\medskip%
|
adamc@36
|
1178
|
adamc@36
|
1179 We can perform a similar manual proof of injectivity of the constructor [S]. I leave a walk-through of the details to curious readers who want to run the proof script interactively. *)
|
adamc@36
|
1180
|
adamc@36
|
1181 Theorem S_inj' : forall n m : nat, S n = S m -> n = m.
|
adamc@41
|
1182 (* begin thide *)
|
adamc@36
|
1183 intros n m H.
|
adamc@36
|
1184 change (pred (S n) = pred (S m)).
|
adamc@36
|
1185 rewrite H.
|
adamc@36
|
1186 reflexivity.
|
adamc@36
|
1187 Qed.
|
adamc@41
|
1188 (* end thide *)
|
adamc@36
|
1189
|
adam@400
|
1190 (** The key piece of creativity in this theorem comes in the use of the natural number predecessor function [pred]. Embodied in the implementation of [injection] is a generic recipe for writing such type-specific functions.
|
adam@317
|
1191
|
adam@317
|
1192 The examples in this section illustrate an important aspect of the design philosophy behind Coq. We could certainly design a Gallina replacement that built in rules for constructor discrimination and injectivity, but a simpler alternative is to include a few carefully chosen rules that enable the desired reasoning patterns and many others. A key benefit of this philosophy is that the complexity of proof checking is minimized, which bolsters our confidence that proved theorems are really true. *)
|