comparison 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
comparison
equal deleted inserted replaced
402:df7cd24383ae 403:1edeec5d5d0c
13 Require Import CpdtTactics. 13 Require Import CpdtTactics.
14 14
15 Set Implicit Arguments. 15 Set Implicit Arguments.
16 (* end hide *) 16 (* end hide *)
17 17
18 (** printing <-- $\longleftarrow$ *)
19
18 20
19 (** %\part{Programming with Dependent Types} 21 (** %\part{Programming with Dependent Types}
20 22
21 \chapter{Subset Types and Variations}% *) 23 \chapter{Subset Types and Variations}% *)
22 24
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 %\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. *) 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. *)
24 26
25 27
26 (** * Introducing Subset Types *) 28 (** * Introducing Subset Types *)
27 29
28 (** Let us consider several ways of implementing the natural number predecessor function. We start by displaying the definition from the standard library: *) 30 (** Let us consider several ways of implementing the natural number predecessor function. We start by displaying the definition from the standard library: *)
112 match n return n > 0 -> nat with 114 match n return n > 0 -> nat with
113 | O => fun pf : 0 > 0 => match zgtz pf with end 115 | O => fun pf : 0 > 0 => match zgtz pf with end
114 | S n' => fun _ => n' 116 | S n' => fun _ => n'
115 end. 117 end.
116 118
117 (** 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. 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.
118 120
119 Let us now take a look at the OCaml code Coq generates for [pred_strong1]. *) 121 Let us now take a look at the OCaml code Coq generates for [pred_strong1]. *)
120 122
121 Extraction pred_strong1. 123 Extraction pred_strong1.
122 124
136 | S n' -> n' 138 | S n' -> n'
137 </pre># *) 139 </pre># *)
138 140
139 (** 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. 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.
140 142
141 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]. *) 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]. *)
142 144
143 Print sig. 145 Print sig.
144 (** %\vspace{-.15in}% [[ 146 (** %\vspace{-.15in}% [[
145 Inductive sig (A : Type) (P : A -> Prop) : Type := 147 Inductive sig (A : Type) (P : A -> Prop) : Type :=
146 exist : forall x : A, P x -> sig P 148 exist : forall x : A, P x -> sig P
266 | S n' => fun _ => exist _ n' _ 268 | S n' => fun _ => exist _ n' _
267 end); crush. 269 end); crush.
268 (* end thide *) 270 (* end thide *)
269 Defined. 271 Defined.
270 272
271 (** 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. *) 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. *)
272 274
273 Print pred_strong4. 275 Print pred_strong4.
274 (** %\vspace{-.15in}% [[ 276 (** %\vspace{-.15in}% [[
275 pred_strong4 = 277 pred_strong4 =
276 fun n : nat => 278 fun n : nat =>
543 | Nil -> false 545 | Nil -> false
544 | Cons (x', ls') -> 546 | Cons (x', ls') ->
545 (match a_eq_dec x x' with 547 (match a_eq_dec x x' with
546 | true -> true 548 | true -> true
547 | false -> in_dec a_eq_dec x ls') 549 | false -> in_dec a_eq_dec x ls')
548 </pre># *) 550 </pre>#
551
552 This is more or the less code for the corresponding function from the OCaml standard library. *)
549 553
550 554
551 (** * Partial Subset Types *) 555 (** * Partial Subset Types *)
552 556
553 (** 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. *) 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. *)
650 [|(m1, m2)|]); tauto. 654 [|(m1, m2)|]); tauto.
651 Defined. 655 Defined.
652 (* end thide *) 656 (* end thide *)
653 657
654 (** 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.% *) 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.% *)
655
656 (** printing <-- $\longleftarrow$ *)
657 659
658 Notation "x <-- e1 ; e2" := (match e1 with 660 Notation "x <-- e1 ; e2" := (match e1 with
659 | inright _ => !! 661 | inright _ => !!
660 | inleft (exist x _) => e2 662 | inleft (exist x _) => e2
661 end) 663 end)
858 induction 1; inversion 1; crush. 860 induction 1; inversion 1; crush.
859 Qed. 861 Qed.
860 862
861 (** Now we can define the type-checker. Its type expresses that it only fails on untypable expressions. *) 863 (** Now we can define the type-checker. Its type expresses that it only fails on untypable expressions. *)
862 864
863 (** printing <-- $\longleftarrow$ *)
864
865 (* end thide *) 865 (* end thide *)
866 Definition typeCheck' : forall e : exp, {t : type | hasType e t} + {forall t, ~ hasType e t}. 866 Definition typeCheck' : forall e : exp, {t : type | hasType e t} + {forall t, ~ hasType e t}.
867 (* begin thide *) 867 (* begin thide *)
868 Hint Constructors hasType. 868 Hint Constructors hasType.
869 (** We register all of the typing rules as hints. *) 869 (** We register all of the typing rules as hints. *)