DepList.v 6.55 KB
Newer Older
1
(* Copyright (c) 2008-2009, Adam Chlipala
Adam Chlipala's avatar
Adam Chlipala committed
2 3 4 5 6 7 8 9
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)

Adam Chlipala's avatar
Adam Chlipala committed
10
(* Dependent list types presented in Chapter 9 *)
Adam Chlipala's avatar
Adam Chlipala committed
11

12
Require Import Arith List CpdtTactics.
Adam Chlipala's avatar
Adam Chlipala committed
13 14 15 16 17

Set Implicit Arguments.


Section ilist.
Adam Chlipala's avatar
Adam Chlipala committed
18
  Variable A : Type.
Adam Chlipala's avatar
Adam Chlipala committed
19

20 21 22 23 24 25 26 27 28 29 30
  Inductive ilist : nat -> Type :=
  | INil : ilist O
  | ICons : forall n, A -> ilist n -> ilist (S n).

  Definition hd n (ls : ilist (S n)) :=
    match ls in ilist n' return match n' with
                                  | O => unit
                                  | S _ => A
                                end with
      | INil => tt
      | ICons _ x _ => x
Adam Chlipala's avatar
Adam Chlipala committed
31
    end.
32 33 34 35 36 37 38 39 40 41 42 43
  Definition tl n (ls : ilist (S n)) :=
    match ls in ilist n' return match n' with
                                  | O => unit
                                  | S n => ilist n
                                end with
      | INil => tt
      | ICons _ _ ls' => ls'
    end.

  Inductive fin : nat -> Set :=
  | First : forall n, fin (S n)
  | Next : forall n, fin n -> fin (S n).
Adam Chlipala's avatar
Adam Chlipala committed
44

45 46 47 48 49 50 51 52 53
  Fixpoint get n (ls : ilist n) : fin n -> A :=
    match ls with
      | INil => fun idx =>
        match idx in fin n' return (match n' with
                                        | O => A
                                        | S _ => unit
                                      end) with
          | First _ => tt
          | Next _ _ => tt
Adam Chlipala's avatar
Adam Chlipala committed
54
        end
55 56 57 58 59
      | ICons _ x ls' => fun idx =>
        match idx in fin n' return (fin (pred n') -> A) -> A with
          | First _ => fun _ => x
          | Next _ idx' => fun get_ls' => get_ls' idx'
        end (get ls')
Adam Chlipala's avatar
Adam Chlipala committed
60
    end.
Adam Chlipala's avatar
Adam Chlipala committed
61 62 63 64 65

  Section everywhere.
    Variable x : A.

    Fixpoint everywhere (n : nat) : ilist n :=
66 67 68
      match n with
        | O => INil
        | S n' => ICons x (everywhere n')
Adam Chlipala's avatar
Adam Chlipala committed
69 70 71 72 73 74
      end.
  End everywhere.

  Section singleton.
    Variables x default : A.

75 76 77
    Fixpoint singleton (n m : nat) : ilist n :=
      match n with
        | O => INil
Adam Chlipala's avatar
Adam Chlipala committed
78 79
        | S n' =>
          match m with
80 81
            | O => ICons x (everywhere default n')
            | S m' => ICons default (singleton n' m')
Adam Chlipala's avatar
Adam Chlipala committed
82 83 84 85 86 87 88
          end
      end.
  End singleton.

  Section map2.
    Variable f : A -> A -> A.

89 90 91 92
    Fixpoint map2 n (il1 : ilist n) : ilist n -> ilist n :=
      match il1 in ilist n return ilist n -> ilist n with
        | INil => fun _ => INil
        | ICons _ x il1' => fun il2 => ICons (f x (hd il2)) (map2 il1' (tl il2))
Adam Chlipala's avatar
Adam Chlipala committed
93 94
      end.
  End map2.
Adam Chlipala's avatar
Adam Chlipala committed
95 96 97 98 99 100

  Section fold.
    Variable B : Type.
    Variable f : A -> B -> B.
    Variable i : B.

101 102 103 104
    Fixpoint foldr n (il : ilist n) : B :=
      match il with
        | INil => i
        | ICons _ x il' => f x (foldr il')
Adam Chlipala's avatar
Adam Chlipala committed
105 106
      end.
  End fold.
Adam Chlipala's avatar
Adam Chlipala committed
107 108
End ilist.

109 110
Implicit Arguments INil [A].
Implicit Arguments First [n].
Adam Chlipala's avatar
Adam Chlipala committed
111

Adam Chlipala's avatar
Adam Chlipala committed
112 113 114 115
Section imap.
  Variables A B : Type.
  Variable f : A -> B.

116 117 118 119
  Fixpoint imap n (il : ilist A n) : ilist B n :=
    match il with
      | INil => INil
      | ICons _ x il' => ICons (f x) (imap il')
Adam Chlipala's avatar
Adam Chlipala committed
120 121 122
    end.
End imap.

Adam Chlipala's avatar
Adam Chlipala committed
123 124 125 126
Section hlist.
  Variable A : Type.
  Variable B : A -> Type.

