annotate src/InductiveTypes.v @ 316:2aaff91f5258

Pass through InductiveTypes, through end of reflexive types
author Adam Chlipala <adam@chlipala.net>
date Sun, 11 Sep 2011 17:22:36 -0400
parents 72bffb046797
children 50db9a6e2742
rev   line source
adam@315 1 (* Copyright (c) 2008-2011, 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@315 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.
adam@315 31
adam@315 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.
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@315 41 (** 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@315 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.
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@315 53 (** 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@315 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. *)
adam@315 59
adam@315 60 Check (fun x : False => x).
adam@315 61 (** [: False -> False] *)
adam@315 62
adam@315 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! *)
adam@315 64
adam@315 65 Check (fun x : False => match x with end : True).
adam@315 66 (** [: False -> True] *)
adam@315 67
adam@315 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.
adam@315 69
adam@315 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). *)
adamc@26 71
adamc@26 72
adamc@26 73 (** * Enumerations *)
adamc@26 74
adam@315 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.
adamc@26 76
adam@315 77 The singleton type [unit] is an inductive type:%\index{Gallina terms!unit}\index{Gallina terms!tt}% *)
adamc@26 78
adamc@26 79 Inductive unit : Set :=
adamc@26 80 | tt.
adamc@26 81
adamc@26 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: *)
adamc@26 83
adamc@26 84 Check unit.
adamc@208 85 (** [unit : Set] *)
adamc@26 86
adamc@26 87 Check tt.
adamc@208 88 (** [tt : unit] *)
adamc@26 89
adamc@26 90 (** We can prove that [unit] is a genuine singleton type. *)
adamc@26 91
adamc@26 92 Theorem unit_singleton : forall x : unit, x = tt.
adamc@208 93
adam@315 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}% *)
adamc@208 95
adamc@41 96 (* begin thide *)
adamc@26 97 induction x.
adamc@26 98
adamc@208 99 (** The goal changes to:
adamc@208 100 [[
adamc@26 101 tt = tt
adam@302 102 ]]
adam@302 103 *)
adamc@208 104
adamc@26 105 (** ...which we can discharge trivially. *)
adamc@208 106
adamc@26 107 reflexivity.
adamc@26 108 Qed.
adamc@41 109 (* end thide *)
adamc@26 110
adam@315 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}% [[
adamc@26 112
adamc@26 113 destruct x.
adamc@205 114
adamc@205 115 ]]
adamc@205 116
adam@292 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.
adamc@26 118
adam@315 119 What exactly %\textit{%#<i>#is#</i>#%}% the %\index{induction principles}%induction principle for [unit]? We can ask Coq: *)
adamc@26 120
adamc@26 121 Check unit_ind.
adamc@208 122 (** [unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u] *)
adamc@26 123
adam@315 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.
adamc@26 125
adam@315 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.
adamc@26 127
adam@315 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)].
adam@315 129
adam@315 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.
adamc@26 131
adamc@26 132 %\medskip%
adamc@26 133
adam@315 134 We can define an inductive type even simpler than [unit]:%\index{Gallina terms!Empty\_set}% *)
adamc@26 135
adamc@26 136 Inductive Empty_set : Set := .
adamc@26 137
adamc@26 138 (** [Empty_set] has no elements. We can prove fun theorems about it: *)
adamc@26 139
adamc@26 140 Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5.
adamc@41 141 (* begin thide *)
adamc@26 142 destruct 1.
adamc@26 143 Qed.
adamc@41 144 (* end thide *)
adamc@26 145
adamc@32 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.)
adamc@26 147
adamc@26 148 We can see the induction principle that made this proof so easy: *)
adamc@26 149
adamc@26 150 Check Empty_set_ind.
adam@315 151 (** [Empty_set_ind : forall (][P : Empty_set -> Prop) (e : Empty_set), P e] *)
adamc@26 152
adam@315 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)].
adamc@26 154
adamc@26 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]: *)
adamc@26 156
adamc@26 157 Definition e2u (e : Empty_set) : unit := match e with end.
adamc@26 158
adam@315 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.
adamc@26 160
adamc@26 161 %\medskip%
adamc@26 162
adam@315 163 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 164
adamc@26 165 Inductive bool : Set :=
adamc@26 166 | true
adamc@26 167 | false.
adamc@26 168
adam@316 169 (** We can use less vacuous pattern matching to define boolean negation.%\index{Gallina terms!negb}% *)
adamc@26 170
adam@279 171 Definition negb (b : bool) : bool :=
adamc@26 172 match b with
adamc@26 173 | true => false
adamc@26 174 | false => true
adamc@26 175 end.
adamc@26 176
adamc@27 177 (** An alternative definition desugars to the above: *)
adamc@27 178
adam@279 179 Definition negb' (b : bool) : bool :=
adamc@27 180 if b then false else true.
adamc@27 181
adam@279 182 (** We might want to prove that [negb] is its own inverse operation. *)
adamc@26 183
adam@279 184 Theorem negb_inverse : forall b : bool, negb (negb b) = b.
adamc@41 185 (* begin thide *)
adamc@26 186 destruct b.
adamc@26 187
adamc@208 188 (** After we case-analyze on [b], we are left with one subgoal for each constructor of [bool].
adamc@26 189
adam@315 190 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@315 191
adamc@26 192 [[
adamc@26 193 ============================
adam@279 194 negb (negb true) = true
adamc@26 195 ]]
adam@315 196 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adamc@26 197 [[
adam@279 198 negb (negb false) = false
adamc@208 199
adamc@26 200 ]]
adamc@26 201
adamc@26 202 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
adamc@26 203
adamc@26 204 reflexivity.
adamc@26 205
adam@315 206 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification.%\index{Vernacular commands!Restart}% *)
adamc@26 207
adam@315 208 (* begin hide *)
adamc@26 209 Restart.
adam@315 210 (* end hide *)
adam@315 211 (** %\noindent \coqdockw{Restart}%#<tt>Restart</tt>#. *)
adam@315 212
adamc@26 213 destruct b; reflexivity.
adamc@26 214 Qed.
adamc@41 215 (* end thide *)
adamc@27 216
adam@315 217 (** Another theorem about booleans illustrates another useful tactic.%\index{tactics!discriminate}% *)
adamc@27 218
adam@279 219 Theorem negb_ineq : forall b : bool, negb b <> b.
adamc@41 220 (* begin thide *)
adamc@27 221 destruct b; discriminate.
adamc@27 222 Qed.
adamc@41 223 (* end thide *)
adamc@27 224
adamc@27 225 (** [discriminate] 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 226
adamc@27 227 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *)
adamc@27 228
adamc@27 229 Check bool_ind.
adamc@208 230 (** [bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b] *)
adamc@28 231
adam@315 232 (** That is, to prove that a property describes all [bool]s, prove that it describes both [true] and [false].
adam@315 233
adam@315 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. *)
adam@315 235
adamc@28 236
adamc@28 237 (** * Simple Recursive Types *)
adamc@28 238
adam@315 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}% *)
adamc@28 240
adamc@28 241 Inductive nat : Set :=
adamc@28 242 | O : nat
adamc@28 243 | S : nat -> nat.
adamc@28 244
adam@315 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.
adamc@28 246
adam@316 247 Pattern matching works as we demonstrated in the last chapter:%\index{Gallina terms!pred}% *)
adamc@28 248
adamc@28 249 Definition isZero (n : nat) : bool :=
adamc@28 250 match n with
adamc@28 251 | O => true
adamc@28 252 | S _ => false
adamc@28 253 end.
adamc@28 254
adamc@28 255 Definition pred (n : nat) : nat :=
adamc@28 256 match n with
adamc@28 257 | O => O
adamc@28 258 | S n' => n'
adamc@28 259 end.
adamc@28 260
adamc@28 261 (** We can prove theorems by case analysis: *)
adamc@28 262
adamc@28 263 Theorem S_isZero : forall n : nat, isZero (pred (S (S n))) = false.
adamc@41 264 (* begin thide *)
adamc@28 265 destruct n; reflexivity.
adamc@28 266 Qed.
adamc@41 267 (* end thide *)
adamc@28 268
adam@316 269 (** 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 270
adamc@208 271 Fixpoint plus (n m : nat) : nat :=
adamc@28 272 match n with
adamc@28 273 | O => m
adamc@28 274 | S n' => S (plus n' m)
adamc@28 275 end.
adamc@28 276
adamc@208 277 (** Recall that [Fixpoint] is Coq's mechanism for recursive function definitions. Some theorems about [plus] can be proved without induction. *)
adamc@28 278
adamc@28 279 Theorem O_plus_n : forall n : nat, plus O n = n.
adamc@41 280 (* begin thide *)
adamc@28 281 intro; reflexivity.
adamc@28 282 Qed.
adamc@41 283 (* end thide *)
adamc@28 284
adamc@208 285 (** 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 286
adamc@28 287 Theorem n_plus_O : forall n : nat, plus n O = n.
adamc@41 288 (* begin thide *)
adamc@28 289 induction n.
adamc@28 290
adamc@28 291 (** Our first subgoal is [plus O O = O], which %\textit{%#<i>#is#</i>#%}% trivial by computation. *)
adamc@28 292
adamc@28 293 reflexivity.
adamc@28 294
adamc@28 295 (** Our second subgoal is more work and also demonstrates our first inductive hypothesis.
adamc@28 296
adamc@28 297 [[
adamc@28 298 n : nat
adamc@28 299 IHn : plus n O = n
adamc@28 300 ============================
adamc@28 301 plus (S n) O = S n
adamc@208 302
adamc@28 303 ]]
adamc@28 304
adam@315 305 We can start out by using computation to simplify the goal as far as we can.%\index{tactics!simpl}% *)
adamc@28 306
adamc@28 307 simpl.
adamc@28 308
adam@315 309 (** Now the conclusion is [S (][plus n O) = S n]. Using our inductive hypothesis: *)
adamc@28 310
adamc@28 311 rewrite IHn.
adamc@28 312
adamc@28 313 (** ...we get a trivial conclusion [S n = S n]. *)
adamc@28 314
adamc@28 315 reflexivity.
adamc@28 316
adam@315 317 (** Not much really went on in this proof, so the [crush] tactic from the [CpdtTactics] module can prove this theorem automatically. *)
adamc@28 318
adam@315 319 (* begin hide *)
adamc@28 320 Restart.
adam@315 321 (* end hide *)
adam@315 322 (** %\noindent \coqdockw{Restart}%#<tt>Restart</tt>#. *)
adam@315 323
adamc@28 324 induction n; crush.
adamc@28 325 Qed.
adamc@41 326 (* end thide *)
adamc@28 327
adamc@28 328 (** We can check out the induction principle at work here: *)
adamc@28 329
adamc@28 330 Check nat_ind.
adamc@208 331 (** %\vspace{-.15in}% [[
adamc@208 332 nat_ind : forall P : nat -> Prop,
adamc@208 333 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@208 334
adamc@208 335 ]]
adamc@28 336
adam@315 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.
adamc@28 338
adam@315 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}% *)
adamc@28 340
adamc@28 341 Theorem S_inj : forall n m : nat, S n = S m -> n = m.
adamc@41 342 (* begin thide *)
adamc@28 343 injection 1; trivial.
adamc@28 344 Qed.
adamc@41 345 (* end thide *)
adamc@28 346
adamc@28 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.
adamc@28 348
adam@315 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.
adamc@29 350
adamc@29 351 %\medskip%
adamc@29 352
adamc@29 353 We can define a type of lists of natural numbers. *)
adamc@29 354
adamc@29 355 Inductive nat_list : Set :=
adamc@29 356 | NNil : nat_list
adamc@29 357 | NCons : nat -> nat_list -> nat_list.
adamc@29 358
adamc@29 359 (** Recursive definitions are straightforward extensions of what we have seen before. *)
adamc@29 360
adamc@29 361 Fixpoint nlength (ls : nat_list) : nat :=
adamc@29 362 match ls with
adamc@29 363 | NNil => O
adamc@29 364 | NCons _ ls' => S (nlength ls')
adamc@29 365 end.
adamc@29 366
adamc@208 367 Fixpoint napp (ls1 ls2 : nat_list) : nat_list :=
adamc@29 368 match ls1 with
adamc@29 369 | NNil => ls2
adamc@29 370 | NCons n ls1' => NCons n (napp ls1' ls2)
adamc@29 371 end.
adamc@29 372
adamc@29 373 (** Inductive theorem proving can again be automated quite effectively. *)
adamc@29 374
adamc@29 375 Theorem nlength_napp : forall ls1 ls2 : nat_list, nlength (napp ls1 ls2)
adamc@29 376 = plus (nlength ls1) (nlength ls2).
adamc@41 377 (* begin thide *)
adamc@29 378 induction ls1; crush.
adamc@29 379 Qed.
adamc@41 380 (* end thide *)
adamc@29 381
adamc@29 382 Check nat_list_ind.
adamc@208 383 (** %\vspace{-.15in}% [[
adamc@208 384 nat_list_ind
adamc@29 385 : forall P : nat_list -> Prop,
adamc@29 386 P NNil ->
adamc@29 387 (forall (n : nat) (n0 : nat_list), P n0 -> P (NCons n n0)) ->
adamc@29 388 forall n : nat_list, P n
adamc@29 389 ]]
adamc@29 390
adamc@29 391 %\medskip%
adamc@29 392
adam@292 393 In general, we can implement any %``%#"#tree#"#%''% types as inductive types. For example, here are binary trees of naturals. *)
adamc@29 394
adamc@29 395 Inductive nat_btree : Set :=
adamc@29 396 | NLeaf : nat_btree
adamc@29 397 | NNode : nat_btree -> nat -> nat_btree -> nat_btree.
adamc@29 398
adamc@29 399 Fixpoint nsize (tr : nat_btree) : nat :=
adamc@29 400 match tr with
adamc@35 401 | NLeaf => S O
adamc@29 402 | NNode tr1 _ tr2 => plus (nsize tr1) (nsize tr2)
adamc@29 403 end.
adamc@29 404
adamc@208 405 Fixpoint nsplice (tr1 tr2 : nat_btree) : nat_btree :=
adamc@29 406 match tr1 with
adamc@35 407 | NLeaf => NNode tr2 O NLeaf
adamc@29 408 | NNode tr1' n tr2' => NNode (nsplice tr1' tr2) n tr2'
adamc@29 409 end.
adamc@29 410
adamc@29 411 Theorem plus_assoc : forall n1 n2 n3 : nat, plus (plus n1 n2) n3 = plus n1 (plus n2 n3).
adamc@41 412 (* begin thide *)
adamc@29 413 induction n1; crush.
adamc@29 414 Qed.
adamc@41 415 (* end thide *)
adamc@29 416
adamc@29 417 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2)
adamc@29 418 = plus (nsize tr2) (nsize tr1).
adamc@41 419 (* begin thide *)
adam@315 420 (* begin hide *)
adamc@29 421 Hint Rewrite n_plus_O plus_assoc : cpdt.
adam@315 422 (* end hide *)
adam@315 423 (** [Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [n_plus_O plus_assoc : cpdt.] *)
adamc@29 424
adamc@29 425 induction tr1; crush.
adamc@29 426 Qed.
adamc@41 427 (* end thide *)
adamc@29 428
adam@315 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. *)
adam@315 430
adamc@29 431 Check nat_btree_ind.
adamc@208 432 (** %\vspace{-.15in}% [[
adamc@208 433 nat_btree_ind
adamc@29 434 : forall P : nat_btree -> Prop,
adamc@29 435 P NLeaf ->
adamc@29 436 (forall n : nat_btree,
adamc@29 437 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) ->
adamc@29 438 forall n : nat_btree, P n
adam@302 439 ]]
adam@315 440
adam@315 441 We have the usual two cases, one for each constructor of [nat_btree]. *)
adamc@30 442
adamc@30 443
adamc@30 444 (** * Parameterized Types *)
adamc@30 445
adam@316 446 (** 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 447
adamc@30 448 Inductive list (T : Set) : Set :=
adamc@30 449 | Nil : list T
adamc@30 450 | Cons : T -> list T -> list T.
adamc@30 451
adamc@30 452 Fixpoint length T (ls : list T) : nat :=
adamc@30 453 match ls with
adamc@30 454 | Nil => O
adamc@30 455 | Cons _ ls' => S (length ls')
adamc@30 456 end.
adamc@30 457
adamc@208 458 Fixpoint app T (ls1 ls2 : list T) : list T :=
adamc@30 459 match ls1 with
adamc@30 460 | Nil => ls2
adamc@30 461 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 462 end.
adamc@30 463
adamc@30 464 Theorem length_app : forall T (ls1 ls2 : list T), length (app ls1 ls2)
adamc@30 465 = plus (length ls1) (length ls2).
adamc@41 466 (* begin thide *)
adamc@30 467 induction ls1; crush.
adamc@30 468 Qed.
adamc@41 469 (* end thide *)
adamc@30 470
adam@316 471 (** 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}\textit{%#<i>#section#</i>#%}% mechanism. The following block of code is equivalent to the above: *)
adamc@30 472
adamc@30 473 (* begin hide *)
adamc@30 474 Reset list.
adamc@30 475 (* end hide *)
adamc@30 476
adamc@30 477 Section list.
adamc@30 478 Variable T : Set.
adamc@30 479
adamc@30 480 Inductive list : Set :=
adamc@30 481 | Nil : list
adamc@30 482 | Cons : T -> list -> list.
adamc@30 483
adamc@30 484 Fixpoint length (ls : list) : nat :=
adamc@30 485 match ls with
adamc@30 486 | Nil => O
adamc@30 487 | Cons _ ls' => S (length ls')
adamc@30 488 end.
adamc@30 489
adamc@208 490 Fixpoint app (ls1 ls2 : list) : list :=
adamc@30 491 match ls1 with
adamc@30 492 | Nil => ls2
adamc@30 493 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 494 end.
adamc@30 495
adamc@30 496 Theorem length_app : forall ls1 ls2 : list, length (app ls1 ls2)
adamc@30 497 = plus (length ls1) (length ls2).
adamc@41 498 (* begin thide *)
adamc@30 499 induction ls1; crush.
adamc@30 500 Qed.
adamc@41 501 (* end thide *)
adamc@30 502 End list.
adamc@30 503
adamc@35 504 (* begin hide *)
adamc@35 505 Implicit Arguments Nil [T].
adamc@35 506 (* end hide *)
adamc@35 507
adamc@210 508 (** After we end the section, the [Variable]s we used are added as extra function parameters for each defined identifier, as needed. We verify that this has happened using the [Print] command, a cousin of [Check] which shows the definition of a symbol, rather than just its type. *)
adamc@30 509
adamc@202 510 Print list.
adamc@208 511 (** %\vspace{-.15in}% [[
adamc@208 512 Inductive list (T : Set) : Set :=
adam@316 513 Nil : list T | Cons : T -> list T -> list T
adamc@208 514
adamc@202 515 ]]
adamc@30 516
adamc@202 517 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 518
adamc@30 519 Check length.
adamc@208 520 (** %\vspace{-.15in}% [[
adamc@208 521 length
adamc@30 522 : forall T : Set, list T -> nat
adamc@30 523 ]]
adamc@30 524
adamc@202 525 The parameter [T] is treated as a new argument to the induction principle, too. *)
adamc@30 526
adamc@30 527 Check list_ind.
adamc@208 528 (** %\vspace{-.15in}% [[
adamc@208 529 list_ind
adamc@30 530 : forall (T : Set) (P : list T -> Prop),
adamc@30 531 P (Nil T) ->
adamc@30 532 (forall (t : T) (l : list T), P l -> P (Cons t l)) ->
adamc@30 533 forall l : list T, P l
adamc@30 534 ]]
adamc@30 535
adamc@30 536 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 537
adamc@31 538
adamc@31 539 (** * Mutually Inductive Types *)
adamc@31 540
adamc@31 541 (** We can define inductive types that refer to each other: *)
adamc@31 542
adamc@31 543 Inductive even_list : Set :=
adamc@31 544 | ENil : even_list
adamc@31 545 | ECons : nat -> odd_list -> even_list
adamc@31 546
adamc@31 547 with odd_list : Set :=
adamc@31 548 | OCons : nat -> even_list -> odd_list.
adamc@31 549
adamc@31 550 Fixpoint elength (el : even_list) : nat :=
adamc@31 551 match el with
adamc@31 552 | ENil => O
adamc@31 553 | ECons _ ol => S (olength ol)
adamc@31 554 end
adamc@31 555
adamc@31 556 with olength (ol : odd_list) : nat :=
adamc@31 557 match ol with
adamc@31 558 | OCons _ el => S (elength el)
adamc@31 559 end.
adamc@31 560
adamc@208 561 Fixpoint eapp (el1 el2 : even_list) : even_list :=
adamc@31 562 match el1 with
adamc@31 563 | ENil => el2
adamc@31 564 | ECons n ol => ECons n (oapp ol el2)
adamc@31 565 end
adamc@31 566
adamc@208 567 with oapp (ol : odd_list) (el : even_list) : odd_list :=
adamc@31 568 match ol with
adamc@31 569 | OCons n el' => OCons n (eapp el' el)
adamc@31 570 end.
adamc@31 571
adamc@31 572 (** 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 573
adamc@31 574 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 575 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@41 576 (* begin thide *)
adamc@31 577 induction el1; crush.
adamc@31 578
adamc@31 579 (** One goal remains: [[
adamc@31 580
adamc@31 581 n : nat
adamc@31 582 o : odd_list
adamc@31 583 el2 : even_list
adamc@31 584 ============================
adamc@31 585 S (olength (oapp o el2)) = S (plus (olength o) (elength el2))
adamc@31 586 ]]
adamc@31 587
adamc@31 588 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 589
adamc@31 590 Abort.
adamc@31 591 Check even_list_ind.
adamc@208 592 (** %\vspace{-.15in}% [[
adamc@208 593 even_list_ind
adamc@31 594 : forall P : even_list -> Prop,
adamc@31 595 P ENil ->
adamc@31 596 (forall (n : nat) (o : odd_list), P (ECons n o)) ->
adamc@31 597 forall e : even_list, P e
adamc@208 598
adamc@31 599 ]]
adamc@31 600
adam@316 601 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 602
adamc@31 603 Scheme even_list_mut := Induction for even_list Sort Prop
adamc@31 604 with odd_list_mut := Induction for odd_list Sort Prop.
adamc@31 605
adam@316 606 (** 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 607
adamc@31 608 Check even_list_mut.
adamc@208 609 (** %\vspace{-.15in}% [[
adamc@208 610 even_list_mut
adamc@31 611 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@31 612 P ENil ->
adamc@31 613 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@31 614 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@31 615 forall e : even_list, P e
adamc@208 616
adamc@31 617 ]]
adamc@31 618
adam@316 619 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 620
adamc@31 621 Theorem n_plus_O' : forall n : nat, plus n O = n.
adamc@31 622 apply (nat_ind (fun n => plus n O = n)); crush.
adamc@31 623 Qed.
adamc@31 624
adamc@31 625 (** 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 626
adamc@31 627 This technique generalizes to our mutual example: *)
adamc@31 628
adamc@31 629 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 630 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@41 631
adamc@31 632 apply (even_list_mut
adamc@31 633 (fun el1 : even_list => forall el2 : even_list,
adamc@31 634 elength (eapp el1 el2) = plus (elength el1) (elength el2))
adamc@31 635 (fun ol : odd_list => forall el : even_list,
adamc@31 636 olength (oapp ol el) = plus (olength ol) (elength el))); crush.
adamc@31 637 Qed.
adamc@41 638 (* end thide *)
adamc@31 639
adamc@31 640 (** 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 641
adamc@33 642
adamc@33 643 (** * Reflexive Types *)
adamc@33 644
adam@316 645 (** A kind of inductive type called a %\textit{%#<i>#reflexive type#</i>#%}% is defined in terms of functions that have the type being defined as their range. 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. *)
adam@316 646
adam@316 647 Inductive pformula : Set :=
adam@316 648 | Truth : pformula
adam@316 649 | Falsehood : pformula
adam@316 650 | Conjunction : pformula -> pformula -> pformula.
adam@316 651
adam@316 652 (** A key distinction here is between, for instance, the %\emph{%#<i>#syntax#</i>#%}% [Truth] and its %\emph{%#<i>#semantics#</i>#%}% [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!conj}%[conj] from the standard library. The family [conj] implements conjunction (i.e., %``%#"#and#"#%''%), 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 653
adam@316 654 Fixpoint pformulaDenote (f : pformula) : Prop :=
adam@316 655 match f with
adam@316 656 | Truth => True
adam@316 657 | Falsehood => False
adam@316 658 | Conjunction f1 f2 => pformulaDenote f1 /\ pformulaDenote f2
adam@316 659 end.
adam@316 660
adam@316 661 (** This is a 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 662
adamc@33 663 Inductive formula : Set :=
adamc@33 664 | Eq : nat -> nat -> formula
adamc@33 665 | And : formula -> formula -> formula
adamc@33 666 | Forall : (nat -> formula) -> formula.
adamc@33 667
adam@316 668 (** 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 quantification. For instance, here is the encoding of [forall x : nat, x = x]:%\index{Vernacular commands!Example}% *)
adamc@33 669
adamc@33 670 Example forall_refl : formula := Forall (fun x => Eq x x).
adamc@33 671
adamc@33 672 (** We can write recursive functions over reflexive types quite naturally. Here is one translating our formulas into native Coq propositions. *)
adamc@33 673
adamc@33 674 Fixpoint formulaDenote (f : formula) : Prop :=
adamc@33 675 match f with
adamc@33 676 | Eq n1 n2 => n1 = n2
adamc@33 677 | And f1 f2 => formulaDenote f1 /\ formulaDenote f2
adamc@33 678 | Forall f' => forall n : nat, formulaDenote (f' n)
adamc@33 679 end.
adamc@33 680
adamc@33 681 (** We can also encode a trivial formula transformation that swaps the order of equality and conjunction operands. *)
adamc@33 682
adamc@33 683 Fixpoint swapper (f : formula) : formula :=
adamc@33 684 match f with
adamc@33 685 | Eq n1 n2 => Eq n2 n1
adamc@33 686 | And f1 f2 => And (swapper f2) (swapper f1)
adamc@33 687 | Forall f' => Forall (fun n => swapper (f' n))
adamc@33 688 end.
adamc@33 689
adamc@33 690 (** It is helpful to prove that this transformation does not make true formulas false. *)
adamc@33 691
adamc@33 692 Theorem swapper_preserves_truth : forall f, formulaDenote f -> formulaDenote (swapper f).
adamc@41 693 (* begin thide *)
adamc@33 694 induction f; crush.
adamc@33 695 Qed.
adamc@41 696 (* end thide *)
adamc@33 697
adamc@33 698 (** We can take a look at the induction principle behind this proof. *)
adamc@33 699
adamc@33 700 Check formula_ind.
adamc@208 701 (** %\vspace{-.15in}% [[
adamc@208 702 formula_ind
adamc@33 703 : forall P : formula -> Prop,
adamc@33 704 (forall n n0 : nat, P (Eq n n0)) ->
adamc@33 705 (forall f0 : formula,
adamc@33 706 P f0 -> forall f1 : formula, P f1 -> P (And f0 f1)) ->
adamc@33 707 (forall f1 : nat -> formula,
adamc@33 708 (forall n : nat, P (f1 n)) -> P (Forall f1)) ->
adamc@33 709 forall f2 : formula, P f2
adamc@208 710
adamc@208 711 ]]
adamc@33 712
adamc@208 713 Focusing on the [Forall] case, which comes third, we see that we are allowed to assume that the theorem holds %\textit{%#<i>#for any application of the argument function [f1]#</i>#%}%. 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 714
adamc@33 715 %\medskip%
adamc@33 716
adam@316 717 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.
adamc@33 718
adam@316 719 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}}\textit{%#<i>#higher-order abstract syntax (HOAS)#</i>#%}~\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: *)
adamc@33 720 (** [[
adamc@33 721 Inductive term : Set :=
adamc@33 722 | App : term -> term -> term
adamc@33 723 | Abs : (term -> term) -> term.
adamc@33 724 ]]
adamc@33 725
adam@316 726 <<
adam@316 727 Error: Non strictly positive occurrence of "term" in "(term -> term) -> term"
adam@316 728 >>
adam@316 729
adam@316 730 We have run afoul of the %\index{strict positivity requirement}\index{positivity requirement}\textit{%#<i>#strict positivity requirement#</i>#%}% 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 731
adamc@33 732 Why must Coq enforce this restriction? Imagine that our last definition had been accepted, allowing us to write this function:
adamc@33 733
adamc@33 734 [[
adamc@33 735 Definition uhoh (t : term) : term :=
adamc@33 736 match t with
adamc@33 737 | Abs f => f t
adamc@33 738 | _ => t
adamc@33 739 end.
adamc@33 740
adamc@205 741 ]]
adamc@205 742
adamc@33 743 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 744
adam@316 745 %\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 746
adamc@33 747 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 later chapters on programming language syntax and semantics. *)
adamc@34 748
adamc@34 749
adamc@34 750 (** * An Interlude on Proof Terms *)
adamc@34 751
adamc@34 752 (** 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 induction principles we have used. *)
adamc@34 753
adamc@34 754 Print unit_ind.
adamc@208 755 (** %\vspace{-.15in}% [[
adamc@208 756 unit_ind =
adamc@208 757 fun P : unit -> Prop => unit_rect P
adamc@34 758 : forall P : unit -> Prop, P tt -> forall u : unit, P u
adamc@208 759
adamc@34 760 ]]
adamc@34 761
adamc@34 762 We see that this induction principle is defined in terms of a more general principle, [unit_rect]. *)
adamc@34 763
adamc@34 764 Check unit_rect.
adamc@208 765 (** %\vspace{-.15in}% [[
adamc@208 766 unit_rect
adamc@34 767 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@208 768
adamc@34 769 ]]
adamc@34 770
adamc@34 771 [unit_rect] gives [P] type [unit -> Type] instead of [unit -> Prop]. [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 [unit] automatically: *)
adamc@34 772
adamc@34 773 Print unit_rec.
adamc@208 774 (** %\vspace{-.15in}% [[
adamc@208 775 unit_rec =
adamc@208 776 fun P : unit -> Set => unit_rect P
adamc@34 777 : forall P : unit -> Set, P tt -> forall u : unit, P u
adamc@208 778
adamc@34 779 ]]
adamc@34 780
adamc@34 781 This is identical to the definition for [unit_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 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 782
adamc@34 783 Definition always_O (u : unit) : nat :=
adamc@34 784 match u with
adamc@34 785 | tt => O
adamc@34 786 end.
adamc@34 787
adamc@34 788 Definition always_O' (u : unit) : nat :=
adamc@34 789 unit_rec (fun _ : unit => nat) O u.
adamc@34 790
adamc@34 791 (** Going even further down the rabbit hole, [unit_rect] itself is not even a primitive. It is a functional program that we can write manually. *)
adamc@34 792
adamc@34 793 Print unit_rect.
adamc@208 794 (** %\vspace{-.15in}% [[
adamc@208 795 unit_rect =
adamc@208 796 fun (P : unit -> Type) (f : P tt) (u : unit) =>
adamc@208 797 match u as u0 return (P u0) with
adamc@208 798 | tt => f
adamc@208 799 end
adamc@34 800 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@208 801
adamc@34 802 ]]
adamc@34 803
adamc@34 804 The only new feature we see is an [as] clause for a [match], which is used in concert with the [return] clause that we saw in the introduction. Since the type of the [match] is dependent on the value of the object being analyzed, we must give that object a name so that we can refer to it in the [return] clause.
adamc@34 805
adamc@34 806 To prove that [unit_rect] is nothing special, we can reimplement it manually. *)
adamc@34 807
adamc@34 808 Definition unit_rect' (P : unit -> Type) (f : P tt) (u : unit) :=
adamc@208 809 match u with
adamc@34 810 | tt => f
adamc@34 811 end.
adamc@34 812
adamc@208 813 (** We rely on Coq's heuristics for inferring [match] annotations.
adamc@34 814
adamc@208 815 We can check the implementation of [nat_rect] as well: *)
adamc@34 816
adamc@34 817 Print nat_rect.
adamc@208 818 (** %\vspace{-.15in}% [[
adamc@208 819 nat_rect =
adamc@208 820 fun (P : nat -> Type) (f : P O) (f0 : forall n : nat, P n -> P (S n)) =>
adamc@208 821 fix F (n : nat) : P n :=
adamc@208 822 match n as n0 return (P n0) with
adamc@208 823 | O => f
adamc@208 824 | S n0 => f0 n0 (F n0)
adamc@208 825 end
adamc@208 826 : forall P : nat -> Type,
adamc@208 827 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@208 828 ]]
adamc@34 829
adamc@208 830 Now we have an actual recursive definition. [fix] expressions are an anonymous form of [Fixpoint], just as [fun] expressions stand for anonymous non-recursive functions. Beyond that, the syntax of [fix] mirrors that of [Fixpoint]. We can understand the definition of [nat_rect] better by reimplementing [nat_ind] using sections. *)
adamc@34 831
adamc@208 832 Section nat_ind'.
adamc@208 833 (** First, we have the property of natural numbers that we aim to prove. *)
adamc@34 834
adamc@208 835 Variable P : nat -> Prop.
adamc@34 836
adamc@208 837 (** Then we require a proof of the [O] case. *)
adamc@34 838
adamc@208 839 Hypothesis O_case : P O.
adamc@34 840
adamc@208 841 (** Next is a proof of the [S] case, which may assume an inductive hypothesis. *)
adamc@34 842
adamc@208 843 Hypothesis S_case : forall n : nat, P n -> P (S n).
adamc@34 844
adamc@208 845 (** Finally, we define a recursive function to tie the pieces together. *)
adamc@34 846
adamc@208 847 Fixpoint nat_ind' (n : nat) : P n :=
adamc@208 848 match n with
adamc@208 849 | O => O_case
adamc@208 850 | S n' => S_case (nat_ind' n')
adamc@208 851 end.
adamc@208 852 End nat_ind'.
adamc@34 853
adamc@208 854 (** 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 855
adamc@208 856 %\medskip%
adamc@34 857
adamc@208 858 We can also examine the definition of [even_list_mut], which we generated with [Scheme] for a mutually-recursive type. *)
adamc@34 859
adamc@208 860 Print even_list_mut.
adamc@208 861 (** %\vspace{-.15in}% [[
adamc@208 862 even_list_mut =
adamc@208 863 fun (P : even_list -> Prop) (P0 : odd_list -> Prop)
adamc@208 864 (f : P ENil) (f0 : forall (n : nat) (o : odd_list), P0 o -> P (ECons n o))
adamc@208 865 (f1 : forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) =>
adamc@208 866 fix F (e : even_list) : P e :=
adamc@208 867 match e as e0 return (P e0) with
adamc@208 868 | ENil => f
adamc@208 869 | ECons n o => f0 n o (F0 o)
adamc@208 870 end
adamc@208 871 with F0 (o : odd_list) : P0 o :=
adamc@208 872 match o as o0 return (P0 o0) with
adamc@208 873 | OCons n e => f1 n e (F e)
adamc@208 874 end
adamc@208 875 for F
adamc@208 876 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@208 877 P ENil ->
adamc@208 878 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@208 879 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@208 880 forall e : even_list, P e
adamc@34 881
adamc@208 882 ]]
adamc@34 883
adamc@208 884 We see a mutually-recursive [fix], with the different functions separated by [with] in the same way that they would be separated by [and] in ML. A final [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 885
adamc@208 886 Section even_list_mut'.
adamc@208 887 (** First, we need the properties that we are proving. *)
adamc@208 888
adamc@208 889 Variable Peven : even_list -> Prop.
adamc@208 890 Variable Podd : odd_list -> Prop.
adamc@208 891
adamc@208 892 (** Next, we need proofs of the three cases. *)
adamc@208 893
adamc@208 894 Hypothesis ENil_case : Peven ENil.
adamc@208 895 Hypothesis ECons_case : forall (n : nat) (o : odd_list), Podd o -> Peven (ECons n o).
adamc@208 896 Hypothesis OCons_case : forall (n : nat) (e : even_list), Peven e -> Podd (OCons n e).
adamc@208 897
adamc@208 898 (** Finally, we define the recursive functions. *)
adamc@208 899
adamc@208 900 Fixpoint even_list_mut' (e : even_list) : Peven e :=
adamc@208 901 match e with
adamc@208 902 | ENil => ENil_case
adamc@208 903 | ECons n o => ECons_case n (odd_list_mut' o)
adamc@208 904 end
adamc@208 905 with odd_list_mut' (o : odd_list) : Podd o :=
adamc@208 906 match o with
adamc@208 907 | OCons n e => OCons_case n (even_list_mut' e)
adamc@208 908 end.
adamc@34 909 End even_list_mut'.
adamc@34 910
adamc@34 911 (** 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 912
adamc@34 913 Section formula_ind'.
adamc@34 914 Variable P : formula -> Prop.
adamc@38 915 Hypothesis Eq_case : forall n1 n2 : nat, P (Eq n1 n2).
adamc@38 916 Hypothesis And_case : forall f1 f2 : formula,
adamc@34 917 P f1 -> P f2 -> P (And f1 f2).
adamc@38 918 Hypothesis Forall_case : forall f : nat -> formula,
adamc@34 919 (forall n : nat, P (f n)) -> P (Forall f).
adamc@34 920
adamc@34 921 Fixpoint formula_ind' (f : formula) : P f :=
adamc@208 922 match f with
adamc@34 923 | Eq n1 n2 => Eq_case n1 n2
adamc@34 924 | And f1 f2 => And_case (formula_ind' f1) (formula_ind' f2)
adamc@34 925 | Forall f' => Forall_case f' (fun n => formula_ind' (f' n))
adamc@34 926 end.
adamc@34 927 End formula_ind'.
adamc@34 928
adamc@35 929
adamc@35 930 (** * Nested Inductive Types *)
adamc@35 931
adamc@35 932 (** 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 933
adamc@35 934 Inductive nat_tree : Set :=
adamc@35 935 | NLeaf' : nat_tree
adamc@35 936 | NNode' : nat -> list nat_tree -> nat_tree.
adamc@35 937
adamc@35 938 (** This is an example of a %\textit{%#<i>#nested#</i>#%}% inductive type definition, because we use the type we are defining as an argument to a parametrized 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 939
adamc@35 940 Like we encountered for mutual inductive types, we find that the automatically-generated induction principle for [nat_tree] is too weak. *)
adamc@35 941
adamc@35 942 Check nat_tree_ind.
adamc@208 943 (** %\vspace{-.15in}% [[
adamc@208 944 nat_tree_ind
adamc@35 945 : forall P : nat_tree -> Prop,
adamc@35 946 P NLeaf' ->
adamc@35 947 (forall (n : nat) (l : list nat_tree), P (NNode' n l)) ->
adamc@35 948 forall n : nat_tree, P n
adamc@208 949
adamc@35 950 ]]
adamc@35 951
adamc@35 952 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 to different type families. 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 953
adamc@35 954 First, we will need an auxiliary definition, characterizing what it means for a property to hold of every element of a list. *)
adamc@35 955
adamc@35 956 Section All.
adamc@35 957 Variable T : Set.
adamc@35 958 Variable P : T -> Prop.
adamc@35 959
adamc@35 960 Fixpoint All (ls : list T) : Prop :=
adamc@35 961 match ls with
adamc@35 962 | Nil => True
adamc@35 963 | Cons h t => P h /\ All t
adamc@35 964 end.
adamc@35 965 End All.
adamc@35 966
adamc@35 967 (** It will be useful to look at the definitions of [True] and [/\], since we will want to write manual proofs of them below. *)
adamc@35 968
adamc@35 969 Print True.
adamc@208 970 (** %\vspace{-.15in}% [[
adamc@208 971 Inductive True : Prop := I : True
adamc@208 972
adamc@208 973 ]]
adamc@35 974
adamc@35 975 That is, [True] is a proposition with exactly one proof, [I], which we may always supply trivially.
adamc@35 976
adamc@35 977 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 [Locate] command. *)
adamc@35 978
adamc@35 979 Locate "/\".
adamc@208 980 (** %\vspace{-.15in}% [[
adamc@208 981 Notation Scope
adamc@208 982 "A /\ B" := and A B : type_scope
adamc@208 983 (default interpretation)
adam@302 984 ]]
adam@302 985 *)
adamc@35 986
adamc@35 987 Print and.
adamc@208 988 (** %\vspace{-.15in}% [[
adamc@208 989 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@208 990 For conj: Arguments A, B are implicit
adamc@208 991 For and: Argument scopes are [type_scope type_scope]
adamc@208 992 For conj: Argument scopes are [type_scope type_scope _ _]
adamc@208 993
adamc@35 994 ]]
adamc@35 995
adamc@35 996 In addition to the definition of [and] itself, we get information on implicit arguments and parsing rules for [and] and its constructor [conj]. We will ignore the parsing information for now. 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 997
adamc@35 998 %\medskip%
adamc@35 999
adamc@35 1000 Now we create a section for our induction principle, following the same basic plan as in the last section of this chapter. *)
adamc@35 1001
adamc@35 1002 Section nat_tree_ind'.
adamc@35 1003 Variable P : nat_tree -> Prop.
adamc@35 1004
adamc@38 1005 Hypothesis NLeaf'_case : P NLeaf'.
adamc@38 1006 Hypothesis NNode'_case : forall (n : nat) (ls : list nat_tree),
adamc@35 1007 All P ls -> P (NNode' n ls).
adamc@35 1008
adamc@35 1009 (** 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 1010
adamc@35 1011 [[
adamc@35 1012 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@208 1013 match tr with
adamc@35 1014 | NLeaf' => NLeaf'_case
adamc@35 1015 | NNode' n ls => NNode'_case n ls (list_nat_tree_ind ls)
adamc@35 1016 end
adamc@35 1017
adamc@35 1018 with list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@208 1019 match ls with
adamc@35 1020 | Nil => I
adamc@35 1021 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 1022 end.
adamc@35 1023
adamc@205 1024 ]]
adamc@205 1025
adam@292 1026 Coq rejects this definition, saying %``%#"#Recursive call to nat_tree_ind' has principal argument equal to "tr" instead of rest.#"#%''% 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 1027
adamc@35 1028 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@208 1029 match tr with
adamc@35 1030 | NLeaf' => NLeaf'_case
adamc@35 1031 | NNode' n ls => NNode'_case n ls
adamc@35 1032 ((fix list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@208 1033 match ls with
adamc@35 1034 | Nil => I
adamc@35 1035 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 1036 end) ls)
adamc@35 1037 end.
adamc@35 1038
adam@279 1039 (** We include an anonymous [fix] version of [list_nat_tree_ind] that is literally %\textit{%#<i>#nested#</i>#%}% inside the definition of the recursive function corresponding to the inductive definition that had the nested use of [list]. *)
adamc@35 1040
adamc@35 1041 End nat_tree_ind'.
adamc@35 1042
adamc@35 1043 (** 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 1044
adamc@35 1045 Section map.
adamc@35 1046 Variables T T' : Set.
adamc@35 1047 Variable f : T -> T'.
adamc@35 1048
adamc@35 1049 Fixpoint map (ls : list T) : list T' :=
adamc@35 1050 match ls with
adamc@35 1051 | Nil => Nil
adamc@35 1052 | Cons h t => Cons (f h) (map t)
adamc@35 1053 end.
adamc@35 1054 End map.
adamc@35 1055
adamc@35 1056 Fixpoint sum (ls : list nat) : nat :=
adamc@35 1057 match ls with
adamc@35 1058 | Nil => O
adamc@35 1059 | Cons h t => plus h (sum t)
adamc@35 1060 end.
adamc@35 1061
adamc@35 1062 (** Now we can define a size function over our trees. *)
adamc@35 1063
adamc@35 1064 Fixpoint ntsize (tr : nat_tree) : nat :=
adamc@35 1065 match tr with
adamc@35 1066 | NLeaf' => S O
adamc@35 1067 | NNode' _ trs => S (sum (map ntsize trs))
adamc@35 1068 end.
adamc@35 1069
adamc@35 1070 (** 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 1071
adamc@208 1072 Fixpoint ntsplice (tr1 tr2 : nat_tree) : nat_tree :=
adamc@35 1073 match tr1 with
adamc@35 1074 | NLeaf' => NNode' O (Cons tr2 Nil)
adamc@35 1075 | NNode' n Nil => NNode' n (Cons tr2 Nil)
adamc@35 1076 | NNode' n (Cons tr trs) => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1077 end.
adamc@35 1078
adamc@35 1079 (** 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 1080
adamc@41 1081 (* begin thide *)
adamc@35 1082 Lemma plus_S : forall n1 n2 : nat,
adamc@35 1083 plus n1 (S n2) = S (plus n1 n2).
adamc@35 1084 induction n1; crush.
adamc@35 1085 Qed.
adamc@41 1086 (* end thide *)
adamc@35 1087
adamc@35 1088 (** Now we begin the proof of the theorem, adding the lemma [plus_S] as a hint. *)
adamc@35 1089
adamc@35 1090 Theorem ntsize_ntsplice : forall tr1 tr2 : nat_tree, ntsize (ntsplice tr1 tr2)
adamc@35 1091 = plus (ntsize tr2) (ntsize tr1).
adamc@41 1092 (* begin thide *)
adamc@35 1093 Hint Rewrite plus_S : cpdt.
adamc@35 1094
adamc@35 1095 (** We know that the standard induction principle is insufficient for the task, so we need to provide a [using] clause for the [induction] tactic to specify our alternate principle. *)
adamc@208 1096
adamc@35 1097 induction tr1 using nat_tree_ind'; crush.
adamc@35 1098
adamc@35 1099 (** One subgoal remains: [[
adamc@35 1100 n : nat
adamc@35 1101 ls : list nat_tree
adamc@35 1102 H : All
adamc@35 1103 (fun tr1 : nat_tree =>
adamc@35 1104 forall tr2 : nat_tree,
adamc@35 1105 ntsize (ntsplice tr1 tr2) = plus (ntsize tr2) (ntsize tr1)) ls
adamc@35 1106 tr2 : nat_tree
adamc@35 1107 ============================
adamc@35 1108 ntsize
adamc@35 1109 match ls with
adamc@35 1110 | Nil => NNode' n (Cons tr2 Nil)
adamc@35 1111 | Cons tr trs => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1112 end = S (plus (ntsize tr2) (sum (map ntsize ls)))
adamc@208 1113
adamc@35 1114 ]]
adamc@35 1115
adamc@35 1116 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 1117
adamc@35 1118 destruct ls; crush.
adamc@35 1119
adamc@36 1120 (** We can go further in automating the proof by exploiting the hint mechanism. *)
adamc@35 1121
adamc@35 1122 Restart.
adamc@35 1123 Hint Extern 1 (ntsize (match ?LS with Nil => _ | Cons _ _ => _ end) = _) =>
adamc@35 1124 destruct LS; crush.
adamc@35 1125 induction tr1 using nat_tree_ind'; crush.
adamc@35 1126 Qed.
adamc@41 1127 (* end thide *)
adamc@35 1128
adamc@35 1129 (** 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 1130
adamc@40 1131 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 on, and the hint can continue working even if the rest of the proof structure changes significantly. *)
adamc@36 1132
adamc@36 1133
adamc@36 1134 (** * Manual Proofs About Constructors *)
adamc@36 1135
adamc@36 1136 (** It can be useful to understand how tactics like [discriminate] and [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 1137
adamc@36 1138 Theorem true_neq_false : true <> false.
adamc@208 1139
adamc@41 1140 (* begin thide *)
adam@292 1141 (** We begin with the tactic [red], which is short for %``%#"#one step of reduction,#"#%''% to unfold the definition of logical negation. *)
adamc@36 1142
adamc@36 1143 red.
adamc@36 1144 (** [[
adamc@36 1145 ============================
adamc@36 1146 true = false -> False
adamc@208 1147
adamc@36 1148 ]]
adamc@36 1149
adamc@36 1150 The negation is replaced with an implication of falsehood. We use the tactic [intro H] to change the assumption of the implication into a hypothesis named [H]. *)
adamc@36 1151
adamc@36 1152 intro H.
adamc@36 1153 (** [[
adamc@36 1154 H : true = false
adamc@36 1155 ============================
adamc@36 1156 False
adamc@208 1157
adamc@36 1158 ]]
adamc@36 1159
adamc@36 1160 This is the point in the proof where we apply some creativity. We define a function whose utility will become clear soon. *)
adamc@36 1161
adamc@36 1162 Definition f (b : bool) := if b then True else False.
adamc@36 1163
adamc@36 1164 (** 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 [f] such that our conclusion of [False] is computationally equivalent to [f false]. Thus, the [change] tactic will let us change the conclusion to [f false]. *)
adamc@36 1165
adamc@36 1166 change (f false).
adamc@36 1167 (** [[
adamc@36 1168 H : true = false
adamc@36 1169 ============================
adamc@36 1170 f false
adamc@208 1171
adamc@36 1172 ]]
adamc@36 1173
adamc@202 1174 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 the equality with the lefthand side. *)
adamc@36 1175
adamc@36 1176 rewrite <- H.
adamc@36 1177 (** [[
adamc@36 1178 H : true = false
adamc@36 1179 ============================
adamc@36 1180 f true
adamc@208 1181
adamc@36 1182 ]]
adamc@36 1183
adamc@36 1184 We are almost done. Just how close we are to done is revealed by computational simplification. *)
adamc@36 1185
adamc@36 1186 simpl.
adamc@36 1187 (** [[
adamc@36 1188 H : true = false
adamc@36 1189 ============================
adamc@36 1190 True
adamc@208 1191
adam@302 1192 ]]
adam@302 1193 *)
adamc@36 1194
adamc@36 1195 trivial.
adamc@36 1196 Qed.
adamc@41 1197 (* end thide *)
adamc@36 1198
adamc@36 1199 (** I have no trivial automated version of this proof to suggest, beyond using [discriminate] or [congruence] in the first place.
adamc@36 1200
adamc@36 1201 %\medskip%
adamc@36 1202
adamc@36 1203 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 1204
adamc@36 1205 Theorem S_inj' : forall n m : nat, S n = S m -> n = m.
adamc@41 1206 (* begin thide *)
adamc@36 1207 intros n m H.
adamc@36 1208 change (pred (S n) = pred (S m)).
adamc@36 1209 rewrite H.
adamc@36 1210 reflexivity.
adamc@36 1211 Qed.
adamc@41 1212 (* end thide *)
adamc@36 1213
adamc@37 1214
adamc@37 1215 (** * Exercises *)
adamc@37 1216
adamc@37 1217 (** %\begin{enumerate}%#<ol>#
adamc@37 1218
adam@292 1219 %\item%#<li># Define an inductive type [truth] with three constructors, [Yes], [No], and [Maybe]. [Yes] stands for certain truth, [No] for certain falsehood, and [Maybe] for an unknown situation. Define %``%#"#not,#"#%''% %``%#"#and,#"#%''% and %``%#"#or#"#%''% for this replacement boolean algebra. Prove that your implementation of %``%#"#and#"#%''% is commutative and distributes over your implementation of %``%#"#or.#"#%''%#</li>#
adamc@37 1220
adamc@39 1221 %\item%#<li># Modify the first example language of Chapter 2 to include variables, where variables are represented with [nat]. Extend the syntax and semantics of expressions to accommodate the change. Your new [expDenote] function should take as a new extra first argument a value of type [var -> nat], where [var] is a synonym for naturals-as-variables, and the function assigns a value to each variable. Define a constant folding function which does a bottom-up pass over an expression, at each stage replacing every binary operation on constants with an equivalent constant. Prove that constant folding preserves the meanings of expressions.#</li>#
adamc@38 1222
adam@292 1223 %\item%#<li># Reimplement the second example language of Chapter 2 to use mutually-inductive types instead of dependent types. That is, define two separate (non-dependent) inductive types [nat_exp] and [bool_exp] for expressions of the two different types, rather than a single indexed type. To keep things simple, you may consider only the binary operators that take naturals as operands. Add natural number variables to the language, as in the last exercise, and add an %``%#"#if#"#%''% expression form taking as arguments one boolean expression and two natural number expressions. Define semantics and constant-folding functions for this new language. Your constant folding should simplify not just binary operations (returning naturals or booleans) with known arguments, but also %``%#"#if#"#%''% expressions with known values for their test expressions but possibly undetermined %``%#"#then#"#%''% and %``%#"#else#"#%''% cases. Prove that constant-folding a natural number expression preserves its meaning.#</li>#
adamc@38 1224
adamc@38 1225 %\item%#<li># Using a reflexive inductive definition, define a type [nat_tree] of infinitary trees, with natural numbers at their leaves and a countable infinity of new trees branching out of each internal node. Define a function [increment] that increments the number in every leaf of a [nat_tree]. Define a function [leapfrog] over a natural [i] and a tree [nt]. [leapfrog] should recurse into the [i]th child of [nt], the [i+1]st child of that node, the [i+2]nd child of the next node, and so on, until reaching a leaf, in which case [leapfrog] should return the number at that leaf. Prove that the result of any call to [leapfrog] is incremented by one by calling [increment] on the tree.#</li>#
adamc@38 1226
adamc@38 1227 %\item%#<li># Define a type of trees of trees of trees of (repeat to infinity). That is, define an inductive type [trexp], whose members are either base cases containing natural numbers or binary trees of [trexp]s. Base your definition on a parameterized binary tree type [btree] that you will also define, so that [trexp] is defined as a nested inductive type. Define a function [total] that sums all of the naturals at the leaves of a [trexp]. Define a function [increment] that increments every leaf of a [trexp] by one. Prove that, for all [tr], [total (increment tr) >= total tr]. On the way to finishing this proof, you will probably want to prove a lemma and add it as a hint using the syntax [Hint Resolve name_of_lemma.].#</li>#
adamc@38 1228
adamc@38 1229 %\item%#<li># Prove discrimination and injectivity theorems for the [nat_btree] type defined earlier in this chapter. In particular, without using the tactics [discriminate], [injection], or [congruence], prove that no leaf equals any node, and prove that two equal nodes carry the same natural number.#</li>#
adamc@37 1230
adamc@37 1231 #</ol>#%\end{enumerate}% *)