annotate src/DataStruct.v @ 105:ec0fb0f00f46

Start of DataStruct
author Adam Chlipala <adamc@hcoop.net>
date Mon, 13 Oct 2008 10:47:21 -0400
parents
children 1d2746eae1a6
rev   line source
adamc@105 1 (* Copyright (c) 2008, Adam Chlipala
adamc@105 2 *
adamc@105 3 * This work is licensed under a
adamc@105 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@105 5 * Unported License.
adamc@105 6 * The license text is available at:
adamc@105 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@105 8 *)
adamc@105 9
adamc@105 10 (* begin hide *)
adamc@105 11 Require Import List.
adamc@105 12
adamc@105 13 Require Import Tactics.
adamc@105 14
adamc@105 15 Set Implicit Arguments.
adamc@105 16 (* end hide *)
adamc@105 17
adamc@105 18
adamc@105 19 (** %\chapter{Dependent Data Structures}% *)
adamc@105 20
adamc@105 21 (** Our red-black tree example from the last chapter illustrated how dependent types enable static enforcement of data structure invariants. To find interesting uses of dependent data structures, however, we need not look to the favorite examples of data structures and algorithms textbooks. More basic examples like known-length and heterogeneous lists prop up again and again as the building blocks of dependent programs. There is a surprisingly large design space for this class of data structure, and we will spend this chapter exploring it. *)
adamc@105 22
adamc@105 23
adamc@105 24 (** * More Known-Length Lists *)
adamc@105 25
adamc@105 26 Section ilist.
adamc@105 27 Variable A : Set.
adamc@105 28
adamc@105 29 Inductive ilist : nat -> Set :=
adamc@105 30 | Nil : ilist O
adamc@105 31 | Cons : forall n, A -> ilist n -> ilist (S n).
adamc@105 32
adamc@105 33 Inductive index : nat -> Set :=
adamc@105 34 | First : forall n, index (S n)
adamc@105 35 | Next : forall n, index n -> index (S n).
adamc@105 36
adamc@105 37 Fixpoint get n (ls : ilist n) {struct ls} : index n -> A :=
adamc@105 38 match ls in ilist n return index n -> A with
adamc@105 39 | Nil => fun idx =>
adamc@105 40 match idx in index n' return (match n' with
adamc@105 41 | O => A
adamc@105 42 | S _ => unit
adamc@105 43 end) with
adamc@105 44 | First _ => tt
adamc@105 45 | Next _ _ => tt
adamc@105 46 end
adamc@105 47 | Cons _ x ls' => fun idx =>
adamc@105 48 match idx in index n' return (index (pred n') -> A) -> A with
adamc@105 49 | First _ => fun _ => x
adamc@105 50 | Next _ idx' => fun get_ls' => get_ls' idx'
adamc@105 51 end (get ls')
adamc@105 52 end.
adamc@105 53 End ilist.
adamc@105 54
adamc@105 55 Implicit Arguments Nil [A].
adamc@105 56
adamc@105 57 Section ilist_map.
adamc@105 58 Variables A B : Set.
adamc@105 59 Variable f : A -> B.
adamc@105 60
adamc@105 61 Fixpoint imap n (ls : ilist A n) {struct ls} : ilist B n :=
adamc@105 62 match ls in ilist _ n return ilist B n with
adamc@105 63 | Nil => Nil
adamc@105 64 | Cons _ x ls' => Cons (f x) (imap ls')
adamc@105 65 end.
adamc@105 66
adamc@105 67 Theorem get_imap : forall n (ls : ilist A n) (idx : index n),
adamc@105 68 get (imap ls) idx = f (get ls idx).
adamc@105 69 induction ls; crush;
adamc@105 70 match goal with
adamc@105 71 | [ |- context[match ?IDX with First _ => _ | Next _ _ => _ end] ] =>
adamc@105 72 dep_destruct IDX; crush
adamc@105 73 end.
adamc@105 74 Qed.
adamc@105 75 End ilist_map.