annotate src/MoreDep.v @ 566:1a231194d164

Two new courses
author Adam Chlipala <adam@chlipala.net>
date Fri, 19 Oct 2018 10:27:35 -0400
parents af97676583f3
children 81d63d9c1cc5
rev   line source
adam@534 1 (* Copyright (c) 2008-2012, 2015, Adam Chlipala
adamc@83 2 *
adamc@83 3 * This work is licensed under a
adamc@83 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@83 5 * Unported License.
adamc@83 6 * The license text is available at:
adamc@83 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@83 8 *)
adamc@83 9
adamc@83 10 (* begin hide *)
adam@534 11 Require Import Arith Bool List Omega.
adamc@83 12
adam@534 13 Require Import Cpdt.CpdtTactics Cpdt.MoreSpecif.
adamc@83 14
adam@563 15 Require Extraction.
adam@563 16
adamc@83 17 Set Implicit Arguments.
adam@534 18 Set Asymmetric Patterns.
adamc@83 19 (* end hide *)
adamc@83 20
adamc@83 21
adamc@83 22 (** %\chapter{More Dependent Types}% *)
adamc@83 23
adam@425 24 (** Subset types and their relatives help us integrate verification with programming. Though they reorganize the certified programmer's workflow, they tend not to have deep effects on proofs. We write largely the same proofs as we would for classical verification, with some of the structure moved into the programs themselves. It turns out that, when we use dependent types to their full potential, we warp the development and proving process even more than that, picking up "free theorems" to the extent that often a certified program is hardly more complex than its uncertified counterpart in Haskell or ML.
adamc@83 25
adam@476 26 In particular, we have only scratched the tip of the iceberg that is Coq's inductive definition mechanism. The inductive types we have seen so far have their counterparts in the other proof assistants that we surveyed in Chapter 1. This chapter explores the strange new world of dependent inductive datatypes outside [Prop], a possibility that sets Coq apart from all of the competition not based on type theory. *)
adamc@83 27
adamc@84 28
adamc@84 29 (** * Length-Indexed Lists *)
adamc@84 30
adam@338 31 (** Many introductions to dependent types start out by showing how to use them to eliminate array bounds checks%\index{array bounds checks}%. When the type of an array tells you how many elements it has, your compiler can detect out-of-bounds dereferences statically. Since we are working in a pure functional language, the next best thing is length-indexed lists%\index{length-indexed lists}%, which the following code defines. *)
adamc@84 32
adamc@84 33 Section ilist.
adamc@84 34 Variable A : Set.
adamc@84 35
adamc@84 36 Inductive ilist : nat -> Set :=
adamc@84 37 | Nil : ilist O
adamc@84 38 | Cons : forall n, A -> ilist n -> ilist (S n).
adamc@84 39
adamc@84 40 (** We see that, within its section, [ilist] is given type [nat -> Set]. Previously, every inductive type we have seen has either had plain [Set] as its type or has been a predicate with some type ending in [Prop]. The full generality of inductive definitions lets us integrate the expressivity of predicates directly into our normal programming.
adamc@84 41
adam@405 42 The [nat] argument to [ilist] tells us the length of the list. The types of [ilist]'s constructors tell us that a [Nil] list has length [O] and that a [Cons] list has length one greater than the length of its tail. We may apply [ilist] to any natural number, even natural numbers that are only known at runtime. It is this breaking of the%\index{phase distinction}% _phase distinction_ that characterizes [ilist] as _dependently typed_.
adamc@84 43
adamc@213 44 In expositions of list types, we usually see the length function defined first, but here that would not be a very productive function to code. Instead, let us implement list concatenation. *)
adamc@84 45
adamc@213 46 Fixpoint app n1 (ls1 : ilist n1) n2 (ls2 : ilist n2) : ilist (n1 + n2) :=
adamc@213 47 match ls1 with
adamc@213 48 | Nil => ls2
adamc@213 49 | Cons _ x ls1' => Cons x (app ls1' ls2)
adamc@213 50 end.
adamc@84 51
adam@338 52 (** Past Coq versions signalled an error for this definition. The code is still invalid within Coq's core language, but current Coq versions automatically add annotations to the original program, producing a valid core program. These are the annotations on [match] discriminees that we began to study in the previous chapter. We can rewrite [app] to give the annotations explicitly. *)
adamc@100 53
adamc@100 54 (* begin thide *)
adam@338 55 Fixpoint app' n1 (ls1 : ilist n1) n2 (ls2 : ilist n2) : ilist (n1 + n2) :=
adam@338 56 match ls1 in (ilist n1) return (ilist (n1 + n2)) with
adam@338 57 | Nil => ls2
adam@338 58 | Cons _ x ls1' => Cons x (app' ls1' ls2)
adam@338 59 end.
adamc@100 60 (* end thide *)
adamc@84 61
adam@398 62 (** Using [return] alone allowed us to express a dependency of the [match] result type on the _value_ of the discriminee. What %\index{Gallina terms!in}%[in] adds to our arsenal is a way of expressing a dependency on the _type_ of the discriminee. Specifically, the [n1] in the [in] clause above is a _binding occurrence_ whose scope is the [return] clause.
adamc@84 63
adam@398 64 We may use [in] clauses only to bind names for the arguments of an inductive type family. That is, each [in] clause must be an inductive type family name applied to a sequence of underscores and variable names of the proper length. The positions for _parameters_ to the type family must all be underscores. Parameters are those arguments declared with section variables or with entries to the left of the first colon in an inductive definition. They cannot vary depending on which constructor was used to build the discriminee, so Coq prohibits pointless matches on them. It is those arguments defined in the type to the right of the colon that we may name with [in] clauses.
adamc@84 65
adam@484 66 Our [app] function could be typed in so-called%\index{stratified type systems}% _stratified_ type systems, which avoid true dependency. That is, we could consider the length indices to lists to live in a separate, compile-time-only universe from the lists themselves. Compile-time data may be _erased_ such that we can still execute a program. As an example where erasure would not work, consider an injection function from regular lists to length-indexed lists. Here the run-time computation actually depends on details of the compile-time argument, if we decide that the list to inject can be considered compile-time. More commonly, we think of lists as run-time data. Neither case will work with %\%naive%{}% erasure. (It is not too important to grasp the details of this run-time/compile-time distinction, since Coq's expressive power comes from avoiding such restrictions.) *)
adamc@84 67
adamc@100 68 (* EX: Implement injection from normal lists *)
adamc@100 69
adamc@100 70 (* begin thide *)
adam@454 71 Fixpoint inject (ls : list A) : ilist (length ls) :=
adam@454 72 match ls with
adam@454 73 | nil => Nil
adam@454 74 | h :: t => Cons h (inject t)
adam@454 75 end.
adamc@84 76
adamc@84 77 (** We can define an inverse conversion and prove that it really is an inverse. *)
adamc@84 78
adam@454 79 Fixpoint unject n (ls : ilist n) : list A :=
adam@454 80 match ls with
adam@454 81 | Nil => nil
adam@454 82 | Cons _ h t => h :: unject t
adam@454 83 end.
adamc@84 84
adam@454 85 Theorem inject_inverse : forall ls, unject (inject ls) = ls.
adam@454 86 induction ls; crush.
adam@454 87 Qed.
adamc@100 88 (* end thide *)
adamc@100 89
adam@338 90 (* EX: Implement statically checked "car"/"hd" *)
adamc@84 91
adam@425 92 (** Now let us attempt a function that is surprisingly tricky to write. In ML, the list head function raises an exception when passed an empty list. With length-indexed lists, we can rule out such invalid calls statically, and here is a first attempt at doing so. We write [???] as a placeholder for a term that we do not know how to write, not for any real Coq notation like those introduced two chapters ago.
adamc@84 93 [[
adam@454 94 Definition hd n (ls : ilist (S n)) : A :=
adam@454 95 match ls with
adam@454 96 | Nil => ???
adam@454 97 | Cons _ h _ => h
adam@454 98 end.
adamc@213 99 ]]
adamc@84 100 It is not clear what to write for the [Nil] case, so we are stuck before we even turn our function over to the type checker. We could try omitting the [Nil] case:
adamc@84 101 [[
adam@454 102 Definition hd n (ls : ilist (S n)) : A :=
adam@454 103 match ls with
adam@454 104 | Cons _ h _ => h
adam@454 105 end.
adam@338 106 ]]
adamc@84 107
adam@338 108 <<
adamc@84 109 Error: Non exhaustive pattern-matching: no clause found for pattern Nil
adam@338 110 >>
adamc@84 111
adam@480 112 Unlike in ML, we cannot use inexhaustive pattern matching, because there is no conception of a <<Match>> exception to be thrown. In fact, recent versions of Coq _do_ allow this, by implicit translation to a [match] that considers all constructors; the error message above was generated by an older Coq version. It is educational to discover for ourselves the encoding that the most recent Coq versions use. We might try using an [in] clause somehow.
adamc@84 113
adamc@84 114 [[
adam@454 115 Definition hd n (ls : ilist (S n)) : A :=
adam@454 116 match ls in (ilist (S n)) with
adam@454 117 | Cons _ h _ => h
adam@454 118 end.
adamc@84 119 ]]
adamc@84 120
adam@338 121 <<
adam@338 122 Error: The reference n was not found in the current environment
adam@338 123 >>
adam@338 124
adam@398 125 In this and other cases, we feel like we want [in] clauses with type family arguments that are not variables. Unfortunately, Coq only supports variables in those positions. A completely general mechanism could only be supported with a solution to the problem of higher-order unification%~\cite{HOU}%, which is undecidable. There _are_ useful heuristics for handling non-variable indices which are gradually making their way into Coq, but we will spend some time in this and the next few chapters on effective pattern matching on dependent types using only the primitive [match] annotations.
adamc@84 126
adamc@84 127 Our final, working attempt at [hd] uses an auxiliary function and a surprising [return] annotation. *)
adamc@84 128
adamc@100 129 (* begin thide *)
adam@454 130 Definition hd' n (ls : ilist n) :=
adam@454 131 match ls in (ilist n) return (match n with O => unit | S _ => A end) with
adam@454 132 | Nil => tt
adam@454 133 | Cons _ h _ => h
adam@454 134 end.
adamc@84 135
adam@454 136 Check hd'.
adam@283 137 (** %\vspace{-.15in}% [[
adam@283 138 hd'
adam@283 139 : forall n : nat, ilist n -> match n with
adam@283 140 | 0 => unit
adam@283 141 | S _ => A
adam@283 142 end
adam@302 143 ]]
adam@302 144 *)
adam@283 145
adam@454 146 Definition hd n (ls : ilist (S n)) : A := hd' ls.
adamc@100 147 (* end thide *)
adamc@84 148
adam@338 149 End ilist.
adam@338 150
adamc@84 151 (** We annotate our main [match] with a type that is itself a [match]. We write that the function [hd'] returns [unit] when the list is empty and returns the carried type [A] in all other cases. In the definition of [hd], we just call [hd']. Because the index of [ls] is known to be nonzero, the type checker reduces the [match] in the type of [hd'] to [A]. *)
adamc@84 152
adamc@85 153
adam@371 154 (** * The One Rule of Dependent Pattern Matching in Coq *)
adam@371 155
adam@405 156 (** The rest of this chapter will demonstrate a few other elegant applications of dependent types in Coq. Readers encountering such ideas for the first time often feel overwhelmed, concluding that there is some magic at work whereby Coq sometimes solves the halting problem for the programmer and sometimes does not, applying automated program understanding in a way far beyond what is found in conventional languages. The point of this section is to cut off that sort of thinking right now! Dependent type-checking in Coq follows just a few algorithmic rules. Chapters 10 and 12 introduce many of those rules more formally, and the main additional rule is centered on%\index{dependent pattern matching}% _dependent pattern matching_ of the kind we met in the previous section.
adam@371 157
adam@405 158 A dependent pattern match is a [match] expression where the type of the overall [match] is a function of the value and/or the type of the%\index{discriminee}% _discriminee_, the value being matched on. In other words, the [match] type _depends_ on the discriminee.
adam@371 159
adam@398 160 When exactly will Coq accept a dependent pattern match as well-typed? Some other dependently typed languages employ fancy decision procedures to determine when programs satisfy their very expressive types. The situation in Coq is just the opposite. Only very straightforward symbolic rules are applied. Such a design choice has its drawbacks, as it forces programmers to do more work to convince the type checker of program validity. However, the great advantage of a simple type checking algorithm is that its action on _invalid_ programs is easier to understand!
adam@371 161
adam@371 162 We come now to the one rule of dependent pattern matching in Coq. A general dependent pattern match assumes this form (with unnecessary parentheses included to make the syntax easier to parse):
adam@371 163 [[
adam@480 164 match E as y in (T x1 ... xn) return U with
adam@371 165 | C z1 ... zm => B
adam@371 166 | ...
adam@371 167 end
adam@371 168 ]]
adam@371 169
adam@480 170 The discriminee is a term [E], a value in some inductive type family [T], which takes [n] arguments. An %\index{as clause}%[as] clause binds the name [y] to refer to the discriminee [E]. An %\index{in clause}%[in] clause binds an explicit name [xi] for the [i]th argument passed to [T] in the type of [E].
adam@371 171
adam@480 172 We bind these new variables [y] and [xi] so that they may be referred to in [U], a type given in the %\index{return clause}%[return] clause. The overall type of the [match] will be [U], with [E] substituted for [y], and with each [xi] substituted by the actual argument appearing in that position within [E]'s type.
adam@371 173
adam@371 174 In general, each case of a [match] may have a pattern built up in several layers from the constructors of various inductive type families. To keep this exposition simple, we will focus on patterns that are just single applications of inductive type constructors to lists of variables. Coq actually compiles the more general kind of pattern matching into this more restricted kind automatically, so understanding the typing of [match] requires understanding the typing of [match]es lowered to match one constructor at a time.
adam@371 175
adam@371 176 The last piece of the typing rule tells how to type-check a [match] case. A generic constructor application [C z1 ... zm] has some type [T x1' ... xn'], an application of the type family used in [E]'s type, probably with occurrences of the [zi] variables. From here, a simple recipe determines what type we will require for the case body [B]. The type of [B] should be [U] with the following two substitutions applied: we replace [y] (the [as] clause variable) with [C z1 ... zm], and we replace each [xi] (the [in] clause variables) with [xi']. In other words, we specialize the result type based on what we learn based on which pattern has matched the discriminee.
adam@371 177
adam@371 178 This is an exhaustive description of the ways to specify how to take advantage of which pattern has matched! No other mechanisms come into play. For instance, there is no way to specify that the types of certain free variables should be refined based on which pattern has matched. In the rest of the book, we will learn design patterns for achieving similar effects, where each technique leads to an encoding only in terms of [in], [as], and [return] clauses.
adam@371 179
adam@425 180 A few details have been omitted above. In Chapter 3, we learned that inductive type families may have both%\index{parameters}% _parameters_ and regular arguments. Within an [in] clause, a parameter position must have the wildcard [_] written, instead of a variable. (In general, Coq uses wildcard [_]'s either to indicate pattern variables that will not be mentioned again or to indicate positions where we would like type inference to infer the appropriate terms.) Furthermore, recent Coq versions are adding more and more heuristics to infer dependent [match] annotations in certain conditions. The general annotation inference problem is undecidable, so there will always be serious limitations on how much work these heuristics can do. When in doubt about why a particular dependent [match] is failing to type-check, add an explicit [return] annotation! At that point, the mechanical rule sketched in this section will provide a complete account of "what the type checker is thinking." Be sure to avoid the common pitfall of writing a [return] annotation that does not mention any variables bound by [in] or [as]; such a [match] will never refine typing requirements based on which pattern has matched. (One simple exception to this rule is that, when the discriminee is a variable, that same variable may be treated as if it were repeated as an [as] clause.) *)
adam@371 181
adam@371 182
adamc@85 183 (** * A Tagless Interpreter *)
adamc@85 184
adam@405 185 (** A favorite example for motivating the power of functional programming is implementation of a simple expression language interpreter. In ML and Haskell, such interpreters are often implemented using an algebraic datatype of values, where at many points it is checked that a value was built with the right constructor of the value type. With dependent types, we can implement a%\index{tagless interpreters}% _tagless_ interpreter that both removes this source of runtime inefficiency and gives us more confidence that our implementation is correct. *)
adamc@85 186
adamc@85 187 Inductive type : Set :=
adamc@85 188 | Nat : type
adamc@85 189 | Bool : type
adamc@85 190 | Prod : type -> type -> type.
adamc@85 191
adamc@85 192 Inductive exp : type -> Set :=
adamc@85 193 | NConst : nat -> exp Nat
adamc@85 194 | Plus : exp Nat -> exp Nat -> exp Nat
adamc@85 195 | Eq : exp Nat -> exp Nat -> exp Bool
adamc@85 196
adamc@85 197 | BConst : bool -> exp Bool
adamc@85 198 | And : exp Bool -> exp Bool -> exp Bool
adamc@85 199 | If : forall t, exp Bool -> exp t -> exp t -> exp t
adamc@85 200
adamc@85 201 | Pair : forall t1 t2, exp t1 -> exp t2 -> exp (Prod t1 t2)
adamc@85 202 | Fst : forall t1 t2, exp (Prod t1 t2) -> exp t1
adamc@85 203 | Snd : forall t1 t2, exp (Prod t1 t2) -> exp t2.
adamc@85 204
adam@448 205 (** We have a standard algebraic datatype [type], defining a type language of naturals, Booleans, and product (pair) types. Then we have the indexed inductive type [exp], where the argument to [exp] tells us the encoded type of an expression. In effect, we are defining the typing rules for expressions simultaneously with the syntax.
adamc@85 206
adam@398 207 We can give types and expressions semantics in a new style, based critically on the chance for _type-level computation_. *)
adamc@85 208
adamc@85 209 Fixpoint typeDenote (t : type) : Set :=
adamc@85 210 match t with
adamc@85 211 | Nat => nat
adamc@85 212 | Bool => bool
adamc@85 213 | Prod t1 t2 => typeDenote t1 * typeDenote t2
adamc@85 214 end%type.
adamc@85 215
adam@465 216 (** The [typeDenote] function compiles types of our object language into "native" Coq types. It is deceptively easy to implement. The only new thing we see is the [%]%\coqdocvar{%#<tt>#type#</tt>#%}% annotation, which tells Coq to parse the [match] expression using the notations associated with types. Without this annotation, the [*] would be interpreted as multiplication on naturals, rather than as the product type constructor. The token %\coqdocvar{%#<tt>#type#</tt>#%}% is one example of an identifier bound to a%\index{notation scope delimiter}% _notation scope delimiter_. In this book, we will not go into more detail on notation scopes, but the Coq manual can be consulted for more information.
adamc@85 217
adamc@85 218 We can define a function [expDenote] that is typed in terms of [typeDenote]. *)
adamc@85 219
adamc@213 220 Fixpoint expDenote t (e : exp t) : typeDenote t :=
adamc@213 221 match e with
adamc@85 222 | NConst n => n
adamc@85 223 | Plus e1 e2 => expDenote e1 + expDenote e2
adamc@85 224 | Eq e1 e2 => if eq_nat_dec (expDenote e1) (expDenote e2) then true else false
adamc@85 225
adamc@85 226 | BConst b => b
adamc@85 227 | And e1 e2 => expDenote e1 && expDenote e2
adamc@85 228 | If _ e' e1 e2 => if expDenote e' then expDenote e1 else expDenote e2
adamc@85 229
adamc@85 230 | Pair _ _ e1 e2 => (expDenote e1, expDenote e2)
adamc@85 231 | Fst _ _ e' => fst (expDenote e')
adamc@85 232 | Snd _ _ e' => snd (expDenote e')
adamc@85 233 end.
adamc@85 234
adam@437 235 (* begin hide *)
adam@437 236 (* begin thide *)
adam@437 237 Definition sumboool := sumbool.
adam@437 238 (* end thide *)
adam@437 239 (* end hide *)
adam@437 240
adam@448 241 (** Despite the fancy type, the function definition is routine. In fact, it is less complicated than what we would write in ML or Haskell 98, since we do not need to worry about pushing final values in and out of an algebraic datatype. The only unusual thing is the use of an expression of the form [if E then true else false] in the [Eq] case. Remember that [eq_nat_dec] has a rich dependent type, rather than a simple Boolean type. Coq's native [if] is overloaded to work on a test of any two-constructor type, so we can use [if] to build a simple Boolean from the [sumbool] that [eq_nat_dec] returns.
adamc@85 242
adamc@85 243 We can implement our old favorite, a constant folding function, and prove it correct. It will be useful to write a function [pairOut] that checks if an [exp] of [Prod] type is a pair, returning its two components if so. Unsurprisingly, a first attempt leads to a type error.
adamc@85 244 [[
adamc@85 245 Definition pairOut t1 t2 (e : exp (Prod t1 t2)) : option (exp t1 * exp t2) :=
adamc@85 246 match e in (exp (Prod t1 t2)) return option (exp t1 * exp t2) with
adamc@85 247 | Pair _ _ e1 e2 => Some (e1, e2)
adamc@85 248 | _ => None
adamc@85 249 end.
adam@338 250 ]]
adamc@85 251
adam@338 252 <<
adamc@85 253 Error: The reference t2 was not found in the current environment
adam@338 254 >>
adamc@85 255
adamc@85 256 We run again into the problem of not being able to specify non-variable arguments in [in] clauses. The problem would just be hopeless without a use of an [in] clause, though, since the result type of the [match] depends on an argument to [exp]. Our solution will be to use a more general type, as we did for [hd]. First, we define a type-valued function to use in assigning a type to [pairOut]. *)
adamc@85 257
adamc@100 258 (* EX: Define a function [pairOut : forall t1 t2, exp (Prod t1 t2) -> option (exp t1 * exp t2)] *)
adamc@100 259
adamc@100 260 (* begin thide *)
adam@499 261 Definition pairOutType (t : type) := option (match t with
adam@499 262 | Prod t1 t2 => exp t1 * exp t2
adam@499 263 | _ => unit
adam@499 264 end).
adamc@85 265
adam@499 266 (** When passed a type that is a product, [pairOutType] returns our final desired type. On any other input type, [pairOutType] returns the harmless [option unit], since we do not care about extracting components of non-pairs. Now [pairOut] is easy to write. *)
adamc@85 267
adamc@85 268 Definition pairOut t (e : exp t) :=
adamc@85 269 match e in (exp t) return (pairOutType t) with
adamc@85 270 | Pair _ _ e1 e2 => Some (e1, e2)
adam@499 271 | _ => None
adamc@85 272 end.
adamc@100 273 (* end thide *)
adamc@85 274
adam@499 275 (** With [pairOut] available, we can write [cfold] in a straightforward way. There are really no surprises beyond that Coq verifies that this code has such an expressive type, given the small annotation burden. In some places, we see that Coq's [match] annotation inference is too smart for its own good, and we have to turn that inference off with explicit [return] clauses. *)
adamc@85 276
adamc@204 277 Fixpoint cfold t (e : exp t) : exp t :=
adamc@204 278 match e with
adamc@85 279 | NConst n => NConst n
adamc@85 280 | Plus e1 e2 =>
adamc@85 281 let e1' := cfold e1 in
adamc@85 282 let e2' := cfold e2 in
adam@417 283 match e1', e2' return exp Nat with
adamc@85 284 | NConst n1, NConst n2 => NConst (n1 + n2)
adamc@85 285 | _, _ => Plus e1' e2'
adamc@85 286 end
adamc@85 287 | Eq e1 e2 =>
adamc@85 288 let e1' := cfold e1 in
adamc@85 289 let e2' := cfold e2 in
adam@417 290 match e1', e2' return exp Bool with
adamc@85 291 | NConst n1, NConst n2 => BConst (if eq_nat_dec n1 n2 then true else false)
adamc@85 292 | _, _ => Eq e1' e2'
adamc@85 293 end
adamc@85 294
adamc@85 295 | BConst b => BConst b
adamc@85 296 | And e1 e2 =>
adamc@85 297 let e1' := cfold e1 in
adamc@85 298 let e2' := cfold e2 in
adam@417 299 match e1', e2' return exp Bool with
adamc@85 300 | BConst b1, BConst b2 => BConst (b1 && b2)
adamc@85 301 | _, _ => And e1' e2'
adamc@85 302 end
adamc@85 303 | If _ e e1 e2 =>
adamc@85 304 let e' := cfold e in
adamc@85 305 match e' with
adamc@85 306 | BConst true => cfold e1
adamc@85 307 | BConst false => cfold e2
adamc@85 308 | _ => If e' (cfold e1) (cfold e2)
adamc@85 309 end
adamc@85 310
adamc@85 311 | Pair _ _ e1 e2 => Pair (cfold e1) (cfold e2)
adamc@85 312 | Fst _ _ e =>
adamc@85 313 let e' := cfold e in
adamc@85 314 match pairOut e' with
adamc@85 315 | Some p => fst p
adamc@85 316 | None => Fst e'
adamc@85 317 end
adamc@85 318 | Snd _ _ e =>
adamc@85 319 let e' := cfold e in
adamc@85 320 match pairOut e' with
adamc@85 321 | Some p => snd p
adamc@85 322 | None => Snd e'
adamc@85 323 end
adamc@85 324 end.
adamc@85 325
adamc@85 326 (** The correctness theorem for [cfold] turns out to be easy to prove, once we get over one serious hurdle. *)
adamc@85 327
adamc@85 328 Theorem cfold_correct : forall t (e : exp t), expDenote e = expDenote (cfold e).
adamc@100 329 (* begin thide *)
adamc@85 330 induction e; crush.
adamc@85 331
adamc@85 332 (** The first remaining subgoal is:
adamc@85 333
adamc@85 334 [[
adamc@85 335 expDenote (cfold e1) + expDenote (cfold e2) =
adamc@85 336 expDenote
adamc@85 337 match cfold e1 with
adamc@85 338 | NConst n1 =>
adamc@85 339 match cfold e2 with
adamc@85 340 | NConst n2 => NConst (n1 + n2)
adamc@85 341 | Plus _ _ => Plus (cfold e1) (cfold e2)
adamc@85 342 | Eq _ _ => Plus (cfold e1) (cfold e2)
adamc@85 343 | BConst _ => Plus (cfold e1) (cfold e2)
adamc@85 344 | And _ _ => Plus (cfold e1) (cfold e2)
adamc@85 345 | If _ _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 346 | Pair _ _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 347 | Fst _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 348 | Snd _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 349 end
adamc@85 350 | Plus _ _ => Plus (cfold e1) (cfold e2)
adamc@85 351 | Eq _ _ => Plus (cfold e1) (cfold e2)
adamc@85 352 | BConst _ => Plus (cfold e1) (cfold e2)
adamc@85 353 | And _ _ => Plus (cfold e1) (cfold e2)
adamc@85 354 | If _ _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 355 | Pair _ _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 356 | Fst _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 357 | Snd _ _ _ => Plus (cfold e1) (cfold e2)
adamc@85 358 end
adamc@213 359
adamc@85 360 ]]
adamc@85 361
adam@454 362 We would like to do a case analysis on [cfold e1], and we attempt to do so in the way that has worked so far.
adamc@85 363 [[
adamc@85 364 destruct (cfold e1).
adam@338 365 ]]
adamc@85 366
adam@338 367 <<
adamc@85 368 User error: e1 is used in hypothesis e
adam@338 369 >>
adamc@85 370
adamc@85 371 Coq gives us another cryptic error message. Like so many others, this one basically means that Coq is not able to build some proof about dependent types. It is hard to generate helpful and specific error messages for problems like this, since that would require some kind of understanding of the dependency structure of a piece of code. We will encounter many examples of case-specific tricks for recovering from errors like this one.
adamc@85 372
adam@499 373 For our current proof, we can use a tactic [dep_destruct]%\index{tactics!dep\_destruct}% defined in the book's [CpdtTactics] module. General elimination/inversion of dependently typed hypotheses is undecidable, as witnessed by a simple reduction from the known-undecidable problem of higher-order unification, which has come up a few times already. The tactic [dep_destruct] makes a best effort to handle some common cases, relying upon the more primitive %\index{tactics!dependent destruction}%[dependent destruction] tactic that comes with Coq. In a future chapter, we will learn about the explicit manipulation of equality proofs that is behind [dependent destruction]'s implementation, but for now, we treat it as a useful black box. (In Chapter 12, we will also see how [dependent destruction] forces us to make a larger philosophical commitment about our logic than we might like, and we will see some workarounds.) *)
adamc@85 374
adamc@85 375 dep_destruct (cfold e1).
adamc@85 376
adamc@85 377 (** This successfully breaks the subgoal into 5 new subgoals, one for each constructor of [exp] that could produce an [exp Nat]. Note that [dep_destruct] is successful in ruling out the other cases automatically, in effect automating some of the work that we have done manually in implementing functions like [hd] and [pairOut].
adamc@85 378
adam@480 379 This is the only new trick we need to learn to complete the proof. We can back up and give a short, automated proof (which again is safe to skip and uses Ltac features not introduced yet). *)
adamc@85 380
adamc@85 381 Restart.
adamc@85 382
adamc@85 383 induction e; crush;
adamc@85 384 repeat (match goal with
adam@405 385 | [ |- context[match cfold ?E with NConst _ => _ | _ => _ end] ] =>
adamc@213 386 dep_destruct (cfold E)
adamc@213 387 | [ |- context[match pairOut (cfold ?E) with Some _ => _
adamc@213 388 | None => _ end] ] =>
adamc@213 389 dep_destruct (cfold E)
adamc@85 390 | [ |- (if ?E then _ else _) = _ ] => destruct E
adamc@85 391 end; crush).
adamc@85 392 Qed.
adamc@100 393 (* end thide *)
adamc@86 394
adam@405 395 (** With this example, we get a first taste of how to build automated proofs that adapt automatically to changes in function definitions. *)
adam@405 396
adamc@86 397
adam@338 398 (** * Dependently Typed Red-Black Trees *)
adamc@94 399
adam@475 400 (** Red-black trees are a favorite purely functional data structure with an interesting invariant. We can use dependent types to guarantee that operations on red-black trees preserve the invariant. For simplicity, we specialize our red-black trees to represent sets of [nat]s. *)
adamc@100 401
adamc@94 402 Inductive color : Set := Red | Black.
adamc@94 403
adamc@94 404 Inductive rbtree : color -> nat -> Set :=
adamc@94 405 | Leaf : rbtree Black 0
adamc@214 406 | RedNode : forall n, rbtree Black n -> nat -> rbtree Black n -> rbtree Red n
adamc@94 407 | BlackNode : forall c1 c2 n, rbtree c1 n -> nat -> rbtree c2 n -> rbtree Black (S n).
adamc@94 408
adam@476 409 (** A value of type [rbtree c d] is a red-black tree whose root has color [c] and that has black depth [d]. The latter property means that there are exactly [d] black-colored nodes on any path from the root to a leaf. *)
adamc@214 410
adamc@214 411 (** At first, it can be unclear that this choice of type indices tracks any useful property. To convince ourselves, we will prove that every red-black tree is balanced. We will phrase our theorem in terms of a depth calculating function that ignores the extra information in the types. It will be useful to parameterize this function over a combining operation, so that we can re-use the same code to calculate the minimum or maximum height among all paths from root to leaf. *)
adamc@214 412
adamc@100 413 (* EX: Prove that every [rbtree] is balanced. *)
adamc@100 414
adamc@100 415 (* begin thide *)
adamc@95 416 Require Import Max Min.
adamc@95 417
adamc@95 418 Section depth.
adamc@95 419 Variable f : nat -> nat -> nat.
adamc@95 420
adamc@214 421 Fixpoint depth c n (t : rbtree c n) : nat :=
adamc@95 422 match t with
adamc@95 423 | Leaf => 0
adamc@95 424 | RedNode _ t1 _ t2 => S (f (depth t1) (depth t2))
adamc@95 425 | BlackNode _ _ _ t1 _ t2 => S (f (depth t1) (depth t2))
adamc@95 426 end.
adamc@95 427 End depth.
adamc@95 428
adam@338 429 (** Our proof of balanced-ness decomposes naturally into a lower bound and an upper bound. We prove the lower bound first. Unsurprisingly, a tree's black depth provides such a bound on the minimum path length. We use the richly typed procedure [min_dec] to do case analysis on whether [min X Y] equals [X] or [Y]. *)
adamc@214 430
adam@283 431 Check min_dec.
adam@283 432 (** %\vspace{-.15in}% [[
adam@283 433 min_dec
adam@283 434 : forall n m : nat, {min n m = n} + {min n m = m}
adam@302 435 ]]
adam@302 436 *)
adam@283 437
adamc@95 438 Theorem depth_min : forall c n (t : rbtree c n), depth min t >= n.
adamc@95 439 induction t; crush;
adamc@95 440 match goal with
adamc@95 441 | [ |- context[min ?X ?Y] ] => destruct (min_dec X Y)
adamc@95 442 end; crush.
adamc@95 443 Qed.
adamc@95 444
adamc@214 445 (** There is an analogous upper-bound theorem based on black depth. Unfortunately, a symmetric proof script does not suffice to establish it. *)
adamc@214 446
adamc@214 447 Theorem depth_max : forall c n (t : rbtree c n), depth max t <= 2 * n + 1.
adamc@214 448 induction t; crush;
adamc@214 449 match goal with
adamc@214 450 | [ |- context[max ?X ?Y] ] => destruct (max_dec X Y)
adamc@214 451 end; crush.
adamc@214 452
adamc@214 453 (** Two subgoals remain. One of them is: [[
adamc@214 454 n : nat
adamc@214 455 t1 : rbtree Black n
adamc@214 456 n0 : nat
adamc@214 457 t2 : rbtree Black n
adamc@214 458 IHt1 : depth max t1 <= n + (n + 0) + 1
adamc@214 459 IHt2 : depth max t2 <= n + (n + 0) + 1
adamc@214 460 e : max (depth max t1) (depth max t2) = depth max t1
adamc@214 461 ============================
adamc@214 462 S (depth max t1) <= n + (n + 0) + 1
adamc@214 463
adamc@214 464 ]]
adamc@214 465
adam@398 466 We see that [IHt1] is _almost_ the fact we need, but it is not quite strong enough. We will need to strengthen our induction hypothesis to get the proof to go through. *)
adamc@214 467
adamc@214 468 Abort.
adamc@214 469
adamc@214 470 (** In particular, we prove a lemma that provides a stronger upper bound for trees with black root nodes. We got stuck above in a case about a red root node. Since red nodes have only black children, our IH strengthening will enable us to finish the proof. *)
adamc@214 471
adamc@95 472 Lemma depth_max' : forall c n (t : rbtree c n), match c with
adamc@95 473 | Red => depth max t <= 2 * n + 1
adamc@95 474 | Black => depth max t <= 2 * n
adamc@95 475 end.
adamc@95 476 induction t; crush;
adamc@95 477 match goal with
adamc@95 478 | [ |- context[max ?X ?Y] ] => destruct (max_dec X Y)
adamc@100 479 end; crush;
adamc@100 480 repeat (match goal with
adamc@214 481 | [ H : context[match ?C with Red => _ | Black => _ end] |- _ ] =>
adamc@214 482 destruct C
adamc@100 483 end; crush).
adamc@95 484 Qed.
adamc@95 485
adam@338 486 (** The original theorem follows easily from the lemma. We use the tactic %\index{tactics!generalize}%[generalize pf], which, when [pf] proves the proposition [P], changes the goal from [Q] to [P -> Q]. This transformation is useful because it makes the truth of [P] manifest syntactically, so that automation machinery can rely on [P], even if that machinery is not smart enough to establish [P] on its own. *)
adamc@214 487
adamc@95 488 Theorem depth_max : forall c n (t : rbtree c n), depth max t <= 2 * n + 1.
adamc@95 489 intros; generalize (depth_max' t); destruct c; crush.
adamc@95 490 Qed.
adamc@95 491
adamc@214 492 (** The final balance theorem establishes that the minimum and maximum path lengths of any tree are within a factor of two of each other. *)
adamc@214 493
adamc@95 494 Theorem balanced : forall c n (t : rbtree c n), 2 * depth min t + 1 >= depth max t.
adamc@95 495 intros; generalize (depth_min t); generalize (depth_max t); crush.
adamc@95 496 Qed.
adamc@100 497 (* end thide *)
adamc@95 498
adamc@214 499 (** Now we are ready to implement an example operation on our trees, insertion. Insertion can be thought of as breaking the tree invariants locally but then rebalancing. In particular, in intermediate states we find red nodes that may have red children. The type [rtree] captures the idea of such a node, continuing to track black depth as a type index. *)
adamc@95 500
adamc@94 501 Inductive rtree : nat -> Set :=
adamc@94 502 | RedNode' : forall c1 c2 n, rbtree c1 n -> nat -> rbtree c2 n -> rtree n.
adamc@94 503
adam@338 504 (** Before starting to define [insert], we define predicates capturing when a data value is in the set represented by a normal or possibly invalid tree. *)
adamc@214 505
adamc@96 506 Section present.
adamc@96 507 Variable x : nat.
adamc@96 508
adamc@214 509 Fixpoint present c n (t : rbtree c n) : Prop :=
adamc@96 510 match t with
adamc@96 511 | Leaf => False
adamc@96 512 | RedNode _ a y b => present a \/ x = y \/ present b
adamc@96 513 | BlackNode _ _ _ a y b => present a \/ x = y \/ present b
adamc@96 514 end.
adamc@96 515
adamc@96 516 Definition rpresent n (t : rtree n) : Prop :=
adamc@96 517 match t with
adamc@96 518 | RedNode' _ _ _ a y b => present a \/ x = y \/ present b
adamc@96 519 end.
adamc@96 520 End present.
adamc@96 521
adam@338 522 (** Insertion relies on two balancing operations. It will be useful to give types to these operations using a relative of the subset types from last chapter. While subset types let us pair a value with a proof about that value, here we want to pair a value with another non-proof dependently typed value. The %\index{Gallina terms!sigT}%[sigT] type fills this role. *)
adamc@214 523
adamc@100 524 Locate "{ _ : _ & _ }".
adam@443 525 (** %\vspace{-.15in}%[[
adamc@214 526 Notation Scope
adamc@214 527 "{ x : A & P }" := sigT (fun x : A => P)
adam@302 528 ]]
adam@302 529 *)
adamc@214 530
adamc@100 531 Print sigT.
adam@443 532 (** %\vspace{-.15in}%[[
adamc@214 533 Inductive sigT (A : Type) (P : A -> Type) : Type :=
adamc@214 534 existT : forall x : A, P x -> sigT P
adam@302 535 ]]
adam@302 536 *)
adamc@214 537
adamc@214 538 (** It will be helpful to define a concise notation for the constructor of [sigT]. *)
adamc@100 539
adamc@94 540 Notation "{< x >}" := (existT _ _ x).
adamc@94 541
adamc@214 542 (** Each balance function is used to construct a new tree whose keys include the keys of two input trees, as well as a new key. One of the two input trees may violate the red-black alternation invariant (that is, it has an [rtree] type), while the other tree is known to be valid. Crucially, the two input trees have the same black depth.
adamc@214 543
adam@338 544 A balance operation may return a tree whose root is of either color. Thus, we use a [sigT] type to package the result tree with the color of its root. Here is the definition of the first balance operation, which applies when the possibly invalid [rtree] belongs to the left of the valid [rbtree].
adam@338 545
adam@425 546 A quick word of encouragement: After writing this code, even I do not understand the precise details of how balancing works! I consulted Chris Okasaki's paper "Red-Black Trees in a Functional Setting" %\cite{Okasaki} %and transcribed the code to use dependent types. Luckily, the details are not so important here; types alone will tell us that insertion preserves balanced-ness, and we will prove that insertion produces trees containing the right keys.*)
adamc@214 547
adamc@94 548 Definition balance1 n (a : rtree n) (data : nat) c2 :=
adamc@214 549 match a in rtree n return rbtree c2 n
adamc@214 550 -> { c : color & rbtree c (S n) } with
adam@380 551 | RedNode' _ c0 _ t1 y t2 =>
adam@380 552 match t1 in rbtree c n return rbtree c0 n -> rbtree c2 n
adamc@214 553 -> { c : color & rbtree c (S n) } with
adamc@214 554 | RedNode _ a x b => fun c d =>
adamc@214 555 {<RedNode (BlackNode a x b) y (BlackNode c data d)>}
adamc@94 556 | t1' => fun t2 =>
adam@380 557 match t2 in rbtree c n return rbtree Black n -> rbtree c2 n
adamc@214 558 -> { c : color & rbtree c (S n) } with
adamc@214 559 | RedNode _ b x c => fun a d =>
adamc@214 560 {<RedNode (BlackNode a y b) x (BlackNode c data d)>}
adamc@95 561 | b => fun a t => {<BlackNode (RedNode a y b) data t>}
adamc@94 562 end t1'
adamc@94 563 end t2
adamc@94 564 end.
adamc@94 565
adam@405 566 (** We apply a trick that I call the%\index{convoy pattern}% _convoy pattern_. Recall that [match] annotations only make it possible to describe a dependence of a [match] _result type_ on the discriminee. There is no automatic refinement of the types of free variables. However, it is possible to effect such a refinement by finding a way to encode free variable type dependencies in the [match] result type, so that a [return] clause can express the connection.
adamc@214 567
adam@425 568 In particular, we can extend the [match] to return _functions over the free variables whose types we want to refine_. In the case of [balance1], we only find ourselves wanting to refine the type of one tree variable at a time. We match on one subtree of a node, and we want the type of the other subtree to be refined based on what we learn. We indicate this with a [return] clause starting like [rbtree _ n -> ...], where [n] is bound in an [in] pattern. Such a [match] expression is applied immediately to the "old version" of the variable to be refined, and the type checker is happy.
adamc@214 569
adam@338 570 Here is the symmetric function [balance2], for cases where the possibly invalid tree appears on the right rather than on the left. *)
adamc@214 571
adamc@94 572 Definition balance2 n (a : rtree n) (data : nat) c2 :=
adamc@94 573 match a in rtree n return rbtree c2 n -> { c : color & rbtree c (S n) } with
adam@380 574 | RedNode' _ c0 _ t1 z t2 =>
adam@380 575 match t1 in rbtree c n return rbtree c0 n -> rbtree c2 n
adamc@214 576 -> { c : color & rbtree c (S n) } with
adamc@214 577 | RedNode _ b y c => fun d a =>
adamc@214 578 {<RedNode (BlackNode a data b) y (BlackNode c z d)>}
adamc@94 579 | t1' => fun t2 =>
adam@380 580 match t2 in rbtree c n return rbtree Black n -> rbtree c2 n
adamc@214 581 -> { c : color & rbtree c (S n) } with
adamc@214 582 | RedNode _ c z' d => fun b a =>
adamc@214 583 {<RedNode (BlackNode a data b) z (BlackNode c z' d)>}
adamc@95 584 | b => fun a t => {<BlackNode t data (RedNode a z b)>}
adamc@94 585 end t1'
adamc@94 586 end t2
adamc@94 587 end.
adamc@94 588
adamc@214 589 (** Now we are almost ready to get down to the business of writing an [insert] function. First, we enter a section that declares a variable [x], for the key we want to insert. *)
adamc@214 590
adamc@94 591 Section insert.
adamc@94 592 Variable x : nat.
adamc@94 593
adamc@214 594 (** Most of the work of insertion is done by a helper function [ins], whose return types are expressed using a type-level function [insResult]. *)
adamc@214 595
adamc@94 596 Definition insResult c n :=
adamc@94 597 match c with
adamc@94 598 | Red => rtree n
adamc@94 599 | Black => { c' : color & rbtree c' n }
adamc@94 600 end.
adamc@94 601
adam@338 602 (** That is, inserting into a tree with root color [c] and black depth [n], the variety of tree we get out depends on [c]. If we started with a red root, then we get back a possibly invalid tree of depth [n]. If we started with a black root, we get back a valid tree of depth [n] with a root node of an arbitrary color.
adamc@214 603
adamc@214 604 Here is the definition of [ins]. Again, we do not want to dwell on the functional details. *)
adamc@214 605
adamc@214 606 Fixpoint ins c n (t : rbtree c n) : insResult c n :=
adamc@214 607 match t with
adamc@94 608 | Leaf => {< RedNode Leaf x Leaf >}
adamc@94 609 | RedNode _ a y b =>
adamc@94 610 if le_lt_dec x y
adamc@94 611 then RedNode' (projT2 (ins a)) y b
adamc@94 612 else RedNode' a y (projT2 (ins b))
adamc@94 613 | BlackNode c1 c2 _ a y b =>
adamc@94 614 if le_lt_dec x y
adamc@94 615 then
adamc@94 616 match c1 return insResult c1 _ -> _ with
adamc@94 617 | Red => fun ins_a => balance1 ins_a y b
adamc@94 618 | _ => fun ins_a => {< BlackNode (projT2 ins_a) y b >}
adamc@94 619 end (ins a)
adamc@94 620 else
adamc@94 621 match c2 return insResult c2 _ -> _ with
adamc@94 622 | Red => fun ins_b => balance2 ins_b y a
adamc@94 623 | _ => fun ins_b => {< BlackNode a y (projT2 ins_b) >}
adamc@94 624 end (ins b)
adamc@94 625 end.
adamc@94 626
adam@479 627 (** The one new trick is a variation of the convoy pattern. In each of the last two pattern matches, we want to take advantage of the typing connection between the trees [a] and [b]. We might %\%naive%{}%ly apply the convoy pattern directly on [a] in the first [match] and on [b] in the second. This satisfies the type checker per se, but it does not satisfy the termination checker. Inside each [match], we would be calling [ins] recursively on a locally bound variable. The termination checker is not smart enough to trace the dataflow into that variable, so the checker does not know that this recursive argument is smaller than the original argument. We make this fact clearer by applying the convoy pattern on _the result of a recursive call_, rather than just on that call's argument.
adamc@214 628
adamc@214 629 Finally, we are in the home stretch of our effort to define [insert]. We just need a few more definitions of non-recursive functions. First, we need to give the final characterization of [insert]'s return type. Inserting into a red-rooted tree gives a black-rooted tree where black depth has increased, and inserting into a black-rooted tree gives a tree where black depth has stayed the same and where the root is an arbitrary color. *)
adamc@214 630
adamc@94 631 Definition insertResult c n :=
adamc@94 632 match c with
adamc@94 633 | Red => rbtree Black (S n)
adamc@94 634 | Black => { c' : color & rbtree c' n }
adamc@94 635 end.
adamc@94 636
adamc@214 637 (** A simple clean-up procedure translates [insResult]s into [insertResult]s. *)
adamc@214 638
adamc@97 639 Definition makeRbtree c n : insResult c n -> insertResult c n :=
adamc@214 640 match c with
adamc@94 641 | Red => fun r =>
adamc@214 642 match r with
adamc@94 643 | RedNode' _ _ _ a x b => BlackNode a x b
adamc@94 644 end
adamc@94 645 | Black => fun r => r
adamc@94 646 end.
adamc@94 647
adamc@214 648 (** We modify Coq's default choice of implicit arguments for [makeRbtree], so that we do not need to specify the [c] and [n] arguments explicitly in later calls. *)
adamc@214 649
adamc@97 650 Implicit Arguments makeRbtree [c n].
adamc@94 651
adamc@214 652 (** Finally, we define [insert] as a simple composition of [ins] and [makeRbtree]. *)
adamc@214 653
adamc@94 654 Definition insert c n (t : rbtree c n) : insertResult c n :=
adamc@97 655 makeRbtree (ins t).
adamc@94 656
adamc@214 657 (** As we noted earlier, the type of [insert] guarantees that it outputs balanced trees whose depths have not increased too much. We also want to know that [insert] operates correctly on trees interpreted as finite sets, so we finish this section with a proof of that fact. *)
adamc@214 658
adamc@95 659 Section present.
adamc@95 660 Variable z : nat.
adamc@95 661
adamc@214 662 (** The variable [z] stands for an arbitrary key. We will reason about [z]'s presence in particular trees. As usual, outside the section the theorems we prove will quantify over all possible keys, giving us the facts we wanted.
adamc@214 663
adam@367 664 We start by proving the correctness of the balance operations. It is useful to define a custom tactic [present_balance] that encapsulates the reasoning common to the two proofs. We use the keyword %\index{Vernacular commands!Ltac}%[Ltac] to assign a name to a proof script. This particular script just iterates between [crush] and identification of a tree that is being pattern-matched on and should be destructed. *)
adamc@214 665
adamc@98 666 Ltac present_balance :=
adamc@98 667 crush;
adamc@98 668 repeat (match goal with
adam@425 669 | [ _ : context[match ?T with Leaf => _ | _ => _ end] |- _ ] =>
adam@425 670 dep_destruct T
adam@405 671 | [ |- context[match ?T with Leaf => _ | _ => _ end] ] => dep_destruct T
adamc@98 672 end; crush).
adamc@98 673
adamc@214 674 (** The balance correctness theorems are simple first-order logic equivalences, where we use the function [projT2] to project the payload of a [sigT] value. *)
adamc@214 675
adam@294 676 Lemma present_balance1 : forall n (a : rtree n) (y : nat) c2 (b : rbtree c2 n),
adamc@95 677 present z (projT2 (balance1 a y b))
adamc@95 678 <-> rpresent z a \/ z = y \/ present z b.
adamc@98 679 destruct a; present_balance.
adamc@95 680 Qed.
adamc@95 681
adamc@213 682 Lemma present_balance2 : forall n (a : rtree n) (y : nat) c2 (b : rbtree c2 n),
adamc@95 683 present z (projT2 (balance2 a y b))
adamc@95 684 <-> rpresent z a \/ z = y \/ present z b.
adamc@98 685 destruct a; present_balance.
adamc@95 686 Qed.
adamc@95 687
adamc@214 688 (** To state the theorem for [ins], it is useful to define a new type-level function, since [ins] returns different result types based on the type indices passed to it. Recall that [x] is the section variable standing for the key we are inserting. *)
adamc@214 689
adamc@95 690 Definition present_insResult c n :=
adamc@95 691 match c return (rbtree c n -> insResult c n -> Prop) with
adamc@95 692 | Red => fun t r => rpresent z r <-> z = x \/ present z t
adamc@95 693 | Black => fun t r => present z (projT2 r) <-> z = x \/ present z t
adamc@95 694 end.
adamc@95 695
adamc@214 696 (** Now the statement and proof of the [ins] correctness theorem are straightforward, if verbose. We proceed by induction on the structure of a tree, followed by finding case analysis opportunities on expressions we see being analyzed in [if] or [match] expressions. After that, we pattern-match to find opportunities to use the theorems we proved about balancing. Finally, we identify two variables that are asserted by some hypothesis to be equal, and we use that hypothesis to replace one variable with the other everywhere. *)
adamc@214 697
adamc@95 698 Theorem present_ins : forall c n (t : rbtree c n),
adamc@95 699 present_insResult t (ins t).
adamc@95 700 induction t; crush;
adamc@95 701 repeat (match goal with
adam@338 702 | [ _ : context[if ?E then _ else _] |- _ ] => destruct E
adamc@95 703 | [ |- context[if ?E then _ else _] ] => destruct E
adam@338 704 | [ _ : context[match ?C with Red => _ | Black => _ end]
adamc@214 705 |- _ ] => destruct C
adamc@95 706 end; crush);
adamc@95 707 try match goal with
adam@338 708 | [ _ : context[balance1 ?A ?B ?C] |- _ ] =>
adamc@95 709 generalize (present_balance1 A B C)
adamc@95 710 end;
adamc@95 711 try match goal with
adam@338 712 | [ _ : context[balance2 ?A ?B ?C] |- _ ] =>
adamc@95 713 generalize (present_balance2 A B C)
adamc@95 714 end;
adamc@95 715 try match goal with
adamc@95 716 | [ |- context[balance1 ?A ?B ?C] ] =>
adamc@95 717 generalize (present_balance1 A B C)
adamc@95 718 end;
adamc@95 719 try match goal with
adamc@95 720 | [ |- context[balance2 ?A ?B ?C] ] =>
adamc@95 721 generalize (present_balance2 A B C)
adamc@95 722 end;
adamc@214 723 crush;
adamc@95 724 match goal with
adamc@95 725 | [ z : nat, x : nat |- _ ] =>
adamc@95 726 match goal with
adamc@95 727 | [ H : z = x |- _ ] => rewrite H in *; clear H
adamc@95 728 end
adamc@95 729 end;
adamc@95 730 tauto.
adamc@95 731 Qed.
adamc@95 732
adamc@214 733 (** The hard work is done. The most readable way to state correctness of [insert] involves splitting the property into two color-specific theorems. We write a tactic to encapsulate the reasoning steps that work to establish both facts. *)
adamc@214 734
adamc@213 735 Ltac present_insert :=
adamc@213 736 unfold insert; intros n t; inversion t;
adamc@97 737 generalize (present_ins t); simpl;
adamc@97 738 dep_destruct (ins t); tauto.
adamc@97 739
adamc@95 740 Theorem present_insert_Red : forall n (t : rbtree Red n),
adamc@95 741 present z (insert t)
adamc@95 742 <-> (z = x \/ present z t).
adamc@213 743 present_insert.
adamc@95 744 Qed.
adamc@95 745
adamc@95 746 Theorem present_insert_Black : forall n (t : rbtree Black n),
adamc@95 747 present z (projT2 (insert t))
adamc@95 748 <-> (z = x \/ present z t).
adamc@213 749 present_insert.
adamc@95 750 Qed.
adamc@95 751 End present.
adamc@94 752 End insert.
adamc@94 753
adam@454 754 (** We can generate executable OCaml code with the command %\index{Vernacular commands!Recursive Extraction}%[Recursive Extraction insert], which also automatically outputs the OCaml versions of all of [insert]'s dependencies. In our previous extractions, we wound up with clean OCaml code. Here, we find uses of %\index{Obj.magic}%<<Obj.magic>>, OCaml's unsafe cast operator for tweaking the apparent type of an expression in an arbitrary way. Casts appear for this example because the return type of [insert] depends on the _value_ of the function's argument, a pattern that OCaml cannot handle. Since Coq's type system is much more expressive than OCaml's, such casts are unavoidable in general. Since the OCaml type-checker is no longer checking full safety of programs, we must rely on Coq's extractor to use casts only in provably safe ways. *)
adam@338 755
adam@338 756 (* begin hide *)
adam@338 757 Recursive Extraction insert.
adam@338 758 (* end hide *)
adam@283 759
adamc@94 760
adamc@86 761 (** * A Certified Regular Expression Matcher *)
adamc@86 762
adamc@93 763 (** Another interesting example is regular expressions with dependent types that express which predicates over strings particular regexps implement. We can then assign a dependent type to a regular expression matching function, guaranteeing that it always decides the string property that we expect it to decide.
adamc@93 764
adam@425 765 Before defining the syntax of expressions, it is helpful to define an inductive type capturing the meaning of the Kleene star. That is, a string [s] matches regular expression [star e] if and only if [s] can be decomposed into a sequence of substrings that all match [e]. We use Coq's string support, which comes through a combination of the [String] library and some parsing notations built into Coq. Operators like [++] and functions like [length] that we know from lists are defined again for strings. Notation scopes help us control which versions we want to use in particular contexts.%\index{Vernacular commands!Open Scope}% *)
adamc@93 766
adamc@86 767 Require Import Ascii String.
adamc@86 768 Open Scope string_scope.
adamc@86 769
adamc@91 770 Section star.
adamc@91 771 Variable P : string -> Prop.
adamc@91 772
adamc@91 773 Inductive star : string -> Prop :=
adamc@91 774 | Empty : star ""
adamc@91 775 | Iter : forall s1 s2,
adamc@91 776 P s1
adamc@91 777 -> star s2
adamc@91 778 -> star (s1 ++ s2).
adamc@91 779 End star.
adamc@91 780
adam@480 781 (** Now we can make our first attempt at defining a [regexp] type that is indexed by predicates on strings, such that the index of a [regexp] tells us which language (string predicate) it recognizes. Here is a reasonable-looking definition that is restricted to constant characters and concatenation. We use the constructor [String], which is the analogue of list cons for the type [string], where [""] is like list nil.
adamc@93 782 [[
adamc@93 783 Inductive regexp : (string -> Prop) -> Set :=
adamc@93 784 | Char : forall ch : ascii,
adamc@93 785 regexp (fun s => s = String ch "")
adamc@93 786 | Concat : forall (P1 P2 : string -> Prop) (r1 : regexp P1) (r2 : regexp P2),
adamc@93 787 regexp (fun s => exists s1, exists s2, s = s1 ++ s2 /\ P1 s1 /\ P2 s2).
adamc@93 788 ]]
adamc@93 789
adam@338 790 <<
adam@338 791 User error: Large non-propositional inductive types must be in Type
adam@338 792 >>
adam@338 793
adam@454 794 What is a %\index{large inductive types}%large inductive type? In Coq, it is an inductive type that has a constructor that quantifies over some type of type [Type]. We have not worked with [Type] very much to this point. Every term of CIC has a type, including [Set] and [Prop], which are assigned type [Type]. The type [string -> Prop] from the failed definition also has type [Type].
adamc@93 795
adamc@93 796 It turns out that allowing large inductive types in [Set] leads to contradictions when combined with certain kinds of classical logic reasoning. Thus, by default, such types are ruled out. There is a simple fix for our [regexp] definition, which is to place our new type in [Type]. While fixing the problem, we also expand the list of constructors to cover the remaining regular expression operators. *)
adamc@93 797
adamc@89 798 Inductive regexp : (string -> Prop) -> Type :=
adamc@86 799 | Char : forall ch : ascii,
adamc@86 800 regexp (fun s => s = String ch "")
adamc@86 801 | Concat : forall P1 P2 (r1 : regexp P1) (r2 : regexp P2),
adamc@87 802 regexp (fun s => exists s1, exists s2, s = s1 ++ s2 /\ P1 s1 /\ P2 s2)
adamc@87 803 | Or : forall P1 P2 (r1 : regexp P1) (r2 : regexp P2),
adamc@91 804 regexp (fun s => P1 s \/ P2 s)
adamc@91 805 | Star : forall P (r : regexp P),
adamc@91 806 regexp (star P).
adamc@86 807
adam@425 808 (** Many theorems about strings are useful for implementing a certified regexp matcher, and few of them are in the [String] library. The book source includes statements, proofs, and hint commands for a handful of such omitted theorems. Since they are orthogonal to our use of dependent types, we hide them in the rendered versions of this book. *)
adamc@93 809
adamc@93 810 (* begin hide *)
adamc@86 811 Open Scope specif_scope.
adamc@86 812
adamc@86 813 Lemma length_emp : length "" <= 0.
adamc@86 814 crush.
adamc@86 815 Qed.
adamc@86 816
adamc@86 817 Lemma append_emp : forall s, s = "" ++ s.
adamc@86 818 crush.
adamc@86 819 Qed.
adamc@86 820
adamc@86 821 Ltac substring :=
adamc@86 822 crush;
adamc@86 823 repeat match goal with
adamc@86 824 | [ |- context[match ?N with O => _ | S _ => _ end] ] => destruct N; crush
adamc@86 825 end.
adamc@86 826
adamc@86 827 Lemma substring_le : forall s n m,
adamc@86 828 length (substring n m s) <= m.
adamc@86 829 induction s; substring.
adamc@86 830 Qed.
adamc@86 831
adamc@86 832 Lemma substring_all : forall s,
adamc@86 833 substring 0 (length s) s = s.
adamc@86 834 induction s; substring.
adamc@86 835 Qed.
adamc@86 836
adamc@86 837 Lemma substring_none : forall s n,
adamc@93 838 substring n 0 s = "".
adamc@86 839 induction s; substring.
adamc@86 840 Qed.
adamc@86 841
adam@375 842 Hint Rewrite substring_all substring_none.
adamc@86 843
adamc@86 844 Lemma substring_split : forall s m,
adamc@86 845 substring 0 m s ++ substring m (length s - m) s = s.
adamc@86 846 induction s; substring.
adamc@86 847 Qed.
adamc@86 848
adamc@86 849 Lemma length_app1 : forall s1 s2,
adamc@86 850 length s1 <= length (s1 ++ s2).
adamc@86 851 induction s1; crush.
adamc@86 852 Qed.
adamc@86 853
adamc@86 854 Hint Resolve length_emp append_emp substring_le substring_split length_app1.
adamc@86 855
adamc@86 856 Lemma substring_app_fst : forall s2 s1 n,
adamc@86 857 length s1 = n
adamc@86 858 -> substring 0 n (s1 ++ s2) = s1.
adamc@86 859 induction s1; crush.
adamc@86 860 Qed.
adamc@86 861
adamc@86 862 Lemma substring_app_snd : forall s2 s1 n,
adamc@86 863 length s1 = n
adamc@86 864 -> substring n (length (s1 ++ s2) - n) (s1 ++ s2) = s2.
adam@375 865 Hint Rewrite <- minus_n_O.
adamc@86 866
adamc@86 867 induction s1; crush.
adamc@86 868 Qed.
adamc@86 869
adam@375 870 Hint Rewrite substring_app_fst substring_app_snd using solve [trivial].
adamc@93 871 (* end hide *)
adamc@93 872
adamc@93 873 (** A few auxiliary functions help us in our final matcher definition. The function [split] will be used to implement the regexp concatenation case. *)
adamc@86 874
adamc@86 875 Section split.
adamc@86 876 Variables P1 P2 : string -> Prop.
adamc@214 877 Variable P1_dec : forall s, {P1 s} + {~ P1 s}.
adamc@214 878 Variable P2_dec : forall s, {P2 s} + {~ P2 s}.
adamc@93 879 (** We require a choice of two arbitrary string predicates and functions for deciding them. *)
adamc@86 880
adamc@86 881 Variable s : string.
adamc@93 882 (** Our computation will take place relative to a single fixed string, so it is easiest to make it a [Variable], rather than an explicit argument to our functions. *)
adamc@93 883
adam@338 884 (** The function [split'] is the workhorse behind [split]. It searches through the possible ways of splitting [s] into two pieces, checking the two predicates against each such pair. The execution of [split'] progresses right-to-left, from splitting all of [s] into the first piece to splitting all of [s] into the second piece. It takes an extra argument, [n], which specifies how far along we are in this search process. *)
adamc@86 885
adam@297 886 Definition split' : forall n : nat, n <= length s
adamc@86 887 -> {exists s1, exists s2, length s1 <= n /\ s1 ++ s2 = s /\ P1 s1 /\ P2 s2}
adamc@214 888 + {forall s1 s2, length s1 <= n -> s1 ++ s2 = s -> ~ P1 s1 \/ ~ P2 s2}.
adamc@86 889 refine (fix F (n : nat) : n <= length s
adamc@86 890 -> {exists s1, exists s2, length s1 <= n /\ s1 ++ s2 = s /\ P1 s1 /\ P2 s2}
adamc@214 891 + {forall s1 s2, length s1 <= n -> s1 ++ s2 = s -> ~ P1 s1 \/ ~ P2 s2} :=
adamc@214 892 match n with
adamc@86 893 | O => fun _ => Reduce (P1_dec "" && P2_dec s)
adamc@93 894 | S n' => fun _ => (P1_dec (substring 0 (S n') s)
adamc@93 895 && P2_dec (substring (S n') (length s - S n') s))
adamc@86 896 || F n' _
adamc@86 897 end); clear F; crush; eauto 7;
adamc@86 898 match goal with
adamc@86 899 | [ _ : length ?S <= 0 |- _ ] => destruct S
adam@338 900 | [ _ : length ?S' <= S ?N |- _ ] => destruct (eq_nat_dec (length S') (S N))
adamc@86 901 end; crush.
adamc@86 902 Defined.
adamc@86 903
adam@338 904 (** There is one subtle point in the [split'] code that is worth mentioning. The main body of the function is a [match] on [n]. In the case where [n] is known to be [S n'], we write [S n'] in several places where we might be tempted to write [n]. However, without further work to craft proper [match] annotations, the type-checker does not use the equality between [n] and [S n']. Thus, it is common to see patterns repeated in [match] case bodies in dependently typed Coq code. We can at least use a [let] expression to avoid copying the pattern more than once, replacing the first case body with:
adamc@93 905 [[
adamc@93 906 | S n' => fun _ => let n := S n' in
adamc@93 907 (P1_dec (substring 0 n s)
adamc@93 908 && P2_dec (substring n (length s - n) s))
adamc@93 909 || F n' _
adamc@93 910 ]]
adamc@93 911
adam@338 912 The [split] function itself is trivial to implement in terms of [split']. We just ask [split'] to begin its search with [n = length s]. *)
adamc@93 913
adamc@86 914 Definition split : {exists s1, exists s2, s = s1 ++ s2 /\ P1 s1 /\ P2 s2}
adamc@214 915 + {forall s1 s2, s = s1 ++ s2 -> ~ P1 s1 \/ ~ P2 s2}.
adamc@86 916 refine (Reduce (split' (n := length s) _)); crush; eauto.
adamc@86 917 Defined.
adamc@86 918 End split.
adamc@86 919
adamc@86 920 Implicit Arguments split [P1 P2].
adamc@86 921
adamc@93 922 (* begin hide *)
adamc@91 923 Lemma app_empty_end : forall s, s ++ "" = s.
adamc@91 924 induction s; crush.
adamc@91 925 Qed.
adamc@91 926
adam@375 927 Hint Rewrite app_empty_end.
adamc@91 928
adamc@91 929 Lemma substring_self : forall s n,
adamc@91 930 n <= 0
adamc@91 931 -> substring n (length s - n) s = s.
adamc@91 932 induction s; substring.
adamc@91 933 Qed.
adamc@91 934
adamc@91 935 Lemma substring_empty : forall s n m,
adamc@91 936 m <= 0
adamc@91 937 -> substring n m s = "".
adamc@91 938 induction s; substring.
adamc@91 939 Qed.
adamc@91 940
adam@375 941 Hint Rewrite substring_self substring_empty using omega.
adamc@91 942
adamc@91 943 Lemma substring_split' : forall s n m,
adamc@91 944 substring n m s ++ substring (n + m) (length s - (n + m)) s
adamc@91 945 = substring n (length s - n) s.
adam@375 946 Hint Rewrite substring_split.
adamc@91 947
adamc@91 948 induction s; substring.
adamc@91 949 Qed.
adamc@91 950
adamc@91 951 Lemma substring_stack : forall s n2 m1 m2,
adamc@91 952 m1 <= m2
adamc@91 953 -> substring 0 m1 (substring n2 m2 s)
adamc@91 954 = substring n2 m1 s.
adamc@91 955 induction s; substring.
adamc@91 956 Qed.
adamc@91 957
adamc@91 958 Ltac substring' :=
adamc@91 959 crush;
adamc@91 960 repeat match goal with
adamc@91 961 | [ |- context[match ?N with O => _ | S _ => _ end] ] => case_eq N; crush
adamc@91 962 end.
adamc@91 963
adamc@91 964 Lemma substring_stack' : forall s n1 n2 m1 m2,
adamc@91 965 n1 + m1 <= m2
adamc@91 966 -> substring n1 m1 (substring n2 m2 s)
adamc@91 967 = substring (n1 + n2) m1 s.
adamc@91 968 induction s; substring';
adamc@91 969 match goal with
adamc@91 970 | [ |- substring ?N1 _ _ = substring ?N2 _ _ ] =>
adamc@91 971 replace N1 with N2; crush
adamc@91 972 end.
adamc@91 973 Qed.
adamc@91 974
adamc@91 975 Lemma substring_suffix : forall s n,
adamc@91 976 n <= length s
adamc@91 977 -> length (substring n (length s - n) s) = length s - n.
adamc@91 978 induction s; substring.
adamc@91 979 Qed.
adamc@91 980
adamc@91 981 Lemma substring_suffix_emp' : forall s n m,
adamc@91 982 substring n (S m) s = ""
adamc@91 983 -> n >= length s.
adamc@91 984 induction s; crush;
adamc@91 985 match goal with
adamc@91 986 | [ |- ?N >= _ ] => destruct N; crush
adamc@91 987 end;
adamc@91 988 match goal with
adamc@91 989 [ |- S ?N >= S ?E ] => assert (N >= E); [ eauto | omega ]
adamc@91 990 end.
adamc@91 991 Qed.
adamc@91 992
adamc@91 993 Lemma substring_suffix_emp : forall s n m,
adamc@92 994 substring n m s = ""
adamc@92 995 -> m > 0
adamc@91 996 -> n >= length s.
adam@335 997 destruct m as [ | m]; [crush | intros; apply substring_suffix_emp' with m; assumption].
adamc@91 998 Qed.
adamc@91 999
adamc@91 1000 Hint Rewrite substring_stack substring_stack' substring_suffix
adam@375 1001 using omega.
adamc@91 1002
adamc@91 1003 Lemma minus_minus : forall n m1 m2,
adamc@91 1004 m1 + m2 <= n
adamc@91 1005 -> n - m1 - m2 = n - (m1 + m2).
adamc@91 1006 intros; omega.
adamc@91 1007 Qed.
adamc@91 1008
adamc@91 1009 Lemma plus_n_Sm' : forall n m : nat, S (n + m) = m + S n.
adamc@91 1010 intros; omega.
adamc@91 1011 Qed.
adamc@91 1012
adam@375 1013 Hint Rewrite minus_minus using omega.
adamc@93 1014 (* end hide *)
adamc@93 1015
adamc@93 1016 (** One more helper function will come in handy: [dec_star], for implementing another linear search through ways of splitting a string, this time for implementing the Kleene star. *)
adamc@91 1017
adamc@91 1018 Section dec_star.
adamc@91 1019 Variable P : string -> Prop.
adamc@214 1020 Variable P_dec : forall s, {P s} + {~ P s}.
adamc@91 1021
adam@338 1022 (** Some new lemmas and hints about the [star] type family are useful. We omit them here; they are included in the book source at this point. *)
adamc@93 1023
adamc@93 1024 (* begin hide *)
adamc@91 1025 Hint Constructors star.
adamc@91 1026
adamc@91 1027 Lemma star_empty : forall s,
adamc@91 1028 length s = 0
adamc@91 1029 -> star P s.
adamc@91 1030 destruct s; crush.
adamc@91 1031 Qed.
adamc@91 1032
adamc@91 1033 Lemma star_singleton : forall s, P s -> star P s.
adamc@91 1034 intros; rewrite <- (app_empty_end s); auto.
adamc@91 1035 Qed.
adamc@91 1036
adamc@91 1037 Lemma star_app : forall s n m,
adamc@91 1038 P (substring n m s)
adamc@91 1039 -> star P (substring (n + m) (length s - (n + m)) s)
adamc@91 1040 -> star P (substring n (length s - n) s).
adamc@91 1041 induction n; substring;
adamc@91 1042 match goal with
adamc@91 1043 | [ H : P (substring ?N ?M ?S) |- _ ] =>
adamc@91 1044 solve [ rewrite <- (substring_split S M); auto
adamc@91 1045 | rewrite <- (substring_split' S N M); auto ]
adamc@91 1046 end.
adamc@91 1047 Qed.
adamc@91 1048
adamc@91 1049 Hint Resolve star_empty star_singleton star_app.
adamc@91 1050
adamc@91 1051 Variable s : string.
adamc@91 1052
adamc@91 1053 Lemma star_inv : forall s,
adamc@91 1054 star P s
adamc@91 1055 -> s = ""
adamc@91 1056 \/ exists i, i < length s
adamc@91 1057 /\ P (substring 0 (S i) s)
adamc@91 1058 /\ star P (substring (S i) (length s - S i) s).
adamc@91 1059 Hint Extern 1 (exists i : nat, _) =>
adamc@91 1060 match goal with
adamc@91 1061 | [ H : P (String _ ?S) |- _ ] => exists (length S); crush
adamc@91 1062 end.
adamc@91 1063
adamc@91 1064 induction 1; [
adamc@91 1065 crush
adamc@91 1066 | match goal with
adamc@91 1067 | [ _ : P ?S |- _ ] => destruct S; crush
adamc@91 1068 end
adamc@91 1069 ].
adamc@91 1070 Qed.
adamc@91 1071
adamc@91 1072 Lemma star_substring_inv : forall n,
adamc@91 1073 n <= length s
adamc@91 1074 -> star P (substring n (length s - n) s)
adamc@91 1075 -> substring n (length s - n) s = ""
adamc@91 1076 \/ exists l, l < length s - n
adamc@91 1077 /\ P (substring n (S l) s)
adamc@91 1078 /\ star P (substring (n + S l) (length s - (n + S l)) s).
adam@375 1079 Hint Rewrite plus_n_Sm'.
adamc@91 1080
adamc@91 1081 intros;
adamc@91 1082 match goal with
adamc@91 1083 | [ H : star _ _ |- _ ] => generalize (star_inv H); do 3 crush; eauto
adamc@91 1084 end.
adamc@91 1085 Qed.
adamc@93 1086 (* end hide *)
adamc@93 1087
adamc@93 1088 (** The function [dec_star''] implements a single iteration of the star. That is, it tries to find a string prefix matching [P], and it calls a parameter function on the remainder of the string. *)
adamc@91 1089
adamc@91 1090 Section dec_star''.
adamc@91 1091 Variable n : nat.
adam@454 1092 (** Variable [n] is the length of the prefix of [s] that we have already processed. *)
adamc@91 1093
adamc@91 1094 Variable P' : string -> Prop.
adamc@91 1095 Variable P'_dec : forall n' : nat, n' > n
adamc@91 1096 -> {P' (substring n' (length s - n') s)}
adamc@214 1097 + {~ P' (substring n' (length s - n') s)}.
adam@475 1098
adamc@93 1099 (** When we use [dec_star''], we will instantiate [P'_dec] with a function for continuing the search for more instances of [P] in [s]. *)
adamc@93 1100
adamc@93 1101 (** Now we come to [dec_star''] itself. It takes as an input a natural [l] that records how much of the string has been searched so far, as we did for [split']. The return type expresses that [dec_star''] is looking for an index into [s] that splits [s] into a nonempty prefix and a suffix, such that the prefix satisfies [P] and the suffix satisfies [P']. *)
adamc@91 1102
adam@297 1103 Definition dec_star'' : forall l : nat,
adam@297 1104 {exists l', S l' <= l
adamc@91 1105 /\ P (substring n (S l') s) /\ P' (substring (n + S l') (length s - (n + S l')) s)}
adamc@91 1106 + {forall l', S l' <= l
adamc@214 1107 -> ~ P (substring n (S l') s)
adamc@214 1108 \/ ~ P' (substring (n + S l') (length s - (n + S l')) s)}.
adamc@91 1109 refine (fix F (l : nat) : {exists l', S l' <= l
adam@480 1110 /\ P (substring n (S l') s) /\ P' (substring (n + S l') (length s - (n + S l')) s)}
adam@480 1111 + {forall l', S l' <= l
adam@480 1112 -> ~ P (substring n (S l') s)
adam@480 1113 \/ ~ P' (substring (n + S l') (length s - (n + S l')) s)} :=
adam@480 1114 match l with
adam@480 1115 | O => _
adam@480 1116 | S l' =>
adam@480 1117 (P_dec (substring n (S l') s) && P'_dec (n' := n + S l') _)
adam@480 1118 || F l'
adam@480 1119 end); clear F; crush; eauto 7;
adam@480 1120 match goal with
adam@480 1121 | [ H : ?X <= S ?Y |- _ ] => destruct (eq_nat_dec X (S Y)); crush
adam@480 1122 end.
adamc@91 1123 Defined.
adamc@91 1124 End dec_star''.
adamc@91 1125
adamc@93 1126 (* begin hide *)
adamc@92 1127 Lemma star_length_contra : forall n,
adamc@92 1128 length s > n
adamc@92 1129 -> n >= length s
adamc@92 1130 -> False.
adamc@92 1131 crush.
adamc@92 1132 Qed.
adamc@92 1133
adamc@92 1134 Lemma star_length_flip : forall n n',
adamc@92 1135 length s - n <= S n'
adamc@92 1136 -> length s > n
adamc@92 1137 -> length s - n > 0.
adamc@92 1138 crush.
adamc@92 1139 Qed.
adamc@92 1140
adamc@92 1141 Hint Resolve star_length_contra star_length_flip substring_suffix_emp.
adamc@93 1142 (* end hide *)
adamc@92 1143
adamc@93 1144 (** The work of [dec_star''] is nested inside another linear search by [dec_star'], which provides the final functionality we need, but for arbitrary suffixes of [s], rather than just for [s] overall. *)
adamc@93 1145
adam@297 1146 Definition dec_star' : forall n n' : nat, length s - n' <= n
adamc@91 1147 -> {star P (substring n' (length s - n') s)}
adamc@214 1148 + {~ star P (substring n' (length s - n') s)}.
adamc@214 1149 refine (fix F (n n' : nat) : length s - n' <= n
adamc@91 1150 -> {star P (substring n' (length s - n') s)}
adamc@214 1151 + {~ star P (substring n' (length s - n') s)} :=
adamc@214 1152 match n with
adamc@91 1153 | O => fun _ => Yes
adamc@91 1154 | S n'' => fun _ =>
adamc@91 1155 le_gt_dec (length s) n'
adam@338 1156 || dec_star'' (n := n') (star P)
adam@338 1157 (fun n0 _ => Reduce (F n'' n0 _)) (length s - n')
adamc@92 1158 end); clear F; crush; eauto;
adamc@92 1159 match goal with
adamc@92 1160 | [ H : star _ _ |- _ ] => apply star_substring_inv in H; crush; eauto
adamc@92 1161 end;
adamc@92 1162 match goal with
adamc@92 1163 | [ H1 : _ < _ - _, H2 : forall l' : nat, _ <= _ - _ -> _ |- _ ] =>
adamc@92 1164 generalize (H2 _ (lt_le_S _ _ H1)); tauto
adamc@92 1165 end.
adamc@91 1166 Defined.
adamc@91 1167
adam@380 1168 (** Finally, we have [dec_star], defined by straightforward reduction from [dec_star']. *)
adamc@93 1169
adamc@214 1170 Definition dec_star : {star P s} + {~ star P s}.
adam@380 1171 refine (Reduce (dec_star' (n := length s) 0 _)); crush.
adamc@91 1172 Defined.
adamc@91 1173 End dec_star.
adamc@91 1174
adamc@93 1175 (* begin hide *)
adamc@86 1176 Lemma app_cong : forall x1 y1 x2 y2,
adamc@86 1177 x1 = x2
adamc@86 1178 -> y1 = y2
adamc@86 1179 -> x1 ++ y1 = x2 ++ y2.
adamc@86 1180 congruence.
adamc@86 1181 Qed.
adamc@86 1182
adamc@86 1183 Hint Resolve app_cong.
adamc@93 1184 (* end hide *)
adamc@93 1185
adamc@93 1186 (** With these helper functions completed, the implementation of our [matches] function is refreshingly straightforward. We only need one small piece of specific tactic work beyond what [crush] does for us. *)
adamc@86 1187
adam@297 1188 Definition matches : forall P (r : regexp P) s, {P s} + {~ P s}.
adamc@214 1189 refine (fix F P (r : regexp P) s : {P s} + {~ P s} :=
adamc@86 1190 match r with
adamc@86 1191 | Char ch => string_dec s (String ch "")
adamc@86 1192 | Concat _ _ r1 r2 => Reduce (split (F _ r1) (F _ r2) s)
adamc@87 1193 | Or _ _ r1 r2 => F _ r1 s || F _ r2 s
adamc@91 1194 | Star _ r => dec_star _ _ _
adamc@86 1195 end); crush;
adamc@86 1196 match goal with
adam@426 1197 | [ H : _ |- _ ] => generalize (H _ _ (eq_refl _))
adamc@93 1198 end; tauto.
adamc@86 1199 Defined.
adamc@86 1200
adam@283 1201 (** It is interesting to pause briefly to consider alternate implementations of [matches]. Dependent types give us much latitude in how specific correctness properties may be encoded with types. For instance, we could have made [regexp] a non-indexed inductive type, along the lines of what is possible in traditional ML and Haskell. We could then have implemented a recursive function to map [regexp]s to their intended meanings, much as we have done with types and programs in other examples. That style is compatible with the [refine]-based approach that we have used here, and it might be an interesting exercise to redo the code from this subsection in that alternate style or some further encoding of the reader's choice. The main advantage of indexed inductive types is that they generally lead to the smallest amount of code. *)
adam@283 1202
adamc@93 1203 (* begin hide *)
adamc@86 1204 Example hi := Concat (Char "h"%char) (Char "i"%char).
adam@380 1205 Eval hnf in matches hi "hi".
adam@380 1206 Eval hnf in matches hi "bye".
adamc@87 1207
adamc@87 1208 Example a_b := Or (Char "a"%char) (Char "b"%char).
adam@380 1209 Eval hnf in matches a_b "".
adam@380 1210 Eval hnf in matches a_b "a".
adam@380 1211 Eval hnf in matches a_b "aa".
adam@380 1212 Eval hnf in matches a_b "b".
adam@283 1213 (* end hide *)
adam@283 1214
adam@454 1215 (** Many regular expression matching problems are easy to test. The reader may run each of the following queries to verify that it gives the correct answer. We use evaluation strategy %\index{tactics!hnf}%[hnf] to reduce each term to%\index{head-normal form}% _head-normal form_, where the datatype constructor used to build its value is known. (Further reduction would involve wasteful simplification of proof terms justifying the answers of our procedures.) *)
adamc@91 1216
adamc@91 1217 Example a_star := Star (Char "a"%char).
adam@380 1218 Eval hnf in matches a_star "".
adam@380 1219 Eval hnf in matches a_star "a".
adam@380 1220 Eval hnf in matches a_star "b".
adam@380 1221 Eval hnf in matches a_star "aa".
adam@283 1222
adam@283 1223 (** Evaluation inside Coq does not scale very well, so it is easy to build other tests that run for hours or more. Such cases are better suited to execution with the extracted OCaml code. *)