annotate src/Subset.v @ 448:2740b8a23cce

Proofreading pass through Chapter 3
author Adam Chlipala <adam@chlipala.net>
date Fri, 17 Aug 2012 14:19:59 -0400
parents 89c67796754e
children 822442bf6d7f
rev   line source
adam@380 1 (* Copyright (c) 2008-2012, Adam Chlipala
adamc@70 2 *
adamc@70 3 * This work is licensed under a
adamc@70 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@70 5 * Unported License.
adamc@70 6 * The license text is available at:
adamc@70 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@70 8 *)
adamc@70 9
adamc@70 10 (* begin hide *)
adamc@70 11 Require Import List.
adamc@70 12
adam@314 13 Require Import CpdtTactics.
adamc@70 14
adamc@70 15 Set Implicit Arguments.
adamc@70 16 (* end hide *)
adamc@70 17
adam@403 18 (** printing <-- $\longleftarrow$ *)
adam@403 19
adamc@70 20
adamc@74 21 (** %\part{Programming with Dependent Types}
adamc@74 22
adamc@74 23 \chapter{Subset Types and Variations}% *)
adamc@70 24
adam@423 25 (** So far, we have seen many examples of what we might call "classical program verification." We write programs, write their specifications, and then prove that the programs satisfy their specifications. The programs that we have written in Coq have been normal functional programs that we could just as well have written in Haskell or ML. In this chapter, we start investigating uses of%\index{dependent types}% _dependent types_ to integrate programming, specification, and proving into a single phase. The techniques we will learn make it possible to reduce the cost of program verification dramatically. *)
adamc@70 26
adamc@70 27
adamc@70 28 (** * Introducing Subset Types *)
adamc@70 29
adamc@70 30 (** Let us consider several ways of implementing the natural number predecessor function. We start by displaying the definition from the standard library: *)
adamc@70 31
adamc@70 32 Print pred.
adamc@212 33 (** %\vspace{-.15in}% [[
adamc@70 34 pred = fun n : nat => match n with
adamc@70 35 | 0 => 0
adamc@70 36 | S u => u
adamc@70 37 end
adamc@70 38 : nat -> nat
adamc@212 39
adamc@212 40 ]]
adamc@70 41
adam@335 42 We can use a new command, %\index{Vernacular commands!Extraction}\index{program extraction}\index{extraction|see{program extraction}}%[Extraction], to produce an %\index{OCaml}%OCaml version of this function. *)
adamc@70 43
adamc@70 44 Extraction pred.
adamc@70 45
adamc@70 46 (** %\begin{verbatim}
adamc@70 47 (** val pred : nat -> nat **)
adamc@70 48
adamc@70 49 let pred = function
adamc@70 50 | O -> O
adamc@70 51 | S u -> u
adamc@70 52 \end{verbatim}%
adamc@70 53
adamc@70 54 #<pre>
adamc@70 55 (** val pred : nat -> nat **)
adamc@70 56
adamc@70 57 let pred = function
adamc@70 58 | O -> O
adamc@70 59 | S u -> u
adamc@70 60 </pre># *)
adamc@70 61
adamc@70 62 (** Returning 0 as the predecessor of 0 can come across as somewhat of a hack. In some situations, we might like to be sure that we never try to take the predecessor of 0. We can enforce this by giving [pred] a stronger, dependent type. *)
adamc@70 63
adamc@70 64 Lemma zgtz : 0 > 0 -> False.
adamc@70 65 crush.
adamc@70 66 Qed.
adamc@70 67
adamc@70 68 Definition pred_strong1 (n : nat) : n > 0 -> nat :=
adamc@212 69 match n with
adamc@70 70 | O => fun pf : 0 > 0 => match zgtz pf with end
adamc@70 71 | S n' => fun _ => n'
adamc@70 72 end.
adamc@70 73
adam@398 74 (** We expand the type of [pred] to include a _proof_ that its argument [n] is greater than 0. When [n] is 0, we use the proof to derive a contradiction, which we can use to build a value of any type via a vacuous pattern match. When [n] is a successor, we have no need for the proof and just return the answer. The proof argument can be said to have a _dependent_ type, because its type depends on the _value_ of the argument [n].
adamc@70 75
adam@398 76 Coq's [Eval] command can execute particular invocations of [pred_strong1] just as easily as it can execute more traditional functional programs. Note that Coq has decided that argument [n] of [pred_strong1] can be made _implicit_, since it can be deduced from the type of the second argument, so we need not write [n] in function calls. *)
adam@282 77
adam@282 78 Theorem two_gt0 : 2 > 0.
adam@282 79 crush.
adam@282 80 Qed.
adam@282 81
adam@282 82 Eval compute in pred_strong1 two_gt0.
adam@282 83 (** %\vspace{-.15in}% [[
adam@282 84 = 1
adam@282 85 : nat
adam@282 86 ]]
adam@282 87
adam@442 88 One aspect in particular of the definition of [pred_strong1] may be surprising. We took advantage of [Definition]'s syntactic sugar for defining function arguments in the case of [n], but we bound the proofs later with explicit [fun] expressions. Let us see what happens if we write this function in the way that at first seems most natural.
adamc@70 89
adam@440 90 %\vspace{-.15in}%[[
adamc@70 91 Definition pred_strong1' (n : nat) (pf : n > 0) : nat :=
adamc@70 92 match n with
adamc@70 93 | O => match zgtz pf with end
adamc@70 94 | S n' => n'
adamc@70 95 end.
adam@335 96 ]]
adamc@70 97
adam@335 98 <<
adamc@70 99 Error: In environment
adamc@70 100 n : nat
adamc@70 101 pf : n > 0
adamc@70 102 The term "pf" has type "n > 0" while it is expected to have type
adamc@70 103 "0 > 0"
adam@335 104 >>
adamc@70 105
adamc@212 106 The term [zgtz pf] fails to type-check. Somehow the type checker has failed to take into account information that follows from which [match] branch that term appears in. The problem is that, by default, [match] does not let us use such implied information. To get refined typing, we must always rely on [match] annotations, either written explicitly or inferred.
adamc@70 107
adam@398 108 In this case, we must use a [return] annotation to declare the relationship between the _value_ of the [match] discriminee and the _type_ of the result. There is no annotation that lets us declare a relationship between the discriminee and the type of a variable that is already in scope; hence, we delay the binding of [pf], so that we can use the [return] annotation to express the needed relationship.
adamc@70 109
adam@335 110 We are lucky that Coq's heuristics infer the [return] clause (specifically, [return n > 0 -> nat]) for us in this case. *)
adam@335 111
adam@335 112 Definition pred_strong1' (n : nat) : n > 0 -> nat :=
adam@335 113 match n return n > 0 -> nat with
adam@335 114 | O => fun pf : 0 > 0 => match zgtz pf with end
adam@335 115 | S n' => fun _ => n'
adam@335 116 end.
adam@335 117
adam@403 118 (** By making explicit the functional relationship between value [n] and the result type of the [match], we guide Coq toward proper type checking. The clause for this example follows by simple copying of the original annotation on the definition. In general, however, the [match] annotation inference problem is undecidable. The known undecidable problem of%\index{higher-order unification}% _higher-order unification_ %\cite{HOU}% reduces to the [match] type inference problem. Over time, Coq is enhanced with more and more heuristics to get around this problem, but there must always exist [match]es whose types Coq cannot infer without annotations.
adamc@70 119
adamc@70 120 Let us now take a look at the OCaml code Coq generates for [pred_strong1]. *)
adamc@70 121
adamc@70 122 Extraction pred_strong1.
adamc@70 123
adamc@70 124 (** %\begin{verbatim}
adamc@70 125 (** val pred_strong1 : nat -> nat **)
adamc@70 126
adamc@70 127 let pred_strong1 = function
adamc@70 128 | O -> assert false (* absurd case *)
adamc@70 129 | S n' -> n'
adamc@70 130 \end{verbatim}%
adamc@70 131
adamc@70 132 #<pre>
adamc@70 133 (** val pred_strong1 : nat -> nat **)
adamc@70 134
adamc@70 135 let pred_strong1 = function
adamc@70 136 | O -> assert false (* absurd case *)
adamc@70 137 | S n' -> n'
adamc@70 138 </pre># *)
adamc@70 139
adamc@70 140 (** The proof argument has disappeared! We get exactly the OCaml code we would have written manually. This is our first demonstration of the main technically interesting feature of Coq program extraction: program components of type [Prop] are erased systematically.
adamc@70 141
adam@403 142 We can reimplement our dependently typed [pred] based on%\index{subset types}% _subset types_, defined in the standard library with the type family %\index{Gallina terms!sig}%[sig]. *)
adamc@70 143
adam@423 144 (* begin hide *)
adam@437 145 (* begin thide *)
adam@437 146 Definition bar := ex.
adam@437 147 (* end thide *)
adam@423 148 (* end hide *)
adam@423 149
adamc@70 150 Print sig.
adamc@212 151 (** %\vspace{-.15in}% [[
adamc@70 152 Inductive sig (A : Type) (P : A -> Prop) : Type :=
adamc@70 153 exist : forall x : A, P x -> sig P
adamc@70 154 ]]
adamc@70 155
adam@442 156 The family [sig] is a Curry-Howard twin of [ex], except that [sig] is in [Type], while [ex] is in [Prop]. That means that [sig] values can survive extraction, while [ex] proofs will always be erased. The actual details of extraction of [sig]s are more subtle, as we will see shortly.
adamc@70 157
adamc@70 158 We rewrite [pred_strong1], using some syntactic sugar for subset types. *)
adamc@70 159
adamc@70 160 Locate "{ _ : _ | _ }".
adamc@212 161 (** %\vspace{-.15in}% [[
adam@335 162 Notation
adamc@70 163 "{ x : A | P }" := sig (fun x : A => P)
adam@302 164 ]]
adam@302 165 *)
adamc@70 166
adamc@70 167 Definition pred_strong2 (s : {n : nat | n > 0}) : nat :=
adamc@70 168 match s with
adamc@70 169 | exist O pf => match zgtz pf with end
adamc@70 170 | exist (S n') _ => n'
adamc@70 171 end.
adamc@70 172
adam@335 173 (** To build a value of a subset type, we use the [exist] constructor, and the details of how to do that follow from the output of our earlier [Print sig] command (where we elided the extra information that parameter [A] is implicit). *)
adam@282 174
adam@282 175 Eval compute in pred_strong2 (exist _ 2 two_gt0).
adam@282 176 (** %\vspace{-.15in}% [[
adam@282 177 = 1
adam@282 178 : nat
adam@302 179 ]]
adam@302 180 *)
adam@282 181
adamc@70 182 Extraction pred_strong2.
adamc@70 183
adamc@70 184 (** %\begin{verbatim}
adamc@70 185 (** val pred_strong2 : nat -> nat **)
adamc@70 186
adamc@70 187 let pred_strong2 = function
adamc@70 188 | O -> assert false (* absurd case *)
adamc@70 189 | S n' -> n'
adamc@70 190 \end{verbatim}%
adamc@70 191
adamc@70 192 #<pre>
adamc@70 193 (** val pred_strong2 : nat -> nat **)
adamc@70 194
adamc@70 195 let pred_strong2 = function
adamc@70 196 | O -> assert false (* absurd case *)
adamc@70 197 | S n' -> n'
adamc@70 198 </pre>#
adamc@70 199
adamc@70 200 We arrive at the same OCaml code as was extracted from [pred_strong1], which may seem surprising at first. The reason is that a value of [sig] is a pair of two pieces, a value and a proof about it. Extraction erases the proof, which reduces the constructor [exist] of [sig] to taking just a single argument. An optimization eliminates uses of datatypes with single constructors taking single arguments, and we arrive back where we started.
adamc@70 201
adamc@70 202 We can continue on in the process of refining [pred]'s type. Let us change its result type to capture that the output is really the predecessor of the input. *)
adamc@70 203
adamc@70 204 Definition pred_strong3 (s : {n : nat | n > 0}) : {m : nat | proj1_sig s = S m} :=
adamc@70 205 match s return {m : nat | proj1_sig s = S m} with
adamc@70 206 | exist 0 pf => match zgtz pf with end
adam@426 207 | exist (S n') pf => exist _ n' (eq_refl _)
adamc@70 208 end.
adamc@70 209
adam@282 210 Eval compute in pred_strong3 (exist _ 2 two_gt0).
adam@282 211 (** %\vspace{-.15in}% [[
adam@426 212 = exist (fun m : nat => 2 = S m) 1 (eq_refl 2)
adam@282 213 : {m : nat | proj1_sig (exist (lt 0) 2 two_gt0) = S m}
adam@335 214 ]]
adam@302 215 *)
adam@282 216
adam@423 217 (* begin hide *)
adam@437 218 (* begin thide *)
adam@423 219 Definition pred_strong := 0.
adam@437 220 (* end thide *)
adam@423 221 (* end hide *)
adam@423 222
adam@335 223 (** The function %\index{Gallina terms!proj1\_sig}%[proj1_sig] extracts the base value from a subset type. It turns out that we need to include an explicit [return] clause here, since Coq's heuristics are not smart enough to propagate the result type that we wrote earlier.
adamc@70 224
adamc@70 225 By now, the reader is probably ready to believe that the new [pred_strong] leads to the same OCaml code as we have seen several times so far, and Coq does not disappoint. *)
adamc@70 226
adamc@70 227 Extraction pred_strong3.
adamc@70 228
adamc@70 229 (** %\begin{verbatim}
adamc@70 230 (** val pred_strong3 : nat -> nat **)
adamc@70 231
adamc@70 232 let pred_strong3 = function
adamc@70 233 | O -> assert false (* absurd case *)
adamc@70 234 | S n' -> n'
adamc@70 235 \end{verbatim}%
adamc@70 236
adamc@70 237 #<pre>
adamc@70 238 (** val pred_strong3 : nat -> nat **)
adamc@70 239
adamc@70 240 let pred_strong3 = function
adamc@70 241 | O -> assert false (* absurd case *)
adamc@70 242 | S n' -> n'
adamc@70 243 </pre>#
adamc@70 244
adam@335 245 We have managed to reach a type that is, in a formal sense, the most expressive possible for [pred]. Any other implementation of the same type must have the same input-output behavior. However, there is still room for improvement in making this kind of code easier to write. Here is a version that takes advantage of tactic-based theorem proving. We switch back to passing a separate proof argument instead of using a subset type for the function's input, because this leads to cleaner code. (Recall that [False_rec] is the [Set]-level induction principle for [False], which can be used to produce a value in any [Set] given a proof of [False].) *)
adamc@70 246
adam@297 247 Definition pred_strong4 : forall n : nat, n > 0 -> {m : nat | n = S m}.
adamc@70 248 refine (fun n =>
adamc@212 249 match n with
adamc@70 250 | O => fun _ => False_rec _ _
adamc@70 251 | S n' => fun _ => exist _ n' _
adamc@70 252 end).
adamc@212 253
adamc@77 254 (* begin thide *)
adam@335 255 (** We build [pred_strong4] using tactic-based proving, beginning with a [Definition] command that ends in a period before a definition is given. Such a command enters the interactive proving mode, with the type given for the new identifier as our proof goal. It may seem strange to change perspective so implicitly between programming and proving, but recall that programs and proofs are two sides of the same coin in Coq, thanks to the Curry-Howard correspondence.
adamc@70 256
adam@423 257 We do most of the work with the %\index{tactics!refine}%[refine] tactic, to which we pass a partial "proof" of the type we are trying to prove. There may be some pieces left to fill in, indicated by underscores. Any underscore that Coq cannot reconstruct with type inference is added as a proof subgoal. In this case, we have two subgoals:
adam@335 258
adam@335 259 [[
adam@439 260 2 subgoals
adamc@70 261
adamc@70 262 n : nat
adamc@70 263 _ : 0 > 0
adamc@70 264 ============================
adamc@70 265 False
adam@439 266
adam@439 267 subgoal 2 is
adam@439 268
adamc@70 269 S n' = S n'
adamc@70 270 ]]
adamc@70 271
adamc@70 272 We can see that the first subgoal comes from the second underscore passed to [False_rec], and the second subgoal comes from the second underscore passed to [exist]. In the first case, we see that, though we bound the proof variable with an underscore, it is still available in our proof context. It is hard to refer to underscore-named variables in manual proofs, but automation makes short work of them. Both subgoals are easy to discharge that way, so let us back up and ask to prove all subgoals automatically. *)
adamc@70 273
adamc@70 274 Undo.
adamc@70 275 refine (fun n =>
adamc@212 276 match n with
adamc@70 277 | O => fun _ => False_rec _ _
adamc@70 278 | S n' => fun _ => exist _ n' _
adamc@70 279 end); crush.
adamc@77 280 (* end thide *)
adamc@70 281 Defined.
adamc@70 282
adam@423 283 (** We end the "proof" with %\index{Vernacular commands!Defined}%[Defined] instead of [Qed], so that the definition we constructed remains visible. This contrasts to the case of ending a proof with [Qed], where the details of the proof are hidden afterward. (More formally, [Defined] marks an identifier as%\index{transparent}% _transparent_, allowing it to be unfolded; while [Qed] marks an identifier as%\index{opaque}% _opaque_, preventing unfolding.) Let us see what our proof script constructed. *)
adamc@70 284
adamc@70 285 Print pred_strong4.
adamc@212 286 (** %\vspace{-.15in}% [[
adamc@70 287 pred_strong4 =
adamc@70 288 fun n : nat =>
adamc@70 289 match n as n0 return (n0 > 0 -> {m : nat | n0 = S m}) with
adamc@70 290 | 0 =>
adamc@70 291 fun _ : 0 > 0 =>
adamc@70 292 False_rec {m : nat | 0 = S m}
adamc@70 293 (Bool.diff_false_true
adamc@70 294 (Bool.absurd_eq_true false
adamc@70 295 (Bool.diff_false_true
adamc@70 296 (Bool.absurd_eq_true false (pred_strong4_subproof n _)))))
adamc@70 297 | S n' =>
adamc@70 298 fun _ : S n' > 0 =>
adam@426 299 exist (fun m : nat => S n' = S m) n' (eq_refl (S n'))
adamc@70 300 end
adamc@70 301 : forall n : nat, n > 0 -> {m : nat | n = S m}
adamc@70 302 ]]
adamc@70 303
adam@442 304 We see the code we entered, with some proofs filled in. The first proof obligation, the second argument to [False_rec], is filled in with a nasty-looking proof term that we can be glad we did not enter by hand. The second proof obligation is a simple reflexivity proof. *)
adamc@70 305
adam@282 306 Eval compute in pred_strong4 two_gt0.
adam@282 307 (** %\vspace{-.15in}% [[
adam@426 308 = exist (fun m : nat => 2 = S m) 1 (eq_refl 2)
adam@282 309 : {m : nat | 2 = S m}
adam@282 310 ]]
adam@282 311
adam@442 312 A tactic modifier called %\index{tactics!abstract}%[abstract] can be helpful for producing shorter terms, by automatically abstracting subgoals into named lemmas. *)
adam@335 313
adam@335 314 (* begin thide *)
adam@335 315 Definition pred_strong4' : forall n : nat, n > 0 -> {m : nat | n = S m}.
adam@335 316 refine (fun n =>
adam@335 317 match n with
adam@335 318 | O => fun _ => False_rec _ _
adam@335 319 | S n' => fun _ => exist _ n' _
adam@335 320 end); abstract crush.
adam@335 321 Defined.
adam@335 322
adam@335 323 Print pred_strong4'.
adam@335 324 (* end thide *)
adam@335 325
adam@335 326 (** %\vspace{-.15in}% [[
adam@335 327 pred_strong4' =
adam@335 328 fun n : nat =>
adam@335 329 match n as n0 return (n0 > 0 -> {m : nat | n0 = S m}) with
adam@335 330 | 0 =>
adam@335 331 fun _H : 0 > 0 =>
adam@335 332 False_rec {m : nat | 0 = S m} (pred_strong4'_subproof n _H)
adam@335 333 | S n' =>
adam@335 334 fun _H : S n' > 0 =>
adam@335 335 exist (fun m : nat => S n' = S m) n' (pred_strong4'_subproof0 n _H)
adam@335 336 end
adam@335 337 : forall n : nat, n > 0 -> {m : nat | n = S m}
adam@335 338 ]]
adam@335 339
adam@338 340 We are almost done with the ideal implementation of dependent predecessor. We can use Coq's syntax extension facility to arrive at code with almost no complexity beyond a Haskell or ML program with a complete specification in a comment. In this book, we will not dwell on the details of syntax extensions; the Coq manual gives a straightforward introduction to them. *)
adamc@70 341
adamc@70 342 Notation "!" := (False_rec _ _).
adamc@70 343 Notation "[ e ]" := (exist _ e _).
adamc@70 344
adam@297 345 Definition pred_strong5 : forall n : nat, n > 0 -> {m : nat | n = S m}.
adamc@70 346 refine (fun n =>
adamc@212 347 match n with
adamc@70 348 | O => fun _ => !
adamc@70 349 | S n' => fun _ => [n']
adamc@70 350 end); crush.
adamc@70 351 Defined.
adamc@71 352
adam@282 353 (** By default, notations are also used in pretty-printing terms, including results of evaluation. *)
adam@282 354
adam@282 355 Eval compute in pred_strong5 two_gt0.
adam@282 356 (** %\vspace{-.15in}% [[
adam@282 357 = [1]
adam@282 358 : {m : nat | 2 = S m}
adam@282 359 ]]
adam@282 360
adam@442 361 One other alternative is worth demonstrating. Recent Coq versions include a facility called %\index{Program}%[Program] that streamlines this style of definition. Here is a complete implementation using [Program].%\index{Vernacular commands!Obligation Tactic}\index{Vernacular commands!Program Definition}% *)
adamc@212 362
adamc@212 363 Obligation Tactic := crush.
adamc@212 364
adamc@212 365 Program Definition pred_strong6 (n : nat) (_ : n > 0) : {m : nat | n = S m} :=
adamc@212 366 match n with
adamc@212 367 | O => _
adamc@212 368 | S n' => n'
adamc@212 369 end.
adamc@212 370
adam@335 371 (** Printing the resulting definition of [pred_strong6] yields a term very similar to what we built with [refine]. [Program] can save time in writing programs that use subset types. Nonetheless, [refine] is often just as effective, and [refine] gives you more control over the form the final term takes, which can be useful when you want to prove additional theorems about your definition. [Program] will sometimes insert type casts that can complicate theorem proving. *)
adamc@212 372
adam@282 373 Eval compute in pred_strong6 two_gt0.
adam@282 374 (** %\vspace{-.15in}% [[
adam@282 375 = [1]
adam@282 376 : {m : nat | 2 = S m}
adam@302 377 ]]
adam@335 378
adam@442 379 In this case, we see that the new definition yields the same computational behavior as before. *)
adam@282 380
adamc@71 381
adamc@71 382 (** * Decidable Proposition Types *)
adamc@71 383
adam@335 384 (** There is another type in the standard library which captures the idea of program values that indicate which of two propositions is true.%\index{Gallina terms!sumbool}% *)
adamc@71 385
adamc@71 386 Print sumbool.
adamc@212 387 (** %\vspace{-.15in}% [[
adamc@71 388 Inductive sumbool (A : Prop) (B : Prop) : Set :=
adamc@71 389 left : A -> {A} + {B} | right : B -> {A} + {B}
adamc@212 390 ]]
adamc@71 391
adamc@212 392 We can define some notations to make working with [sumbool] more convenient. *)
adamc@71 393
adamc@71 394 Notation "'Yes'" := (left _ _).
adamc@71 395 Notation "'No'" := (right _ _).
adamc@71 396 Notation "'Reduce' x" := (if x then Yes else No) (at level 50).
adamc@71 397
adam@436 398 (** The %\coqdocnotation{%#<tt>#Reduce#</tt>#%}% notation is notable because it demonstrates how [if] is overloaded in Coq. The [if] form actually works when the test expression has any two-constructor inductive type. Moreover, in the [then] and [else] branches, the appropriate constructor arguments are bound. This is important when working with [sumbool]s, when we want to have the proof stored in the test expression available when proving the proof obligations generated in the appropriate branch.
adamc@71 399
adamc@71 400 Now we can write [eq_nat_dec], which compares two natural numbers, returning either a proof of their equality or a proof of their inequality. *)
adamc@71 401
adam@297 402 Definition eq_nat_dec : forall n m : nat, {n = m} + {n <> m}.
adamc@212 403 refine (fix f (n m : nat) : {n = m} + {n <> m} :=
adamc@212 404 match n, m with
adamc@71 405 | O, O => Yes
adamc@71 406 | S n', S m' => Reduce (f n' m')
adamc@71 407 | _, _ => No
adamc@71 408 end); congruence.
adamc@71 409 Defined.
adamc@71 410
adam@282 411 Eval compute in eq_nat_dec 2 2.
adam@282 412 (** %\vspace{-.15in}% [[
adam@282 413 = Yes
adam@282 414 : {2 = 2} + {2 <> 2}
adam@302 415 ]]
adam@302 416 *)
adam@282 417
adam@282 418 Eval compute in eq_nat_dec 2 3.
adam@282 419 (** %\vspace{-.15in}% [[
adam@282 420 = No
adam@341 421 : {2 = 3} + {2 <> 3}
adam@302 422 ]]
adam@282 423
adam@442 424 Note that the %\coqdocnotation{%#<tt>#Yes#</tt>#%}% and %\coqdocnotation{%#<tt>#No#</tt>#%}% notations are hiding proofs establishing the correctness of the outputs.
adam@335 425
adam@335 426 Our definition extracts to reasonable OCaml code. *)
adamc@71 427
adamc@71 428 Extraction eq_nat_dec.
adamc@71 429
adamc@71 430 (** %\begin{verbatim}
adamc@71 431 (** val eq_nat_dec : nat -> nat -> sumbool **)
adamc@71 432
adamc@71 433 let rec eq_nat_dec n m =
adamc@71 434 match n with
adamc@71 435 | O -> (match m with
adamc@71 436 | O -> Left
adamc@71 437 | S n0 -> Right)
adamc@71 438 | S n' -> (match m with
adamc@71 439 | O -> Right
adamc@71 440 | S m' -> eq_nat_dec n' m')
adamc@71 441 \end{verbatim}%
adamc@71 442
adamc@71 443 #<pre>
adamc@71 444 (** val eq_nat_dec : nat -> nat -> sumbool **)
adamc@71 445
adamc@71 446 let rec eq_nat_dec n m =
adamc@71 447 match n with
adamc@71 448 | O -> (match m with
adamc@71 449 | O -> Left
adamc@71 450 | S n0 -> Right)
adamc@71 451 | S n' -> (match m with
adamc@71 452 | O -> Right
adamc@71 453 | S m' -> eq_nat_dec n' m')
adamc@71 454 </pre>#
adamc@71 455
adam@335 456 Proving this kind of decidable equality result is so common that Coq comes with a tactic for automating it.%\index{tactics!decide equality}% *)
adamc@71 457
adamc@71 458 Definition eq_nat_dec' (n m : nat) : {n = m} + {n <> m}.
adamc@71 459 decide equality.
adamc@71 460 Defined.
adamc@71 461
adam@448 462 (** Curious readers can verify that the [decide equality] version extracts to the same OCaml code as our more manual version does. That OCaml code had one undesirable property, which is that it uses <<Left>> and <<Right>> constructors instead of the Boolean values built into OCaml. We can fix this, by using Coq's facility for mapping Coq inductive types to OCaml variant types.%\index{Vernacular commands!Extract Inductive}% *)
adamc@71 463
adamc@71 464 Extract Inductive sumbool => "bool" ["true" "false"].
adamc@71 465 Extraction eq_nat_dec'.
adamc@71 466
adamc@71 467 (** %\begin{verbatim}
adamc@71 468 (** val eq_nat_dec' : nat -> nat -> bool **)
adamc@71 469
adamc@71 470 let rec eq_nat_dec' n m0 =
adamc@71 471 match n with
adamc@71 472 | O -> (match m0 with
adamc@71 473 | O -> true
adamc@71 474 | S n0 -> false)
adamc@71 475 | S n0 -> (match m0 with
adamc@71 476 | O -> false
adamc@71 477 | S n1 -> eq_nat_dec' n0 n1)
adamc@71 478 \end{verbatim}%
adamc@71 479
adamc@71 480 #<pre>
adamc@71 481 (** val eq_nat_dec' : nat -> nat -> bool **)
adamc@71 482
adamc@71 483 let rec eq_nat_dec' n m0 =
adamc@71 484 match n with
adamc@71 485 | O -> (match m0 with
adamc@71 486 | O -> true
adamc@71 487 | S n0 -> false)
adamc@71 488 | S n0 -> (match m0 with
adamc@71 489 | O -> false
adamc@71 490 | S n1 -> eq_nat_dec' n0 n1)
adamc@71 491 </pre># *)
adamc@72 492
adamc@72 493 (** %\smallskip%
adamc@72 494
adam@448 495 We can build "smart" versions of the usual Boolean operators and put them to good use in certified programming. For instance, here is a [sumbool] version of Boolean "or." *)
adamc@72 496
adam@337 497 (* EX: Write a function that decides if an element belongs to a list. *)
adam@337 498
adamc@77 499 (* begin thide *)
adamc@204 500 Notation "x || y" := (if x then Yes else Reduce y).
adamc@72 501
adamc@72 502 (** Let us use it for building a function that decides list membership. We need to assume the existence of an equality decision procedure for the type of list elements. *)
adamc@72 503
adamc@72 504 Section In_dec.
adamc@72 505 Variable A : Set.
adamc@72 506 Variable A_eq_dec : forall x y : A, {x = y} + {x <> y}.
adamc@72 507
adamc@72 508 (** The final function is easy to write using the techniques we have developed so far. *)
adamc@72 509
adamc@212 510 Definition In_dec : forall (x : A) (ls : list A), {In x ls} + {~ In x ls}.
adamc@212 511 refine (fix f (x : A) (ls : list A) : {In x ls} + {~ In x ls} :=
adamc@212 512 match ls with
adamc@72 513 | nil => No
adamc@72 514 | x' :: ls' => A_eq_dec x x' || f x ls'
adamc@72 515 end); crush.
adam@282 516 Defined.
adamc@72 517 End In_dec.
adamc@72 518
adam@282 519 Eval compute in In_dec eq_nat_dec 2 (1 :: 2 :: nil).
adam@282 520 (** %\vspace{-.15in}% [[
adam@282 521 = Yes
adam@282 522 : {In 2 (1 :: 2 :: nil)} + {~ In 2 (1 :: 2 :: nil)}
adam@302 523 ]]
adam@302 524 *)
adam@282 525
adam@282 526 Eval compute in In_dec eq_nat_dec 3 (1 :: 2 :: nil).
adam@282 527 (** %\vspace{-.15in}% [[
adam@282 528 = No
adam@282 529 : {In 3 (1 :: 2 :: nil)} + {~ In 3 (1 :: 2 :: nil)}
adam@302 530 ]]
adam@282 531
adam@442 532 [In_dec] has a reasonable extraction to OCaml. *)
adamc@72 533
adamc@72 534 Extraction In_dec.
adamc@77 535 (* end thide *)
adamc@72 536
adamc@72 537 (** %\begin{verbatim}
adamc@72 538 (** val in_dec : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 list -> bool **)
adamc@72 539
adamc@72 540 let rec in_dec a_eq_dec x = function
adamc@72 541 | Nil -> false
adamc@72 542 | Cons (x', ls') ->
adamc@72 543 (match a_eq_dec x x' with
adamc@72 544 | true -> true
adamc@72 545 | false -> in_dec a_eq_dec x ls')
adamc@72 546 \end{verbatim}%
adamc@72 547
adamc@72 548 #<pre>
adamc@72 549 (** val in_dec : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 list -> bool **)
adamc@72 550
adamc@72 551 let rec in_dec a_eq_dec x = function
adamc@72 552 | Nil -> false
adamc@72 553 | Cons (x', ls') ->
adamc@72 554 (match a_eq_dec x x' with
adamc@72 555 | true -> true
adamc@72 556 | false -> in_dec a_eq_dec x ls')
adam@403 557 </pre>#
adam@403 558
adam@403 559 This is more or the less code for the corresponding function from the OCaml standard library. *)
adamc@72 560
adamc@72 561
adamc@72 562 (** * Partial Subset Types *)
adamc@72 563
adam@335 564 (** Our final implementation of dependent predecessor used a very specific argument type to ensure that execution could always complete normally. Sometimes we want to allow execution to fail, and we want a more principled way of signaling failure than returning a default value, as [pred] does for [0]. One approach is to define this type family %\index{Gallina terms!maybe}%[maybe], which is a version of [sig] that allows obligation-free failure. *)
adamc@73 565
adamc@89 566 Inductive maybe (A : Set) (P : A -> Prop) : Set :=
adamc@72 567 | Unknown : maybe P
adamc@72 568 | Found : forall x : A, P x -> maybe P.
adamc@72 569
adamc@73 570 (** We can define some new notations, analogous to those we defined for subset types. *)
adamc@73 571
adamc@72 572 Notation "{{ x | P }}" := (maybe (fun x => P)).
adamc@72 573 Notation "??" := (Unknown _).
adam@335 574 Notation "[| x |]" := (Found _ x _).
adamc@72 575
adamc@73 576 (** Now our next version of [pred] is trivial to write. *)
adamc@73 577
adam@297 578 Definition pred_strong7 : forall n : nat, {{m | n = S m}}.
adamc@73 579 refine (fun n =>
adam@380 580 match n return {{m | n = S m}} with
adamc@73 581 | O => ??
adam@335 582 | S n' => [|n'|]
adamc@73 583 end); trivial.
adamc@73 584 Defined.
adamc@73 585
adam@282 586 Eval compute in pred_strong7 2.
adam@282 587 (** %\vspace{-.15in}% [[
adam@335 588 = [|1|]
adam@282 589 : {{m | 2 = S m}}
adam@335 590 ]]
adam@302 591 *)
adam@282 592
adam@282 593 Eval compute in pred_strong7 0.
adam@282 594 (** %\vspace{-.15in}% [[
adam@282 595 = ??
adam@282 596 : {{m | 0 = S m}}
adam@282 597 ]]
adam@282 598
adam@442 599 Because we used [maybe], one valid implementation of the type we gave [pred_strong7] would return [??] in every case. We can strengthen the type to rule out such vacuous implementations, and the type family %\index{Gallina terms!sumor}%[sumor] from the standard library provides the easiest starting point. For type [A] and proposition [B], [A + {B}] desugars to [sumor A B], whose values are either values of [A] or proofs of [B]. *)
adamc@73 600
adamc@73 601 Print sumor.
adamc@212 602 (** %\vspace{-.15in}% [[
adamc@73 603 Inductive sumor (A : Type) (B : Prop) : Type :=
adamc@73 604 inleft : A -> A + {B} | inright : B -> A + {B}
adam@302 605 ]]
adamc@73 606
adam@442 607 We add notations for easy use of the [sumor] constructors. The second notation is specialized to [sumor]s whose [A] parameters are instantiated with regular subset types, since this is how we will use [sumor] below. *)
adamc@73 608
adamc@73 609 Notation "!!" := (inright _ _).
adam@335 610 Notation "[|| x ||]" := (inleft _ [x]).
adamc@73 611
adam@335 612 (** Now we are ready to give the final version of possibly failing predecessor. The [sumor]-based type that we use is maximally expressive; any implementation of the type has the same input-output behavior. *)
adamc@73 613
adam@297 614 Definition pred_strong8 : forall n : nat, {m : nat | n = S m} + {n = 0}.
adamc@73 615 refine (fun n =>
adamc@212 616 match n with
adamc@73 617 | O => !!
adam@335 618 | S n' => [||n'||]
adamc@73 619 end); trivial.
adamc@73 620 Defined.
adamc@73 621
adam@282 622 Eval compute in pred_strong8 2.
adam@282 623 (** %\vspace{-.15in}% [[
adam@335 624 = [||1||]
adam@282 625 : {m : nat | 2 = S m} + {2 = 0}
adam@302 626 ]]
adam@302 627 *)
adam@282 628
adam@282 629 Eval compute in pred_strong8 0.
adam@282 630 (** %\vspace{-.15in}% [[
adam@282 631 = !!
adam@282 632 : {m : nat | 0 = S m} + {0 = 0}
adam@302 633 ]]
adam@302 634 *)
adam@282 635
adam@335 636 (** As with our other maximally expressive [pred] function, we arrive at quite simple output values, thanks to notations. *)
adam@335 637
adamc@73 638
adamc@73 639 (** * Monadic Notations *)
adamc@73 640
adam@423 641 (** We can treat [maybe] like a monad%~\cite{Monads}\index{monad}\index{failure monad}%, in the same way that the Haskell <<Maybe>> type is interpreted as a failure monad. Our [maybe] has the wrong type to be a literal monad, but a "bind"-like notation will still be helpful. *)
adamc@73 642
adamc@72 643 Notation "x <- e1 ; e2" := (match e1 with
adamc@72 644 | Unknown => ??
adamc@72 645 | Found x _ => e2
adamc@72 646 end)
adamc@72 647 (right associativity, at level 60).
adamc@72 648
adam@398 649 (** The meaning of [x <- e1; e2] is: First run [e1]. If it fails to find an answer, then announce failure for our derived computation, too. If [e1] _does_ find an answer, pass that answer on to [e2] to find the final result. The variable [x] can be considered bound in [e2].
adamc@73 650
adam@335 651 This notation is very helpful for composing richly typed procedures. For instance, here is a very simple implementation of a function to take the predecessors of two naturals at once. *)
adamc@73 652
adam@337 653 (* EX: Write a function that tries to compute predecessors of two [nat]s at once. *)
adam@337 654
adam@337 655 (* begin thide *)
adam@297 656 Definition doublePred : forall n1 n2 : nat, {{p | n1 = S (fst p) /\ n2 = S (snd p)}}.
adamc@73 657 refine (fun n1 n2 =>
adamc@212 658 m1 <- pred_strong7 n1;
adamc@212 659 m2 <- pred_strong7 n2;
adam@335 660 [|(m1, m2)|]); tauto.
adamc@73 661 Defined.
adam@337 662 (* end thide *)
adamc@73 663
adam@423 664 (** We can build a [sumor] version of the "bind" notation and use it to write a similarly straightforward version of this function. %The operator rendered here as $\longleftarrow$ is noted in the source as a less-than character followed by two hyphens.% *)
adamc@73 665
adamc@73 666 Notation "x <-- e1 ; e2" := (match e1 with
adamc@73 667 | inright _ => !!
adamc@73 668 | inleft (exist x _) => e2
adamc@73 669 end)
adamc@73 670 (right associativity, at level 60).
adamc@73 671
adamc@73 672 (** printing * $\times$ *)
adamc@73 673
adam@337 674 (* EX: Write a more expressively typed version of the last exercise. *)
adam@337 675
adam@337 676 (* begin thide *)
adam@297 677 Definition doublePred' : forall n1 n2 : nat,
adam@297 678 {p : nat * nat | n1 = S (fst p) /\ n2 = S (snd p)}
adamc@73 679 + {n1 = 0 \/ n2 = 0}.
adamc@73 680 refine (fun n1 n2 =>
adamc@212 681 m1 <-- pred_strong8 n1;
adamc@212 682 m2 <-- pred_strong8 n2;
adam@335 683 [||(m1, m2)||]); tauto.
adamc@73 684 Defined.
adam@337 685 (* end thide *)
adamc@72 686
adam@392 687 (** This example demonstrates how judicious selection of notations can hide complexities in the rich types of programs. *)
adam@392 688
adamc@72 689
adamc@72 690 (** * A Type-Checking Example *)
adamc@72 691
adam@335 692 (** We can apply these specification types to build a certified type checker for a simple expression language. *)
adamc@75 693
adamc@72 694 Inductive exp : Set :=
adamc@72 695 | Nat : nat -> exp
adamc@72 696 | Plus : exp -> exp -> exp
adamc@72 697 | Bool : bool -> exp
adamc@72 698 | And : exp -> exp -> exp.
adamc@72 699
adamc@75 700 (** We define a simple language of types and its typing rules, in the style introduced in Chapter 4. *)
adamc@75 701
adamc@72 702 Inductive type : Set := TNat | TBool.
adamc@72 703
adamc@72 704 Inductive hasType : exp -> type -> Prop :=
adamc@72 705 | HtNat : forall n,
adamc@72 706 hasType (Nat n) TNat
adamc@72 707 | HtPlus : forall e1 e2,
adamc@72 708 hasType e1 TNat
adamc@72 709 -> hasType e2 TNat
adamc@72 710 -> hasType (Plus e1 e2) TNat
adamc@72 711 | HtBool : forall b,
adamc@72 712 hasType (Bool b) TBool
adamc@72 713 | HtAnd : forall e1 e2,
adamc@72 714 hasType e1 TBool
adamc@72 715 -> hasType e2 TBool
adamc@72 716 -> hasType (And e1 e2) TBool.
adamc@72 717
adamc@75 718 (** It will be helpful to have a function for comparing two types. We build one using [decide equality]. *)
adamc@75 719
adamc@77 720 (* begin thide *)
adamc@75 721 Definition eq_type_dec : forall t1 t2 : type, {t1 = t2} + {t1 <> t2}.
adamc@72 722 decide equality.
adamc@72 723 Defined.
adamc@72 724
adam@423 725 (** Another notation complements the monadic notation for [maybe] that we defined earlier. Sometimes we want to include "assertions" in our procedures. That is, we want to run a decision procedure and fail if it fails; otherwise, we want to continue, with the proof that it produced made available to us. This infix notation captures that idea, for a procedure that returns an arbitrary two-constructor type. *)
adamc@75 726
adamc@73 727 Notation "e1 ;; e2" := (if e1 then e2 else ??)
adamc@73 728 (right associativity, at level 60).
adamc@73 729
adam@335 730 (** With that notation defined, we can implement a [typeCheck] function, whose code is only more complex than what we would write in ML because it needs to include some extra type annotations. Every [[|e|]] expression adds a [hasType] proof obligation, and [crush] makes short work of them when we add [hasType]'s constructors as hints. *)
adamc@77 731 (* end thide *)
adamc@75 732
adam@297 733 Definition typeCheck : forall e : exp, {{t | hasType e t}}.
adamc@77 734 (* begin thide *)
adamc@72 735 Hint Constructors hasType.
adamc@72 736
adamc@72 737 refine (fix F (e : exp) : {{t | hasType e t}} :=
adam@380 738 match e return {{t | hasType e t}} with
adam@335 739 | Nat _ => [|TNat|]
adamc@72 740 | Plus e1 e2 =>
adamc@72 741 t1 <- F e1;
adamc@72 742 t2 <- F e2;
adamc@72 743 eq_type_dec t1 TNat;;
adamc@72 744 eq_type_dec t2 TNat;;
adam@335 745 [|TNat|]
adam@335 746 | Bool _ => [|TBool|]
adamc@72 747 | And e1 e2 =>
adamc@72 748 t1 <- F e1;
adamc@72 749 t2 <- F e2;
adamc@72 750 eq_type_dec t1 TBool;;
adamc@72 751 eq_type_dec t2 TBool;;
adam@335 752 [|TBool|]
adamc@72 753 end); crush.
adamc@77 754 (* end thide *)
adamc@72 755 Defined.
adamc@72 756
adamc@75 757 (** Despite manipulating proofs, our type checker is easy to run. *)
adamc@75 758
adamc@72 759 Eval simpl in typeCheck (Nat 0).
adamc@212 760 (** %\vspace{-.15in}% [[
adam@335 761 = [|TNat|]
adamc@75 762 : {{t | hasType (Nat 0) t}}
adam@302 763 ]]
adam@302 764 *)
adamc@75 765
adamc@72 766 Eval simpl in typeCheck (Plus (Nat 1) (Nat 2)).
adamc@212 767 (** %\vspace{-.15in}% [[
adam@335 768 = [|TNat|]
adamc@75 769 : {{t | hasType (Plus (Nat 1) (Nat 2)) t}}
adam@302 770 ]]
adam@302 771 *)
adamc@75 772
adamc@72 773 Eval simpl in typeCheck (Plus (Nat 1) (Bool false)).
adamc@212 774 (** %\vspace{-.15in}% [[
adamc@75 775 = ??
adamc@75 776 : {{t | hasType (Plus (Nat 1) (Bool false)) t}}
adam@302 777 ]]
adamc@75 778
adam@442 779 The type checker also extracts to some reasonable OCaml code. *)
adamc@75 780
adamc@75 781 Extraction typeCheck.
adamc@75 782
adamc@75 783 (** %\begin{verbatim}
adamc@75 784 (** val typeCheck : exp -> type0 maybe **)
adamc@75 785
adamc@75 786 let rec typeCheck = function
adamc@75 787 | Nat n -> Found TNat
adamc@75 788 | Plus (e1, e2) ->
adamc@75 789 (match typeCheck e1 with
adamc@75 790 | Unknown -> Unknown
adamc@75 791 | Found t1 ->
adamc@75 792 (match typeCheck e2 with
adamc@75 793 | Unknown -> Unknown
adamc@75 794 | Found t2 ->
adamc@75 795 (match eq_type_dec t1 TNat with
adamc@75 796 | true ->
adamc@75 797 (match eq_type_dec t2 TNat with
adamc@75 798 | true -> Found TNat
adamc@75 799 | false -> Unknown)
adamc@75 800 | false -> Unknown)))
adamc@75 801 | Bool b -> Found TBool
adamc@75 802 | And (e1, e2) ->
adamc@75 803 (match typeCheck e1 with
adamc@75 804 | Unknown -> Unknown
adamc@75 805 | Found t1 ->
adamc@75 806 (match typeCheck e2 with
adamc@75 807 | Unknown -> Unknown
adamc@75 808 | Found t2 ->
adamc@75 809 (match eq_type_dec t1 TBool with
adamc@75 810 | true ->
adamc@75 811 (match eq_type_dec t2 TBool with
adamc@75 812 | true -> Found TBool
adamc@75 813 | false -> Unknown)
adamc@75 814 | false -> Unknown)))
adamc@75 815 \end{verbatim}%
adamc@75 816
adamc@75 817 #<pre>
adamc@75 818 (** val typeCheck : exp -> type0 maybe **)
adamc@75 819
adamc@75 820 let rec typeCheck = function
adamc@75 821 | Nat n -> Found TNat
adamc@75 822 | Plus (e1, e2) ->
adamc@75 823 (match typeCheck e1 with
adamc@75 824 | Unknown -> Unknown
adamc@75 825 | Found t1 ->
adamc@75 826 (match typeCheck e2 with
adamc@75 827 | Unknown -> Unknown
adamc@75 828 | Found t2 ->
adamc@75 829 (match eq_type_dec t1 TNat with
adamc@75 830 | true ->
adamc@75 831 (match eq_type_dec t2 TNat with
adamc@75 832 | true -> Found TNat
adamc@75 833 | false -> Unknown)
adamc@75 834 | false -> Unknown)))
adamc@75 835 | Bool b -> Found TBool
adamc@75 836 | And (e1, e2) ->
adamc@75 837 (match typeCheck e1 with
adamc@75 838 | Unknown -> Unknown
adamc@75 839 | Found t1 ->
adamc@75 840 (match typeCheck e2 with
adamc@75 841 | Unknown -> Unknown
adamc@75 842 | Found t2 ->
adamc@75 843 (match eq_type_dec t1 TBool with
adamc@75 844 | true ->
adamc@75 845 (match eq_type_dec t2 TBool with
adamc@75 846 | true -> Found TBool
adamc@75 847 | false -> Unknown)
adamc@75 848 | false -> Unknown)))
adamc@75 849 </pre># *)
adamc@75 850
adamc@75 851 (** %\smallskip%
adamc@75 852
adam@423 853 We can adapt this implementation to use [sumor], so that we know our type-checker only fails on ill-typed inputs. First, we define an analogue to the "assertion" notation. *)
adamc@73 854
adamc@77 855 (* begin thide *)
adamc@73 856 Notation "e1 ;;; e2" := (if e1 then e2 else !!)
adamc@73 857 (right associativity, at level 60).
adamc@73 858
adamc@75 859 (** Next, we prove a helpful lemma, which states that a given expression can have at most one type. *)
adamc@75 860
adamc@75 861 Lemma hasType_det : forall e t1,
adamc@73 862 hasType e t1
adam@335 863 -> forall t2, hasType e t2
adamc@73 864 -> t1 = t2.
adamc@73 865 induction 1; inversion 1; crush.
adamc@73 866 Qed.
adamc@73 867
adamc@75 868 (** Now we can define the type-checker. Its type expresses that it only fails on untypable expressions. *)
adamc@75 869
adamc@77 870 (* end thide *)
adam@297 871 Definition typeCheck' : forall e : exp, {t : type | hasType e t} + {forall t, ~ hasType e t}.
adamc@77 872 (* begin thide *)
adamc@73 873 Hint Constructors hasType.
adamc@75 874 (** We register all of the typing rules as hints. *)
adamc@75 875
adamc@73 876 Hint Resolve hasType_det.
adam@335 877 (** The lemma [hasType_det] will also be useful for proving proof obligations with contradictory contexts. Since its statement includes [forall]-bound variables that do not appear in its conclusion, only [eauto] will apply this hint. *)
adamc@73 878
adamc@75 879 (** Finally, the implementation of [typeCheck] can be transcribed literally, simply switching notations as needed. *)
adamc@212 880
adamc@212 881 refine (fix F (e : exp) : {t : type | hasType e t} + {forall t, ~ hasType e t} :=
adam@380 882 match e return {t : type | hasType e t} + {forall t, ~ hasType e t} with
adam@335 883 | Nat _ => [||TNat||]
adamc@73 884 | Plus e1 e2 =>
adamc@73 885 t1 <-- F e1;
adamc@73 886 t2 <-- F e2;
adamc@73 887 eq_type_dec t1 TNat;;;
adamc@73 888 eq_type_dec t2 TNat;;;
adam@335 889 [||TNat||]
adam@335 890 | Bool _ => [||TBool||]
adamc@73 891 | And e1 e2 =>
adamc@73 892 t1 <-- F e1;
adamc@73 893 t2 <-- F e2;
adamc@73 894 eq_type_dec t1 TBool;;;
adamc@73 895 eq_type_dec t2 TBool;;;
adam@335 896 [||TBool||]
adamc@73 897 end); clear F; crush' tt hasType; eauto.
adamc@75 898
adam@335 899 (** We clear [F], the local name for the recursive function, to avoid strange proofs that refer to recursive calls that we never make. The [crush] variant %\index{tactics!crush'}%[crush'] helps us by performing automatic inversion on instances of the predicates specified in its second argument. Once we throw in [eauto] to apply [hasType_det] for us, we have discharged all the subgoals. *)
adamc@77 900 (* end thide *)
adamc@212 901
adamc@212 902
adamc@73 903 Defined.
adamc@73 904
adamc@75 905 (** The short implementation here hides just how time-saving automation is. Every use of one of the notations adds a proof obligation, giving us 12 in total. Most of these obligations require multiple inversions and either uses of [hasType_det] or applications of [hasType] rules.
adamc@75 906
adam@335 907 Our new function remains easy to test: *)
adamc@75 908
adamc@73 909 Eval simpl in typeCheck' (Nat 0).
adamc@212 910 (** %\vspace{-.15in}% [[
adam@335 911 = [||TNat||]
adamc@75 912 : {t : type | hasType (Nat 0) t} +
adamc@75 913 {(forall t : type, ~ hasType (Nat 0) t)}
adam@302 914 ]]
adam@302 915 *)
adamc@75 916
adamc@73 917 Eval simpl in typeCheck' (Plus (Nat 1) (Nat 2)).
adamc@212 918 (** %\vspace{-.15in}% [[
adam@335 919 = [||TNat||]
adamc@75 920 : {t : type | hasType (Plus (Nat 1) (Nat 2)) t} +
adamc@75 921 {(forall t : type, ~ hasType (Plus (Nat 1) (Nat 2)) t)}
adam@302 922 ]]
adam@302 923 *)
adamc@75 924
adamc@73 925 Eval simpl in typeCheck' (Plus (Nat 1) (Bool false)).
adamc@212 926 (** %\vspace{-.15in}% [[
adamc@75 927 = !!
adamc@75 928 : {t : type | hasType (Plus (Nat 1) (Bool false)) t} +
adamc@75 929 {(forall t : type, ~ hasType (Plus (Nat 1) (Bool false)) t)}
adam@302 930 ]]
adam@335 931
adam@442 932 The results of simplifying calls to [typeCheck'] look deceptively similar to the results for [typeCheck], but now the types of the results provide more information. *)