annotate src/Subset.v @ 403:1edeec5d5d0c

Typesetting pass over Subset
author Adam Chlipala <adam@chlipala.net>
date Fri, 08 Jun 2012 13:41:33 -0400
parents 05efde66559d
children d3a40c044ab4
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@403 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@282 88
adam@294 89 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 90
adamc@212 91 [[
adamc@70 92 Definition pred_strong1' (n : nat) (pf : n > 0) : nat :=
adamc@70 93 match n with
adamc@70 94 | O => match zgtz pf with end
adamc@70 95 | S n' => n'
adamc@70 96 end.
adam@335 97 ]]
adamc@70 98
adam@335 99 <<
adamc@70 100 Error: In environment
adamc@70 101 n : nat
adamc@70 102 pf : n > 0
adamc@70 103 The term "pf" has type "n > 0" while it is expected to have type
adamc@70 104 "0 > 0"
adam@335 105 >>
adamc@70 106
adamc@212 107 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 108
adam@398 109 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 110
adam@335 111 We are lucky that Coq's heuristics infer the [return] clause (specifically, [return n > 0 -> nat]) for us in this case. *)
adam@335 112
adam@335 113 Definition pred_strong1' (n : nat) : n > 0 -> nat :=
adam@335 114 match n return n > 0 -> nat with
adam@335 115 | O => fun pf : 0 > 0 => match zgtz pf with end
adam@335 116 | S n' => fun _ => n'
adam@335 117 end.
adam@335 118
adam@403 119 (** 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 120
adamc@70 121 Let us now take a look at the OCaml code Coq generates for [pred_strong1]. *)
adamc@70 122
adamc@70 123 Extraction pred_strong1.
adamc@70 124
adamc@70 125 (** %\begin{verbatim}
adamc@70 126 (** val pred_strong1 : nat -> nat **)
adamc@70 127
adamc@70 128 let pred_strong1 = function
adamc@70 129 | O -> assert false (* absurd case *)
adamc@70 130 | S n' -> n'
adamc@70 131 \end{verbatim}%
adamc@70 132
adamc@70 133 #<pre>
adamc@70 134 (** val pred_strong1 : nat -> nat **)
adamc@70 135
adamc@70 136 let pred_strong1 = function
adamc@70 137 | O -> assert false (* absurd case *)
adamc@70 138 | S n' -> n'
adamc@70 139 </pre># *)
adamc@70 140
adamc@70 141 (** 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 142
adam@403 143 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 144
adamc@70 145 Print sig.
adamc@212 146 (** %\vspace{-.15in}% [[
adamc@70 147 Inductive sig (A : Type) (P : A -> Prop) : Type :=
adamc@70 148 exist : forall x : A, P x -> sig P
adamc@212 149
adamc@70 150 ]]
adamc@70 151
adam@335 152 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 153
adamc@70 154 We rewrite [pred_strong1], using some syntactic sugar for subset types. *)
adamc@70 155
adamc@70 156 Locate "{ _ : _ | _ }".
adamc@212 157 (** %\vspace{-.15in}% [[
adam@335 158 Notation
adamc@70 159 "{ x : A | P }" := sig (fun x : A => P)
adam@302 160 ]]
adam@302 161 *)
adamc@70 162
adamc@70 163 Definition pred_strong2 (s : {n : nat | n > 0}) : nat :=
adamc@70 164 match s with
adamc@70 165 | exist O pf => match zgtz pf with end
adamc@70 166 | exist (S n') _ => n'
adamc@70 167 end.
adamc@70 168
adam@335 169 (** 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 170
adam@282 171 Eval compute in pred_strong2 (exist _ 2 two_gt0).
adam@282 172 (** %\vspace{-.15in}% [[
adam@282 173 = 1
adam@282 174 : nat
adam@302 175 ]]
adam@302 176 *)
adam@282 177
adamc@70 178 Extraction pred_strong2.
adamc@70 179
adamc@70 180 (** %\begin{verbatim}
adamc@70 181 (** val pred_strong2 : nat -> nat **)
adamc@70 182
adamc@70 183 let pred_strong2 = function
adamc@70 184 | O -> assert false (* absurd case *)
adamc@70 185 | S n' -> n'
adamc@70 186 \end{verbatim}%
adamc@70 187
adamc@70 188 #<pre>
adamc@70 189 (** val pred_strong2 : nat -> nat **)
adamc@70 190
adamc@70 191 let pred_strong2 = function
adamc@70 192 | O -> assert false (* absurd case *)
adamc@70 193 | S n' -> n'
adamc@70 194 </pre>#
adamc@70 195
adamc@70 196 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 197
adamc@70 198 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 199
adamc@70 200 Definition pred_strong3 (s : {n : nat | n > 0}) : {m : nat | proj1_sig s = S m} :=
adamc@70 201 match s return {m : nat | proj1_sig s = S m} with
adamc@70 202 | exist 0 pf => match zgtz pf with end
adamc@212 203 | exist (S n') pf => exist _ n' (refl_equal _)
adamc@70 204 end.
adamc@70 205
adam@282 206 Eval compute in pred_strong3 (exist _ 2 two_gt0).
adam@282 207 (** %\vspace{-.15in}% [[
adam@282 208 = exist (fun m : nat => 2 = S m) 1 (refl_equal 2)
adam@282 209 : {m : nat | proj1_sig (exist (lt 0) 2 two_gt0) = S m}
adam@335 210 ]]
adam@302 211 *)
adam@282 212
adam@335 213 (** 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 214
adamc@70 215 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 216
adamc@70 217 Extraction pred_strong3.
adamc@70 218
adamc@70 219 (** %\begin{verbatim}
adamc@70 220 (** val pred_strong3 : nat -> nat **)
adamc@70 221
adamc@70 222 let pred_strong3 = function
adamc@70 223 | O -> assert false (* absurd case *)
adamc@70 224 | S n' -> n'
adamc@70 225 \end{verbatim}%
adamc@70 226
adamc@70 227 #<pre>
adamc@70 228 (** val pred_strong3 : nat -> nat **)
adamc@70 229
adamc@70 230 let pred_strong3 = function
adamc@70 231 | O -> assert false (* absurd case *)
adamc@70 232 | S n' -> n'
adamc@70 233 </pre>#
adamc@70 234
adam@335 235 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 236
adam@297 237 Definition pred_strong4 : forall n : nat, n > 0 -> {m : nat | n = S m}.
adamc@70 238 refine (fun n =>
adamc@212 239 match n with
adamc@70 240 | O => fun _ => False_rec _ _
adamc@70 241 | S n' => fun _ => exist _ n' _
adamc@70 242 end).
adamc@212 243
adamc@77 244 (* begin thide *)
adam@335 245 (** 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 246
adam@335 247 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 248
adam@335 249 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@335 250 [[
adamc@70 251
adamc@70 252 n : nat
adamc@70 253 _ : 0 > 0
adamc@70 254 ============================
adamc@70 255 False
adam@335 256 ]]
adam@335 257 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@335 258 [[
adamc@70 259 S n' = S n'
adamc@70 260 ]]
adamc@70 261
adamc@70 262 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 263
adamc@70 264 Undo.
adamc@70 265 refine (fun n =>
adamc@212 266 match n with
adamc@70 267 | O => fun _ => False_rec _ _
adamc@70 268 | S n' => fun _ => exist _ n' _
adamc@70 269 end); crush.
adamc@77 270 (* end thide *)
adamc@70 271 Defined.
adamc@70 272
adam@403 273 (** 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 274
adamc@70 275 Print pred_strong4.
adamc@212 276 (** %\vspace{-.15in}% [[
adamc@70 277 pred_strong4 =
adamc@70 278 fun n : nat =>
adamc@70 279 match n as n0 return (n0 > 0 -> {m : nat | n0 = S m}) with
adamc@70 280 | 0 =>
adamc@70 281 fun _ : 0 > 0 =>
adamc@70 282 False_rec {m : nat | 0 = S m}
adamc@70 283 (Bool.diff_false_true
adamc@70 284 (Bool.absurd_eq_true false
adamc@70 285 (Bool.diff_false_true
adamc@70 286 (Bool.absurd_eq_true false (pred_strong4_subproof n _)))))
adamc@70 287 | S n' =>
adamc@70 288 fun _ : S n' > 0 =>
adamc@70 289 exist (fun m : nat => S n' = S m) n' (refl_equal (S n'))
adamc@70 290 end
adamc@70 291 : forall n : nat, n > 0 -> {m : nat | n = S m}
adamc@212 292
adamc@70 293 ]]
adamc@70 294
adam@282 295 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 296
adam@282 297 Eval compute in pred_strong4 two_gt0.
adam@282 298 (** %\vspace{-.15in}% [[
adam@282 299 = exist (fun m : nat => 2 = S m) 1 (refl_equal 2)
adam@282 300 : {m : nat | 2 = S m}
adam@282 301 ]]
adam@282 302
adam@335 303 A tactic modifier called %\index{tactics!abstract}%[abstract] can be helpful for producing shorter terms, by automatically abstracting subgoals into named lemmas. *)
adam@335 304
adam@335 305 (* begin thide *)
adam@335 306 Definition pred_strong4' : forall n : nat, n > 0 -> {m : nat | n = S m}.
adam@335 307 refine (fun n =>
adam@335 308 match n with
adam@335 309 | O => fun _ => False_rec _ _
adam@335 310 | S n' => fun _ => exist _ n' _
adam@335 311 end); abstract crush.
adam@335 312 Defined.
adam@335 313
adam@335 314 Print pred_strong4'.
adam@335 315 (* end thide *)
adam@335 316
adam@335 317 (** %\vspace{-.15in}% [[
adam@335 318 pred_strong4' =
adam@335 319 fun n : nat =>
adam@335 320 match n as n0 return (n0 > 0 -> {m : nat | n0 = S m}) with
adam@335 321 | 0 =>
adam@335 322 fun _H : 0 > 0 =>
adam@335 323 False_rec {m : nat | 0 = S m} (pred_strong4'_subproof n _H)
adam@335 324 | S n' =>
adam@335 325 fun _H : S n' > 0 =>
adam@335 326 exist (fun m : nat => S n' = S m) n' (pred_strong4'_subproof0 n _H)
adam@335 327 end
adam@335 328 : forall n : nat, n > 0 -> {m : nat | n = S m}
adam@335 329 ]]
adam@335 330
adam@338 331 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 332
adamc@70 333 Notation "!" := (False_rec _ _).
adamc@70 334 Notation "[ e ]" := (exist _ e _).
adamc@70 335
adam@297 336 Definition pred_strong5 : forall n : nat, n > 0 -> {m : nat | n = S m}.
adamc@70 337 refine (fun n =>
adamc@212 338 match n with
adamc@70 339 | O => fun _ => !
adamc@70 340 | S n' => fun _ => [n']
adamc@70 341 end); crush.
adamc@70 342 Defined.
adamc@71 343
adam@282 344 (** By default, notations are also used in pretty-printing terms, including results of evaluation. *)
adam@282 345
adam@282 346 Eval compute in pred_strong5 two_gt0.
adam@282 347 (** %\vspace{-.15in}% [[
adam@282 348 = [1]
adam@282 349 : {m : nat | 2 = S m}
adam@282 350 ]]
adam@282 351
adam@335 352 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 353
adamc@212 354 Obligation Tactic := crush.
adamc@212 355
adamc@212 356 Program Definition pred_strong6 (n : nat) (_ : n > 0) : {m : nat | n = S m} :=
adamc@212 357 match n with
adamc@212 358 | O => _
adamc@212 359 | S n' => n'
adamc@212 360 end.
adamc@212 361
adam@335 362 (** 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 363
adam@282 364 Eval compute in pred_strong6 two_gt0.
adam@282 365 (** %\vspace{-.15in}% [[
adam@282 366 = [1]
adam@282 367 : {m : nat | 2 = S m}
adam@302 368 ]]
adam@335 369
adam@335 370 In this case, we see that the new definition yields the same computational behavior as before. *)
adam@282 371
adamc@71 372
adamc@71 373 (** * Decidable Proposition Types *)
adamc@71 374
adam@335 375 (** 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 376
adamc@71 377 Print sumbool.
adamc@212 378 (** %\vspace{-.15in}% [[
adamc@71 379 Inductive sumbool (A : Prop) (B : Prop) : Set :=
adamc@71 380 left : A -> {A} + {B} | right : B -> {A} + {B}
adamc@212 381 ]]
adamc@71 382
adamc@212 383 We can define some notations to make working with [sumbool] more convenient. *)
adamc@71 384
adamc@71 385 Notation "'Yes'" := (left _ _).
adamc@71 386 Notation "'No'" := (right _ _).
adamc@71 387 Notation "'Reduce' x" := (if x then Yes else No) (at level 50).
adamc@71 388
adamc@71 389 (** The [Reduce] 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 390
adamc@71 391 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 392
adam@297 393 Definition eq_nat_dec : forall n m : nat, {n = m} + {n <> m}.
adamc@212 394 refine (fix f (n m : nat) : {n = m} + {n <> m} :=
adamc@212 395 match n, m with
adamc@71 396 | O, O => Yes
adamc@71 397 | S n', S m' => Reduce (f n' m')
adamc@71 398 | _, _ => No
adamc@71 399 end); congruence.
adamc@71 400 Defined.
adamc@71 401
adam@282 402 Eval compute in eq_nat_dec 2 2.
adam@282 403 (** %\vspace{-.15in}% [[
adam@282 404 = Yes
adam@282 405 : {2 = 2} + {2 <> 2}
adam@302 406 ]]
adam@302 407 *)
adam@282 408
adam@282 409 Eval compute in eq_nat_dec 2 3.
adam@282 410 (** %\vspace{-.15in}% [[
adam@282 411 = No
adam@341 412 : {2 = 3} + {2 <> 3}
adam@302 413 ]]
adam@302 414 *)
adam@282 415
adam@335 416 (** Note that the [Yes] and [No] notations are hiding proofs establishing the correctness of the outputs.
adam@335 417
adam@335 418 Our definition extracts to reasonable OCaml code. *)
adamc@71 419
adamc@71 420 Extraction eq_nat_dec.
adamc@71 421
adamc@71 422 (** %\begin{verbatim}
adamc@71 423 (** val eq_nat_dec : nat -> nat -> sumbool **)
adamc@71 424
adamc@71 425 let rec eq_nat_dec n m =
adamc@71 426 match n with
adamc@71 427 | O -> (match m with
adamc@71 428 | O -> Left
adamc@71 429 | S n0 -> Right)
adamc@71 430 | S n' -> (match m with
adamc@71 431 | O -> Right
adamc@71 432 | S m' -> eq_nat_dec n' m')
adamc@71 433 \end{verbatim}%
adamc@71 434
adamc@71 435 #<pre>
adamc@71 436 (** val eq_nat_dec : nat -> nat -> sumbool **)
adamc@71 437
adamc@71 438 let rec eq_nat_dec n m =
adamc@71 439 match n with
adamc@71 440 | O -> (match m with
adamc@71 441 | O -> Left
adamc@71 442 | S n0 -> Right)
adamc@71 443 | S n' -> (match m with
adamc@71 444 | O -> Right
adamc@71 445 | S m' -> eq_nat_dec n' m')
adamc@71 446 </pre>#
adamc@71 447
adam@335 448 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 449
adamc@71 450 Definition eq_nat_dec' (n m : nat) : {n = m} + {n <> m}.
adamc@71 451 decide equality.
adamc@71 452 Defined.
adamc@71 453
adam@335 454 (** 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 %\texttt{%#<tt>#Left#</tt>#%}% and %\texttt{%#<tt>#Right#</tt>#%}% 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 455
adamc@71 456 Extract Inductive sumbool => "bool" ["true" "false"].
adamc@71 457 Extraction eq_nat_dec'.
adamc@71 458
adamc@71 459 (** %\begin{verbatim}
adamc@71 460 (** val eq_nat_dec' : nat -> nat -> bool **)
adamc@71 461
adamc@71 462 let rec eq_nat_dec' n m0 =
adamc@71 463 match n with
adamc@71 464 | O -> (match m0 with
adamc@71 465 | O -> true
adamc@71 466 | S n0 -> false)
adamc@71 467 | S n0 -> (match m0 with
adamc@71 468 | O -> false
adamc@71 469 | S n1 -> eq_nat_dec' n0 n1)
adamc@71 470 \end{verbatim}%
adamc@71 471
adamc@71 472 #<pre>
adamc@71 473 (** val eq_nat_dec' : nat -> nat -> bool **)
adamc@71 474
adamc@71 475 let rec eq_nat_dec' n m0 =
adamc@71 476 match n with
adamc@71 477 | O -> (match m0 with
adamc@71 478 | O -> true
adamc@71 479 | S n0 -> false)
adamc@71 480 | S n0 -> (match m0 with
adamc@71 481 | O -> false
adamc@71 482 | S n1 -> eq_nat_dec' n0 n1)
adamc@71 483 </pre># *)
adamc@72 484
adamc@72 485 (** %\smallskip%
adamc@72 486
adam@292 487 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 488
adam@337 489 (* EX: Write a function that decides if an element belongs to a list. *)
adam@337 490
adamc@77 491 (* begin thide *)
adamc@204 492 Notation "x || y" := (if x then Yes else Reduce y).
adamc@72 493
adamc@72 494 (** 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 495
adamc@72 496 Section In_dec.
adamc@72 497 Variable A : Set.
adamc@72 498 Variable A_eq_dec : forall x y : A, {x = y} + {x <> y}.
adamc@72 499
adamc@72 500 (** The final function is easy to write using the techniques we have developed so far. *)
adamc@72 501
adamc@212 502 Definition In_dec : forall (x : A) (ls : list A), {In x ls} + {~ In x ls}.
adamc@212 503 refine (fix f (x : A) (ls : list A) : {In x ls} + {~ In x ls} :=
adamc@212 504 match ls with
adamc@72 505 | nil => No
adamc@72 506 | x' :: ls' => A_eq_dec x x' || f x ls'
adamc@72 507 end); crush.
adam@282 508 Defined.
adamc@72 509 End In_dec.
adamc@72 510
adam@282 511 Eval compute in In_dec eq_nat_dec 2 (1 :: 2 :: nil).
adam@282 512 (** %\vspace{-.15in}% [[
adam@282 513 = Yes
adam@282 514 : {In 2 (1 :: 2 :: nil)} + {~ In 2 (1 :: 2 :: nil)}
adam@302 515 ]]
adam@302 516 *)
adam@282 517
adam@282 518 Eval compute in In_dec eq_nat_dec 3 (1 :: 2 :: nil).
adam@282 519 (** %\vspace{-.15in}% [[
adam@282 520 = No
adam@282 521 : {In 3 (1 :: 2 :: nil)} + {~ In 3 (1 :: 2 :: nil)}
adam@302 522 ]]
adam@302 523 *)
adam@282 524
adamc@72 525 (** [In_dec] has a reasonable extraction to OCaml. *)
adamc@72 526
adamc@72 527 Extraction In_dec.
adamc@77 528 (* end thide *)
adamc@72 529
adamc@72 530 (** %\begin{verbatim}
adamc@72 531 (** val in_dec : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 list -> bool **)
adamc@72 532
adamc@72 533 let rec in_dec a_eq_dec x = function
adamc@72 534 | Nil -> false
adamc@72 535 | Cons (x', ls') ->
adamc@72 536 (match a_eq_dec x x' with
adamc@72 537 | true -> true
adamc@72 538 | false -> in_dec a_eq_dec x ls')
adamc@72 539 \end{verbatim}%
adamc@72 540
adamc@72 541 #<pre>
adamc@72 542 (** val in_dec : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 list -> bool **)
adamc@72 543
adamc@72 544 let rec in_dec a_eq_dec x = function
adamc@72 545 | Nil -> false
adamc@72 546 | Cons (x', ls') ->
adamc@72 547 (match a_eq_dec x x' with
adamc@72 548 | true -> true
adamc@72 549 | false -> in_dec a_eq_dec x ls')
adam@403 550 </pre>#
adam@403 551
adam@403 552 This is more or the less code for the corresponding function from the OCaml standard library. *)
adamc@72 553
adamc@72 554
adamc@72 555 (** * Partial Subset Types *)
adamc@72 556
adam@335 557 (** 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 558
adamc@89 559 Inductive maybe (A : Set) (P : A -> Prop) : Set :=
adamc@72 560 | Unknown : maybe P
adamc@72 561 | Found : forall x : A, P x -> maybe P.
adamc@72 562
adamc@73 563 (** We can define some new notations, analogous to those we defined for subset types. *)
adamc@73 564
adamc@72 565 Notation "{{ x | P }}" := (maybe (fun x => P)).
adamc@72 566 Notation "??" := (Unknown _).
adam@335 567 Notation "[| x |]" := (Found _ x _).
adamc@72 568
adamc@73 569 (** Now our next version of [pred] is trivial to write. *)
adamc@73 570
adam@297 571 Definition pred_strong7 : forall n : nat, {{m | n = S m}}.
adamc@73 572 refine (fun n =>
adam@380 573 match n return {{m | n = S m}} with
adamc@73 574 | O => ??
adam@335 575 | S n' => [|n'|]
adamc@73 576 end); trivial.
adamc@73 577 Defined.
adamc@73 578
adam@282 579 Eval compute in pred_strong7 2.
adam@282 580 (** %\vspace{-.15in}% [[
adam@335 581 = [|1|]
adam@282 582 : {{m | 2 = S m}}
adam@335 583 ]]
adam@302 584 *)
adam@282 585
adam@282 586 Eval compute in pred_strong7 0.
adam@282 587 (** %\vspace{-.15in}% [[
adam@282 588 = ??
adam@282 589 : {{m | 0 = S m}}
adam@282 590 ]]
adam@282 591
adam@335 592 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 593
adamc@73 594 Print sumor.
adamc@212 595 (** %\vspace{-.15in}% [[
adamc@73 596 Inductive sumor (A : Type) (B : Prop) : Type :=
adamc@73 597 inleft : A -> A + {B} | inright : B -> A + {B}
adam@302 598 ]]
adam@302 599 *)
adamc@73 600
adamc@73 601 (** 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 602
adamc@73 603 Notation "!!" := (inright _ _).
adam@335 604 Notation "[|| x ||]" := (inleft _ [x]).
adamc@73 605
adam@335 606 (** 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 607
adam@297 608 Definition pred_strong8 : forall n : nat, {m : nat | n = S m} + {n = 0}.
adamc@73 609 refine (fun n =>
adamc@212 610 match n with
adamc@73 611 | O => !!
adam@335 612 | S n' => [||n'||]
adamc@73 613 end); trivial.
adamc@73 614 Defined.
adamc@73 615
adam@282 616 Eval compute in pred_strong8 2.
adam@282 617 (** %\vspace{-.15in}% [[
adam@335 618 = [||1||]
adam@282 619 : {m : nat | 2 = S m} + {2 = 0}
adam@302 620 ]]
adam@302 621 *)
adam@282 622
adam@282 623 Eval compute in pred_strong8 0.
adam@282 624 (** %\vspace{-.15in}% [[
adam@282 625 = !!
adam@282 626 : {m : nat | 0 = S m} + {0 = 0}
adam@302 627 ]]
adam@302 628 *)
adam@282 629
adam@335 630 (** As with our other maximally expressive [pred] function, we arrive at quite simple output values, thanks to notations. *)
adam@335 631
adamc@73 632
adamc@73 633 (** * Monadic Notations *)
adamc@73 634
adam@335 635 (** 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 636
adamc@72 637 Notation "x <- e1 ; e2" := (match e1 with
adamc@72 638 | Unknown => ??
adamc@72 639 | Found x _ => e2
adamc@72 640 end)
adamc@72 641 (right associativity, at level 60).
adamc@72 642
adam@398 643 (** 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 644
adam@335 645 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 646
adam@337 647 (* EX: Write a function that tries to compute predecessors of two [nat]s at once. *)
adam@337 648
adam@337 649 (* begin thide *)
adam@297 650 Definition doublePred : forall n1 n2 : nat, {{p | n1 = S (fst p) /\ n2 = S (snd p)}}.
adamc@73 651 refine (fun n1 n2 =>
adamc@212 652 m1 <- pred_strong7 n1;
adamc@212 653 m2 <- pred_strong7 n2;
adam@335 654 [|(m1, m2)|]); tauto.
adamc@73 655 Defined.
adam@337 656 (* end thide *)
adamc@73 657
adam@392 658 (** 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 659
adamc@73 660 Notation "x <-- e1 ; e2" := (match e1 with
adamc@73 661 | inright _ => !!
adamc@73 662 | inleft (exist x _) => e2
adamc@73 663 end)
adamc@73 664 (right associativity, at level 60).
adamc@73 665
adamc@73 666 (** printing * $\times$ *)
adamc@73 667
adam@337 668 (* EX: Write a more expressively typed version of the last exercise. *)
adam@337 669
adam@337 670 (* begin thide *)
adam@297 671 Definition doublePred' : forall n1 n2 : nat,
adam@297 672 {p : nat * nat | n1 = S (fst p) /\ n2 = S (snd p)}
adamc@73 673 + {n1 = 0 \/ n2 = 0}.
adamc@73 674 refine (fun n1 n2 =>
adamc@212 675 m1 <-- pred_strong8 n1;
adamc@212 676 m2 <-- pred_strong8 n2;
adam@335 677 [||(m1, m2)||]); tauto.
adamc@73 678 Defined.
adam@337 679 (* end thide *)
adamc@72 680
adam@392 681 (** This example demonstrates how judicious selection of notations can hide complexities in the rich types of programs. *)
adam@392 682
adamc@72 683
adamc@72 684 (** * A Type-Checking Example *)
adamc@72 685
adam@335 686 (** We can apply these specification types to build a certified type checker for a simple expression language. *)
adamc@75 687
adamc@72 688 Inductive exp : Set :=
adamc@72 689 | Nat : nat -> exp
adamc@72 690 | Plus : exp -> exp -> exp
adamc@72 691 | Bool : bool -> exp
adamc@72 692 | And : exp -> exp -> exp.
adamc@72 693
adamc@75 694 (** We define a simple language of types and its typing rules, in the style introduced in Chapter 4. *)
adamc@75 695
adamc@72 696 Inductive type : Set := TNat | TBool.
adamc@72 697
adamc@72 698 Inductive hasType : exp -> type -> Prop :=
adamc@72 699 | HtNat : forall n,
adamc@72 700 hasType (Nat n) TNat
adamc@72 701 | HtPlus : forall e1 e2,
adamc@72 702 hasType e1 TNat
adamc@72 703 -> hasType e2 TNat
adamc@72 704 -> hasType (Plus e1 e2) TNat
adamc@72 705 | HtBool : forall b,
adamc@72 706 hasType (Bool b) TBool
adamc@72 707 | HtAnd : forall e1 e2,
adamc@72 708 hasType e1 TBool
adamc@72 709 -> hasType e2 TBool
adamc@72 710 -> hasType (And e1 e2) TBool.
adamc@72 711
adamc@75 712 (** It will be helpful to have a function for comparing two types. We build one using [decide equality]. *)
adamc@75 713
adamc@77 714 (* begin thide *)
adamc@75 715 Definition eq_type_dec : forall t1 t2 : type, {t1 = t2} + {t1 <> t2}.
adamc@72 716 decide equality.
adamc@72 717 Defined.
adamc@72 718
adam@292 719 (** 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 720
adamc@73 721 Notation "e1 ;; e2" := (if e1 then e2 else ??)
adamc@73 722 (right associativity, at level 60).
adamc@73 723
adam@335 724 (** 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 725 (* end thide *)
adamc@75 726
adam@297 727 Definition typeCheck : forall e : exp, {{t | hasType e t}}.
adamc@77 728 (* begin thide *)
adamc@72 729 Hint Constructors hasType.
adamc@72 730
adamc@72 731 refine (fix F (e : exp) : {{t | hasType e t}} :=
adam@380 732 match e return {{t | hasType e t}} with
adam@335 733 | Nat _ => [|TNat|]
adamc@72 734 | Plus e1 e2 =>
adamc@72 735 t1 <- F e1;
adamc@72 736 t2 <- F e2;
adamc@72 737 eq_type_dec t1 TNat;;
adamc@72 738 eq_type_dec t2 TNat;;
adam@335 739 [|TNat|]
adam@335 740 | Bool _ => [|TBool|]
adamc@72 741 | And e1 e2 =>
adamc@72 742 t1 <- F e1;
adamc@72 743 t2 <- F e2;
adamc@72 744 eq_type_dec t1 TBool;;
adamc@72 745 eq_type_dec t2 TBool;;
adam@335 746 [|TBool|]
adamc@72 747 end); crush.
adamc@77 748 (* end thide *)
adamc@72 749 Defined.
adamc@72 750
adamc@75 751 (** Despite manipulating proofs, our type checker is easy to run. *)
adamc@75 752
adamc@72 753 Eval simpl in typeCheck (Nat 0).
adamc@212 754 (** %\vspace{-.15in}% [[
adam@335 755 = [|TNat|]
adamc@75 756 : {{t | hasType (Nat 0) t}}
adam@302 757 ]]
adam@302 758 *)
adamc@75 759
adamc@72 760 Eval simpl in typeCheck (Plus (Nat 1) (Nat 2)).
adamc@212 761 (** %\vspace{-.15in}% [[
adam@335 762 = [|TNat|]
adamc@75 763 : {{t | hasType (Plus (Nat 1) (Nat 2)) t}}
adam@302 764 ]]
adam@302 765 *)
adamc@75 766
adamc@72 767 Eval simpl in typeCheck (Plus (Nat 1) (Bool false)).
adamc@212 768 (** %\vspace{-.15in}% [[
adamc@75 769 = ??
adamc@75 770 : {{t | hasType (Plus (Nat 1) (Bool false)) t}}
adam@302 771 ]]
adam@302 772 *)
adamc@75 773
adam@335 774 (** The type checker also extracts to some reasonable OCaml code. *)
adamc@75 775
adamc@75 776 Extraction typeCheck.
adamc@75 777
adamc@75 778 (** %\begin{verbatim}
adamc@75 779 (** val typeCheck : exp -> type0 maybe **)
adamc@75 780
adamc@75 781 let rec typeCheck = function
adamc@75 782 | Nat n -> Found TNat
adamc@75 783 | Plus (e1, e2) ->
adamc@75 784 (match typeCheck e1 with
adamc@75 785 | Unknown -> Unknown
adamc@75 786 | Found t1 ->
adamc@75 787 (match typeCheck e2 with
adamc@75 788 | Unknown -> Unknown
adamc@75 789 | Found t2 ->
adamc@75 790 (match eq_type_dec t1 TNat with
adamc@75 791 | true ->
adamc@75 792 (match eq_type_dec t2 TNat with
adamc@75 793 | true -> Found TNat
adamc@75 794 | false -> Unknown)
adamc@75 795 | false -> Unknown)))
adamc@75 796 | Bool b -> Found TBool
adamc@75 797 | And (e1, e2) ->
adamc@75 798 (match typeCheck e1 with
adamc@75 799 | Unknown -> Unknown
adamc@75 800 | Found t1 ->
adamc@75 801 (match typeCheck e2 with
adamc@75 802 | Unknown -> Unknown
adamc@75 803 | Found t2 ->
adamc@75 804 (match eq_type_dec t1 TBool with
adamc@75 805 | true ->
adamc@75 806 (match eq_type_dec t2 TBool with
adamc@75 807 | true -> Found TBool
adamc@75 808 | false -> Unknown)
adamc@75 809 | false -> Unknown)))
adamc@75 810 \end{verbatim}%
adamc@75 811
adamc@75 812 #<pre>
adamc@75 813 (** val typeCheck : exp -> type0 maybe **)
adamc@75 814
adamc@75 815 let rec typeCheck = function
adamc@75 816 | Nat n -> Found TNat
adamc@75 817 | Plus (e1, e2) ->
adamc@75 818 (match typeCheck e1 with
adamc@75 819 | Unknown -> Unknown
adamc@75 820 | Found t1 ->
adamc@75 821 (match typeCheck e2 with
adamc@75 822 | Unknown -> Unknown
adamc@75 823 | Found t2 ->
adamc@75 824 (match eq_type_dec t1 TNat with
adamc@75 825 | true ->
adamc@75 826 (match eq_type_dec t2 TNat with
adamc@75 827 | true -> Found TNat
adamc@75 828 | false -> Unknown)
adamc@75 829 | false -> Unknown)))
adamc@75 830 | Bool b -> Found TBool
adamc@75 831 | And (e1, e2) ->
adamc@75 832 (match typeCheck e1 with
adamc@75 833 | Unknown -> Unknown
adamc@75 834 | Found t1 ->
adamc@75 835 (match typeCheck e2 with
adamc@75 836 | Unknown -> Unknown
adamc@75 837 | Found t2 ->
adamc@75 838 (match eq_type_dec t1 TBool with
adamc@75 839 | true ->
adamc@75 840 (match eq_type_dec t2 TBool with
adamc@75 841 | true -> Found TBool
adamc@75 842 | false -> Unknown)
adamc@75 843 | false -> Unknown)))
adamc@75 844 </pre># *)
adamc@75 845
adamc@75 846 (** %\smallskip%
adamc@75 847
adam@292 848 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 849
adamc@77 850 (* begin thide *)
adamc@73 851 Notation "e1 ;;; e2" := (if e1 then e2 else !!)
adamc@73 852 (right associativity, at level 60).
adamc@73 853
adamc@75 854 (** Next, we prove a helpful lemma, which states that a given expression can have at most one type. *)
adamc@75 855
adamc@75 856 Lemma hasType_det : forall e t1,
adamc@73 857 hasType e t1
adam@335 858 -> forall t2, hasType e t2
adamc@73 859 -> t1 = t2.
adamc@73 860 induction 1; inversion 1; crush.
adamc@73 861 Qed.
adamc@73 862
adamc@75 863 (** Now we can define the type-checker. Its type expresses that it only fails on untypable expressions. *)
adamc@75 864
adamc@77 865 (* end thide *)
adam@297 866 Definition typeCheck' : forall e : exp, {t : type | hasType e t} + {forall t, ~ hasType e t}.
adamc@77 867 (* begin thide *)
adamc@73 868 Hint Constructors hasType.
adamc@75 869 (** We register all of the typing rules as hints. *)
adamc@75 870
adamc@73 871 Hint Resolve hasType_det.
adam@335 872 (** 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 873
adamc@75 874 (** Finally, the implementation of [typeCheck] can be transcribed literally, simply switching notations as needed. *)
adamc@212 875
adamc@212 876 refine (fix F (e : exp) : {t : type | hasType e t} + {forall t, ~ hasType e t} :=
adam@380 877 match e return {t : type | hasType e t} + {forall t, ~ hasType e t} with
adam@335 878 | Nat _ => [||TNat||]
adamc@73 879 | Plus e1 e2 =>
adamc@73 880 t1 <-- F e1;
adamc@73 881 t2 <-- F e2;
adamc@73 882 eq_type_dec t1 TNat;;;
adamc@73 883 eq_type_dec t2 TNat;;;
adam@335 884 [||TNat||]
adam@335 885 | Bool _ => [||TBool||]
adamc@73 886 | And e1 e2 =>
adamc@73 887 t1 <-- F e1;
adamc@73 888 t2 <-- F e2;
adamc@73 889 eq_type_dec t1 TBool;;;
adamc@73 890 eq_type_dec t2 TBool;;;
adam@335 891 [||TBool||]
adamc@73 892 end); clear F; crush' tt hasType; eauto.
adamc@75 893
adam@335 894 (** 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 895 (* end thide *)
adamc@212 896
adamc@212 897
adamc@73 898 Defined.
adamc@73 899
adamc@75 900 (** 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 901
adam@335 902 Our new function remains easy to test: *)
adamc@75 903
adamc@73 904 Eval simpl in typeCheck' (Nat 0).
adamc@212 905 (** %\vspace{-.15in}% [[
adam@335 906 = [||TNat||]
adamc@75 907 : {t : type | hasType (Nat 0) t} +
adamc@75 908 {(forall t : type, ~ hasType (Nat 0) t)}
adam@302 909 ]]
adam@302 910 *)
adamc@75 911
adamc@73 912 Eval simpl in typeCheck' (Plus (Nat 1) (Nat 2)).
adamc@212 913 (** %\vspace{-.15in}% [[
adam@335 914 = [||TNat||]
adamc@75 915 : {t : type | hasType (Plus (Nat 1) (Nat 2)) t} +
adamc@75 916 {(forall t : type, ~ hasType (Plus (Nat 1) (Nat 2)) t)}
adam@302 917 ]]
adam@302 918 *)
adamc@75 919
adamc@73 920 Eval simpl in typeCheck' (Plus (Nat 1) (Bool false)).
adamc@212 921 (** %\vspace{-.15in}% [[
adamc@75 922 = !!
adamc@75 923 : {t : type | hasType (Plus (Nat 1) (Bool false)) t} +
adamc@75 924 {(forall t : type, ~ hasType (Plus (Nat 1) (Bool false)) t)}
adam@302 925 ]]
adam@335 926
adam@335 927 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. *)