Skip to content
Snippets Groups Projects
Commit 36f903b6 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

More hlist stuff.

parent 46db392a
No related branches found
No related tags found
No related merge requests found
From stdpp Require Import base. From stdpp Require Import tactics.
(* Not using [list Type] in order to avoid universe inconsistencies *) (* Not using [list Type] in order to avoid universe inconsistencies *)
Inductive tlist := tnil : tlist | tcons : Type tlist tlist. Inductive tlist := tnil : tlist | tcons : Type tlist tlist.
...@@ -7,22 +7,53 @@ Inductive hlist : tlist → Type := ...@@ -7,22 +7,53 @@ Inductive hlist : tlist → Type :=
| hnil : hlist tnil | hnil : hlist tnil
| hcons {A As} : A hlist As hlist (tcons A As). | hcons {A As} : A hlist As hlist (tcons A As).
Fixpoint tapp (As Bs : tlist) : tlist :=
match As with tnil => Bs | tcons A As => tcons A (tapp As Bs) end.
Fixpoint happ {As Bs} (xs : hlist As) (ys : hlist Bs) : hlist (tapp As Bs) :=
match xs with hnil => ys | hcons _ _ x xs => hcons x (happ xs ys) end.
Fixpoint hhead {A As} (xs : hlist (tcons A As)) : A :=
match xs with hnil => () | hcons _ _ x _ => x end.
Fixpoint htail {A As} (xs : hlist (tcons A As)) : hlist As :=
match xs with hnil => () | hcons _ _ _ xs => xs end.
Fixpoint hheads {As Bs} : hlist (tapp As Bs) hlist As :=
match As with
| tnil => λ _, hnil
| tcons A As => λ xs, hcons (hhead xs) (hheads (htail xs))
end.
Fixpoint htails {As Bs} : hlist (tapp As Bs) hlist Bs :=
match As with
| tnil => id
| tcons A As => λ xs, htails (htail xs)
end.
Fixpoint himpl (As : tlist) (B : Type) : Type := Fixpoint himpl (As : tlist) (B : Type) : Type :=
match As with tnil => B | tcons A As => A himpl As B end. match As with tnil => B | tcons A As => A himpl As B end.
Definition happly {As B} (f : himpl As B) (xs : hlist As) : B := Definition hinit {B} (y : B) : himpl tnil B := y.
Definition hlam {A As B} (f : A himpl As B) : himpl (tcons A As) B := f.
Arguments hlam _ _ _ _ _/.
Definition hcurry {As B} (f : himpl As B) (xs : hlist As) : B :=
(fix go As xs := (fix go As xs :=
match xs in hlist As return himpl As B B with match xs in hlist As return himpl As B B with
| hnil => λ f, f | hnil => λ f, f
| hcons A As x xs => λ f, go As xs (f x) | hcons A As x xs => λ f, go As xs (f x)
end) _ xs f. end) _ xs f.
Coercion happly : himpl >-> Funclass. Coercion hcurry : himpl >-> Funclass.
Fixpoint huncurry {As B} : (hlist As B) himpl As B :=
match As with
| tnil => λ f, f hnil
| tcons x xs => λ f, hlam (λ x, huncurry (f hcons x))
end.
Lemma hcurry_uncurry {As B} (f : hlist As B) xs : huncurry f xs = f xs.
Proof. by induction xs as [|A As x xs IH]; simpl; rewrite ?IH. Qed.
Fixpoint hcompose {As B C} (f : B C) {struct As} : himpl As B himpl As C := Fixpoint hcompose {As B C} (f : B C) {struct As} : himpl As B himpl As C :=
match As with match As with
| tnil => f | tnil => f
| tcons A As => λ g x, hcompose f (g x) | tcons A As => λ g, hlam (λ x, hcompose f (g x))
end. end.
Definition hinit {B} (y : B) : himpl tnil B := y.
Definition hlam {A As B} (f : A himpl As B) : himpl (tcons A As) B := f.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment