annotate src/Generic.v @ 219:dbac52f5bce1

Port Generic
author Adam Chlipala <adamc@hcoop.net>
date Fri, 13 Nov 2009 12:03:08 -0500
parents d1464997078d
children 540a09187193
rev   line source
adamc@216 1 (* Copyright (c) 2008-2009, Adam Chlipala
adamc@193 2 *
adamc@193 3 * This work is licensed under a
adamc@193 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@193 5 * Unported License.
adamc@193 6 * The license text is available at:
adamc@193 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@193 8 *)
adamc@193 9
adamc@193 10 (* begin hide *)
adamc@195 11 Require Import String List.
adamc@193 12
adamc@193 13 Require Import Tactics DepList.
adamc@193 14
adamc@193 15 Set Implicit Arguments.
adamc@193 16 (* end hide *)
adamc@193 17
adamc@193 18
adamc@219 19 (** %\chapter{Generic Programming}% *)
adamc@193 20
adamc@219 21 (** %\textit{%#<i>#Generic programming#</i>#%}% makes it possible to write functions that operate over different types of data. Parametric polymorphism in ML and Haskell is one of the simplest examples. ML-style module systems and Haskell type classes are more flexible cases. These language features are often not as powerful so we would like. For instance, while Haskell includes a type class classifying those types whose values can be pretty-printed, per-type pretty-printing is usually either implemented manually or implemented via a [deriving] clause, which triggers ad-hoc code generation. Some clever encoding tricks have been used to achieve better within Haskell and other languages, but we can do datatype-generic programming much more cleanly with dependent types. Thanks to the expressive power of CIC, we need no special language support.
adamc@193 22
adamc@219 23 Generic programming can often be very useful in Coq developments, so we devote this chapter to studying it. In a proof assistant, there is the new possibility of generic proofs about generic programs, which we also devote some space to. *)
adamc@193 24
adamc@195 25 (** * Reflecting Datatype Definitions *)
adamc@193 26
adamc@219 27 (** The key to generic programming with dependent types is %\textit{%#<i>#universe types#</i>#%}%. This concept should not be confused with the idea of %\textit{%#<i>#universes#</i>#%}% from the metatheory of CIC and related languages. Rather, the idea of universe types is to define inductive types that provide %\textit{%#<i>#syntactic representations#</i>#%}% of Coq types. We cannot directly write CIC programs that do case analysis on types, but we %\textit{%#<i>#can#</i>#%}% case analyze on reflected syntactic versions of those types.
adamc@219 28
adamc@219 29 Thus, to begin, we must define a syntactic representation of some class of datatypes. In this chapter, our running example will have to do with basic algebraic datatypes, of the kind found in ML and Haskell, but without additional bells and whistles like type parameters and mutually-recursive definitions.
adamc@219 30
adamc@219 31 The first step is to define a representation for constructors of our datatypes. *)
adamc@219 32
adamc@198 33 (* EX: Define a reflected representation of simple algebraic datatypes. *)
adamc@198 34
adamc@198 35 (* begin thide *)
adamc@193 36 Record constructor : Type := Con {
adamc@193 37 nonrecursive : Type;
adamc@193 38 recursive : nat
adamc@193 39 }.
adamc@193 40
adamc@219 41 (** The idea is that a constructor represented as [Con T n] has [n] arguments of the type that we are defining. Additionally, all of the other, non-recursive arguments can be encoded in the type [T]. When there are no non-recursive arguments, [T] can be [unit]. When there are two non-recursive arguments, of types [A] and [B], [T] can be [A * B]. We can generalizer to any number of arguments via tupling.
adamc@219 42
adamc@219 43 With this definition, it as easy to define a datatype representation in terms of lists of constructors. *)
adamc@219 44
adamc@193 45 Definition datatype := list constructor.
adamc@193 46
adamc@219 47 (** Here are a few example encodings for some common types from the Coq standard library. While our syntax type does not support type parameters directly, we can implement them at the meta level, via functions from types to [datatype]s. *)
adamc@219 48
adamc@193 49 Definition Empty_set_dt : datatype := nil.
adamc@193 50 Definition unit_dt : datatype := Con unit 0 :: nil.
adamc@193 51 Definition bool_dt : datatype := Con unit 0 :: Con unit 0 :: nil.
adamc@193 52 Definition nat_dt : datatype := Con unit 0 :: Con unit 1 :: nil.
adamc@193 53 Definition list_dt (A : Type) : datatype := Con unit 0 :: Con A 1 :: nil.
adamc@219 54
adamc@219 55 (** [Empty_set] has no constructors, so its representation is the empty list. [unit] has one constructor with no arguments, so its one reflected constructor indicates no non-recursive data and [0] recursive arguments. The representation for [bool] just duplicates this single argumentless constructor. We get from [bool] to [nat] by changing one of the constructors to indicate 1 recursive argument. We get from [nat] to [list] by adding a non-recursive argument of a parameter type [A].
adamc@219 56
adamc@219 57 As a further example, we can do the same encoding for a generic binary tree type. *)
adamc@219 58
adamc@198 59 (* end thide *)
adamc@193 60
adamc@193 61 Section tree.
adamc@193 62 Variable A : Type.
adamc@193 63
adamc@193 64 Inductive tree : Type :=
adamc@193 65 | Leaf : A -> tree
adamc@193 66 | Node : tree -> tree -> tree.
adamc@193 67 End tree.
adamc@193 68
adamc@198 69 (* begin thide *)
adamc@193 70 Definition tree_dt (A : Type) : datatype := Con A 0 :: Con unit 2 :: nil.
adamc@193 71
adamc@219 72 (** Each datatype representation stands for a family of inductive types. For a specific real datatype and a reputed representation for it, it is useful to define a type of %\textit{%#<i>#evidence#</i>#%}% that the datatype is compatible with the encoding. *)
adamc@219 73
adamc@193 74 Section denote.
adamc@193 75 Variable T : Type.
adamc@219 76 (** This variable stands for the concrete datatype that we are interested in. *)
adamc@193 77
adamc@193 78 Definition constructorDenote (c : constructor) :=
adamc@193 79 nonrecursive c -> ilist T (recursive c) -> T.
adamc@219 80 (** We write that a constructor is represented as a function returning a [T]. Such a function takes two arguments, which pack together the non-recursive and recursive arguments of the constructor. We represent a tuple of all recursive arguments using the length-indexed list type [ilist] that we met in Chapter 7. *)
adamc@193 81
adamc@193 82 Definition datatypeDenote := hlist constructorDenote.
adamc@219 83 (** Finally, the evidence for type [T] is a hetergeneous list, including a constructor denotation for every constructor encoding in a datatype encoding. Recall that, since we are inside a section binding [T] as a variable, [constructorDenote] is automatically parameterized by [T]. *)
adamc@219 84
adamc@193 85 End denote.
adamc@198 86 (* end thide *)
adamc@193 87
adamc@219 88 (** Some example pieces of evidence should help clarify the convention. First, we define some helpful notations, providing different ways of writing constructor denotations. There is really just one notation, but we need several versions of it to cover different choices of which variables will be used in the body of a definition. %The ASCII \texttt{\textasciitilde{}>} from the notation will be rendered later as $\leadsto$.% *)
adamc@219 89
adamc@219 90 (** printing ~> $\leadsto$ *)
adamc@219 91
adamc@193 92 Notation "[ ! , ! ~> x ]" := ((fun _ _ => x) : constructorDenote _ (Con _ _)).
adamc@193 93 Notation "[ v , ! ~> x ]" := ((fun v _ => x) : constructorDenote _ (Con _ _)).
adamc@219 94 Notation "[ ! , r ~> x ]" := ((fun _ r => x) : constructorDenote _ (Con _ _)).
adamc@219 95 Notation "[ v , r ~> x ]" := ((fun v r => x) : constructorDenote _ (Con _ _)).
adamc@193 96
adamc@198 97 (* begin thide *)
adamc@193 98 Definition Empty_set_den : datatypeDenote Empty_set Empty_set_dt :=
adamc@216 99 HNil.
adamc@193 100 Definition unit_den : datatypeDenote unit unit_dt :=
adamc@216 101 [!, ! ~> tt] ::: HNil.
adamc@193 102 Definition bool_den : datatypeDenote bool bool_dt :=
adamc@216 103 [!, ! ~> true] ::: [!, ! ~> false] ::: HNil.
adamc@193 104 Definition nat_den : datatypeDenote nat nat_dt :=
adamc@219 105 [!, ! ~> O] ::: [!, r ~> S (hd r)] ::: HNil.
adamc@193 106 Definition list_den (A : Type) : datatypeDenote (list A) (list_dt A) :=
adamc@219 107 [!, ! ~> nil] ::: [x, r ~> x :: hd r] ::: HNil.
adamc@193 108 Definition tree_den (A : Type) : datatypeDenote (tree A) (tree_dt A) :=
adamc@219 109 [v, ! ~> Leaf v] ::: [!, r ~> Node (hd r) (hd (tl r))] ::: HNil.
adamc@198 110 (* end thide *)
adamc@194 111
adamc@195 112
adamc@195 113 (** * Recursive Definitions *)
adamc@195 114
adamc@198 115 (* EX: Define a generic [size] function. *)
adamc@198 116
adamc@219 117 (** We built these encodings of datatypes to help us write datatype-generic recursive functions. To do so, we will want a reflected representation of a %\textit{%#<i>#recursion scheme#</i>#%}% for each type, similar to the [T_rect] principle generated automatically for an inductive definition of [T]. A clever reuse of [datatypeDenote] yields a short definition. *)
adamc@219 118
adamc@198 119 (* begin thide *)
adamc@194 120 Definition fixDenote (T : Type) (dt : datatype) :=
adamc@194 121 forall (R : Type), datatypeDenote R dt -> (T -> R).
adamc@194 122
adamc@219 123 (** The idea of a recursion scheme is parameterized by a type and a reputed encoding of it. The principle itself is polymorphic in a type [R], which is the return type of the recursive function that we mean to write. The next argument is a hetergeneous list of one case of the recursive function definition for each datatype constructor. The [datatypeDenote] function turns out to have just the right definition to express the type we need; a set of function cases is just like an alternate set of constructors where we replace the original type [T] with the function result type [R]. Given such a reflected definition, a [fixDenote] invocation returns a function from [T] to [R], which is just what we wanted.
adamc@219 124
adamc@219 125 We are ready to write some example functions now. It will be useful to use one new function from the [DepList] library included in the book source. *)
adamc@219 126
adamc@219 127 Check hmake.
adamc@219 128 (** %\vspace{-.15in}% [[
adamc@219 129 hmake
adamc@219 130 : forall (A : Type) (B : A -> Type),
adamc@219 131 (forall x : A, B x) -> forall ls : list A, hlist B l
adamc@219 132
adamc@219 133 ]]
adamc@219 134
adamc@219 135 [hmake] is a kind of [map] alternative that goes from a regular [list] to an [hlist]. We can use it to define a generic size function which counts the number of constructors used to build a value in a datatype. *)
adamc@219 136
adamc@194 137 Definition size T dt (fx : fixDenote T dt) : T -> nat :=
adamc@194 138 fx nat (hmake (B := constructorDenote nat) (fun _ _ r => foldr plus 1 r) dt).
adamc@194 139
adamc@219 140 (** Our definition is parameterized over a recursion scheme [fx]. We instantiate [fx] by passing it the function result type and a set of function cases, where we build the latter with [hmake]. The function argument to [hmake] takes three arguments: the representation of a constructor, its non-recursive arguments, and the results of recursive calls on all of its recursive arguments. We only need the recursive call results here, so we call them [r] and bind the other two inputs with wildcards. The actual case body is simple: we add together the recursive call results and increment the result by one (to account for the current constructor). This [foldr] function is an [hlist]-specific version defined in the [DepList] module.
adamc@219 141
adamc@219 142 It is instructive to build [fixDenote] values for our example types and see what specialized [size] functions result from them. *)
adamc@219 143
adamc@194 144 Definition Empty_set_fix : fixDenote Empty_set Empty_set_dt :=
adamc@194 145 fun R _ emp => match emp with end.
adamc@194 146 Eval compute in size Empty_set_fix.
adamc@219 147 (** %\vspace{-.15in}% [[
adamc@219 148 = fun emp : Empty_set => match emp return nat with
adamc@219 149 end
adamc@219 150 : Empty_set -> nat
adamc@219 151
adamc@219 152 ]]
adamc@219 153
adamc@219 154 Despite all the fanciness of the generic [size] function, CIC's standard computation rules suffice to normalize the generic function specialization to exactly what we would have written manually. *)
adamc@194 155
adamc@194 156 Definition unit_fix : fixDenote unit unit_dt :=
adamc@216 157 fun R cases _ => (hhd cases) tt INil.
adamc@194 158 Eval compute in size unit_fix.
adamc@219 159 (** %\vspace{-.15in}% [[
adamc@219 160 = fun _ : unit => 1
adamc@219 161 : unit -> nat
adamc@219 162
adamc@219 163 ]]
adamc@219 164
adamc@219 165 Again normalization gives us the natural function definition. We see this pattern repeated for our other example types. *)
adamc@194 166
adamc@194 167 Definition bool_fix : fixDenote bool bool_dt :=
adamc@194 168 fun R cases b => if b
adamc@216 169 then (hhd cases) tt INil
adamc@216 170 else (hhd (htl cases)) tt INil.
adamc@194 171 Eval compute in size bool_fix.
adamc@219 172 (** %\vspace{-.15in}% [[
adamc@219 173 = fun b : bool => if b then 1 else 1
adamc@219 174 : bool -> nat
adamc@219 175 ]] *)
adamc@194 176
adamc@194 177 Definition nat_fix : fixDenote nat nat_dt :=
adamc@194 178 fun R cases => fix F (n : nat) : R :=
adamc@194 179 match n with
adamc@216 180 | O => (hhd cases) tt INil
adamc@216 181 | S n' => (hhd (htl cases)) tt (ICons (F n') INil)
adamc@194 182 end.
adamc@219 183
adamc@219 184 (** To peek at the [size] function for [nat], it is useful to avoid full computation, so that the recursive definition of addition is not expanded inline. We can accomplish this with proper flags for the [cbv] reduction strategy. *)
adamc@219 185
adamc@194 186 Eval cbv beta iota delta -[plus] in size nat_fix.
adamc@219 187 (** %\vspace{-.15in}% [[
adamc@219 188 = fix F (n : nat) : nat := match n with
adamc@219 189 | 0 => 1
adamc@219 190 | S n' => F n' + 1
adamc@219 191 end
adamc@219 192 : nat -> nat
adamc@219 193 ]] *)
adamc@194 194
adamc@194 195 Definition list_fix (A : Type) : fixDenote (list A) (list_dt A) :=
adamc@194 196 fun R cases => fix F (ls : list A) : R :=
adamc@194 197 match ls with
adamc@216 198 | nil => (hhd cases) tt INil
adamc@216 199 | x :: ls' => (hhd (htl cases)) x (ICons (F ls') INil)
adamc@194 200 end.
adamc@194 201 Eval cbv beta iota delta -[plus] in fun A => size (@list_fix A).
adamc@219 202 (** %\vspace{-.15in}% [[
adamc@219 203 = fun A : Type =>
adamc@219 204 fix F (ls : list A) : nat :=
adamc@219 205 match ls with
adamc@219 206 | nil => 1
adamc@219 207 | _ :: ls' => F ls' + 1
adamc@219 208 end
adamc@219 209 : forall A : Type, list A -> nat
adamc@219 210 ]] *)
adamc@194 211
adamc@194 212 Definition tree_fix (A : Type) : fixDenote (tree A) (tree_dt A) :=
adamc@194 213 fun R cases => fix F (t : tree A) : R :=
adamc@194 214 match t with
adamc@216 215 | Leaf x => (hhd cases) x INil
adamc@216 216 | Node t1 t2 => (hhd (htl cases)) tt (ICons (F t1) (ICons (F t2) INil))
adamc@194 217 end.
adamc@194 218 Eval cbv beta iota delta -[plus] in fun A => size (@tree_fix A).
adamc@219 219 (** %\vspace{-.15in}% [[
adamc@219 220 = fun A : Type =>
adamc@219 221 fix F (t : tree A) : nat :=
adamc@219 222 match t with
adamc@219 223 | Leaf _ => 1
adamc@219 224 | Node t1 t2 => F t1 + (F t2 + 1)
adamc@219 225 end
adamc@219 226 : forall A : Type, tree A -> n
adamc@219 227 ]] *)
adamc@198 228 (* end thide *)
adamc@195 229
adamc@195 230
adamc@195 231 (** ** Pretty-Printing *)
adamc@195 232
adamc@198 233 (* EX: Define a generic pretty-printing function. *)
adamc@198 234
adamc@219 235 (** It is also useful to do generic pretty-printing of datatype values, rendering them as human-readable strings. To do so, we will need a bit of metadata for each constructor. Specifically, we need the name to print for the constructor and the function to use to render its non-recursive arguments. Everything else can be done generically. *)
adamc@219 236
adamc@198 237 (* begin thide *)
adamc@195 238 Record print_constructor (c : constructor) : Type := PI {
adamc@195 239 printName : string;
adamc@195 240 printNonrec : nonrecursive c -> string
adamc@195 241 }.
adamc@195 242
adamc@219 243 (** It is useful to define a shorthand for applying the constructor [PI]. By applying it explicitly to an unknown application of the constructor [Con], we help type inference work. *)
adamc@219 244
adamc@195 245 Notation "^" := (PI (Con _ _)).
adamc@195 246
adamc@219 247 (** As in earlier examples, we define the type of metadata for a datatype to be a heterogeneous list type collecting metadata for each constructor. *)
adamc@219 248
adamc@195 249 Definition print_datatype := hlist print_constructor.
adamc@195 250
adamc@219 251 (** We will be doing some string manipulation here, so we import the notations associated with strings. *)
adamc@219 252
adamc@219 253 Local Open Scope string_scope.
adamc@219 254
adamc@219 255 (** Now it is easy to implement our generic printer, using another function from [DepList.] *)
adamc@219 256
adamc@219 257 Check hmap.
adamc@219 258 (** %\vspace{-.15in}% [[
adamc@219 259 hmap
adamc@219 260 : forall (A : Type) (B1 B2 : A -> Type),
adamc@219 261 (forall x : A, B1 x -> B2 x) ->
adamc@219 262 forall ls : list A, hlist B1 ls -> hlist B2 ls
adamc@219 263 ]] *)
adamc@195 264
adamc@195 265 Definition print T dt (pr : print_datatype dt) (fx : fixDenote T dt) : T -> string :=
adamc@195 266 fx string (hmap (B1 := print_constructor) (B2 := constructorDenote string)
adamc@195 267 (fun _ pc x r => printName pc ++ "(" ++ printNonrec pc x
adamc@195 268 ++ foldr (fun s acc => ", " ++ s ++ acc) ")" r) pr).
adamc@198 269 (* end thide *)
adamc@195 270
adamc@219 271 (** Some simple tests establish that [print] gets the job done. *)
adamc@219 272
adamc@216 273 Eval compute in print HNil Empty_set_fix.
adamc@219 274 (** %\vspace{-.15in}% [[
adamc@219 275 = fun emp : Empty_set => match emp return string with
adamc@219 276 end
adamc@219 277 : Empty_set -> string
adamc@219 278 ]] *)
adamc@219 279
adamc@216 280 Eval compute in print (^ "tt" (fun _ => "") ::: HNil) unit_fix.
adamc@219 281 (** %\vspace{-.15in}% [[
adamc@219 282 = fun _ : unit => "tt()"
adamc@219 283 : unit -> string
adamc@219 284 ]] *)
adamc@219 285
adamc@195 286 Eval compute in print (^ "true" (fun _ => "")
adamc@195 287 ::: ^ "false" (fun _ => "")
adamc@216 288 ::: HNil) bool_fix.
adamc@219 289 (** %\vspace{-.15in}% [[
adamc@219 290 = fun b : bool => if b then "true()" else "false()"
adamc@219 291 : bool -> s
adamc@219 292 ]] *)
adamc@195 293
adamc@195 294 Definition print_nat := print (^ "O" (fun _ => "")
adamc@195 295 ::: ^ "S" (fun _ => "")
adamc@216 296 ::: HNil) nat_fix.
adamc@195 297 Eval cbv beta iota delta -[append] in print_nat.
adamc@219 298 (** %\vspace{-.15in}% [[
adamc@219 299 = fix F (n : nat) : string :=
adamc@219 300 match n with
adamc@219 301 | 0%nat => "O" ++ "(" ++ "" ++ ")"
adamc@219 302 | S n' => "S" ++ "(" ++ "" ++ ", " ++ F n' ++ ")"
adamc@219 303 end
adamc@219 304 : nat -> string
adamc@219 305 ]] *)
adamc@219 306
adamc@195 307 Eval simpl in print_nat 0.
adamc@219 308 (** %\vspace{-.15in}% [[
adamc@219 309 = "O()"
adamc@219 310 : string
adamc@219 311 ]] *)
adamc@219 312
adamc@195 313 Eval simpl in print_nat 1.
adamc@219 314 (** %\vspace{-.15in}% [[
adamc@219 315 = "S(, O())"
adamc@219 316 : string
adamc@219 317 ]] *)
adamc@219 318
adamc@195 319 Eval simpl in print_nat 2.
adamc@219 320 (** %\vspace{-.15in}% [[
adamc@219 321 = "S(, S(, O()))"
adamc@219 322 : string
adamc@219 323 ]] *)
adamc@195 324
adamc@195 325 Eval cbv beta iota delta -[append] in fun A (pr : A -> string) =>
adamc@195 326 print (^ "nil" (fun _ => "")
adamc@195 327 ::: ^ "cons" pr
adamc@216 328 ::: HNil) (@list_fix A).
adamc@219 329 (** %\vspace{-.15in}% [[
adamc@219 330 = fun (A : Type) (pr : A -> string) =>
adamc@219 331 fix F (ls : list A) : string :=
adamc@219 332 match ls with
adamc@219 333 | nil => "nil" ++ "(" ++ "" ++ ")"
adamc@219 334 | x :: ls' => "cons" ++ "(" ++ pr x ++ ", " ++ F ls' ++ ")"
adamc@219 335 end
adamc@219 336 : forall A : Type, (A -> string) -> list A -> string
adamc@219 337 ]] *)
adamc@219 338
adamc@195 339 Eval cbv beta iota delta -[append] in fun A (pr : A -> string) =>
adamc@195 340 print (^ "Leaf" pr
adamc@195 341 ::: ^ "Node" (fun _ => "")
adamc@216 342 ::: HNil) (@tree_fix A).
adamc@219 343 (** %\vspace{-.15in}% [[
adamc@219 344 = fun (A : Type) (pr : A -> string) =>
adamc@219 345 fix F (t : tree A) : string :=
adamc@219 346 match t with
adamc@219 347 | Leaf x => "Leaf" ++ "(" ++ pr x ++ ")"
adamc@219 348 | Node t1 t2 =>
adamc@219 349 "Node" ++ "(" ++ "" ++ ", " ++ F t1 ++ ", " ++ F t2 ++ ")"
adamc@219 350 end
adamc@219 351 : forall A : Type, (A -> string) -> tree A -> string
adamc@219 352 ]] *)
adamc@196 353
adamc@196 354
adamc@196 355 (** ** Mapping *)
adamc@196 356
adamc@198 357 (* EX: Define a generic [map] function. *)
adamc@198 358
adamc@219 359 (** By this point, we have developed enough machinery that it is old hat to define a generic function similar to the list [map] function. *)
adamc@219 360
adamc@198 361 (* begin thide *)
adamc@219 362 Definition map T dt (dd : datatypeDenote T dt) (fx : fixDenote T dt) (f : T -> T)
adamc@219 363 : T -> T :=
adamc@196 364 fx T (hmap (B1 := constructorDenote T) (B2 := constructorDenote T)
adamc@196 365 (fun _ c x r => f (c x r)) dd).
adamc@198 366 (* end thide *)
adamc@196 367
adamc@196 368 Eval compute in map Empty_set_den Empty_set_fix.
adamc@219 369 (** %\vspace{-.15in}% [[
adamc@219 370 = fun (_ : Empty_set -> Empty_set) (emp : Empty_set) =>
adamc@219 371 match emp return Empty_set with
adamc@219 372 end
adamc@219 373 : (Empty_set -> Empty_set) -> Empty_set -> Empty_set
adamc@219 374 ]] *)
adamc@219 375
adamc@196 376 Eval compute in map unit_den unit_fix.
adamc@219 377 (** %\vspace{-.15in}% [[
adamc@219 378 = fun (f : unit -> unit) (_ : unit) => f tt
adamc@219 379 : (unit -> unit) -> unit -> unit
adamc@219 380 ]] *)
adamc@219 381
adamc@196 382 Eval compute in map bool_den bool_fix.
adamc@219 383 (** %\vspace{-.15in}% [[
adamc@219 384 = fun (f : bool -> bool) (b : bool) => if b then f true else f false
adamc@219 385 : (bool -> bool) -> bool -> bool
adamc@219 386 ]] *)
adamc@219 387
adamc@196 388 Eval compute in map nat_den nat_fix.
adamc@219 389 (** %\vspace{-.15in}% [[
adamc@219 390 = fun f : nat -> nat =>
adamc@219 391 fix F (n : nat) : nat :=
adamc@219 392 match n with
adamc@219 393 | 0%nat => f 0%nat
adamc@219 394 | S n' => f (S (F n'))
adamc@219 395 end
adamc@219 396 : (nat -> nat) -> nat -> nat
adamc@219 397 ]] *)
adamc@219 398
adamc@196 399 Eval compute in fun A => map (list_den A) (@list_fix A).
adamc@219 400 (** %\vspace{-.15in}% [[
adamc@219 401 = fun (A : Type) (f : list A -> list A) =>
adamc@219 402 fix F (ls : list A) : list A :=
adamc@219 403 match ls with
adamc@219 404 | nil => f nil
adamc@219 405 | x :: ls' => f (x :: F ls')
adamc@219 406 end
adamc@219 407 : forall A : Type, (list A -> list A) -> list A -> list A
adamc@219 408 ]] *)
adamc@219 409
adamc@196 410 Eval compute in fun A => map (tree_den A) (@tree_fix A).
adamc@219 411 (** %\vspace{-.15in}% [[
adamc@219 412 = fun (A : Type) (f : tree A -> tree A) =>
adamc@219 413 fix F (t : tree A) : tree A :=
adamc@219 414 match t with
adamc@219 415 | Leaf x => f (Leaf x)
adamc@219 416 | Node t1 t2 => f (Node (F t1) (F t2))
adamc@219 417 end
adamc@219 418 : forall A : Type, (tree A -> tree A) -> tree A -> tree A
adamc@219 419 ]] *)
adamc@196 420
adamc@196 421 Definition map_nat := map nat_den nat_fix.
adamc@196 422 Eval simpl in map_nat S 0.
adamc@219 423 (** %\vspace{-.15in}% [[
adamc@219 424 = 1%nat
adamc@219 425 : nat
adamc@219 426 ]] *)
adamc@219 427
adamc@196 428 Eval simpl in map_nat S 1.
adamc@219 429 (** %\vspace{-.15in}% [[
adamc@219 430 = 3%nat
adamc@219 431 : nat
adamc@219 432 ]] *)
adamc@219 433
adamc@196 434 Eval simpl in map_nat S 2.
adamc@219 435 (** %\vspace{-.15in}% [[
adamc@219 436 = 5%nat
adamc@219 437 : nat
adamc@219 438 ]] *)
adamc@196 439
adamc@196 440
adamc@196 441 (** * Proving Theorems about Recursive Definitions *)
adamc@196 442
adamc@219 443 (** We would like to be able to prove theorems about our generic functions. To do so, we need to establish additional well-formedness properties that must hold of pieces of evidence. *)
adamc@219 444
adamc@198 445 (* begin thide *)
adamc@196 446 Section ok.
adamc@196 447 Variable T : Type.
adamc@196 448 Variable dt : datatype.
adamc@196 449
adamc@196 450 Variable dd : datatypeDenote T dt.
adamc@196 451 Variable fx : fixDenote T dt.
adamc@196 452
adamc@219 453 (** First, we characterize when a piece of evidence about a datatype is acceptable. The basic idea is that the type [T] should really be an inductive type with the definition given by [dd]. Semantically, inductive types are characterized by the ability to do induction on them. Therefore, we require that the usual induction principle is true, with respect to the constructors given in the encoding [dd]. *)
adamc@219 454
adamc@196 455 Definition datatypeDenoteOk :=
adamc@196 456 forall P : T -> Prop,
adamc@196 457 (forall c (m : member c dt) (x : nonrecursive c) (r : ilist T (recursive c)),
adamc@215 458 (forall i : fin (recursive c), P (get r i))
adamc@196 459 -> P ((hget dd m) x r))
adamc@196 460 -> forall v, P v.
adamc@196 461
adamc@219 462 (** This definition can take a while to digest. The quantifier over [m : member c dt] is considering each constructor in turn; like in normal induction principles, each constructor has an associated proof case. The expression [hget dd m] then names the constructor we have selected. After binding [m], we quantify over all possible arguments (encoded with [x] and [r]) to the constructor that [m] selects. Within each specific case, we quantify further over [i : fin (recursive c)] to consider all of our induction hypotheses, one for each recursive argument of the current constructor.
adamc@219 463
adamc@219 464 We have completed half the burden of defining side conditions. The other half comes in characterizing when a recursion scheme [fx] is valid. The natural condition is that [fx] behaves appropriately when applied to any constructor application. *)
adamc@219 465
adamc@196 466 Definition fixDenoteOk :=
adamc@196 467 forall (R : Type) (cases : datatypeDenote R dt)
adamc@196 468 c (m : member c dt)
adamc@196 469 (x : nonrecursive c) (r : ilist T (recursive c)),
adamc@216 470 fx cases ((hget dd m) x r)
adamc@216 471 = (hget cases m) x (imap (fx cases) r).
adamc@219 472
adamc@219 473 (** As for [datatypeDenoteOk], we consider all constructors and all possible arguments to them by quantifying over [m], [x], and [r]. The lefthand side of the equality that follows shows a call to the recursive function on the specific constructor application that we selected. The righthand side shows an application of the function case associated with constructor [m], applied to the non-recursive arguments and to appropriate recursive calls on the recursive arguments. *)
adamc@219 474
adamc@196 475 End ok.
adamc@196 476
adamc@219 477 (** We are now ready to prove that the [size] function we defined earlier always returns positive results. First, we establish a simple lemma. *)
adamc@196 478
adamc@196 479 Lemma foldr_plus : forall n (ils : ilist nat n),
adamc@196 480 foldr plus 1 ils > 0.
adamc@216 481 induction ils; crush.
adamc@196 482 Qed.
adamc@198 483 (* end thide *)
adamc@196 484
adamc@197 485 Theorem size_positive : forall T dt
adamc@197 486 (dd : datatypeDenote T dt) (fx : fixDenote T dt)
adamc@197 487 (dok : datatypeDenoteOk dd) (fok : fixDenoteOk dd fx)
adamc@196 488 (v : T),
adamc@196 489 size fx v > 0.
adamc@198 490 (* begin thide *)
adamc@219 491 unfold size; intros.
adamc@219 492 (** [[
adamc@219 493 ============================
adamc@219 494 fx nat
adamc@219 495 (hmake
adamc@219 496 (fun (x : constructor) (_ : nonrecursive x)
adamc@219 497 (r : ilist nat (recursive x)) => foldr plus 1%nat r) dt) v > 0
adamc@219 498
adamc@219 499 ]]
adamc@219 500
adamc@219 501 Our goal is an inequality over a particular call to [size], with its definition expanded. How can we proceed here? We cannot use [induction] directly, because there is no way for Coq to know that [T] is an inductive type. Instead, we need to use the induction principle encoded in our hypothesis [dok] of type [datatypeDenoteOk dd]. Let us try applying it directly.
adamc@219 502
adamc@219 503 [[
adamc@219 504 apply dok.
adamc@219 505
adamc@219 506 Error: Impossible to unify "datatypeDenoteOk dd" with
adamc@219 507 "fx nat
adamc@219 508 (hmake
adamc@219 509 (fun (x : constructor) (_ : nonrecursive x)
adamc@219 510 (r : ilist nat (recursive x)) => foldr plus 1%nat r) dt) v > 0".
adamc@219 511
adamc@219 512 ]]
adamc@219 513
adamc@219 514 Matching the type of [dok] with the type of our conclusion requires more than simple first-order unification, so [apply] is not up to the challenge. We can use the [pattern] tactic to get our goal into a form that makes it apparent exactly what the induction hypothesis is. *)
adamc@219 515
adamc@219 516 pattern v.
adamc@219 517
adamc@219 518 (** [[
adamc@219 519 ============================
adamc@219 520 (fun t : T =>
adamc@219 521 fx nat
adamc@219 522 (hmake
adamc@219 523 (fun (x : constructor) (_ : nonrecursive x)
adamc@219 524 (r : ilist nat (recursive x)) => foldr plus 1%nat r) dt) t > 0) v
adamc@219 525
adamc@219 526 ]] *)
adamc@219 527
adamc@219 528 apply dok; crush.
adamc@219 529 (** [[
adamc@219 530 H : forall i : fin (recursive c),
adamc@219 531 fx nat
adamc@219 532 (hmake
adamc@219 533 (fun (x : constructor) (_ : nonrecursive x)
adamc@219 534 (r : ilist nat (recursive x)) => foldr plus 1%nat r) dt)
adamc@219 535 (get r i) > 0
adamc@219 536 ============================
adamc@219 537 hget
adamc@219 538 (hmake
adamc@219 539 (fun (x0 : constructor) (_ : nonrecursive x0)
adamc@219 540 (r0 : ilist nat (recursive x0)) => foldr plus 1%nat r0) dt) m x
adamc@219 541 (imap
adamc@219 542 (fx nat
adamc@219 543 (hmake
adamc@219 544 (fun (x0 : constructor) (_ : nonrecursive x0)
adamc@219 545 (r0 : ilist nat (recursive x0)) =>
adamc@219 546 foldr plus 1%nat r0) dt)) r) > 0
adamc@219 547
adamc@219 548 ]]
adamc@219 549
adamc@219 550 An induction hypothesis [H] is generated, but we turn out not to need it for this example. We can simplify the goal using a library theorem about the composition of [hget] and [hmake]. *)
adamc@219 551
adamc@219 552 rewrite hget_hmake.
adamc@219 553 (** [[
adamc@219 554 ============================
adamc@219 555 foldr plus 1%nat
adamc@219 556 (imap
adamc@219 557 (fx nat
adamc@219 558 (hmake
adamc@219 559 (fun (x0 : constructor) (_ : nonrecursive x0)
adamc@219 560 (r0 : ilist nat (recursive x0)) =>
adamc@219 561 foldr plus 1%nat r0) dt)) r) > 0
adamc@219 562
adamc@219 563 ]]
adamc@219 564
adamc@219 565 The lemma we proved earlier finishes the proof. *)
adamc@219 566
adamc@219 567 apply foldr_plus.
adamc@219 568
adamc@219 569 (** Using hints, we can redo this proof in a nice automated form. *)
adamc@219 570
adamc@219 571 Restart.
adamc@219 572
adamc@196 573 Hint Rewrite hget_hmake : cpdt.
adamc@196 574 Hint Resolve foldr_plus.
adamc@196 575
adamc@197 576 unfold size; intros; pattern v; apply dok; crush.
adamc@196 577 Qed.
adamc@198 578 (* end thide *)
adamc@197 579
adamc@219 580 (** It turned out that, in this example, we only needed to use induction degenerately as case analysis. A more involved theorem may only be proved using induction hypotheses. We will give its proof only in unautomated form and leave effective automation as an exercise for the motivated reader.
adamc@219 581
adamc@219 582 In particular, it ought to be the case that generic [map] applied to an identity function is itself an identity function. *)
adamc@219 583
adamc@197 584 Theorem map_id : forall T dt
adamc@197 585 (dd : datatypeDenote T dt) (fx : fixDenote T dt)
adamc@197 586 (dok : datatypeDenoteOk dd) (fok : fixDenoteOk dd fx)
adamc@197 587 (v : T),
adamc@197 588 map dd fx (fun x => x) v = v.
adamc@198 589 (* begin thide *)
adamc@219 590 (** Let us begin as we did in the last theorem, after adding another useful library equality as a hint. *)
adamc@219 591
adamc@197 592 Hint Rewrite hget_hmap : cpdt.
adamc@197 593
adamc@197 594 unfold map; intros; pattern v; apply dok; crush.
adamc@219 595 (** [[
adamc@219 596 H : forall i : fin (recursive c),
adamc@219 597 fx T
adamc@219 598 (hmap
adamc@219 599 (fun (x : constructor) (c : constructorDenote T x)
adamc@219 600 (x0 : nonrecursive x) (r : ilist T (recursive x)) =>
adamc@219 601 c x0 r) dd) (get r i) = get r i
adamc@219 602 ============================
adamc@219 603 hget dd m x
adamc@219 604 (imap
adamc@219 605 (fx T
adamc@219 606 (hmap
adamc@219 607 (fun (x0 : constructor) (c0 : constructorDenote T x0)
adamc@219 608 (x1 : nonrecursive x0) (r0 : ilist T (recursive x0)) =>
adamc@219 609 c0 x1 r0) dd)) r) = hget dd m x r
adamc@219 610
adamc@219 611 ]]
adamc@197 612
adamc@219 613 Our goal is an equality whose two sides begin with the same function call and initial arguments. We believe that the remaining arguments are in fact equal as well, and the [f_equal] tactic applies this reasoning step for us formally. *)
adamc@219 614
adamc@197 615 f_equal.
adamc@219 616 (** [[
adamc@219 617 ============================
adamc@219 618 imap
adamc@219 619 (fx T
adamc@219 620 (hmap
adamc@219 621 (fun (x0 : constructor) (c0 : constructorDenote T x0)
adamc@219 622 (x1 : nonrecursive x0) (r0 : ilist T (recursive x0)) =>
adamc@219 623 c0 x1 r0) dd)) r = r
adamc@219 624
adamc@219 625 ]]
adamc@219 626
adamc@219 627 At this point, it is helpful to proceed by an inner induction on the heterogeneous list [r] of recursive call results. We could arrive at a cleaner proof by breaking this step out into an explicit lemma, but here we will do the induction inline to save space.*)
adamc@219 628
adamc@219 629 induction r; crush.
adamc@219 630
adamc@219 631 (** The base case is discharged automatically, and the inductive case looks like this, where [H] is the outer IH (for induction over [T] values) and [IHn] is the inner IH (for induction over the recursive arguments).
adamc@219 632
adamc@219 633 [[
adamc@219 634 H : forall i : fin (S n),
adamc@219 635 fx T
adamc@219 636 (hmap
adamc@219 637 (fun (x : constructor) (c : constructorDenote T x)
adamc@219 638 (x0 : nonrecursive x) (r : ilist T (recursive x)) =>
adamc@219 639 c x0 r) dd)
adamc@219 640 (match i in (fin n') return ((fin (pred n') -> T) -> T) with
adamc@219 641 | First n => fun _ : fin n -> T => a
adamc@219 642 | Next n idx' => fun get_ls' : fin n -> T => get_ls' idx'
adamc@219 643 end (get r)) =
adamc@219 644 match i in (fin n') return ((fin (pred n') -> T) -> T) with
adamc@219 645 | First n => fun _ : fin n -> T => a
adamc@219 646 | Next n idx' => fun get_ls' : fin n -> T => get_ls' idx'
adamc@219 647 end (get r)
adamc@219 648 IHr : (forall i : fin n,
adamc@219 649 fx T
adamc@219 650 (hmap
adamc@219 651 (fun (x : constructor) (c : constructorDenote T x)
adamc@219 652 (x0 : nonrecursive x) (r : ilist T (recursive x)) =>
adamc@219 653 c x0 r) dd) (get r i) = get r i) ->
adamc@219 654 imap
adamc@219 655 (fx T
adamc@219 656 (hmap
adamc@219 657 (fun (x : constructor) (c : constructorDenote T x)
adamc@219 658 (x0 : nonrecursive x) (r : ilist T (recursive x)) =>
adamc@219 659 c x0 r) dd)) r = r
adamc@219 660 ============================
adamc@219 661 ICons
adamc@219 662 (fx T
adamc@219 663 (hmap
adamc@219 664 (fun (x0 : constructor) (c0 : constructorDenote T x0)
adamc@219 665 (x1 : nonrecursive x0) (r0 : ilist T (recursive x0)) =>
adamc@219 666 c0 x1 r0) dd) a)
adamc@219 667 (imap
adamc@219 668 (fx T
adamc@219 669 (hmap
adamc@219 670 (fun (x0 : constructor) (c0 : constructorDenote T x0)
adamc@219 671 (x1 : nonrecursive x0) (r0 : ilist T (recursive x0)) =>
adamc@219 672 c0 x1 r0) dd)) r) = ICons a r
adamc@219 673
adamc@219 674 ]]
adamc@219 675
adamc@219 676 We see another opportunity to apply [f_equal], this time to split our goal into two different equalities over corresponding arguments. After that, the form of the first goal matches our outer induction hypothesis [H], when we give type inference some help by specifying the right quantifier instantiation. *)
adamc@219 677
adamc@219 678 f_equal.
adamc@219 679 apply (H First).
adamc@219 680 (** [[
adamc@219 681 ============================
adamc@219 682 imap
adamc@219 683 (fx T
adamc@219 684 (hmap
adamc@219 685 (fun (x0 : constructor) (c0 : constructorDenote T x0)
adamc@219 686 (x1 : nonrecursive x0) (r0 : ilist T (recursive x0)) =>
adamc@219 687 c0 x1 r0) dd)) r = r
adamc@219 688
adamc@219 689 ]]
adamc@219 690
adamc@219 691 Now the goal matches the inner IH [IHr]. *)
adamc@219 692
adamc@219 693 apply IHr; crush.
adamc@219 694
adamc@219 695 (** [[
adamc@219 696 i : fin n
adamc@219 697 ============================
adamc@219 698 fx T
adamc@219 699 (hmap
adamc@219 700 (fun (x0 : constructor) (c0 : constructorDenote T x0)
adamc@219 701 (x1 : nonrecursive x0) (r0 : ilist T (recursive x0)) =>
adamc@219 702 c0 x1 r0) dd) (get r i) = get r i
adamc@219 703
adamc@219 704 ]]
adamc@219 705
adamc@219 706 We can finish the proof by applying the outer IH again, specialized to a different [fin] value. *)
adamc@219 707
adamc@216 708 apply (H (Next i)).
adamc@197 709 Qed.
adamc@198 710 (* end thide *)