comparison src/DataStruct.v @ 215:f8bcd33bdd91

Port DataStruct
author Adam Chlipala <adamc@hcoop.net>
date Wed, 11 Nov 2009 14:00:04 -0500
parents f05514cc6c0d
children d1464997078d
comparison
equal deleted inserted replaced
214:768889c969e9 215:f8bcd33bdd91
1 (* Copyright (c) 2008, Adam Chlipala 1 (* Copyright (c) 2008-2009, Adam Chlipala
2 * 2 *
3 * This work is licensed under a 3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License. 5 * Unported License.
6 * The license text is available at: 6 * The license text is available at:
30 30
31 Inductive ilist : nat -> Set := 31 Inductive ilist : nat -> Set :=
32 | Nil : ilist O 32 | Nil : ilist O
33 | Cons : forall n, A -> ilist n -> ilist (S n). 33 | Cons : forall n, A -> ilist n -> ilist (S n).
34 34
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], where [index n] is isomorphic to [{m : nat | m < n}]. Such a type family is also often called [Fin] or similar, standing for "finite." *) 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 [fin], where [fin n] is isomorphic to [{m : nat | m < n}]. The type family names stands for "finite." *)
36 36
37 (* EX: Define a function [get] for extracting an [ilist] element by position. *) 37 (* EX: Define a function [get] for extracting an [ilist] element by position. *)
38 38
39 (* begin thide *) 39 (* begin thide *)
40 Inductive index : nat -> Set := 40 Inductive fin : nat -> Set :=
41 | First : forall n, index (S n) 41 | First : forall n, fin (S n)
42 | Next : forall n, index n -> index (S n). 42 | Next : forall n, fin n -> fin (S n).
43 43
44 (** [index] essentially makes 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. 44 (** [fin] essentially makes 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.
45 45
46 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. 46 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.
47 47
48 [[ 48 [[
49 Fixpoint get n (ls : ilist n) {struct ls} : index n -> A := 49 Fixpoint get n (ls : ilist n) : fin n -> A :=
50 match ls in ilist n return index n -> A with 50 match ls with
51 | Nil => fun idx => ? 51 | Nil => fun idx => ?
52 | Cons _ x ls' => fun idx => 52 | Cons _ x ls' => fun idx =>
53 match idx with 53 match idx with
54 | First _ => x 54 | First _ => x
55 | Next _ idx' => get ls' idx' 55 | Next _ idx' => get ls' idx'
56 end 56 end
57 end. 57 end.
58 58
59 ]] 59 ]]
60 60
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 [index] 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]. 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].
62 62
63 [[ 63 [[
64 Fixpoint get n (ls : ilist n) {struct ls} : index n -> A := 64 Fixpoint get n (ls : ilist n) : fin n -> A :=
65 match ls in ilist n return index n -> A with 65 match ls with
66 | Nil => fun idx => 66 | Nil => fun idx =>
67 match idx in index n' return (match n' with 67 match idx in fin n' return (match n' with
68 | O => A 68 | O => A
69 | S _ => unit 69 | S _ => unit
70 end) with 70 end) with
71 | First _ => tt 71 | First _ => tt
72 | Next _ _ => tt 72 | Next _ _ => tt
78 end 78 end
79 end. 79 end.
80 80
81 ]] 81 ]]
82 82
83 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']. 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 a trick that we will call "the convoy pattern," introducing a new function and applying it immediately, to satisfy the type checker. 83 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']. 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.
84 84
85 [[ 85 [[
86 Fixpoint get n (ls : ilist n) {struct ls} : index n -> A := 86 Fixpoint get n (ls : ilist n) : fin n -> A :=
87 match ls in ilist n return index n -> A with 87 match ls with
88 | Nil => fun idx => 88 | Nil => fun idx =>
89 match idx in index n' return (match n' with 89 match idx in fin n' return (match n' with
90 | O => A 90 | O => A
91 | S _ => unit 91 | S _ => unit
92 end) with 92 end) with
93 | First _ => tt 93 | First _ => tt
94 | Next _ _ => tt 94 | Next _ _ => tt
95 end 95 end
96 | Cons _ x ls' => fun idx => 96 | Cons _ x ls' => fun idx =>
97 match idx in index n' return ilist (pred n') -> A with 97 match idx in fin n' return ilist (pred n') -> A with
98 | First _ => fun _ => x 98 | First _ => fun _ => x
99 | Next _ idx' => fun ls' => get ls' idx' 99 | Next _ idx' => fun ls' => get ls' idx'
100 end ls' 100 end ls'
101 end. 101 end.
102 102
103 ]] 103 ]]
104 104
105 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. *) 105 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. *)
106 106
107 Fixpoint get n (ls : ilist n) {struct ls} : index n -> A := 107 Fixpoint get n (ls : ilist n) : fin n -> A :=
108 match ls in ilist n return index n -> A with 108 match ls with
109 | Nil => fun idx => 109 | Nil => fun idx =>
110 match idx in index n' return (match n' with 110 match idx in fin n' return (match n' with
111 | O => A 111 | O => A
112 | S _ => unit 112 | S _ => unit
113 end) with 113 end) with
114 | First _ => tt 114 | First _ => tt
115 | Next _ _ => tt 115 | Next _ _ => tt
116 end 116 end
117 | Cons _ x ls' => fun idx => 117 | Cons _ x ls' => fun idx =>
118 match idx in index n' return (index (pred n') -> A) -> A with 118 match idx in fin n' return (fin (pred n') -> A) -> A with
119 | First _ => fun _ => x 119 | First _ => fun _ => x
120 | Next _ idx' => fun get_ls' => get_ls' idx' 120 | Next _ idx' => fun get_ls' => get_ls' idx'
121 end (get ls') 121 end (get ls')
122 end. 122 end.
123 (* end thide *) 123 (* end thide *)
129 (* end thide *) 129 (* end thide *)
130 130
131 (** A few examples show how to make use of these definitions. *) 131 (** A few examples show how to make use of these definitions. *)
132 132
133 Check Cons 0 (Cons 1 (Cons 2 Nil)). 133 Check Cons 0 (Cons 1 (Cons 2 Nil)).
134 (** [[ 134 (** %\vspace{-.15in}% [[
135 135 Cons 0 (Cons 1 (Cons 2 Nil))
136 Cons 0 (Cons 1 (Cons 2 Nil))
137 : ilist nat 3 136 : ilist nat 3
138 ]] *) 137 ]] *)
138
139 (* begin thide *) 139 (* begin thide *)
140 Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) First. 140 Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) First.
141 (** [[ 141 (** %\vspace{-.15in}% [[
142
143 = 0 142 = 0
144 : nat 143 : nat
145 ]] *) 144 ]] *)
145
146 Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) (Next First). 146 Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) (Next First).
147 (** [[ 147 (** %\vspace{-.15in}% [[
148
149 = 1 148 = 1
150 : nat 149 : nat
151 ]] *) 150 ]] *)
151
152 Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) (Next (Next First)). 152 Eval simpl in get (Cons 0 (Cons 1 (Cons 2 Nil))) (Next (Next First)).
153 (** [[ 153 (** %\vspace{-.15in}% [[
154
155 = 2 154 = 2
156 : nat 155 : nat
157 ]] *) 156 ]] *)
158 (* end thide *) 157 (* end thide *)
159 158
161 160
162 Section ilist_map. 161 Section ilist_map.
163 Variables A B : Set. 162 Variables A B : Set.
164 Variable f : A -> B. 163 Variable f : A -> B.
165 164
166 Fixpoint imap n (ls : ilist A n) {struct ls} : ilist B n := 165 Fixpoint imap n (ls : ilist A n) : ilist B n :=
167 match ls in ilist _ n return ilist B n with 166 match ls with
168 | Nil => Nil 167 | Nil => Nil
169 | Cons _ x ls' => Cons (f x) (imap ls') 168 | Cons _ x ls' => Cons (f x) (imap ls')
170 end. 169 end.
171 170
172 (** It is easy to prove that [get] "distributes over" [imap] calls. The only tricky bit is remembering to use the [dep_destruct] tactic in place of plain [destruct] when faced with a baffling tactic error message. *) 171 (** It is easy to prove that [get] "distributes over" [imap] calls. The only tricky bit is remembering to use the [dep_destruct] tactic in place of plain [destruct] when faced with a baffling tactic error message. *)
173 172
174 Theorem get_imap : forall n (idx : index n) (ls : ilist A n), 173 Theorem get_imap : forall n (idx : fin n) (ls : ilist A n),
175 get (imap ls) idx = f (get ls idx). 174 get (imap ls) idx = f (get ls idx).
176 (* begin thide *) 175 (* begin thide *)
177 induction ls; dep_destruct idx; crush. 176 induction ls; dep_destruct idx; crush.
178 Qed. 177 Qed.
179 (* end thide *) 178 (* end thide *)
180 End ilist_map. 179 End ilist_map.
181 180
182 181
183 (** * Heterogeneous Lists *) 182 (** * Heterogeneous Lists *)
184 183
185 (** 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 it we can do it much more cleanly and directly in Coq. *) 184 (** 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. *)
186 185
187 Section hlist. 186 Section hlist.
188 Variable A : Type. 187 Variable A : Type.
189 Variable B : A -> Type. 188 Variable B : A -> Type.
190 189
211 210
212 (** 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. 211 (** 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.
213 212
214 We can use [member] to adapt our definition of [get] to [hlists]. 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. *) 213 We can use [member] to adapt our definition of [get] to [hlists]. 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. *)
215 214
216 Fixpoint hget ls (mls : hlist ls) {struct mls} : member ls -> B elm := 215 Fixpoint hget ls (mls : hlist ls) : member ls -> B elm :=
217 match mls in hlist ls return member ls -> B elm with 216 match mls with
218 | MNil => fun mem => 217 | MNil => fun mem =>
219 match mem in member ls' return (match ls' with 218 match mem in member ls' return (match ls' with
220 | nil => B elm 219 | nil => B elm
221 | _ :: _ => unit 220 | _ :: _ => unit
222 end) with 221 end) with
252 251
253 Example someValues : hlist (fun T : Set => T) someTypes := 252 Example someValues : hlist (fun T : Set => T) someTypes :=
254 MCons 5 (MCons true MNil). 253 MCons 5 (MCons true MNil).
255 254
256 Eval simpl in hget someValues MFirst. 255 Eval simpl in hget someValues MFirst.
257 (** [[ 256 (** %\vspace{-.15in}% [[
258
259 = 5 257 = 5
260 : (fun T : Set => T) nat 258 : (fun T : Set => T) nat
261 ]] *) 259 ]] *)
260
262 Eval simpl in hget someValues (MNext MFirst). 261 Eval simpl in hget someValues (MNext MFirst).
263 (** [[ 262 (** %\vspace{-.15in}% [[
264
265 = true 263 = true
266 : (fun T : Set => T) bool 264 : (fun T : Set => T) bool
267 ]] *) 265 ]] *)
268 266
269 (** We can also build indexed lists of pairs in this way. *) 267 (** We can also build indexed lists of pairs in this way. *)
286 284
287 (** 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 variable representation, which we will discuss in more detail in later chapters. 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. *) 285 (** 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 variable representation, which we will discuss in more detail in later chapters. 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. *)
288 286
289 Inductive exp : list type -> type -> Set := 287 Inductive exp : list type -> type -> Set :=
290 | Const : forall ts, exp ts Unit 288 | Const : forall ts, exp ts Unit
291
292 (* begin thide *) 289 (* begin thide *)
293 | Var : forall ts t, member t ts -> exp ts t 290 | Var : forall ts t, member t ts -> exp ts t
294 | App : forall ts dom ran, exp ts (Arrow dom ran) -> exp ts dom -> exp ts ran 291 | App : forall ts dom ran, exp ts (Arrow dom ran) -> exp ts dom -> exp ts ran
295 | Abs : forall ts dom ran, exp (dom :: ts) ran -> exp ts (Arrow dom ran). 292 | Abs : forall ts dom ran, exp (dom :: ts) ran -> exp ts (Arrow dom ran).
296 (* end thide *) 293 (* end thide *)
308 (** 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. *) 305 (** 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. *)
309 306
310 (* EX: Define an interpreter for [exp]s. *) 307 (* EX: Define an interpreter for [exp]s. *)
311 308
312 (* begin thide *) 309 (* begin thide *)
313 Fixpoint expDenote ts t (e : exp ts t) {struct e} : hlist typeDenote ts -> typeDenote t := 310 Fixpoint expDenote ts t (e : exp ts t) : hlist typeDenote ts -> typeDenote t :=
314 match e in exp ts t return hlist typeDenote ts -> typeDenote t with 311 match e with
315 | Const _ => fun _ => tt 312 | Const _ => fun _ => tt
316 313
317 | Var _ _ mem => fun s => hget s mem 314 | Var _ _ mem => fun s => hget s mem
318 | App _ _ _ e1 e2 => fun s => (expDenote e1 s) (expDenote e2 s) 315 | App _ _ _ e1 e2 => fun s => (expDenote e1 s) (expDenote e2 s)
319 | Abs _ _ _ e' => fun s => fun x => expDenote e' (MCons x s) 316 | Abs _ _ _ e' => fun s => fun x => expDenote e' (MCons x s)
320 end. 317 end.
321 318
322 (** Like for previous examples, our interpreter is easy to run with [simpl]. *) 319 (** Like for previous examples, our interpreter is easy to run with [simpl]. *)
323 320
324 Eval simpl in expDenote Const MNil. 321 Eval simpl in expDenote Const MNil.
325 (** [[ 322 (** %\vspace{-.15in}% [[
326
327 = tt 323 = tt
328 : typeDenote Unit 324 : typeDenote Unit
329 ]] *) 325 ]] *)
326
330 Eval simpl in expDenote (Abs (dom := Unit) (Var MFirst)) MNil. 327 Eval simpl in expDenote (Abs (dom := Unit) (Var MFirst)) MNil.
331 (** [[ 328 (** %\vspace{-.15in}% [[
332
333 = fun x : unit => x 329 = fun x : unit => x
334 : typeDenote (Arrow Unit Unit) 330 : typeDenote (Arrow Unit Unit)
335 ]] *) 331 ]] *)
332
336 Eval simpl in expDenote (Abs (dom := Unit) 333 Eval simpl in expDenote (Abs (dom := Unit)
337 (Abs (dom := Unit) (Var (MNext MFirst)))) MNil. 334 (Abs (dom := Unit) (Var (MNext MFirst)))) MNil.
338 (** [[ 335 (** %\vspace{-.15in}% [[
339
340 = fun x _ : unit => x 336 = fun x _ : unit => x
341 : typeDenote (Arrow Unit (Arrow Unit Unit)) 337 : typeDenote (Arrow Unit (Arrow Unit Unit))
342 ]] *) 338 ]] *)
339
343 Eval simpl in expDenote (Abs (dom := Unit) (Abs (dom := Unit) (Var MFirst))) MNil. 340 Eval simpl in expDenote (Abs (dom := Unit) (Abs (dom := Unit) (Var MFirst))) MNil.
344 (** [[ 341 (** %\vspace{-.15in}% [[
345
346 = fun _ x0 : unit => x0 342 = fun _ x0 : unit => x0
347 : typeDenote (Arrow Unit (Arrow Unit Unit)) 343 : typeDenote (Arrow Unit (Arrow Unit Unit))
348 ]] *) 344 ]] *)
345
349 Eval simpl in expDenote (App (Abs (Var MFirst)) Const) MNil. 346 Eval simpl in expDenote (App (Abs (Var MFirst)) Const) MNil.
350 (** [[ 347 (** %\vspace{-.15in}% [[
351
352 = tt 348 = tt
353 : typeDenote Unit 349 : typeDenote Unit
354 ]] *) 350 ]] *)
355 351
356 (* end thide *) 352 (* end thide *)
374 | S n' => A * filist n' 370 | S n' => A * filist n'
375 end%type. 371 end%type.
376 372
377 (** 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']. *) 373 (** 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']. *)
378 374
379 Fixpoint findex (n : nat) : Set := 375 Fixpoint ffin (n : nat) : Set :=
380 match n with 376 match n with
381 | O => Empty_set 377 | O => Empty_set
382 | S n' => option (findex n') 378 | S n' => option (ffin n')
383 end. 379 end.
384 380
385 (** 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). *) 381 (** 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). *)
386 382
387 Fixpoint fget (n : nat) : filist n -> findex n -> A := 383 Fixpoint fget (n : nat) : filist n -> ffin n -> A :=
388 match n return filist n -> findex n -> A with 384 match n with
389 | O => fun _ idx => match idx with end 385 | O => fun _ idx => match idx with end
390 | S n' => fun ls idx => 386 | S n' => fun ls idx =>
391 match idx with 387 match idx with
392 | None => fst ls 388 | None => fst ls
393 | Some idx' => fget n' (snd ls) idx' 389 | Some idx' => fget n' (snd ls) idx'
394 end 390 end
395 end. 391 end.
396 392
397 (** Our new [get] implementation needs only one dependent [match], which just copies the stated return type of the function. Our choices of data structure implementations lead to just the right typing behavior for this new definition to work out. *) 393 (** 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. *)
398 (* end thide *) 394 (* end thide *)
395
399 End filist. 396 End filist.
400 397
401 (** Heterogeneous lists are a little trickier to define with recursion, but we then reap similar benefits in simplicity of use. *) 398 (** Heterogeneous lists are a little trickier to define with recursion, but we then reap similar benefits in simplicity of use. *)
402 399
403 (* EX: Come up with an alternate [hlist] definition that makes it easier to write [hget]. *) 400 (* EX: Come up with an alternate [hlist] definition that makes it easier to write [hget]. *)
421 match ls with 418 match ls with
422 | nil => Empty_set 419 | nil => Empty_set
423 | x :: ls' => (x = elm) + fmember ls' 420 | x :: ls' => (x = elm) + fmember ls'
424 end%type. 421 end%type.
425 422
426 (** The definition of [fmember] follows the definition of [findex]. 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]. 423 (** 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].
427 424
428 We know all of the tricks needed to write a first attempt at a [get] function for [fhlist]s. 425 We know all of the tricks needed to write a first attempt at a [get] function for [fhlist]s.
429 426
430 [[ 427 [[
431
432 Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm := 428 Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm :=
433 match ls return fhlist ls -> fmember ls -> B elm with 429 match ls with
434 | nil => fun _ idx => match idx with end 430 | nil => fun _ idx => match idx with end
435 | _ :: ls' => fun mls idx => 431 | _ :: ls' => fun mls idx =>
436 match idx with 432 match idx with
437 | inl _ => fst mls 433 | inl _ => fst mls
438 | inr idx' => fhget ls' (snd mls) idx' 434 | inr idx' => fhget ls' (snd mls) idx'
442 ]] 438 ]]
443 439
444 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]. *) 440 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]. *)
445 441
446 Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm := 442 Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm :=
447 match ls return fhlist ls -> fmember ls -> B elm with 443 match ls with
448 | nil => fun _ idx => match idx with end 444 | nil => fun _ idx => match idx with end
449 | _ :: ls' => fun mls idx => 445 | _ :: ls' => fun mls idx =>
450 match idx with 446 match idx with
451 | inl pf => match pf with 447 | inl pf => match pf with
452 | refl_equal => fst mls 448 | refl_equal => fst mls
456 end. 452 end.
457 453
458 (** 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. *) 454 (** 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. *)
459 455
460 Print eq. 456 Print eq.
461 (** [[ 457 (** %\vspace{-.15in}% [[
462
463 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x 458 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
459
464 ]] 460 ]]
465 461
466 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. All examples of similar dependent pattern matching that we have seen before require explicit annotations, but Coq implements a special case of annotation inference for matches on equality proofs. *) 462 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. *)
467 (* end thide *) 463 (* end thide *)
464
468 End fhlist. 465 End fhlist.
469 466
470 Implicit Arguments fhget [A B elm ls]. 467 Implicit Arguments fhget [A B elm ls].
471 468
472 469
487 Section ifoldr. 484 Section ifoldr.
488 Variables A B : Set. 485 Variables A B : Set.
489 Variable f : A -> B -> B. 486 Variable f : A -> B -> B.
490 Variable i : B. 487 Variable i : B.
491 488
492 Fixpoint ifoldr n (ls : ilist A n) {struct ls} : B := 489 Fixpoint ifoldr n (ls : ilist A n) : B :=
493 match ls with 490 match ls with
494 | Nil => i 491 | Nil => i
495 | Cons _ x ls' => f x (ifoldr ls') 492 | Cons _ x ls' => f x (ifoldr ls')
496 end. 493 end.
497 End ifoldr. 494 End ifoldr.
518 n : nat 515 n : nat
519 i : ilist (tree nat) n 516 i : ilist (tree nat) n
520 ============================ 517 ============================
521 ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 (imap inc i) >= 518 ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 (imap inc i) >=
522 ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 i 519 ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 i
520
523 ]] 521 ]]
524 522
525 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 nested inductive types. *) 523 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 nested inductive types. *)
526 524
527 Check tree_ind. 525 Check tree_ind.
528 (** [[ 526 (** %\vspace{-.15in}% [[
529 527 tree_ind
530 tree_ind
531 : forall (A : Set) (P : tree A -> Prop), 528 : forall (A : Set) (P : tree A -> Prop),
532 (forall a : A, P (Leaf a)) -> 529 (forall a : A, P (Leaf a)) ->
533 (forall (n : nat) (i : ilist (tree A) n), P (Node i)) -> 530 (forall (n : nat) (i : ilist (tree A) n), P (Node i)) ->
534 forall t : tree A, P t 531 forall t : tree A, P t
532
535 ]] 533 ]]
536 534
537 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]. *) 535 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]. *)
536
538 Abort. 537 Abort.
539 538
540 Reset tree. 539 Reset tree.
541 540
542 (** First, let us try using our recursive definition of [ilist]s instead of the inductive version. *) 541 (** First, let us try using our recursive definition of [ilist]s instead of the inductive version. *)
543 542
544 Section tree. 543 Section tree.
545 Variable A : Set. 544 Variable A : Set.
546 545
547 (** [[ 546 (** %\vspace{-.15in}% [[
548
549 Inductive tree : Set := 547 Inductive tree : Set :=
550 | Leaf : A -> tree 548 | Leaf : A -> tree
551 | Node : forall n, filist tree n -> tree. 549 | Node : forall n, filist tree n -> tree.
552 550
553 ]]
554
555 [[
556
557 Error: Non strictly positive occurrence of "tree" in 551 Error: Non strictly positive occurrence of "tree" in
558 "forall n : nat, filist tree n -> tree" 552 "forall n : nat, filist tree n -> tree"
553
559 ]] 554 ]]
560 555
561 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. 556 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.
562 557
563 Our final solution uses yet another of the inductive definition techniques introduced in Chapter 3, reflexive types. Instead of merely using [index] to get elements out of [ilist], we can %\textit{%#<i>#define#</i>#%}% [ilist] in terms of [index]. For the reasons outlined above, it turns out to be easier to work with [findex] in place of [index]. *) 558 Our final solution uses yet another of the inductive definition techniques introduced in Chapter 3, 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]. *)
564 559
565 Inductive tree : Set := 560 Inductive tree : Set :=
566 | Leaf : A -> tree 561 | Leaf : A -> tree
567 | Node : forall n, (findex n -> tree) -> tree. 562 | Node : forall n, (ffin n -> tree) -> tree.
568 563
569 (** A [Node] is indexed by a natural number [n], and the node's [n] children are represented as a function from [findex n] to trees, which is isomorphic to the [ilist]-based representation that we used above. *) 564 (** 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. *)
565
570 End tree. 566 End tree.
571 567
572 Implicit Arguments Node [A n]. 568 Implicit Arguments Node [A n].
573 569
574 (** 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 [findex] type, and it folds another function over the results of calling the first function at every possible [findex] value. *) 570 (** 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. *)
575 571
576 Section rifoldr. 572 Section rifoldr.
577 Variables A B : Set. 573 Variables A B : Set.
578 Variable f : A -> B -> B. 574 Variable f : A -> B -> B.
579 Variable i : B. 575 Variable i : B.
580 576
581 Fixpoint rifoldr (n : nat) : (findex n -> A) -> B := 577 Fixpoint rifoldr (n : nat) : (ffin n -> A) -> B :=
582 match n return (findex n -> A) -> B with 578 match n with
583 | O => fun _ => i 579 | O => fun _ => i
584 | S n' => fun get => f (get None) (rifoldr n' (fun idx => get (Some idx))) 580 | S n' => fun get => f (get None) (rifoldr n' (fun idx => get (Some idx)))
585 end. 581 end.
586 End rifoldr. 582 End rifoldr.
587 583
606 -> y1 >= y2 602 -> y1 >= y2
607 -> x1 + y1 >= x2 + y2. 603 -> x1 + y1 >= x2 + y2.
608 crush. 604 crush.
609 Qed. 605 Qed.
610 606
611 Lemma sum_inc' : forall n (f1 f2 : findex n -> nat), 607 Lemma sum_inc' : forall n (f1 f2 : ffin n -> nat),
612 (forall idx, f1 idx >= f2 idx) 608 (forall idx, f1 idx >= f2 idx)
613 -> rifoldr plus 0 f1 >= rifoldr plus 0 f2. 609 -> rifoldr plus 0 f1 >= rifoldr plus 0 f2.
614 Hint Resolve plus_ge. 610 Hint Resolve plus_ge.
615 611
616 induction n; crush. 612 induction n; crush.
637 | Plus : exp' Nat -> exp' Nat -> exp' Nat 633 | Plus : exp' Nat -> exp' Nat -> exp' Nat
638 | Eq : exp' Nat -> exp' Nat -> exp' Bool 634 | Eq : exp' Nat -> exp' Nat -> exp' Bool
639 635
640 | BConst : bool -> exp' Bool 636 | BConst : bool -> exp' Bool
641 (* begin thide *) 637 (* begin thide *)
642 | Cond : forall n t, (findex n -> exp' Bool) 638 | Cond : forall n t, (ffin n -> exp' Bool)
643 -> (findex n -> exp' t) -> exp' t -> exp' t. 639 -> (ffin n -> exp' t) -> exp' t -> exp' t.
644 (* end thide *) 640 (* end thide *)
645 641
646 (** 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 [findex n -> exp' Bool], and the bodies are represented with a function of type [findex n -> exp' t], where [t] is the overall type. The final [exp' t] argument is the default case. 642 (** 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.
647 643
648 We start implementing our interpreter with a standard type denotation function. *) 644 We start implementing our interpreter with a standard type denotation function. *)
649 645
650 Definition type'Denote (t : type') : Set := 646 Definition type'Denote (t : type') : Set :=
651 match t with 647 match t with
658 (* begin thide *) 654 (* begin thide *)
659 Section cond. 655 Section cond.
660 Variable A : Set. 656 Variable A : Set.
661 Variable default : A. 657 Variable default : A.
662 658
663 Fixpoint cond (n : nat) : (findex n -> bool) -> (findex n -> A) -> A := 659 Fixpoint cond (n : nat) : (ffin n -> bool) -> (ffin n -> A) -> A :=
664 match n return (findex n -> bool) -> (findex n -> A) -> A with 660 match n with
665 | O => fun _ _ => default 661 | O => fun _ _ => default
666 | S n' => fun tests bodies => 662 | S n' => fun tests bodies =>
667 if tests None 663 if tests None
668 then bodies None 664 then bodies None
669 else cond n' 665 else cond n'
675 Implicit Arguments cond [A n]. 671 Implicit Arguments cond [A n].
676 (* end thide *) 672 (* end thide *)
677 673
678 (** Now the expression interpreter is straightforward to write. *) 674 (** Now the expression interpreter is straightforward to write. *)
679 675
680 Fixpoint exp'Denote t (e : exp' t) {struct e} : type'Denote t := 676 Fixpoint exp'Denote t (e : exp' t) : type'Denote t :=
681 match e in exp' t return type'Denote t with 677 match e with
682 | NConst n => 678 | NConst n => n
683 n 679 | Plus e1 e2 => exp'Denote e1 + exp'Denote e2
684 | Plus e1 e2 =>
685 exp'Denote e1 + exp'Denote e2
686 | Eq e1 e2 => 680 | Eq e1 e2 =>
687 if eq_nat_dec (exp'Denote e1) (exp'Denote e2) then true else false 681 if eq_nat_dec (exp'Denote e1) (exp'Denote e2) then true else false
688 682
689 | BConst b => 683 | BConst b => b
690 b
691 | Cond _ _ tests bodies default => 684 | Cond _ _ tests bodies default =>
692 (* begin thide *) 685 (* begin thide *)
693 cond 686 cond
694 (exp'Denote default) 687 (exp'Denote default)
695 (fun idx => exp'Denote (tests idx)) 688 (fun idx => exp'Denote (tests idx))
703 Section cfoldCond. 696 Section cfoldCond.
704 Variable t : type'. 697 Variable t : type'.
705 Variable default : exp' t. 698 Variable default : exp' t.
706 699
707 Fixpoint cfoldCond (n : nat) 700 Fixpoint cfoldCond (n : nat)
708 : (findex n -> exp' Bool) -> (findex n -> exp' t) -> exp' t := 701 : (ffin n -> exp' Bool) -> (ffin n -> exp' t) -> exp' t :=
709 match n return (findex n -> exp' Bool) -> (findex n -> exp' t) -> exp' t with 702 match n with
710 | O => fun _ _ => default 703 | O => fun _ _ => default
711 | S n' => fun tests bodies => 704 | S n' => fun tests bodies =>
712 match tests None return _ with 705 match tests None return _ with
713 | BConst true => bodies None 706 | BConst true => bodies None
714 | BConst false => cfoldCond n' 707 | BConst false => cfoldCond n'
745 Implicit Arguments cfoldCond [t n]. 738 Implicit Arguments cfoldCond [t n].
746 (* end thide *) 739 (* end thide *)
747 740
748 (** Like for the interpreters, most of the action was in this helper function, and [cfold] itself is easy to write. *) 741 (** Like for the interpreters, most of the action was in this helper function, and [cfold] itself is easy to write. *)
749 742
750 Fixpoint cfold t (e : exp' t) {struct e} : exp' t := 743 Fixpoint cfold t (e : exp' t) : exp' t :=
751 match e in exp' t return exp' t with 744 match e with
752 | NConst n => NConst n 745 | NConst n => NConst n
753 | Plus e1 e2 => 746 | Plus e1 e2 =>
754 let e1' := cfold e1 in 747 let e1' := cfold e1 in
755 let e2' := cfold e2 in 748 let e2' := cfold e2 in
756 match e1', e2' return _ with 749 match e1', e2' return _ with
777 770
778 (* begin thide *) 771 (* begin thide *)
779 (** 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 instantation of the quantifiers in the induction hypothesis. *) 772 (** 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 instantation of the quantifiers in the induction hypothesis. *)
780 773
781 Lemma cfoldCond_correct : forall t (default : exp' t) 774 Lemma cfoldCond_correct : forall t (default : exp' t)
782 n (tests : findex n -> exp' Bool) (bodies : findex n -> exp' t), 775 n (tests : ffin n -> exp' Bool) (bodies : ffin n -> exp' t),
783 exp'Denote (cfoldCond default tests bodies) 776 exp'Denote (cfoldCond default tests bodies)
784 = exp'Denote (Cond n tests bodies default). 777 = exp'Denote (Cond n tests bodies default).
785 induction n; crush; 778 induction n; crush;
786 match goal with 779 match goal with
787 | [ IHn : forall tests bodies, _, tests : _ -> _, bodies : _ -> _ |- _ ] => 780 | [ IHn : forall tests bodies, _, tests : _ -> _, bodies : _ -> _ |- _ ] =>
800 end; crush). 793 end; crush).
801 Qed. 794 Qed.
802 795
803 (** 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. *) 796 (** 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. *)
804 797
805 Lemma cond_ext : forall (A : Set) (default : A) n (tests tests' : findex n -> bool) 798 Lemma cond_ext : forall (A : Set) (default : A) n (tests tests' : ffin n -> bool)
806 (bodies bodies' : findex n -> A), 799 (bodies bodies' : ffin n -> A),
807 (forall idx, tests idx = tests' idx) 800 (forall idx, tests idx = tests' idx)
808 -> (forall idx, bodies idx = bodies' idx) 801 -> (forall idx, bodies idx = bodies' idx)
809 -> cond default tests bodies 802 -> cond default tests bodies
810 = cond default tests' bodies'. 803 = cond default tests' bodies'.
811 induction n; crush; 804 induction n; crush;
829 end; crush). 822 end; crush).
830 Qed. 823 Qed.
831 (* end thide *) 824 (* end thide *)
832 825
833 826
827 (** * Choosing Between Representations *)
828
829 (** 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.
830
831 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.
832
833 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 function [replace] of type [forall A, filist A n -> fin n -> A -> filist A n], such that [replace l f x] should substitute [x] for the element in position [f] of [l]. A call to [replace] on a variable [l] of type [filist A (S n)] would probably be simplified to an explicit pair, even though we know nothing about the structure of [l] beyond its type. In a proof involving many recursive types, this kind of unhelpful "simplification" can lead to rapid bloat in the sizes of subgoals.
834
835 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.
836
837 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. *)
838
839
834 (** * Exercises *) 840 (** * Exercises *)
835 841
836 (** remove printing * *) 842 (** remove printing * *)
837 843
838 (** Some of the type family definitions from this chapter are duplicated in the [DepList] module of the book source. Only the recursive versions of length-indexed and heterogeneous lists are included, and they are renamed without the [f] prefixes, e.g., [ilist] in place of [filist]. 844 (** Some of the type family definitions from this chapter are duplicated in the [DepList] module of the book source. Only the recursive versions of length-indexed and heterogeneous lists are included, and they are renamed without the [f] prefixes, e.g., [ilist] in place of [filist].
847 853
848 [[ 854 [[
849 t ::= bool | t + t 855 t ::= bool | t + t
850 p ::= x | b | inl p | inr p 856 p ::= x | b | inl p | inr p
851 e ::= x | b | inl e | inr e | case e of [p => e]* | _ => e 857 e ::= x | b | inl e | inr e | case e of [p => e]* | _ => e
858
852 ]] 859 ]]
853 860
854 [x] stands for a variable, and [b] stands for a boolean constant. The production for [case] expressions means that a pattern-match includes zero or more pairs of patterns and expressions, along with a default case. 861 [x] stands for a variable, and [b] stands for a boolean constant. The production for [case] expressions means that a pattern-match includes zero or more pairs of patterns and expressions, along with a default case.
855 862
856 Your interpreter should be implemented in the style demonstrated in this chapter. That is, your definition of expressions should use dependent types and de Bruijn indices to combine syntax and typing rules, such that the type of an expression tells the types of variables that are in scope. You should implement a simple recursive function translating types [t] to [Set], and your interpreter should produce values in the image of this translation.#</li># 863 Your interpreter should be implemented in the style demonstrated in this chapter. That is, your definition of expressions should use dependent types and de Bruijn indices to combine syntax and typing rules, such that the type of an expression tells the types of variables that are in scope. You should implement a simple recursive function translating types [t] to [Set], and your interpreter should produce values in the image of this translation.#</li>#