DataStruct.v 37.9 KB
Newer Older
Adam Chlipala's avatar
Adam Chlipala committed
1
(* Copyright (c) 2008-2011, Adam Chlipala
Adam Chlipala's avatar
Adam Chlipala committed
2 3 4 5 6 7 8 9 10
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)

(* begin hide *)
11
Require Import Arith List.
Adam Chlipala's avatar
Adam Chlipala committed
12

13
Require Import CpdtTactics.
Adam Chlipala's avatar
Adam Chlipala committed
14 15 16 17 18 19 20

Set Implicit Arguments.
(* end hide *)


(** %\chapter{Dependent Data Structures}% *)

Adam Chlipala's avatar
Adam Chlipala committed
21
(** Our red-black tree example from the last chapter illustrated how dependent types enable static enforcement of data structure invariants.  To find interesting uses of dependent data structures, however, we need not look to the favorite examples of data structures and algorithms textbooks.  More basic examples like length-indexed and heterogeneous lists come up again and again as the building blocks of dependent programs.  There is a surprisingly large design space for this class of data structure, and we will spend this chapter exploring it. *)
Adam Chlipala's avatar
Adam Chlipala committed
22 23


Adam Chlipala's avatar
Adam Chlipala committed
24 25
(** * More Length-Indexed Lists *)

Adam Chlipala's avatar
Adam Chlipala committed
26
(** We begin with a deeper look at the length-indexed lists that began the last chapter.%\index{Gallina terms!ilist}% *)
Adam Chlipala's avatar
Adam Chlipala committed
27 28 29 30 31 32 33 34

Section ilist.
  Variable A : Set.

  Inductive ilist : nat -> Set :=
  | Nil : ilist O
  | Cons : forall n, A -> ilist n -> ilist (S n).

Adam Chlipala's avatar
Adam Chlipala committed
35
  (** We might like to have a certified function for selecting an element of an [ilist] by position.  We could do this using subset types and explicit manipulation of proofs, but dependent types let us do it more directly.  It is helpful to define a type family %\index{Gallina terms!fin}%[fin], where [fin n] is isomorphic to [{m : nat | m < n}].  The type family name stands for %``%#"#finite.#"#%''% *)
Adam Chlipala's avatar
Adam Chlipala committed
36

Adam Chlipala's avatar
Adam Chlipala committed
37 38 39
  (* EX: Define a function [get] for extracting an [ilist] element by position. *)

(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
40 41 42
  Inductive fin : nat -> Set :=
  | First : forall n, fin (S n)
  | Next : forall n, fin n -> fin (S n).
Adam Chlipala's avatar
Adam Chlipala committed
43

Adam Chlipala's avatar
Adam Chlipala committed
44
  (** An instance of [fin] is essentially a more richly typed copy of the natural numbers.  Every element is a [First] iterated through applying [Next] a number of times that indicates which number is being selected.  For instance, the three values of type [fin 3] are [First 2], [Next (][First 1)], and [Next (][Next (][First 0))].
Adam Chlipala's avatar
Adam Chlipala committed
45 46 47 48

     Now it is easy to pick a [Prop]-free type for a selection function.  As usual, our first implementation attempt will not convince the type checker, and we will attack the deficiencies one at a time.

     [[
Adam Chlipala's avatar
Adam Chlipala committed
49 50
  Fixpoint get n (ls : ilist n) : fin n -> A :=
    match ls with
Adam Chlipala's avatar
Adam Chlipala committed
51 52 53 54 55 56 57 58
      | Nil => fun idx => ?
      | Cons _ x ls' => fun idx =>
        match idx with
          | First _ => x
          | Next _ idx' => get ls' idx'
        end
    end.

59 60
    ]]

Adam Chlipala's avatar
Adam Chlipala committed
61
    We apply the usual wisdom of delaying arguments in [Fixpoint]s so that they may be included in [return] clauses.  This still leaves us with a quandary in each of the [match] cases.  First, we need to figure out how to take advantage of the contradiction in the [Nil] case.  Every [fin] has a type of the form [S n], which cannot unify with the [O] value that we learn for [n] in the [Nil] case.  The solution we adopt is another case of [match]-within-[return].
Adam Chlipala's avatar
Adam Chlipala committed
62 63

    [[
Adam Chlipala's avatar
Adam Chlipala committed
64 65
  Fixpoint get n (ls : ilist n) : fin n -> A :=
    match ls with
Adam Chlipala's avatar
Adam Chlipala committed
66
      | Nil => fun idx =>
Adam Chlipala's avatar
Adam Chlipala committed
67
        match idx in fin n' return (match n' with
Adam Chlipala's avatar
Adam Chlipala committed
68 69 70 71 72 73 74 75 76 77 78 79 80
                                        | O => A
                                        | S _ => unit
                                      end) with
          | First _ => tt
          | Next _ _ => tt
        end
      | Cons _ x ls' => fun idx =>
        match idx with
          | First _ => x
          | Next _ idx' => get ls' idx'
        end
    end.

81 82
    ]]

83 84 85
    Now the first [match] case type-checks, and we see that the problem with the [Cons] case is that the pattern-bound variable [idx'] does not have an apparent type compatible with [ls'].  In fact, the error message Coq gives for this exact code can be confusing, thanks to an overenthusiastic type inference heuristic.  We are told that the [Nil] case body has type [match X with | 0 => A | S _ => unit end] for a unification variable [X], while it is expected to have type [A].  We can see that setting [X] to [0] resolves the conflict, but Coq is not yet smart enough to do this unification automatically.  Repeating the function's type in a [return] annotation, used with an [in] annotation, leads us to a more informative error message, saying that [idx'] has type [fin n1] while it is expected to have type [fin n0], where [n0] is bound by the [Cons] pattern and [n1] by the [Next] pattern.  As the code is written above, nothing forces these two natural numbers to be equal, though we know intuitively that they must be.

    We need to use [match] annotations to make the relationship explicit.  Unfortunately, the usual trick of postponing argument binding will not help us here.  We need to match on both [ls] and [idx]; one or the other must be matched first.  To get around this, we apply the convoy pattern that we met last chapter.  This application is a little more clever than those we saw before; we use the natural number predecessor function [pred] to express the relationship between the types of these variables.
Adam Chlipala's avatar
Adam Chlipala committed
86 87

    [[
Adam Chlipala's avatar
Adam Chlipala committed
88 89
  Fixpoint get n (ls : ilist n) : fin n -> A :=
    match ls with
Adam Chlipala's avatar
Adam Chlipala committed
90
      | Nil => fun idx =>
Adam Chlipala's avatar
Adam Chlipala committed
91
        match idx in fin n' return (match n' with
Adam Chlipala's avatar
Adam Chlipala committed
92 93 94 95 96 97 98
                                        | O => A
                                        | S _ => unit
                                      end) with
          | First _ => tt
          | Next _ _ => tt
        end
      | Cons _ x ls' => fun idx =>
Adam Chlipala's avatar
Adam Chlipala committed
99
        match idx in fin n' return ilist (pred n') -> A with
Adam Chlipala's avatar
Adam Chlipala committed
100 101 102 103 104
          | First _ => fun _ => x
          | Next _ idx' => fun ls' => get ls' idx'
        end ls'
    end.

105 106
    ]]

Adam Chlipala's avatar
Adam Chlipala committed
107 108
    There is just one problem left with this implementation.  Though we know that the local [ls'] in the [Next] case is equal to the original [ls'], the type-checker is not satisfied that the recursive call to [get] does not introduce non-termination.  We solve the problem by convoy-binding the partial application of [get] to [ls'], rather than [ls'] by itself. *)

Adam Chlipala's avatar
Adam Chlipala committed
109 110
  Fixpoint get n (ls : ilist n) : fin n -> A :=
    match ls with
Adam Chlipala's avatar
Adam Chlipala committed
111
      | Nil => fun idx =>
Adam Chlipala's avatar
Adam Chlipala committed
112
        match idx in fin n' return (match n' with
Adam Chlipala's avatar
Adam Chlipala committed
113 114 115 116 117 118 119
                                        | O => A
                                        | S _ => unit
                                      end) with
          | First _ => tt
          | Next _ _ => tt
        end
      | Cons _ x ls' => fun idx =>
Adam Chlipala's avatar
Adam Chlipala committed
120
        match idx in fin n' return (fin (pred n') -> A) -> A with
Adam Chlipala's avatar
Adam Chlipala committed
121 122 123 124
          | First _ => fun _ => x
          | Next _ idx' => fun get_ls' => get_ls' idx'
        end (get ls')
    end.
Adam Chlipala's avatar
Adam Chlipala committed
125
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
126 127 128
End ilist.

Implicit Arguments Nil [A].
Adam Chlipala's avatar
Adam Chlipala committed
129
Implicit Arguments First [n].
Adam Chlipala's avatar
Adam Chlipala committed
130

Adam Chlipala's avatar
Adam Chlipala committed
131 132 133
(** A few examples show how to make use of these definitions. *)

Check Cons 0 (Cons 1 (Cons 2 Nil)).
Adam Chlipala's avatar
Adam Chlipala committed
134 135
(** %\vspace{-.15in}% [[
  Cons 0 (Cons 1 (Cons 2 Nil))
Adam Chlipala's avatar
Adam Chlipala committed
136
     : ilist nat 3
137 138
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
139

Adam Chlipala's avatar
Adam Chlipala committed
140
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
141
Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) First.
Adam Chlipala's avatar
Adam Chlipala committed
142
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
143 144
     = 0
     : nat
145 146
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
147

Adam Chlipala's avatar
Adam Chlipala committed
148 149
Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) (Next First).
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
150 151
     = 1
     : nat
152 153
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
154

Adam Chlipala's avatar
Adam Chlipala committed
155 156
Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) (Next (Next First)).
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
157 158
     = 2
     : nat
159 160
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
161
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
162 163

(** Our [get] function is also quite easy to reason about.  We show how with a short example about an analogue to the list [map] function. *)
Adam Chlipala's avatar
Adam Chlipala committed
164

Adam Chlipala's avatar
Adam Chlipala committed
165 166 167 168
Section ilist_map.
  Variables A B : Set.
  Variable f : A -> B.

Adam Chlipala's avatar
Adam Chlipala committed
169 170
  Fixpoint imap n (ls : ilist A n) : ilist B n :=
    match ls with
Adam Chlipala's avatar
Adam Chlipala committed
171 172 173 174
      | Nil => Nil
      | Cons _ x ls' => Cons (f x) (imap ls')
    end.

