comparison src/InductiveTypes.v @ 315:72bffb046797

Pass through InductiveTypes, through end of recursive types
author Adam Chlipala <adam@chlipala.net>
date Sun, 11 Sep 2011 16:26:48 -0400
parents d5787b70cf48
children 2aaff91f5258
comparison
equal deleted inserted replaced
314:d5787b70cf48 315:72bffb046797
1 (* Copyright (c) 2008-2010, Adam Chlipala 1 (* Copyright (c) 2008-2011, Adam Chlipala
2 * 2 *
3 * This work is licensed under a 3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License. 5 * Unported License.
6 * The license text is available at: 6 * The license text is available at:
18 18
19 (** %\part{Basic Programming and Proving} 19 (** %\part{Basic Programming and Proving}
20 20
21 \chapter{Introducing Inductive Types}% *) 21 \chapter{Introducing Inductive Types}% *)
22 22
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. *) 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.
24
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. *)
26
27
28 (** * Proof Terms *)
29
30 (** Mainstream presentations of mathematics treat proofs as objects that exist outside of the universe of mathematical objects. However, for a variety of reasoning, 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 typoes 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.
31
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}\emph{%#<i>#Curry-Howard correspondence#</i>#%}~\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.
33
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.
35
36 A short demonstration should explain how this can be. The identity function over the natural numbers is certainly not a controversial program. *)
37
38 Check (fun x : nat => x).
39 (** [: nat -> nat] *)
40
41 (** Consider this alternate program, which is almost identical to the last one. *)
42
43 Check (fun x : True => x).
44 (** [: True -> True] *)
45
46 (** 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.
47
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. *)
49
50 Check I.
51 (** [: True] *)
52
53 (** With [I], we can prove another simple propositional theorem. *)
54
55 Check (fun _ : False => I).
56 (** [: False -> True] *)
57
58 (** 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. *)
59
60 Check (fun x : False => x).
61 (** [: False -> False] *)
62
63 (** In fact, [False] implies anything, and we can take advantage of this fact with an odd looking [match] expression that has no branches. Since there are no rules for deducing [False], there are no cases to consider! *)
64
65 Check (fun x : False => match x with end : True).
66 (** [: False -> True] *)
67
68 (** Every one of these example programs whose type looks like a logical formula is a %\index{proof term}\emph{%#<i>#proof term#</i>#%}%. We use that name for any Gallina term of a logical type, and we will elaborate shortly on what makes a type logical.
69
70 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 %\index{proposition}%propositions (i.e., formulas or theorem statements). *)
24 71
25 72
26 (** * Enumerations *) 73 (** * Enumerations *)
27 74
28 (** Coq inductive types generalize the algebraic datatypes found in Haskell and ML. Confusingly enough, inductive types also generalize 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. 75 (** 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.
29 76
30 The singleton type [unit] is an inductive type: *) 77 The singleton type [unit] is an inductive type:%\index{Gallina terms!unit}\index{Gallina terms!tt}% *)
31 78
32 Inductive unit : Set := 79 Inductive unit : Set :=
33 | tt. 80 | tt.
34 81
35 (** 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: *) 82 (** 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: *)
42 89
43 (** We can prove that [unit] is a genuine singleton type. *) 90 (** We can prove that [unit] is a genuine singleton type. *)
44 91
45 Theorem unit_singleton : forall x : unit, x = tt. 92 Theorem unit_singleton : forall x : unit, x = tt.
46 93
47 (** 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]. *) 94 (** 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}% *)
48 95
49 (* begin thide *) 96 (* begin thide *)
50 induction x. 97 induction x.
51 98
52 (** The goal changes to: 99 (** The goal changes to:
59 106
60 reflexivity. 107 reflexivity.
61 Qed. 108 Qed.
62 (* end thide *) 109 (* end thide *)
63 110
64 (** 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: [[ 111 (** 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}% [[
65 112
66 destruct x. 113 destruct x.
67 114
68 ]] 115 ]]
69 116
70 %\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. 117 %\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.
71 118
72 What exactly %\textit{%#<i>#is#</i>#%}% the induction principle for [unit]? We can ask Coq: *) 119 What exactly %\textit{%#<i>#is#</i>#%}% the %\index{induction principles}%induction principle for [unit]? We can ask Coq: *)
73 120
74 Check unit_ind. 121 Check unit_ind.
75 (** [unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u] *) 122 (** [unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u] *)
76 123
77 (** Every [Inductive] command defining a type [T] also defines an induction principle named [T_ind]. Coq follows the Curry-Howard correspondence and includes the ingredients of programming and proving in the same single syntactic class. Thus, 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 [Prop], which appears in our induction principle; and the type [Set], which we have seen a few times already. 124 (** 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.
78 125
79 The convention goes like this: [Set] is the type of normal types, 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. 126 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.
80 127
81 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)]. 128 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)].
129
130 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.
82 131
83 %\medskip% 132 %\medskip%
84 133
85 We can define an inductive type even simpler than [unit]: *) 134 We can define an inductive type even simpler than [unit]:%\index{Gallina terms!Empty\_set}% *)
86 135
87 Inductive Empty_set : Set := . 136 Inductive Empty_set : Set := .
88 137
89 (** [Empty_set] has no elements. We can prove fun theorems about it: *) 138 (** [Empty_set] has no elements. We can prove fun theorems about it: *)
90 139
97 (** 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. 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.) 146 (** 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. 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.)
98 147
99 We can see the induction principle that made this proof so easy: *) 148 We can see the induction principle that made this proof so easy: *)
100 149
101 Check Empty_set_ind. 150 Check Empty_set_ind.
102 (** [Empty_set_ind : forall (P : Empty_set -> Prop) (e : Empty_set), P e] *) 151 (** [Empty_set_ind : forall (][P : Empty_set -> Prop) (e : Empty_set), P e] *)
103 152
104 (** 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)]. 153 (** 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)].
105 154
106 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]: *) 155 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]: *)
107 156
108 Definition e2u (e : Empty_set) : unit := match e with end. 157 Definition e2u (e : Empty_set) : unit := match e with end.
109 158
110 (** 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. 159 (** 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.
111 160
112 %\medskip% 161 %\medskip%
113 162
114 Moving up the ladder of complexity, we can define the booleans: *) 163 Moving up the ladder of complexity, we can define the booleans:%\index{Gallina terms!bool}\index{Gallina terms!true}\index{Gallina terms!false}% *)
115 164
116 Inductive bool : Set := 165 Inductive bool : Set :=
117 | true 166 | true
118 | false. 167 | false.
119 168
136 (* begin thide *) 185 (* begin thide *)
137 destruct b. 186 destruct b.
138 187
139 (** After we case-analyze on [b], we are left with one subgoal for each constructor of [bool]. 188 (** After we case-analyze on [b], we are left with one subgoal for each constructor of [bool].
140 189
190 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
191
141 [[ 192 [[
142 2 subgoals
143
144 ============================ 193 ============================
145 negb (negb true) = true 194 negb (negb true) = true
146 ]] 195 ]]
147 196 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
148 [[ 197 [[
149 subgoal 2 is:
150 negb (negb false) = false 198 negb (negb false) = false
151 199
152 ]] 200 ]]
153 201
154 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *) 202 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
155 203
156 reflexivity. 204 reflexivity.
157 205
158 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification. *) 206 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification.%\index{Vernacular commands!Restart}% *)
159 207
208 (* begin hide *)
160 Restart. 209 Restart.
210 (* end hide *)
211 (** %\noindent \coqdockw{Restart}%#<tt>Restart</tt>#. *)
212
161 destruct b; reflexivity. 213 destruct b; reflexivity.
162 Qed. 214 Qed.
163 (* end thide *) 215 (* end thide *)
164 216
165 (** Another theorem about booleans illustrates another useful tactic. *) 217 (** Another theorem about booleans illustrates another useful tactic.%\index{tactics!discriminate}% *)
166 218
167 Theorem negb_ineq : forall b : bool, negb b <> b. 219 Theorem negb_ineq : forall b : bool, negb b <> b.
168 (* begin thide *) 220 (* begin thide *)
169 destruct b; discriminate. 221 destruct b; discriminate.
170 Qed. 222 Qed.
175 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *) 227 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *)
176 228
177 Check bool_ind. 229 Check bool_ind.
178 (** [bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b] *) 230 (** [bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b] *)
179 231
232 (** That is, to prove that a property describes all [bool]s, prove that it describes both [true] and [false].
233
234 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 it 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. *)
235
180 236
181 (** * Simple Recursive Types *) 237 (** * Simple Recursive Types *)
182 238
183 (** The natural numbers are the simplest common example of an inductive type that actually deserves the name. *) 239 (** 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}% *)
184 240
185 Inductive nat : Set := 241 Inductive nat : Set :=
186 | O : nat 242 | O : nat
187 | S : nat -> nat. 243 | S : nat -> nat.
188 244
189 (** [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. 245 (** [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.
190 246
191 Pattern matching works as we demonstrated in the last chapter: *) 247 Pattern matching works as we demonstrated in the last chapter: *)
192 248
193 Definition isZero (n : nat) : bool := 249 Definition isZero (n : nat) : bool :=
194 match n with 250 match n with
244 ============================ 300 ============================
245 plus (S n) O = S n 301 plus (S n) O = S n
246 302
247 ]] 303 ]]
248 304
249 We can start out by using computation to simplify the goal as far as we can. *) 305 We can start out by using computation to simplify the goal as far as we can.%\index{tactics!simpl}% *)
250 306
251 simpl. 307 simpl.
252 308
253 (** Now the conclusion is [S (plus n O) = S n]. Using our inductive hypothesis: *) 309 (** Now the conclusion is [S (][plus n O) = S n]. Using our inductive hypothesis: *)
254 310
255 rewrite IHn. 311 rewrite IHn.
256 312
257 (** ...we get a trivial conclusion [S n = S n]. *) 313 (** ...we get a trivial conclusion [S n = S n]. *)
258 314
259 reflexivity. 315 reflexivity.
260 316
261 (** Not much really went on in this proof, so the [crush] tactic from the [Tactics] module can prove this theorem automatically. *) 317 (** Not much really went on in this proof, so the [crush] tactic from the [CpdtTactics] module can prove this theorem automatically. *)
262 318
319 (* begin hide *)
263 Restart. 320 Restart.
321 (* end hide *)
322 (** %\noindent \coqdockw{Restart}%#<tt>Restart</tt>#. *)
323
264 induction n; crush. 324 induction n; crush.
265 Qed. 325 Qed.
266 (* end thide *) 326 (* end thide *)
267 327
268 (** We can check out the induction principle at work here: *) 328 (** We can check out the induction principle at work here: *)
272 nat_ind : forall P : nat -> Prop, 332 nat_ind : forall P : nat -> Prop,
273 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n 333 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
274 334
275 ]] 335 ]]
276 336
277 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. 337 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.
278 338
279 Since [nat] has a constructor that takes an argument, we may sometimes need to know that that constructor is injective. *) 339 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}% *)
280 340
281 Theorem S_inj : forall n m : nat, S n = S m -> n = m. 341 Theorem S_inj : forall n m : nat, S n = S m -> n = m.
282 (* begin thide *) 342 (* begin thide *)
283 injection 1; trivial. 343 injection 1; trivial.
284 Qed. 344 Qed.
285 (* end thide *) 345 (* end thide *)
286 346
287 (** [injection] 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. 347 (** [injection] 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.
288 348
289 There is also a very useful tactic called [congruence] that can prove this theorem immediately. [congruence] 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 %\textit{%#<i>#complete decision procedure for the theory of equality and uninterpreted functions#</i>#%}%, plus some smarts about inductive types. 349 There is also a very useful tactic called %\index{tactics!congruence}%[congruence] that can prove this theorem immediately. [congruence] 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}\textit{%#<i>#complete decision procedure for the theory of equality and uninterpreted functions#</i>#%}%, plus some smarts about inductive types.
290 350
291 %\medskip% 351 %\medskip%
292 352
293 We can define a type of lists of natural numbers. *) 353 We can define a type of lists of natural numbers. *)
294 354
355 (* end thide *) 415 (* end thide *)
356 416
357 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2) 417 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2)
358 = plus (nsize tr2) (nsize tr1). 418 = plus (nsize tr2) (nsize tr1).
359 (* begin thide *) 419 (* begin thide *)
420 (* begin hide *)
360 Hint Rewrite n_plus_O plus_assoc : cpdt. 421 Hint Rewrite n_plus_O plus_assoc : cpdt.
422 (* end hide *)
423 (** [Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [n_plus_O plus_assoc : cpdt.] *)
361 424
362 induction tr1; crush. 425 induction tr1; crush.
363 Qed. 426 Qed.
364 (* end thide *) 427 (* end thide *)
428
429 (** It is convenient that these proofs go through so easily, but it is useful to check that the tree induction principle works as usual. *)
365 430
366 Check nat_btree_ind. 431 Check nat_btree_ind.
367 (** %\vspace{-.15in}% [[ 432 (** %\vspace{-.15in}% [[
368 nat_btree_ind 433 nat_btree_ind
369 : forall P : nat_btree -> Prop, 434 : forall P : nat_btree -> Prop,
370 P NLeaf -> 435 P NLeaf ->
371 (forall n : nat_btree, 436 (forall n : nat_btree,
372 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) -> 437 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) ->
373 forall n : nat_btree, P n 438 forall n : nat_btree, P n
374 ]] 439 ]]
375 *) 440
441 We have the usual two cases, one for each constructor of [nat_btree]. *)
376 442
377 443
378 (** * Parameterized Types *) 444 (** * Parameterized Types *)
379 445
380 (** We can also define polymorphic inductive types, as with algebraic datatypes in Haskell and ML. *) 446 (** We can also define polymorphic inductive types, as with algebraic datatypes in Haskell and ML. *)