127 128 129 130 131 132 133 134 135 136 137 138
  Inductive hlist : list A -> Type :=
  | HNil : hlist nil
  | HCons : forall (x : A) (ls : list A), B x -> hlist ls -> hlist (x :: ls).

  Definition hhd ls (hl : hlist ls) :=
    match hl in hlist ls return match ls with
                                  | nil => unit
                                  | x :: _ => B x
                                end with
      | HNil => tt
      | HCons _ _ v _ => v
    end.
Adam Chlipala's avatar
Adam Chlipala committed
139

140 141 142 143 144 145 146 147
  Definition htl ls (hl : hlist ls) :=
    match hl in hlist ls return match ls with
                                  | nil => unit
                                  | _ :: ls' => hlist ls'
                                end with
      | HNil => tt
      | HCons _ _ _ hl' => hl'
    end.
148

Adam Chlipala's avatar
Adam Chlipala committed
149 150
  Variable elm : A.

151 152 153 154 155 156 157 158 159 160 161 162 163
  Inductive member : list A -> Type :=
  | HFirst : forall ls, member (elm :: ls)
  | HNext : forall x ls, member ls -> member (x :: ls).

  Fixpoint hget ls (mls : hlist ls) : member ls -> B elm :=
    match mls with
      | HNil => fun mem =>
        match mem in member ls' return (match ls' with
                                          | nil => B elm
                                          | _ :: _ => unit
                                        end) with
          | HFirst _ => tt
          | HNext _ _ _ => tt
Adam Chlipala's avatar
Adam Chlipala committed
164
        end
165 166 167 168 169 170 171 172 173
      | HCons _ _ x mls' => fun mem =>
        match mem in member ls' return (match ls' with
                                          | nil => Empty_set
                                          | x' :: ls'' =>
                                            B x' -> (member ls'' -> B elm) -> B elm
                                        end) with
          | HFirst _ => fun x _ => x
          | HNext _ _ mem' => fun _ get_mls' => get_mls' mem'
        end x (hget mls')
Adam Chlipala's avatar
Adam Chlipala committed
174
    end.
175

176 177 178 179
  Fixpoint happ (ls1 : list A) (hl1 : hlist ls1) : forall ls2, hlist ls2 -> hlist (ls1 ++ ls2) :=
    match hl1 in hlist ls1 return forall ls2, hlist ls2 -> hlist (ls1 ++ ls2) with
      | HNil => fun _ hl2 => hl2
      | HCons _ _ x hl1' => fun _ hl2 => HCons x (happ hl1' hl2)
180
    end.
Adam Chlipala's avatar
Adam Chlipala committed
181 182 183 184

  Variable f : forall x, B x.

  Fixpoint hmake (ls : list A) : hlist ls :=
185 186 187
    match ls with
      | nil => HNil
      | x :: ls' => HCons (f x) (hmake ls')
Adam Chlipala's avatar
Adam Chlipala committed
188
    end.
Adam Chlipala's avatar
Adam Chlipala committed
189 190 191

  Theorem hget_hmake : forall ls (m : member ls),
    hget (hmake ls) m = f elm.
192 193 194 195
    induction ls; crush;
      match goal with
        | [ |- context[match ?E with HFirst _ => _ | HNext _ _ _ => _ end] ] => dep_destruct E
      end; crush.
Adam Chlipala's avatar
Adam Chlipala committed
196
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
197 198
End hlist.

199 200
Implicit Arguments HNil [A B].
Implicit Arguments HCons [A B x ls].
Adam Chlipala's avatar
Adam Chlipala committed
201
Implicit Arguments hmake [A B].
202

203 204
Implicit Arguments HFirst [A elm ls].
Implicit Arguments HNext [A elm x ls].
Adam Chlipala's avatar
Adam Chlipala committed
205

206
Infix ":::" := HCons (right associativity, at level 60).
207
Infix "+++" := happ (right associativity, at level 60).
Adam Chlipala's avatar
Adam Chlipala committed
208 209 210 211 212 213 214

Section hmap.
  Variable A : Type.
  Variables B1 B2 : A -> Type.

  Variable f : forall x, B1 x -> B2 x.

215 216 217 218
  Fixpoint hmap (ls : list A) (hl : hlist B1 ls) : hlist B2 ls :=
    match hl with
      | HNil => HNil
      | HCons _ _ x hl' => f x ::: hmap hl'
Adam Chlipala's avatar
Adam Chlipala committed
219
    end.
220 221 222

  Theorem hmap_happ : forall ls2 (h2 : hlist B1 ls2) ls1 (h1 : hlist B1 ls1),
    hmap h1 +++ hmap h2 = hmap (h1 +++ h2).
223
    induction h1; crush.
224
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
225 226 227

  Theorem hget_hmap : forall elm ls (hls : hlist B1 ls) (m : member elm ls),
    hget (hmap hls) m = f (hget hls m).
228 229 230 231
    induction hls; crush;
      match goal with
        | [ |- context[match ?E with HFirst _ => _ | HNext _ _ _ => _ end] ] => dep_destruct E
      end; crush.
Adam Chlipala's avatar
Adam Chlipala committed
232
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
233 234 235
End hmap.

Implicit Arguments hmap [A B1 B2 ls].