175
  (** It is easy to prove that [get] %``%#"#distributes over#"#%''% [imap] calls.  The only tricky bit is remembering to use our [dep_destruct] tactic in place of plain [destruct] when faced with a baffling tactic error message. *)
Adam Chlipala's avatar
Adam Chlipala committed
176

Adam Chlipala's avatar
Adam Chlipala committed
177 178 179
(* EX: Prove that [get] distributes over [imap]. *)

(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
180
  Theorem get_imap : forall n (idx : fin n) (ls : ilist A n),
Adam Chlipala's avatar
Adam Chlipala committed
181
    get (imap ls) idx = f (get ls idx).
Adam Chlipala's avatar
Adam Chlipala committed
182
    induction ls; dep_destruct idx; crush.
Adam Chlipala's avatar
Adam Chlipala committed
183
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
184
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
185
End ilist_map.
Adam Chlipala's avatar
Adam Chlipala committed
186 187 188 189


(** * Heterogeneous Lists *)

Adam Chlipala's avatar
Adam Chlipala committed
190
(** Programmers who move to statically typed functional languages from scripting languages often complain about the requirement that every element of a list have the same type.  With fancy type systems, we can partially lift this requirement.  We can index a list type with a %``%#"#type-level#"#%''% list that explains what type each element of the list should have.  This has been done in a variety of ways in Haskell using type classes, and we can do it much more cleanly and directly in Coq. *)
Adam Chlipala's avatar
Adam Chlipala committed
191 192 193 194 195

Section hlist.
  Variable A : Type.
  Variable B : A -> Type.

Adam Chlipala's avatar
Adam Chlipala committed
196 197
  (* EX: Define a type [hlist] indexed by a [list A], where the type of each element is determined by running [B] on the corresponding element of the index list. *)

Adam Chlipala's avatar
Adam Chlipala committed
198
  (** We parameterize our heterogeneous lists by a type [A] and an [A]-indexed type [B].%\index{Gallina terms!hlist}% *)
Adam Chlipala's avatar
Adam Chlipala committed
199

Adam Chlipala's avatar
Adam Chlipala committed
200
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
201 202 203 204
  Inductive hlist : list A -> Type :=
  | MNil : hlist nil
  | MCons : forall (x : A) (ls : list A), B x -> hlist ls -> hlist (x :: ls).

Adam Chlipala's avatar
Adam Chlipala committed
205
  (** We can implement a variant of the last section's [get] function for [hlist]s.  To get the dependent typing to work out, we will need to index our element selectors by the types of data that they point to.%\index{Gallina terms!member}% *)
Adam Chlipala's avatar
Adam Chlipala committed
206

Adam Chlipala's avatar
Adam Chlipala committed
207 208 209 210
(* end thide *)
  (* EX: Define an analogue to [get] for [hlist]s. *)

(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
211 212 213 214 215 216
  Variable elm : A.

  Inductive member : list A -> Type :=
  | MFirst : forall ls, member (elm :: ls)
  | MNext : forall x ls, member ls -> member (x :: ls).

217
  (** Because the element [elm] that we are %``%#"#searching for#"#%''% in a list does not change across the constructors of [member], we simplify our definitions by making [elm] a local variable.  In the definition of [member], we say that [elm] is found in any list that begins with [elm], and, if removing the first element of a list leaves [elm] present, then [elm] is present in the original list, too.  The form looks much like a predicate for list membership, but we purposely define [member] in [Type] so that we may decompose its values to guide computations.
Adam Chlipala's avatar
Adam Chlipala committed
218

Adam Chlipala's avatar
Adam Chlipala committed
219
     We can use [member] to adapt our definition of [get] to [hlist]s.  The same basic [match] tricks apply.  In the [MCons] case, we form a two-element convoy, passing both the data element [x] and the recursor for the sublist [mls'] to the result of the inner [match].  We did not need to do that in [get]'s definition because the types of list elements were not dependent there. *)
Adam Chlipala's avatar
Adam Chlipala committed
220

Adam Chlipala's avatar
Adam Chlipala committed
221 222
  Fixpoint hget ls (mls : hlist ls) : member ls -> B elm :=
    match mls with
Adam Chlipala's avatar
Adam Chlipala committed
223 224 225 226 227 228 229 230 231 232 233 234 235 236 237 238 239 240
      | MNil => fun mem =>
        match mem in member ls' return (match ls' with
                                          | nil => B elm
                                          | _ :: _ => unit
                                        end) with
          | MFirst _ => tt
          | MNext _ _ _ => tt
        end
      | MCons _ _ x mls' => fun mem =>
        match mem in member ls' return (match ls' with
                                          | nil => Empty_set
                                          | x' :: ls'' =>
                                            B x' -> (member ls'' -> B elm) -> B elm
                                        end) with
          | MFirst _ => fun x _ => x
          | MNext _ _ mem' => fun _ get_mls' => get_mls' mem'
        end x (hget mls')
    end.
Adam Chlipala's avatar
Adam Chlipala committed
241
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
242
End hlist.
Adam Chlipala's avatar
Adam Chlipala committed
243

Adam Chlipala's avatar
Adam Chlipala committed
244
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
245 246 247 248 249
Implicit Arguments MNil [A B].
Implicit Arguments MCons [A B x ls].

Implicit Arguments MFirst [A elm ls].
Implicit Arguments MNext [A elm x ls].
Adam Chlipala's avatar
Adam Chlipala committed
250
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
251 252 253 254 255

(** By putting the parameters [A] and [B] in [Type], we allow some very higher-order uses.  For instance, one use of [hlist] is for the simple heterogeneous lists that we referred to earlier. *)

Definition someTypes : list Set := nat :: bool :: nil.

Adam Chlipala's avatar
Adam Chlipala committed
256 257
(* begin thide *)

Adam Chlipala's avatar
Adam Chlipala committed
258 259 260 261
Example someValues : hlist (fun T : Set => T) someTypes :=
  MCons 5 (MCons true MNil).

Eval simpl in hget someValues MFirst.
Adam Chlipala's avatar
Adam Chlipala committed
262
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
263 264
     = 5
     : (fun T : Set => T) nat
265 266
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
267

Adam Chlipala's avatar
Adam Chlipala committed
268 269
Eval simpl in hget someValues (MNext MFirst).
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
270 271
     = true
     : (fun T : Set => T) bool
272 273
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
274 275 276 277 278 279

(** We can also build indexed lists of pairs in this way. *)

Example somePairs : hlist (fun T : Set => T * T)%type someTypes :=
  MCons (1, 2) (MCons (true, false) MNil).

Adam Chlipala's avatar
Adam Chlipala committed
280 281 282
(* end thide *)


Adam Chlipala's avatar
Adam Chlipala committed
283 284
(** ** A Lambda Calculus Interpreter *)

Adam Chlipala's avatar
Adam Chlipala committed
285
(** Heterogeneous lists are very useful in implementing %\index{interpreters}%interpreters for functional programming languages.  Using the types and operations we have already defined, it is trivial to write an interpreter for simply typed lambda calculus%\index{lambda calculus}%.  Our interpreter can alternatively be thought of as a denotational semantics.
Adam Chlipala's avatar
Adam Chlipala committed
286 287 288 289 290 291 292

   We start with an algebraic datatype for types. *)

Inductive type : Set :=
| Unit : type
| Arrow : type -> type -> type.

Adam Chlipala's avatar
Adam Chlipala committed
293
(** Now we can define a type family for expressions.  An [exp ts t] will stand for an expression that has type [t] and whose free variables have types in the list [ts].  We effectively use the de Bruijn index variable representation%~\cite{DeBruijn}%.  Variables are represented as [member] values; that is, a variable is more or less a constructive proof that a particular type is found in the type environment. *)
Adam Chlipala's avatar
Adam Chlipala committed
294 295 296

Inductive exp : list type -> type -> Set :=
| Const : forall ts, exp ts Unit
Adam Chlipala's avatar
Adam Chlipala committed
297
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
298 299 300
| Var : forall ts t, member t ts -> exp ts t
| App : forall ts dom ran, exp ts (Arrow dom ran) -> exp ts dom -> exp ts ran
| Abs : forall ts dom ran, exp (dom :: ts) ran -> exp ts (Arrow dom ran).
Adam Chlipala's avatar
Adam Chlipala committed
301
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
302 303 304 305 306 307 308 309 310 311 312

Implicit Arguments Const [ts].

(** We write a simple recursive function to translate [type]s into [Set]s. *)

Fixpoint typeDenote (t : type) : Set :=
  match t with
    | Unit => unit
    | Arrow t1 t2 => typeDenote t1 -> typeDenote t2
  end.

Adam Chlipala's avatar
Adam Chlipala committed
313
(** Now it is straightforward to write an expression interpreter.  The type of the function, [expDenote], tells us that we translate expressions into functions from properly typed environments to final values.  An environment for a free variable list [ts] is simply a [hlist typeDenote ts].  That is, for each free variable, the heterogeneous list that is the environment must have a value of the variable's associated type.  We use [hget] to implement the [Var] case, and we use [MCons] to extend the environment in the [Abs] case. *)
Adam Chlipala's avatar
Adam Chlipala committed
314

Adam Chlipala's avatar
Adam Chlipala committed
315 316 317
(* EX: Define an interpreter for [exp]s. *)

(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
318 319
Fixpoint expDenote ts t (e : exp ts t) : hlist typeDenote ts -> typeDenote t :=
  match e with
Adam Chlipala's avatar
Adam Chlipala committed
320 321 322 323 324 325 326 327 328 329
    | Const _ => fun _ => tt

    | Var _ _ mem => fun s => hget s mem
    | App _ _ _ e1 e2 => fun s => (expDenote e1 s) (expDenote e2 s)
    | Abs _ _ _ e' => fun s => fun x => expDenote e' (MCons x s)
  end.

(** Like for previous examples, our interpreter is easy to run with [simpl]. *)

Eval simpl in expDenote Const MNil.
Adam Chlipala's avatar
Adam Chlipala committed
330
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
331 332
    = tt
     : typeDenote Unit
333 334
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
335

Adam Chlipala's avatar
Adam Chlipala committed
336 337
Eval simpl in expDenote (Abs (dom := Unit) (Var MFirst)) MNil.
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
338 339
     = fun x : unit => x
     : typeDenote (Arrow Unit Unit)
340 341
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
342

Adam Chlipala's avatar
Adam Chlipala committed
343 344
Eval simpl in expDenote (Abs (dom := Unit)
  (Abs (dom := Unit) (Var (MNext MFirst)))) MNil.
Adam Chlipala's avatar
Adam Chlipala committed
345
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
346 347
     = fun x _ : unit => x
     : typeDenote (Arrow Unit (Arrow Unit Unit))
348 349
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
350

Adam Chlipala's avatar
Adam Chlipala committed
351 352
Eval simpl in expDenote (Abs (dom := Unit) (Abs (dom := Unit) (Var MFirst))) MNil.
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
353 354
     = fun _ x0 : unit => x0
     : typeDenote (Arrow Unit (Arrow Unit Unit))
355 356
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
357

Adam Chlipala's avatar
Adam Chlipala committed
358 359
Eval simpl in expDenote (App (Abs (Var MFirst)) Const) MNil.
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
360 361
     = tt
     : typeDenote Unit
362 363
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
364

Adam Chlipala's avatar
Adam Chlipala committed
365 366
(* end thide *)

Adam Chlipala's avatar
Adam Chlipala committed
367
(** We are starting to develop the tools behind dependent typing's amazing advantage over alternative approaches in several important areas.  Here, we have implemented complete syntax, typing rules, and evaluation semantics for simply typed lambda calculus without even needing to define a syntactic substitution operation.  We did it all without a single line of proof, and our implementation is manifestly executable.  Other, more common approaches to language formalization often state and prove explicit theorems about type safety of languages.  In the above example, we got type safety, termination, and other meta-theorems for free, by reduction to CIC, which we know has those properties. *)
Adam Chlipala's avatar
Adam Chlipala committed
368 369


370 371
(** * Recursive Type Definitions *)

Adam Chlipala's avatar
Adam Chlipala committed
372
(** %\index{recursive type definition}%There is another style of datatype definition that leads to much simpler definitions of the [get] and [hget] definitions above.  Because Coq supports %``%#"#type-level computation,#"#%''% we can redo our inductive definitions as %\textit{%#<i>#recursive#</i>#%}% definitions. *)
373

Adam Chlipala's avatar
Adam Chlipala committed
374 375
(* EX: Come up with an alternate [ilist] definition that makes it easier to write [get]. *)

376 377 378
Section filist.
  Variable A : Set.

Adam Chlipala's avatar
Adam Chlipala committed
379
(* begin thide *)
380 381 382 383 384 385 386 387
  Fixpoint filist (n : nat) : Set :=
    match n with
      | O => unit
      | S n' => A * filist n'
    end%type.

  (** We say that a list of length 0 has no contents, and a list of length [S n'] is a pair of a data value and a list of length [n']. *)

Adam Chlipala's avatar
Adam Chlipala committed
388
  Fixpoint ffin (n : nat) : Set :=
389 390
    match n with
      | O => Empty_set
Adam Chlipala's avatar
Adam Chlipala committed
391
      | S n' => option (ffin n')
392 393
    end.

Adam Chlipala's avatar
Adam Chlipala committed
394
  (** We express that there are no index values when [n = O], by defining such indices as type [Empty_set]; and we express that, at [n = S n'], there is a choice between picking the first element of the list (represented as [None]) or choosing a later element (represented by [Some idx], where [idx] is an index into the list tail).  For instance, the three values of type [ffin 3] are [None], [Some None], and [Some (][Some None)]. *)
395

Adam Chlipala's avatar
Adam Chlipala committed
396 397
  Fixpoint fget (n : nat) : filist n -> ffin n -> A :=
    match n with
398 399 400 401 402 403 404 405
      | O => fun _ idx => match idx with end
      | S n' => fun ls idx =>
        match idx with
          | None => fst ls
          | Some idx' => fget n' (snd ls) idx'
        end
    end.

Adam Chlipala's avatar
Adam Chlipala committed
406
  (** Our new [get] implementation needs only one dependent [match], and its annotation is inferred for us.  Our choices of data structure implementations lead to just the right typing behavior for this new definition to work out. *)
Adam Chlipala's avatar
Adam Chlipala committed
407
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
408

409 410 411 412
End filist.

(** Heterogeneous lists are a little trickier to define with recursion, but we then reap similar benefits in simplicity of use. *)

Adam Chlipala's avatar
Adam Chlipala committed
413 414
(* EX: Come up with an alternate [hlist] definition that makes it easier to write [hget]. *)

415 416 417 418
Section fhlist.
  Variable A : Type.
  Variable B : A -> Type.

Adam Chlipala's avatar
Adam Chlipala committed
419
(* begin thide *)
420 421 422 423 424 425
  Fixpoint fhlist (ls : list A) : Type :=
    match ls with
      | nil => unit
      | x :: ls' => B x * fhlist ls'
    end%type.

Adam Chlipala's avatar
Adam Chlipala committed
426
  (** The definition of [fhlist] follows the definition of [filist], with the added wrinkle of dependently typed data elements. *)
427 428 429 430 431 432 433 434 435

  Variable elm : A.

  Fixpoint fmember (ls : list A) : Type :=
    match ls with
      | nil => Empty_set
      | x :: ls' => (x = elm) + fmember ls'
    end%type.

Adam Chlipala's avatar
Adam Chlipala committed
436
  (** The definition of [fmember] follows the definition of [ffin].  Empty lists have no members, and member types for nonempty lists are built by adding one new option to the type of members of the list tail.  While for [index] we needed no new information associated with the option that we add, here we need to know that the head of the list equals the element we are searching for.  We express that with a sum type whose left branch is the appropriate equality proposition.  Since we define [fmember] to live in [Type], we can insert [Prop] types as needed, because [Prop] is a subtype of [Type].
437 438 439 440 441

     We know all of the tricks needed to write a first attempt at a [get] function for [fhlist]s.

     [[
  Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm :=
Adam Chlipala's avatar
Adam Chlipala committed
442
    match ls with
443 444 445 446 447 448 449 450
      | nil => fun _ idx => match idx with end
      | _ :: ls' => fun mls idx =>
        match idx with
          | inl _ => fst mls
          | inr idx' => fhget ls' (snd mls) idx'
        end
    end.

451 452
    ]]

453 454 455
    Only one problem remains.  The expression [fst mls] is not known to have the proper type.  To demonstrate that it does, we need to use the proof available in the [inl] case of the inner [match]. *)

  Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm :=
Adam Chlipala's avatar
Adam Chlipala committed
456
    match ls with
457 458 459 460 461 462 463 464 465 466 467 468 469
      | nil => fun _ idx => match idx with end
      | _ :: ls' => fun mls idx =>
        match idx with
          | inl pf => match pf with
                        | refl_equal => fst mls
                      end
          | inr idx' => fhget ls' (snd mls) idx'
        end
    end.

  (** By pattern-matching on the equality proof [pf], we make that equality known to the type-checker.  Exactly why this works can be seen by studying the definition of equality. *)

  Print eq.
Adam Chlipala's avatar
Adam Chlipala committed
470
  (** %\vspace{-.15in}% [[
471
Inductive eq (A : Type) (x : A) : A -> Prop :=  refl_equal : x = x
Adam Chlipala's avatar
Adam Chlipala committed
472
 
473 474
]]

Adam Chlipala's avatar
Adam Chlipala committed
475
In a proposition [x = y], we see that [x] is a parameter and [y] is a regular argument.  The type of the constructor [refl_equal] shows that [y] can only ever be instantiated to [x].  Thus, within a pattern-match with [refl_equal], occurrences of [y] can be replaced with occurrences of [x] for typing purposes. *)
Adam Chlipala's avatar
Adam Chlipala committed
476
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
477

478
End fhlist.
Adam Chlipala's avatar
Adam Chlipala committed
479

480 481
Implicit Arguments fhget [A B elm ls].

Adam Chlipala's avatar
Adam Chlipala committed
482 483 484

(** * Data Structures as Index Functions *)

Adam Chlipala's avatar
Adam Chlipala committed
485
(** %\index{index function}%Indexed lists can be useful in defining other inductive types with constructors that take variable numbers of arguments.  In this section, we consider parameterized trees with arbitrary branching factor. *)
Adam Chlipala's avatar
Adam Chlipala committed
486 487 488 489 490 491 492 493 494 495 496 497 498 499 500 501

Section tree.
  Variable A : Set.

  Inductive tree : Set :=
  | Leaf : A -> tree
  | Node : forall n, ilist tree n -> tree.
End tree.

(** Every [Node] of a [tree] has a natural number argument, which gives the number of child trees in the second argument, typed with [ilist].  We can define two operations on trees of naturals: summing their elements and incrementing their elements.  It is useful to define a generic fold function on [ilist]s first. *)

Section ifoldr.
  Variables A B : Set.
  Variable f : A -> B -> B.
  Variable i : B.

Adam Chlipala's avatar
Adam Chlipala committed
502
  Fixpoint ifoldr n (ls : ilist A n) : B :=
Adam Chlipala's avatar
Adam Chlipala committed
503 504 505 506 507 508 509 510 511 512 513 514 515 516 517 518 519 520 521 522 523
    match ls with
      | Nil => i
      | Cons _ x ls' => f x (ifoldr ls')
    end.
End ifoldr.

Fixpoint sum (t : tree nat) : nat :=
  match t with
    | Leaf n => n
    | Node _ ls => ifoldr (fun t' n => sum t' + n) O ls
  end.

Fixpoint inc (t : tree nat) : tree nat :=
  match t with
    | Leaf n => Leaf (S n)
    | Node _ ls => Node (imap inc ls)
  end.

(** Now we might like to prove that [inc] does not decrease a tree's [sum]. *)

Theorem sum_inc : forall t, sum (inc t) >= sum t.
Adam Chlipala's avatar
Adam Chlipala committed
524
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
525 526 527 528 529 530 531
  induction t; crush.
  (** [[
  n : nat
  i : ilist (tree nat) n
  ============================
   ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 (imap inc i) >=
   ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 i
Adam Chlipala's avatar
Adam Chlipala committed
532
 
Adam Chlipala's avatar
Adam Chlipala committed
533 534
   ]]

Adam Chlipala's avatar
Adam Chlipala committed
535
   We are left with a single subgoal which does not seem provable directly.  This is the same problem that we met in Chapter 3 with other %\index{nested inductive type}%nested inductive types. *)
Adam Chlipala's avatar
Adam Chlipala committed
536 537

  Check tree_ind.
Adam Chlipala's avatar
Adam Chlipala committed
538 539
  (** %\vspace{-.15in}% [[
  tree_ind
Adam Chlipala's avatar
Adam Chlipala committed
540 541 542 543
     : forall (A : Set) (P : tree A -> Prop),
       (forall a : A, P (Leaf a)) ->
       (forall (n : nat) (i : ilist (tree A) n), P (Node i)) ->
       forall t : tree A, P t
Adam Chlipala's avatar
Adam Chlipala committed
544
 
Adam Chlipala's avatar
Adam Chlipala committed
545 546
]]

Adam Chlipala's avatar
Adam Chlipala committed
547
The automatically generated induction principle is too weak.  For the [Node] case, it gives us no inductive hypothesis.  We could write our own induction principle, as we did in Chapter 3, but there is an easier way, if we are willing to alter the definition of [tree]. *)
Adam Chlipala's avatar
Adam Chlipala committed
548

Adam Chlipala's avatar
Adam Chlipala committed
549 550 551 552 553 554 555 556 557
Abort.

Reset tree.

(** First, let us try using our recursive definition of [ilist]s instead of the inductive version. *)

Section tree.
  Variable A : Set.

Adam Chlipala's avatar
Adam Chlipala committed
558
  (** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
559 560 561
  Inductive tree : Set :=
  | Leaf : A -> tree
  | Node : forall n, filist tree n -> tree.
Adam Chlipala's avatar
Adam Chlipala committed
562
]]
Adam Chlipala's avatar
Adam Chlipala committed
563

Adam Chlipala's avatar
Adam Chlipala committed
564
<<
Adam Chlipala's avatar
Adam Chlipala committed
565 566
Error: Non strictly positive occurrence of "tree" in
 "forall n : nat, filist tree n -> tree"
Adam Chlipala's avatar
Adam Chlipala committed
567
>>
Adam Chlipala's avatar
Adam Chlipala committed
568

Adam Chlipala's avatar
Adam Chlipala committed
569
  The special-case rule for nested datatypes only works with nested uses of other inductive types, which could be replaced with uses of new mutually inductive types.  We defined [filist] recursively, so it may not be used for nested recursion.
Adam Chlipala's avatar
Adam Chlipala committed
570

Adam Chlipala's avatar
Adam Chlipala committed
571
  Our final solution uses yet another of the inductive definition techniques introduced in Chapter 3, %\index{reflexive inductive type}%reflexive types.  Instead of merely using [fin] to get elements out of [ilist], we can %\textit{%#<i>#define#</i>#%}% [ilist] in terms of [fin].  For the reasons outlined above, it turns out to be easier to work with [ffin] in place of [fin]. *)
Adam Chlipala's avatar
Adam Chlipala committed
572 573 574

  Inductive tree : Set :=
  | Leaf : A -> tree
Adam Chlipala's avatar
Adam Chlipala committed
575 576 577
  | Node : forall n, (ffin n -> tree) -> tree.

  (** A [Node] is indexed by a natural number [n], and the node's [n] children are represented as a function from [ffin n] to trees, which is isomorphic to the [ilist]-based representation that we used above. *)
Adam Chlipala's avatar
Adam Chlipala committed
578 579 580 581 582

End tree.

Implicit Arguments Node [A n].

Adam Chlipala's avatar
Adam Chlipala committed
583
(** We can redefine [sum] and [inc] for our new [tree] type.  Again, it is useful to define a generic fold function first.  This time, it takes in a function whose range is some [ffin] type, and it folds another function over the results of calling the first function at every possible [ffin] value. *)
Adam Chlipala's avatar
Adam Chlipala committed
584 585 586 587 588 589

Section rifoldr.
  Variables A B : Set.
  Variable f : A -> B -> B.
  Variable i : B.

Adam Chlipala's avatar
Adam Chlipala committed
590 591
  Fixpoint rifoldr (n : nat) : (ffin n -> A) -> B :=
    match n with
Adam Chlipala's avatar
Adam Chlipala committed
592 593 594 595 596 597 598 599 600 601 602 603 604 605 606 607 608 609 610 611 612 613 614 615 616 617 618 619
      | O => fun _ => i
      | S n' => fun get => f (get None) (rifoldr n' (fun idx => get (Some idx)))
    end.
End rifoldr.

Implicit Arguments rifoldr [A B n].

Fixpoint sum (t : tree nat) : nat :=
  match t with
    | Leaf n => n
    | Node _ f => rifoldr plus O (fun idx => sum (f idx))
  end.

Fixpoint inc (t : tree nat) : tree nat :=
  match t with
    | Leaf n => Leaf (S n)
    | Node _ f => Node (fun idx => inc (f idx))
  end.

(** Now we are ready to prove the theorem where we got stuck before.  We will not need to define any new induction principle, but it %\textit{%#<i>#will#</i>#%}% be helpful to prove some lemmas. *)

Lemma plus_ge : forall x1 y1 x2 y2,
  x1 >= x2
  -> y1 >= y2
  -> x1 + y1 >= x2 + y2.
  crush.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
620
Lemma sum_inc' : forall n (f1 f2 : ffin n -> nat),
Adam Chlipala's avatar
Adam Chlipala committed
621 622 623 624 625 626 627 628 629 630 631 632 633
  (forall idx, f1 idx >= f2 idx)
  -> rifoldr plus 0 f1 >= rifoldr plus 0 f2.
  Hint Resolve plus_ge.

  induction n; crush.
Qed.

Theorem sum_inc : forall t, sum (inc t) >= sum t.
  Hint Resolve sum_inc'.

  induction t; crush.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
634 635
(* end thide *)

Adam Chlipala's avatar
Adam Chlipala committed
636
(** Even if Coq would generate complete induction principles automatically for nested inductive definitions like the one we started with, there would still be advantages to using this style of reflexive encoding.  We see one of those advantages in the definition of [inc], where we did not need to use any kind of auxiliary function.  In general, reflexive encodings often admit direct implementations of operations that would require recursion if performed with more traditional inductive data structures. *)
637 638 639

(** ** Another Interpreter Example *)

Adam Chlipala's avatar
Adam Chlipala committed
640
(** We develop another example of variable-arity constructors, in the form of optimization of a small expression language with a construct like Scheme's %\texttt{%#<tt>#cond#</tt>#%}%.  Each of our conditional expressions takes a list of pairs of boolean tests and bodies.  The value of the conditional comes from the body of the first test in the list to evaluate to [true].  To simplify the %\index{interpreters}%interpreter we will write, we force each conditional to include a final, default case. *)
641 642

Inductive type' : Type := Nat | Bool.
643 644

Inductive exp' : type' -> Type :=
645 646 647
| NConst : nat -> exp' Nat
| Plus : exp' Nat -> exp' Nat -> exp' Nat
| Eq : exp' Nat -> exp' Nat -> exp' Bool
648

649
| BConst : bool -> exp' Bool
Adam Chlipala's avatar
Adam Chlipala committed
650
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
651 652
| Cond : forall n t, (ffin n -> exp' Bool)
  -> (ffin n -> exp' t) -> exp' t -> exp' t.
Adam Chlipala's avatar
Adam Chlipala committed
653
(* end thide *)
654

655 656 657 658 659 660 661 662 663 664 665 666 667 668
(** A [Cond] is parameterized by a natural [n], which tells us how many cases this conditional has.  The test expressions are represented with a function of type [ffin n -> exp' Bool], and the bodies are represented with a function of type [ffin n -> exp' t], where [t] is the overall type.  The final [exp' t] argument is the default case.  For example, here is an expression that successively checks whether [2 + 2 = 5] (returning 0 if so) or if [1 + 1 = 2] (returning 1 if so), returning 2 otherwise. *)

Example ex1 := Cond 2
  (fun f => match f with
              | None => Eq (Plus (NConst 2) (NConst 2)) (NConst 5)
              | Some None => Eq (Plus (NConst 1) (NConst 1)) (NConst 2)
              | Some (Some v) => match v with end
            end)
  (fun f => match f with
              | None => NConst 0
              | Some None => NConst 1
              | Some (Some v) => match v with end
            end)
  (NConst 2).
669

670
(** We start implementing our interpreter with a standard type denotation function. *)
671

672 673
Definition type'Denote (t : type') : Set :=
  match t with
674 675
    | Nat => nat
    | Bool => bool
676 677
  end.

678 679
(** To implement the expression interpreter, it is useful to have the following function that implements the functionality of [Cond] without involving any syntax. *)

Adam Chlipala's avatar
Adam Chlipala committed
680
(* begin thide *)
681 682 683 684
Section cond.
  Variable A : Set.
  Variable default : A.

Adam Chlipala's avatar
Adam Chlipala committed
685 686
  Fixpoint cond (n : nat) : (ffin n -> bool) -> (ffin n -> A) -> A :=
    match n with
687 688 689 690 691 692 693 694 695 696 697
      | O => fun _ _ => default
      | S n' => fun tests bodies =>
        if tests None
          then bodies None
          else cond n'
            (fun idx => tests (Some idx))
            (fun idx => bodies (Some idx))
    end.
End cond.

Implicit Arguments cond [A n].
Adam Chlipala's avatar
Adam Chlipala committed
698
(* end thide *)
699

700 701
(** Now the expression interpreter is straightforward to write. *)

Adam Chlipala's avatar
Adam Chlipala committed
702 703 704 705
Fixpoint exp'Denote t (e : exp' t) : type'Denote t :=
  match e with
    | NConst n => n
    | Plus e1 e2 => exp'Denote e1 + exp'Denote e2
706 707 708
    | Eq e1 e2 =>
      if eq_nat_dec (exp'Denote e1) (exp'Denote e2) then true else false

Adam Chlipala's avatar
Adam Chlipala committed
709
    | BConst b => b
710
    | Cond _ _ tests bodies default =>
Adam Chlipala's avatar
Adam Chlipala committed
711
(* begin thide *)
712 713 714 715
      cond
      (exp'Denote default)
      (fun idx => exp'Denote (tests idx))
      (fun idx => exp'Denote (bodies idx))
Adam Chlipala's avatar
Adam Chlipala committed
716
(* end thide *)
717 718
  end.

719 720
(** We will implement a constant-folding function that optimizes conditionals, removing cases with known-[false] tests and cases that come after known-[true] tests.  A function [cfoldCond] implements the heart of this logic.  The convoy pattern is used again near the end of the implementation. *)

Adam Chlipala's avatar
Adam Chlipala committed
721
(* begin thide *)
722 723 724 725
Section cfoldCond.
  Variable t : type'.
  Variable default : exp' t.

726
  Fixpoint cfoldCond (n : nat)
Adam Chlipala's avatar
Adam Chlipala committed
727 728
    : (ffin n -> exp' Bool) -> (ffin n -> exp' t) -> exp' t :=
    match n with
729 730
      | O => fun _ _ => default
      | S n' => fun tests bodies =>
731
        match tests None return _ with
732 733 734 735 736 737 738 739
          | BConst true => bodies None
          | BConst false => cfoldCond n'
            (fun idx => tests (Some idx))
            (fun idx => bodies (Some idx))
          | _ =>
            let e := cfoldCond n'
              (fun idx => tests (Some idx))
              (fun idx => bodies (Some idx)) in
740 741
            match e in exp' t return exp' t -> exp' t with
              | Cond n _ tests' bodies' default' => fun body =>
742 743 744
                Cond
                (S n)
                (fun idx => match idx with
745
                              | None => tests None
746 747 748 749 750 751 752
                              | Some idx => tests' idx
                            end)
                (fun idx => match idx with
                              | None => body
                              | Some idx => bodies' idx
                            end)
                default'
753
              | e => fun body =>
754 755
                Cond
                1
756
                (fun _ => tests None)
757 758
                (fun _ => body)
                e
759
            end (bodies None)
760 761 762 763 764
        end
    end.
End cfoldCond.

Implicit Arguments cfoldCond [t n].
Adam Chlipala's avatar
Adam Chlipala committed
765
(* end thide *)
766

767 768
(** Like for the interpreters, most of the action was in this helper function, and [cfold] itself is easy to write. *)

Adam Chlipala's avatar
Adam Chlipala committed
769 770
Fixpoint cfold t (e : exp' t) : exp' t :=
  match e with
771 772 773 774
    | NConst n => NConst n
    | Plus e1 e2 =>
      let e1' := cfold e1 in
      let e2' := cfold e2 in
775
      match e1', e2' return _ with
776 777 778 779 780 781
        | NConst n1, NConst n2 => NConst (n1 + n2)
        | _, _ => Plus e1' e2'
      end
    | Eq e1 e2 =>
      let e1' := cfold e1 in
      let e2' := cfold e2 in
782
      match e1', e2' return _ with
783 784 785 786 787 788
        | NConst n1, NConst n2 => BConst (if eq_nat_dec n1 n2 then true else false)
        | _, _ => Eq e1' e2'
      end

    | BConst b => BConst b
    | Cond _ _ tests bodies default =>
Adam Chlipala's avatar
Adam Chlipala committed
789
(* begin thide *)
790 791 792 793
      cfoldCond
      (cfold default)
      (fun idx => cfold (tests idx))
      (fun idx => cfold (bodies idx))
Adam Chlipala's avatar
Adam Chlipala committed
794
(* end thide *)
795 796
  end.

Adam Chlipala's avatar
Adam Chlipala committed
797
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
798
(** To prove our final correctness theorem, it is useful to know that [cfoldCond] preserves expression meanings.  This lemma formalizes that property.  The proof is a standard mostly automated one, with the only wrinkle being a guided instantiation of the quantifiers in the induction hypothesis. *)
799

800
Lemma cfoldCond_correct : forall t (default : exp' t)
Adam Chlipala's avatar
Adam Chlipala committed
801
  n (tests : ffin n -> exp' Bool) (bodies : ffin n -> exp' t),
802 803 804 805 806
  exp'Denote (cfoldCond default tests bodies)
  = exp'Denote (Cond n tests bodies default).
  induction n; crush;
    match goal with
      | [ IHn : forall tests bodies, _, tests : _ -> _, bodies : _ -> _ |- _ ] =>
807
        specialize (IHn (fun idx => tests (Some idx)) (fun idx => bodies (Some idx)))
808 809 810 811 812 813 814 815 816 817 818 819 820
    end;
    repeat (match goal with
              | [ |- context[match ?E with
                               | NConst _ => _
                               | Plus _ _ => _
                               | Eq _ _ => _
                               | BConst _ => _
                               | Cond _ _ _ _ _ => _
                             end] ] => dep_destruct E
              | [ |- context[if ?B then _ else _] ] => destruct B
            end; crush).
Qed.

821 822
(** It is also useful to know that the result of a call to [cond] is not changed by substituting new tests and bodies functions, so long as the new functions have the same input-output behavior as the old.  It turns out that, in Coq, it is not possible to prove in general that functions related in this way are equal.  We treat this issue with our discussion of axioms in a later chapter.  For now, it suffices to prove that the particular function [cond] is %\textit{%#<i>#extensional#</i>#%}%; that is, it is unaffected by substitution of functions with input-output equivalents. *)

Adam Chlipala's avatar
Adam Chlipala committed
823 824
Lemma cond_ext : forall (A : Set) (default : A) n (tests tests' : ffin n -> bool)
  (bodies bodies' : ffin n -> A),
825 826 827 828 829 830 831 832 833 834
  (forall idx, tests idx = tests' idx)
  -> (forall idx, bodies idx = bodies' idx)
  -> cond default tests bodies
  = cond default tests' bodies'.
  induction n; crush;
    match goal with
      | [ |- context[if ?E then _ else _] ] => destruct E
    end; crush.
Qed.

835
(** Now the final theorem is easy to prove.  We add our two lemmas as hints and perform standard automation with pattern-matching of subterms to destruct. *)
Adam Chlipala's avatar
Adam Chlipala committed
836
(* end thide *)
837

838 839
Theorem cfold_correct : forall t (e : exp' t),
  exp'Denote (cfold e) = exp'Denote e.
Adam Chlipala's avatar
Adam Chlipala committed
840
(* begin thide *)
841
  Hint Rewrite cfoldCond_correct.
842 843 844 845 846 847 848
  Hint Resolve cond_ext.

  induction e; crush;
    repeat (match goal with
              | [ |- context[cfold ?E] ] => dep_destruct (cfold E)
            end; crush).
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
849
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
850 851


Adam Chlipala's avatar
Adam Chlipala committed
852 853 854 855 856 857
(** * Choosing Between Representations *)

(** It is not always clear which of these representation techniques to apply in a particular situation, but I will try to summarize the pros and cons of each.

   Inductive types are often the most pleasant to work with, after someone has spent the time implementing some basic library functions for them, using fancy [match] annotations.  Many aspects of Coq's logic and tactic support are specialized to deal with inductive types, and you may miss out if you use alternate encodings.

Adam Chlipala's avatar
Adam Chlipala committed
858 859 860
   Recursive types usually involve much less initial effort, but they can be less convenient to use with proof automation.  For instance, the [simpl] tactic (which is among the ingredients in [crush]) will sometimes be overzealous in simplifying uses of functions over recursive types.  Consider a call [get l f], where variable [l] has type [filist A (][S n)].  The type of [l] would be simplified to an explicit pair type.  In a proof involving many recursive types, this kind of unhelpful %``%#"#simplification#"#%''% can lead to rapid bloat in the sizes of subgoals.  Even worse, it can prevent syntactic pattern-matching, like in cases where [filist] is expected but a pair type is found in the %``%#"#simplified#"#%''% version.  The same problem applies to applications of recursive functions to values in recursive types: the recursive function call may %``%#"#simplify#"#%''% when the top-level structure of the type index but not the recursive value is known, because such functions are generally defined by recursion on the index, not the value.

   Another disadvantage of recursive types is that they only apply to type families whose indices determine their %``%#"#skeletons.#"#%''%  This is not true for all data structures; a good counterexample comes from the richly typed programming language syntax types we have used several times so far.  The fact that a piece of syntax has type [Nat] tells us nothing about the tree structure of that syntax.
Adam Chlipala's avatar
Adam Chlipala committed
861

Adam Chlipala's avatar
Adam Chlipala committed
862
   Finally, Coq type inference can be more helpful in constructing values in inductive types.  Application of a particular constructor of that type tells Coq what to expect from the arguments, while, for instance, forming a generic pair does not make clear an intention to interpret the value as belonging to a particular recursive type.  This downside can be mitigated to an extent by writing %``%#"#constructor#"#%''% functions for a recursive type, mirroring the definition of the corresponding inductive type.
Adam Chlipala's avatar
Adam Chlipala committed
863

Adam Chlipala's avatar
Adam Chlipala committed
864
   Reflexive encodings of data types are seen relatively rarely.  As our examples demonstrated, manipulating index values manually can lead to hard-to-read code.  A normal inductive type is generally easier to work with, once someone has gone through the trouble of implementing an induction principle manually with the techniques we studied in Chapter 3.  For small developments, avoiding that kind of coding can justify the use of reflexive data structures.  There are also some useful instances of %\index{co-inductive types}%co-inductive definitions with nested data structures (e.g., lists of values in the co-inductive type) that can only be deconstructed effectively with reflexive encoding of the nested structures. *)