annotate src/Subset.v @ 314:d5787b70cf48

Rename Tactics; change 'principal typing' to 'principal types'
author Adam Chlipala <adam@chlipala.net>
date Wed, 07 Sep 2011 13:47:24 -0400
parents 7b38729be069
children 1f57a8d0ed3d
rev   line source
adam@297 1 (* Copyright (c) 2008-2011, 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
adamc@70 18
adamc@74 19 (** %\part{Programming with Dependent Types}
adamc@74 20
adamc@74 21 \chapter{Subset Types and Variations}% *)
adamc@70 22
adam@292 23 (** 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 %\textit{%#<i>#dependent types#</i>#%}% to integrate programming, specification, and proving into a single phase. *)
adamc@70 24
adamc@70 25
adamc@70 26 (** * Introducing Subset Types *)
adamc@70 27
adamc@70 28 (** Let us consider several ways of implementing the natural number predecessor function. We start by displaying the definition from the standard library: *)
adamc@70 29
adamc@70 30 Print pred.
adamc@212 31 (** %\vspace{-.15in}% [[
adamc@70 32 pred = fun n : nat => match n with
adamc@70 33 | 0 => 0
adamc@70 34 | S u => u
adamc@70 35 end
adamc@70 36 : nat -> nat
adamc@212 37
adamc@212 38 ]]
adamc@70 39
adamc@212 40 We can use a new command, [Extraction], to produce an OCaml version of this function. *)
adamc@70 41
adamc@70 42 Extraction pred.
adamc@70 43
adamc@70 44 (** %\begin{verbatim}
adamc@70 45 (** val pred : nat -> nat **)
adamc@70 46
adamc@70 47 let pred = function
adamc@70 48 | O -> O
adamc@70 49 | S u -> u
adamc@70 50 \end{verbatim}%
adamc@70 51
adamc@70 52 #<pre>
adamc@70 53 (** val pred : nat -> nat **)
adamc@70 54
adamc@70 55 let pred = function
adamc@70 56 | O -> O
adamc@70 57 | S u -> u
adamc@70 58 </pre># *)
adamc@70 59
adamc@70 60 (** 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 61
adamc@70 62 Lemma zgtz : 0 > 0 -> False.
adamc@70 63 crush.
adamc@70 64 Qed.
adamc@70 65
adamc@70 66 Definition pred_strong1 (n : nat) : n > 0 -> nat :=
adamc@212 67 match n with
adamc@70 68 | O => fun pf : 0 > 0 => match zgtz pf with end
adamc@70 69 | S n' => fun _ => n'
adamc@70 70 end.
adamc@70 71
adamc@70 72 (** We expand the type of [pred] to include a %\textit{%#<i>#proof#</i>#%}% 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 %\textit{%#<i>#dependent#</i>#%}% type, because its type depends on the %\textit{%#<i>#value#</i>#%}% of the argument [n].
adamc@70 73
adam@282 74 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 %\textit{%#<i>#implicit#</i>#%}%, since it can be deduced from the type of the second argument, so we need not write [n] in function calls. *)
adam@282 75
adam@282 76 Theorem two_gt0 : 2 > 0.
adam@282 77 crush.
adam@282 78 Qed.
adam@282 79
adam@282 80 Eval compute in pred_strong1 two_gt0.
adam@282 81 (** %\vspace{-.15in}% [[
adam@282 82 = 1
adam@282 83 : nat
adam@282 84
adam@282 85 ]]
adam@282 86
adam@294 87 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 88
adamc@212 89 [[
adamc@70 90 Definition pred_strong1' (n : nat) (pf : n > 0) : nat :=
adamc@70 91 match n with
adamc@70 92 | O => match zgtz pf with end
adamc@70 93 | S n' => n'
adamc@70 94 end.
adamc@70 95
adamc@70 96 Error: In environment
adamc@70 97 n : nat
adamc@70 98 pf : n > 0
adamc@70 99 The term "pf" has type "n > 0" while it is expected to have type
adamc@70 100 "0 > 0"
adamc@212 101
adamc@70 102 ]]
adamc@70 103
adamc@212 104 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 105
adamc@70 106 In this case, we must use a [return] annotation to declare the relationship between the %\textit{%#<i>#value#</i>#%}% of the [match] discriminee and the %\textit{%#<i>#type#</i>#%}% 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 107
adamc@212 108 We are lucky that Coq's heuristics infer the [return] clause (specifically, [return n > 0 -> nat]) for us in this case. In general, however, the inference problem is undecidable. The known undecidable problem of %\textit{%#<i>#higher-order unification#</i>#%}% 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 109
adamc@70 110 Let us now take a look at the OCaml code Coq generates for [pred_strong1]. *)
adamc@70 111
adamc@70 112 Extraction pred_strong1.
adamc@70 113
adamc@70 114 (** %\begin{verbatim}
adamc@70 115 (** val pred_strong1 : nat -> nat **)
adamc@70 116
adamc@70 117 let pred_strong1 = function
adamc@70 118 | O -> assert false (* absurd case *)
adamc@70 119 | S n' -> n'
adamc@70 120 \end{verbatim}%
adamc@70 121
adamc@70 122 #<pre>
adamc@70 123 (** val pred_strong1 : nat -> nat **)
adamc@70 124
adamc@70 125 let pred_strong1 = function
adamc@70 126 | O -> assert false (* absurd case *)
adamc@70 127 | S n' -> n'
adamc@70 128 </pre># *)
adamc@70 129
adamc@70 130 (** 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 131
adamc@70 132 We can reimplement our dependently-typed [pred] based on %\textit{%#<i>#subset types#</i>#%}%, defined in the standard library with the type family [sig]. *)
adamc@70 133
adamc@70 134 Print sig.
adamc@212 135 (** %\vspace{-.15in}% [[
adamc@70 136 Inductive sig (A : Type) (P : A -> Prop) : Type :=
adamc@70 137 exist : forall x : A, P x -> sig P
adamc@70 138 For sig: Argument A is implicit
adamc@70 139 For exist: Argument A is implicit
adamc@212 140
adamc@70 141 ]]
adamc@70 142
adamc@70 143 [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 144
adamc@70 145 We rewrite [pred_strong1], using some syntactic sugar for subset types. *)
adamc@70 146
adamc@70 147 Locate "{ _ : _ | _ }".
adamc@212 148 (** %\vspace{-.15in}% [[
adamc@70 149 Notation Scope
adamc@70 150 "{ x : A | P }" := sig (fun x : A => P)
adamc@70 151 : type_scope
adamc@70 152 (default interpretation)
adam@302 153 ]]
adam@302 154 *)
adamc@70 155
adamc@70 156 Definition pred_strong2 (s : {n : nat | n > 0}) : nat :=
adamc@70 157 match s with
adamc@70 158 | exist O pf => match zgtz pf with end
adamc@70 159 | exist (S n') _ => n'
adamc@70 160 end.
adamc@70 161
adam@282 162 (** 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. *)
adam@282 163
adam@282 164 Eval compute in pred_strong2 (exist _ 2 two_gt0).
adam@282 165 (** %\vspace{-.15in}% [[
adam@282 166 = 1
adam@282 167 : nat
adam@282 168
adam@302 169 ]]
adam@302 170 *)
adam@282 171
adamc@70 172 Extraction pred_strong2.
adamc@70 173
adamc@70 174 (** %\begin{verbatim}
adamc@70 175 (** val pred_strong2 : nat -> nat **)
adamc@70 176
adamc@70 177 let pred_strong2 = function
adamc@70 178 | O -> assert false (* absurd case *)
adamc@70 179 | S n' -> n'
adamc@70 180 \end{verbatim}%
adamc@70 181
adamc@70 182 #<pre>
adamc@70 183 (** val pred_strong2 : nat -> nat **)
adamc@70 184
adamc@70 185 let pred_strong2 = function
adamc@70 186 | O -> assert false (* absurd case *)
adamc@70 187 | S n' -> n'
adamc@70 188 </pre>#
adamc@70 189
adamc@70 190 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 191
adamc@70 192 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 193
adamc@70 194 Definition pred_strong3 (s : {n : nat | n > 0}) : {m : nat | proj1_sig s = S m} :=
adamc@70 195 match s return {m : nat | proj1_sig s = S m} with
adamc@70 196 | exist 0 pf => match zgtz pf with end
adamc@212 197 | exist (S n') pf => exist _ n' (refl_equal _)
adamc@70 198 end.
adamc@70 199
adam@282 200 Eval compute in pred_strong3 (exist _ 2 two_gt0).
adam@282 201 (** %\vspace{-.15in}% [[
adam@282 202 = exist (fun m : nat => 2 = S m) 1 (refl_equal 2)
adam@282 203 : {m : nat | proj1_sig (exist (lt 0) 2 two_gt0) = S m}
adam@282 204
adam@302 205 ]]
adam@302 206 *)
adam@282 207
adam@282 208 (** The function [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 209
adamc@70 210 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 211
adamc@70 212 Extraction pred_strong3.
adamc@70 213
adamc@70 214 (** %\begin{verbatim}
adamc@70 215 (** val pred_strong3 : nat -> nat **)
adamc@70 216
adamc@70 217 let pred_strong3 = function
adamc@70 218 | O -> assert false (* absurd case *)
adamc@70 219 | S n' -> n'
adamc@70 220 \end{verbatim}%
adamc@70 221
adamc@70 222 #<pre>
adamc@70 223 (** val pred_strong3 : nat -> nat **)
adamc@70 224
adamc@70 225 let pred_strong3 = function
adamc@70 226 | O -> assert false (* absurd case *)
adamc@70 227 | S n' -> n'
adamc@70 228 </pre>#
adamc@70 229
adamc@70 230 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. *)
adamc@70 231
adam@297 232 Definition pred_strong4 : forall n : nat, n > 0 -> {m : nat | n = S m}.
adamc@70 233 refine (fun n =>
adamc@212 234 match n with
adamc@70 235 | O => fun _ => False_rec _ _
adamc@70 236 | S n' => fun _ => exist _ n' _
adamc@70 237 end).
adamc@212 238
adamc@77 239 (* begin thide *)
adam@292 240 (** 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. We do most of the work with the [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:
adamc@70 241
adamc@70 242 [[
adamc@70 243 2 subgoals
adamc@70 244
adamc@70 245 n : nat
adamc@70 246 _ : 0 > 0
adamc@70 247 ============================
adamc@70 248 False
adamc@70 249
adamc@70 250 subgoal 2 is:
adamc@70 251 S n' = S n'
adamc@212 252
adamc@70 253 ]]
adamc@70 254
adamc@70 255 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 256
adamc@70 257 Undo.
adamc@70 258 refine (fun n =>
adamc@212 259 match n with
adamc@70 260 | O => fun _ => False_rec _ _
adamc@70 261 | S n' => fun _ => exist _ n' _
adamc@70 262 end); crush.
adamc@77 263 (* end thide *)
adamc@70 264 Defined.
adamc@70 265
adam@292 266 (** We end the %``%#"#proof#"#%''% with [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. Let us see what our proof script constructed. *)
adamc@70 267
adamc@70 268 Print pred_strong4.
adamc@212 269 (** %\vspace{-.15in}% [[
adamc@70 270 pred_strong4 =
adamc@70 271 fun n : nat =>
adamc@70 272 match n as n0 return (n0 > 0 -> {m : nat | n0 = S m}) with
adamc@70 273 | 0 =>
adamc@70 274 fun _ : 0 > 0 =>
adamc@70 275 False_rec {m : nat | 0 = S m}
adamc@70 276 (Bool.diff_false_true
adamc@70 277 (Bool.absurd_eq_true false
adamc@70 278 (Bool.diff_false_true
adamc@70 279 (Bool.absurd_eq_true false (pred_strong4_subproof n _)))))
adamc@70 280 | S n' =>
adamc@70 281 fun _ : S n' > 0 =>
adamc@70 282 exist (fun m : nat => S n' = S m) n' (refl_equal (S n'))
adamc@70 283 end
adamc@70 284 : forall n : nat, n > 0 -> {m : nat | n = S m}
adamc@212 285
adamc@70 286 ]]
adamc@70 287
adam@282 288 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 289
adam@282 290 Eval compute in pred_strong4 two_gt0.
adam@282 291 (** %\vspace{-.15in}% [[
adam@282 292 = exist (fun m : nat => 2 = S m) 1 (refl_equal 2)
adam@282 293 : {m : nat | 2 = S m}
adam@282 294
adam@282 295 ]]
adam@282 296
adam@282 297 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. *)
adamc@70 298
adamc@70 299 Notation "!" := (False_rec _ _).
adamc@70 300 Notation "[ e ]" := (exist _ e _).
adamc@70 301
adam@297 302 Definition pred_strong5 : forall n : nat, n > 0 -> {m : nat | n = S m}.
adamc@70 303 refine (fun n =>
adamc@212 304 match n with
adamc@70 305 | O => fun _ => !
adamc@70 306 | S n' => fun _ => [n']
adamc@70 307 end); crush.
adamc@70 308 Defined.
adamc@71 309
adam@282 310 (** By default, notations are also used in pretty-printing terms, including results of evaluation. *)
adam@282 311
adam@282 312 Eval compute in pred_strong5 two_gt0.
adam@282 313 (** %\vspace{-.15in}% [[
adam@282 314 = [1]
adam@282 315 : {m : nat | 2 = S m}
adam@282 316
adam@282 317 ]]
adam@282 318
adam@282 319 One other alternative is worth demonstrating. Recent Coq versions include a facility called [Program] that streamlines this style of definition. Here is a complete implementation using [Program]. *)
adamc@212 320
adamc@212 321 Obligation Tactic := crush.
adamc@212 322
adamc@212 323 Program Definition pred_strong6 (n : nat) (_ : n > 0) : {m : nat | n = S m} :=
adamc@212 324 match n with
adamc@212 325 | O => _
adamc@212 326 | S n' => n'
adamc@212 327 end.
adamc@212 328
adamc@212 329 (** 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 330
adam@282 331 Eval compute in pred_strong6 two_gt0.
adam@282 332 (** %\vspace{-.15in}% [[
adam@282 333 = [1]
adam@282 334 : {m : nat | 2 = S m}
adam@282 335
adam@302 336 ]]
adam@302 337 *)
adam@282 338
adamc@71 339
adamc@71 340 (** * Decidable Proposition Types *)
adamc@71 341
adamc@71 342 (** There is another type in the standard library which captures the idea of program values that indicate which of two propositions is true. *)
adamc@71 343
adamc@71 344 Print sumbool.
adamc@212 345 (** %\vspace{-.15in}% [[
adamc@71 346 Inductive sumbool (A : Prop) (B : Prop) : Set :=
adamc@71 347 left : A -> {A} + {B} | right : B -> {A} + {B}
adamc@71 348 For left: Argument A is implicit
adamc@71 349 For right: Argument B is implicit
adamc@212 350
adamc@212 351 ]]
adamc@71 352
adamc@212 353 We can define some notations to make working with [sumbool] more convenient. *)
adamc@71 354
adamc@71 355 Notation "'Yes'" := (left _ _).
adamc@71 356 Notation "'No'" := (right _ _).
adamc@71 357 Notation "'Reduce' x" := (if x then Yes else No) (at level 50).
adamc@71 358
adamc@71 359 (** 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 360
adamc@71 361 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 362
adam@297 363 Definition eq_nat_dec : forall n m : nat, {n = m} + {n <> m}.
adamc@212 364 refine (fix f (n m : nat) : {n = m} + {n <> m} :=
adamc@212 365 match n, m with
adamc@71 366 | O, O => Yes
adamc@71 367 | S n', S m' => Reduce (f n' m')
adamc@71 368 | _, _ => No
adamc@71 369 end); congruence.
adamc@71 370 Defined.
adamc@71 371
adam@282 372 Eval compute in eq_nat_dec 2 2.
adam@282 373 (** %\vspace{-.15in}% [[
adam@282 374 = Yes
adam@282 375 : {2 = 2} + {2 <> 2}
adam@282 376
adam@302 377 ]]
adam@302 378 *)
adam@282 379
adam@282 380 Eval compute in eq_nat_dec 2 3.
adam@282 381 (** %\vspace{-.15in}% [[
adam@282 382 = No
adam@282 383 : {2 = 2} + {2 <> 2}
adam@282 384
adam@302 385 ]]
adam@302 386 *)
adam@282 387
adamc@71 388 (** Our definition extracts to reasonable OCaml code. *)
adamc@71 389
adamc@71 390 Extraction eq_nat_dec.
adamc@71 391
adamc@71 392 (** %\begin{verbatim}
adamc@71 393 (** val eq_nat_dec : nat -> nat -> sumbool **)
adamc@71 394
adamc@71 395 let rec eq_nat_dec n m =
adamc@71 396 match n with
adamc@71 397 | O -> (match m with
adamc@71 398 | O -> Left
adamc@71 399 | S n0 -> Right)
adamc@71 400 | S n' -> (match m with
adamc@71 401 | O -> Right
adamc@71 402 | S m' -> eq_nat_dec n' m')
adamc@71 403 \end{verbatim}%
adamc@71 404
adamc@71 405 #<pre>
adamc@71 406 (** val eq_nat_dec : nat -> nat -> sumbool **)
adamc@71 407
adamc@71 408 let rec eq_nat_dec n m =
adamc@71 409 match n with
adamc@71 410 | O -> (match m with
adamc@71 411 | O -> Left
adamc@71 412 | S n0 -> Right)
adamc@71 413 | S n' -> (match m with
adamc@71 414 | O -> Right
adamc@71 415 | S m' -> eq_nat_dec n' m')
adamc@71 416 </pre>#
adamc@71 417
adamc@71 418 Proving this kind of decidable equality result is so common that Coq comes with a tactic for automating it. *)
adamc@71 419
adamc@71 420 Definition eq_nat_dec' (n m : nat) : {n = m} + {n <> m}.
adamc@71 421 decide equality.
adamc@71 422 Defined.
adamc@71 423
adamc@71 424 (** 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. *)
adamc@71 425
adamc@71 426 Extract Inductive sumbool => "bool" ["true" "false"].
adamc@71 427 Extraction eq_nat_dec'.
adamc@71 428
adamc@71 429 (** %\begin{verbatim}
adamc@71 430 (** val eq_nat_dec' : nat -> nat -> bool **)
adamc@71 431
adamc@71 432 let rec eq_nat_dec' n m0 =
adamc@71 433 match n with
adamc@71 434 | O -> (match m0 with
adamc@71 435 | O -> true
adamc@71 436 | S n0 -> false)
adamc@71 437 | S n0 -> (match m0 with
adamc@71 438 | O -> false
adamc@71 439 | S n1 -> eq_nat_dec' n0 n1)
adamc@71 440 \end{verbatim}%
adamc@71 441
adamc@71 442 #<pre>
adamc@71 443 (** val eq_nat_dec' : nat -> nat -> bool **)
adamc@71 444
adamc@71 445 let rec eq_nat_dec' n m0 =
adamc@71 446 match n with
adamc@71 447 | O -> (match m0 with
adamc@71 448 | O -> true
adamc@71 449 | S n0 -> false)
adamc@71 450 | S n0 -> (match m0 with
adamc@71 451 | O -> false
adamc@71 452 | S n1 -> eq_nat_dec' n0 n1)
adamc@71 453 </pre># *)
adamc@72 454
adamc@72 455 (** %\smallskip%
adamc@72 456
adam@292 457 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 458
adamc@77 459 (* begin thide *)
adamc@204 460 Notation "x || y" := (if x then Yes else Reduce y).
adamc@72 461
adamc@72 462 (** 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 463
adamc@72 464 Section In_dec.
adamc@72 465 Variable A : Set.
adamc@72 466 Variable A_eq_dec : forall x y : A, {x = y} + {x <> y}.
adamc@72 467
adamc@72 468 (** The final function is easy to write using the techniques we have developed so far. *)
adamc@72 469
adamc@212 470 Definition In_dec : forall (x : A) (ls : list A), {In x ls} + {~ In x ls}.
adamc@212 471 refine (fix f (x : A) (ls : list A) : {In x ls} + {~ In x ls} :=
adamc@212 472 match ls with
adamc@72 473 | nil => No
adamc@72 474 | x' :: ls' => A_eq_dec x x' || f x ls'
adamc@72 475 end); crush.
adam@282 476 Defined.
adamc@72 477 End In_dec.
adamc@72 478
adam@282 479 Eval compute in In_dec eq_nat_dec 2 (1 :: 2 :: nil).
adam@282 480 (** %\vspace{-.15in}% [[
adam@282 481 = Yes
adam@282 482 : {In 2 (1 :: 2 :: nil)} + {~ In 2 (1 :: 2 :: nil)}
adam@282 483
adam@302 484 ]]
adam@302 485 *)
adam@282 486
adam@282 487 Eval compute in In_dec eq_nat_dec 3 (1 :: 2 :: nil).
adam@282 488 (** %\vspace{-.15in}% [[
adam@282 489 = No
adam@282 490 : {In 3 (1 :: 2 :: nil)} + {~ In 3 (1 :: 2 :: nil)}
adam@282 491
adam@302 492 ]]
adam@302 493 *)
adam@282 494
adamc@72 495 (** [In_dec] has a reasonable extraction to OCaml. *)
adamc@72 496
adamc@72 497 Extraction In_dec.
adamc@77 498 (* end thide *)
adamc@72 499
adamc@72 500 (** %\begin{verbatim}
adamc@72 501 (** val in_dec : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 list -> bool **)
adamc@72 502
adamc@72 503 let rec in_dec a_eq_dec x = function
adamc@72 504 | Nil -> false
adamc@72 505 | Cons (x', ls') ->
adamc@72 506 (match a_eq_dec x x' with
adamc@72 507 | true -> true
adamc@72 508 | false -> in_dec a_eq_dec x ls')
adamc@72 509 \end{verbatim}%
adamc@72 510
adamc@72 511 #<pre>
adamc@72 512 (** val in_dec : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 list -> bool **)
adamc@72 513
adamc@72 514 let rec in_dec a_eq_dec x = function
adamc@72 515 | Nil -> false
adamc@72 516 | Cons (x', ls') ->
adamc@72 517 (match a_eq_dec x x' with
adamc@72 518 | true -> true
adamc@72 519 | false -> in_dec a_eq_dec x ls')
adamc@72 520 </pre># *)
adamc@72 521
adamc@72 522
adamc@72 523 (** * Partial Subset Types *)
adamc@72 524
adamc@73 525 (** 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 that than returning a default value, as [pred] does for [0]. One approach is to define this type family [maybe], which is a version of [sig] that allows obligation-free failure. *)
adamc@73 526
adamc@89 527 Inductive maybe (A : Set) (P : A -> Prop) : Set :=
adamc@72 528 | Unknown : maybe P
adamc@72 529 | Found : forall x : A, P x -> maybe P.
adamc@72 530
adamc@73 531 (** We can define some new notations, analogous to those we defined for subset types. *)
adamc@73 532
adamc@72 533 Notation "{{ x | P }}" := (maybe (fun x => P)).
adamc@72 534 Notation "??" := (Unknown _).
adamc@72 535 Notation "[[ x ]]" := (Found _ x _).
adamc@72 536
adamc@73 537 (** Now our next version of [pred] is trivial to write. *)
adamc@73 538
adam@297 539 Definition pred_strong7 : forall n : nat, {{m | n = S m}}.
adamc@73 540 refine (fun n =>
adamc@212 541 match n with
adamc@73 542 | O => ??
adamc@73 543 | S n' => [[n']]
adamc@73 544 end); trivial.
adamc@73 545 Defined.
adamc@73 546
adam@282 547 Eval compute in pred_strong7 2.
adam@282 548 (** %\vspace{-.15in}% [[
adam@282 549 = [[1]]
adam@282 550 : {{m | 2 = S m}}
adam@282 551
adam@302 552 ]]
adam@302 553 *)
adam@282 554
adam@282 555 Eval compute in pred_strong7 0.
adam@282 556 (** %\vspace{-.15in}% [[
adam@282 557 = ??
adam@282 558 : {{m | 0 = S m}}
adam@282 559
adam@282 560 ]]
adam@282 561
adam@282 562 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 [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 563
adamc@73 564 Print sumor.
adamc@212 565 (** %\vspace{-.15in}% [[
adamc@73 566 Inductive sumor (A : Type) (B : Prop) : Type :=
adamc@73 567 inleft : A -> A + {B} | inright : B -> A + {B}
adamc@73 568 For inleft: Argument A is implicit
adamc@73 569 For inright: Argument B is implicit
adam@302 570 ]]
adam@302 571 *)
adamc@73 572
adamc@73 573 (** 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 574
adamc@73 575 Notation "!!" := (inright _ _).
adamc@73 576 Notation "[[[ x ]]]" := (inleft _ [x]).
adamc@73 577
adamc@73 578 (** 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 579
adam@297 580 Definition pred_strong8 : forall n : nat, {m : nat | n = S m} + {n = 0}.
adamc@73 581 refine (fun n =>
adamc@212 582 match n with
adamc@73 583 | O => !!
adamc@73 584 | S n' => [[[n']]]
adamc@73 585 end); trivial.
adamc@73 586 Defined.
adamc@73 587
adam@282 588 Eval compute in pred_strong8 2.
adam@282 589 (** %\vspace{-.15in}% [[
adam@282 590 = [[[1]]]
adam@282 591 : {m : nat | 2 = S m} + {2 = 0}
adam@282 592
adam@302 593 ]]
adam@302 594 *)
adam@282 595
adam@282 596 Eval compute in pred_strong8 0.
adam@282 597 (** %\vspace{-.15in}% [[
adam@282 598 = !!
adam@282 599 : {m : nat | 0 = S m} + {0 = 0}
adam@282 600
adam@302 601 ]]
adam@302 602 *)
adam@282 603
adamc@73 604
adamc@73 605 (** * Monadic Notations *)
adamc@73 606
adam@292 607 (** We can treat [maybe] like a 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 608
adamc@72 609 Notation "x <- e1 ; e2" := (match e1 with
adamc@72 610 | Unknown => ??
adamc@72 611 | Found x _ => e2
adamc@72 612 end)
adamc@72 613 (right associativity, at level 60).
adamc@72 614
adamc@73 615 (** 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] %\textit{%#<i>#does#</i>#%}% 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 616
adamc@73 617 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 618
adam@297 619 Definition doublePred : forall n1 n2 : nat, {{p | n1 = S (fst p) /\ n2 = S (snd p)}}.
adamc@73 620 refine (fun n1 n2 =>
adamc@212 621 m1 <- pred_strong7 n1;
adamc@212 622 m2 <- pred_strong7 n2;
adamc@73 623 [[(m1, m2)]]); tauto.
adamc@73 624 Defined.
adamc@73 625
adam@292 626 (** We can build a [sumor] version of the %``%#"#bind#"#%''% notation and use it to write a similarly straightforward version of this function. *)
adamc@73 627
adamc@73 628 (** printing <-- $\longleftarrow$ *)
adamc@73 629
adamc@73 630 Notation "x <-- e1 ; e2" := (match e1 with
adamc@73 631 | inright _ => !!
adamc@73 632 | inleft (exist x _) => e2
adamc@73 633 end)
adamc@73 634 (right associativity, at level 60).
adamc@73 635
adamc@73 636 (** printing * $\times$ *)
adamc@73 637
adam@297 638 Definition doublePred' : forall n1 n2 : nat,
adam@297 639 {p : nat * nat | n1 = S (fst p) /\ n2 = S (snd p)}
adamc@73 640 + {n1 = 0 \/ n2 = 0}.
adamc@73 641 refine (fun n1 n2 =>
adamc@212 642 m1 <-- pred_strong8 n1;
adamc@212 643 m2 <-- pred_strong8 n2;
adamc@73 644 [[[(m1, m2)]]]); tauto.
adamc@73 645 Defined.
adamc@72 646
adamc@72 647
adamc@72 648 (** * A Type-Checking Example *)
adamc@72 649
adamc@75 650 (** We can apply these specification types to build a certified type-checker for a simple expression language. *)
adamc@75 651
adamc@72 652 Inductive exp : Set :=
adamc@72 653 | Nat : nat -> exp
adamc@72 654 | Plus : exp -> exp -> exp
adamc@72 655 | Bool : bool -> exp
adamc@72 656 | And : exp -> exp -> exp.
adamc@72 657
adamc@75 658 (** We define a simple language of types and its typing rules, in the style introduced in Chapter 4. *)
adamc@75 659
adamc@72 660 Inductive type : Set := TNat | TBool.
adamc@72 661
adamc@72 662 Inductive hasType : exp -> type -> Prop :=
adamc@72 663 | HtNat : forall n,
adamc@72 664 hasType (Nat n) TNat
adamc@72 665 | HtPlus : forall e1 e2,
adamc@72 666 hasType e1 TNat
adamc@72 667 -> hasType e2 TNat
adamc@72 668 -> hasType (Plus e1 e2) TNat
adamc@72 669 | HtBool : forall b,
adamc@72 670 hasType (Bool b) TBool
adamc@72 671 | HtAnd : forall e1 e2,
adamc@72 672 hasType e1 TBool
adamc@72 673 -> hasType e2 TBool
adamc@72 674 -> hasType (And e1 e2) TBool.
adamc@72 675
adamc@75 676 (** It will be helpful to have a function for comparing two types. We build one using [decide equality]. *)
adamc@75 677
adamc@77 678 (* begin thide *)
adamc@75 679 Definition eq_type_dec : forall t1 t2 : type, {t1 = t2} + {t1 <> t2}.
adamc@72 680 decide equality.
adamc@72 681 Defined.
adamc@72 682
adam@292 683 (** 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 684
adamc@73 685 Notation "e1 ;; e2" := (if e1 then e2 else ??)
adamc@73 686 (right associativity, at level 60).
adamc@73 687
adamc@75 688 (** 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 689 (* end thide *)
adamc@75 690
adam@297 691 Definition typeCheck : forall e : exp, {{t | hasType e t}}.
adamc@77 692 (* begin thide *)
adamc@72 693 Hint Constructors hasType.
adamc@72 694
adamc@72 695 refine (fix F (e : exp) : {{t | hasType e t}} :=
adamc@212 696 match e with
adamc@72 697 | Nat _ => [[TNat]]
adamc@72 698 | Plus e1 e2 =>
adamc@72 699 t1 <- F e1;
adamc@72 700 t2 <- F e2;
adamc@72 701 eq_type_dec t1 TNat;;
adamc@72 702 eq_type_dec t2 TNat;;
adamc@72 703 [[TNat]]
adamc@72 704 | Bool _ => [[TBool]]
adamc@72 705 | And e1 e2 =>
adamc@72 706 t1 <- F e1;
adamc@72 707 t2 <- F e2;
adamc@72 708 eq_type_dec t1 TBool;;
adamc@72 709 eq_type_dec t2 TBool;;
adamc@72 710 [[TBool]]
adamc@72 711 end); crush.
adamc@77 712 (* end thide *)
adamc@72 713 Defined.
adamc@72 714
adamc@75 715 (** Despite manipulating proofs, our type checker is easy to run. *)
adamc@75 716
adamc@72 717 Eval simpl in typeCheck (Nat 0).
adamc@212 718 (** %\vspace{-.15in}% [[
adamc@75 719 = [[TNat]]
adamc@75 720 : {{t | hasType (Nat 0) t}}
adam@302 721 ]]
adam@302 722 *)
adamc@75 723
adamc@72 724 Eval simpl in typeCheck (Plus (Nat 1) (Nat 2)).
adamc@212 725 (** %\vspace{-.15in}% [[
adamc@75 726 = [[TNat]]
adamc@75 727 : {{t | hasType (Plus (Nat 1) (Nat 2)) t}}
adam@302 728 ]]
adam@302 729 *)
adamc@75 730
adamc@72 731 Eval simpl in typeCheck (Plus (Nat 1) (Bool false)).
adamc@212 732 (** %\vspace{-.15in}% [[
adamc@75 733 = ??
adamc@75 734 : {{t | hasType (Plus (Nat 1) (Bool false)) t}}
adam@302 735 ]]
adam@302 736 *)
adamc@75 737
adamc@75 738 (** The type-checker also extracts to some reasonable OCaml code. *)
adamc@75 739
adamc@75 740 Extraction typeCheck.
adamc@75 741
adamc@75 742 (** %\begin{verbatim}
adamc@75 743 (** val typeCheck : exp -> type0 maybe **)
adamc@75 744
adamc@75 745 let rec typeCheck = function
adamc@75 746 | Nat n -> Found TNat
adamc@75 747 | Plus (e1, e2) ->
adamc@75 748 (match typeCheck e1 with
adamc@75 749 | Unknown -> Unknown
adamc@75 750 | Found t1 ->
adamc@75 751 (match typeCheck e2 with
adamc@75 752 | Unknown -> Unknown
adamc@75 753 | Found t2 ->
adamc@75 754 (match eq_type_dec t1 TNat with
adamc@75 755 | true ->
adamc@75 756 (match eq_type_dec t2 TNat with
adamc@75 757 | true -> Found TNat
adamc@75 758 | false -> Unknown)
adamc@75 759 | false -> Unknown)))
adamc@75 760 | Bool b -> Found TBool
adamc@75 761 | And (e1, e2) ->
adamc@75 762 (match typeCheck e1 with
adamc@75 763 | Unknown -> Unknown
adamc@75 764 | Found t1 ->
adamc@75 765 (match typeCheck e2 with
adamc@75 766 | Unknown -> Unknown
adamc@75 767 | Found t2 ->
adamc@75 768 (match eq_type_dec t1 TBool with
adamc@75 769 | true ->
adamc@75 770 (match eq_type_dec t2 TBool with
adamc@75 771 | true -> Found TBool
adamc@75 772 | false -> Unknown)
adamc@75 773 | false -> Unknown)))
adamc@75 774 \end{verbatim}%
adamc@75 775
adamc@75 776 #<pre>
adamc@75 777 (** val typeCheck : exp -> type0 maybe **)
adamc@75 778
adamc@75 779 let rec typeCheck = function
adamc@75 780 | Nat n -> Found TNat
adamc@75 781 | Plus (e1, e2) ->
adamc@75 782 (match typeCheck e1 with
adamc@75 783 | Unknown -> Unknown
adamc@75 784 | Found t1 ->
adamc@75 785 (match typeCheck e2 with
adamc@75 786 | Unknown -> Unknown
adamc@75 787 | Found t2 ->
adamc@75 788 (match eq_type_dec t1 TNat with
adamc@75 789 | true ->
adamc@75 790 (match eq_type_dec t2 TNat with
adamc@75 791 | true -> Found TNat
adamc@75 792 | false -> Unknown)
adamc@75 793 | false -> Unknown)))
adamc@75 794 | Bool b -> Found TBool
adamc@75 795 | And (e1, e2) ->
adamc@75 796 (match typeCheck e1 with
adamc@75 797 | Unknown -> Unknown
adamc@75 798 | Found t1 ->
adamc@75 799 (match typeCheck e2 with
adamc@75 800 | Unknown -> Unknown
adamc@75 801 | Found t2 ->
adamc@75 802 (match eq_type_dec t1 TBool with
adamc@75 803 | true ->
adamc@75 804 (match eq_type_dec t2 TBool with
adamc@75 805 | true -> Found TBool
adamc@75 806 | false -> Unknown)
adamc@75 807 | false -> Unknown)))
adamc@75 808 </pre># *)
adamc@75 809
adamc@75 810 (** %\smallskip%
adamc@75 811
adam@292 812 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 813
adamc@77 814 (* begin thide *)
adamc@73 815 Notation "e1 ;;; e2" := (if e1 then e2 else !!)
adamc@73 816 (right associativity, at level 60).
adamc@73 817
adamc@75 818 (** Next, we prove a helpful lemma, which states that a given expression can have at most one type. *)
adamc@75 819
adamc@75 820 Lemma hasType_det : forall e t1,
adamc@73 821 hasType e t1
adamc@73 822 -> forall t2,
adamc@73 823 hasType e t2
adamc@73 824 -> t1 = t2.
adamc@73 825 induction 1; inversion 1; crush.
adamc@73 826 Qed.
adamc@73 827
adamc@75 828 (** Now we can define the type-checker. Its type expresses that it only fails on untypable expressions. *)
adamc@75 829
adamc@77 830 (* end thide *)
adam@297 831 Definition typeCheck' : forall e : exp, {t : type | hasType e t} + {forall t, ~ hasType e t}.
adamc@77 832 (* begin thide *)
adamc@73 833 Hint Constructors hasType.
adamc@75 834 (** We register all of the typing rules as hints. *)
adamc@75 835
adamc@73 836 Hint Resolve hasType_det.
adamc@75 837 (** [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 838
adamc@75 839 (** Finally, the implementation of [typeCheck] can be transcribed literally, simply switching notations as needed. *)
adamc@212 840
adamc@212 841 refine (fix F (e : exp) : {t : type | hasType e t} + {forall t, ~ hasType e t} :=
adamc@212 842 match e with
adamc@73 843 | Nat _ => [[[TNat]]]
adamc@73 844 | Plus e1 e2 =>
adamc@73 845 t1 <-- F e1;
adamc@73 846 t2 <-- F e2;
adamc@73 847 eq_type_dec t1 TNat;;;
adamc@73 848 eq_type_dec t2 TNat;;;
adamc@73 849 [[[TNat]]]
adamc@73 850 | Bool _ => [[[TBool]]]
adamc@73 851 | And e1 e2 =>
adamc@73 852 t1 <-- F e1;
adamc@73 853 t2 <-- F e2;
adamc@73 854 eq_type_dec t1 TBool;;;
adamc@73 855 eq_type_dec t2 TBool;;;
adamc@73 856 [[[TBool]]]
adamc@73 857 end); clear F; crush' tt hasType; eauto.
adamc@75 858
adamc@75 859 (** 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 [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 860 (* end thide *)
adamc@212 861
adamc@212 862
adamc@73 863 Defined.
adamc@73 864
adamc@75 865 (** 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 866
adamc@75 867 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. *)
adamc@75 868
adamc@73 869 Eval simpl in typeCheck' (Nat 0).
adamc@212 870 (** %\vspace{-.15in}% [[
adamc@75 871 = [[[TNat]]]
adamc@75 872 : {t : type | hasType (Nat 0) t} +
adamc@75 873 {(forall t : type, ~ hasType (Nat 0) t)}
adam@302 874 ]]
adam@302 875 *)
adamc@75 876
adamc@73 877 Eval simpl in typeCheck' (Plus (Nat 1) (Nat 2)).
adamc@212 878 (** %\vspace{-.15in}% [[
adamc@75 879 = [[[TNat]]]
adamc@75 880 : {t : type | hasType (Plus (Nat 1) (Nat 2)) t} +
adamc@75 881 {(forall t : type, ~ hasType (Plus (Nat 1) (Nat 2)) t)}
adam@302 882 ]]
adam@302 883 *)
adamc@75 884
adamc@73 885 Eval simpl in typeCheck' (Plus (Nat 1) (Bool false)).
adamc@212 886 (** %\vspace{-.15in}% [[
adamc@75 887 = !!
adamc@75 888 : {t : type | hasType (Plus (Nat 1) (Bool false)) t} +
adamc@75 889 {(forall t : type, ~ hasType (Plus (Nat 1) (Bool false)) t)}
adam@302 890 ]]
adam@302 891 *)
adamc@82 892
adamc@82 893
adamc@82 894 (** * Exercises *)
adamc@82 895
adamc@82 896 (** All of the notations defined in this chapter, plus some extras, are available for import from the module [MoreSpecif] of the book source.
adamc@82 897
adamc@82 898 %\begin{enumerate}%#<ol>#
adamc@82 899 %\item%#<li># Write a function of type [forall n m : nat, {n <= m} + {n > m}]. That is, this function decides whether one natural is less than another, and its dependent type guarantees that its results are accurate.#</li>#
adamc@82 900
adamc@82 901 %\item%#<li># %\begin{enumerate}%#<ol>#
adamc@82 902 %\item%#<li># Define [var], a type of propositional variables, as a synonym for [nat].#</li>#
adamc@82 903 %\item%#<li># Define an inductive type [prop] of propositional logic formulas, consisting of variables, negation, and binary conjunction and disjunction.#</li>#
adamc@82 904 %\item%#<li># Define a function [propDenote] from variable truth assignments and [prop]s to [Prop], based on the usual meanings of the connectives. Represent truth assignments as functions from [var] to [bool].#</li>#
adamc@82 905 %\item%#<li># Define a function [bool_true_dec] that checks whether a boolean is true, with a maximally expressive dependent type. That is, the function should have type [forall b, {b = true} + {b = true -> False}]. #</li>#
adamc@212 906 %\item%#<li># Define a function [decide] that determines whether a particular [prop] is true under a particular truth assignment. That is, the function should have type [forall (truth : var -> bool) (p : prop), {propDenote truth p} + {~ propDenote truth p}]. This function is probably easiest to write in the usual tactical style, instead of programming with [refine]. [bool_true_dec] may come in handy as a hint.#</li>#
adamc@212 907 %\item%#<li># Define a function [negate] that returns a simplified version of the negation of a [prop]. That is, the function should have type [forall p : prop, {p' : prop | forall truth, propDenote truth p <-> ~ propDenote truth p'}]. To simplify a variable, just negate it. Simplify a negation by returning its argument. Simplify conjunctions and disjunctions using De Morgan's laws, negating the arguments recursively and switching the kind of connective. [decide] may be useful in some of the proof obligations, even if you do not use it in the computational part of [negate]'s definition. Lemmas like [decide] allow us to compensate for the lack of a general Law of the Excluded Middle in CIC.#</li>#
adamc@82 908 #</ol>#%\end{enumerate}% #</li>#
adamc@82 909
adam@292 910 %\item%#<li># Implement the DPLL satisfiability decision procedure for boolean formulas in conjunctive normal form, with a dependent type that guarantees its correctness. An example of a reasonable type for this function would be [forall f : formula, {truth : tvals | formulaTrue truth f} + {forall truth, ~ formulaTrue truth f}]. Implement at least %``%#"#the basic backtracking algorithm#"#%''% as defined here:
adamc@82 911 %\begin{center}\url{http://en.wikipedia.org/wiki/DPLL_algorithm}\end{center}%
adamc@82 912 #<blockquote><a href="http://en.wikipedia.org/wiki/DPLL_algorithm">http://en.wikipedia.org/wiki/DPLL_algorithm</a></blockquote>#
adamc@82 913 It might also be instructive to implement the unit propagation and pure literal elimination optimizations described there or some other optimizations that have been used in modern SAT solvers.#</li>#
adamc@82 914
adamc@82 915 #</ol>#%\end{enumerate}% *)