Library Rustre.Nelist

Require Import Setoid.
Require Import Morphisms.

Non-empty lists

This module re-implements the List library, specialized to the case where the list is necessarily non-empty.

Set Implicit Arguments.
Require List.

Datatype


Inductive nelist (A : Type) : Type :=
  | nebase (e : A)
  | necons (e : A) (l : nelist A).

Operations


Fixpoint length {A} (l : nelist A) :=
  match l with
    | nebase _ ⇒ 1
    | necons _ S (length )
  end.

Fixpoint nelist2list {A} (l : nelist A) : list A :=
  match l with
    | nebase econs e nil
    | necons e cons e (nelist2list )
  end.


Fixpoint map {A B : Type} (f : AB) l :=
  match l with
    | nebase enebase (f e)
    | necons e necons (f e) (map f )
  end.

Definition fold_left {A B : Type} (f : ABA) :=
fix fold_left (l : nelist B) (a0 : A) {struct l} : A :=
  match l with
  | nebase ef a0 e
  | necons b tfold_left t (f a0 b)
  end.

Definition fold_right {A B : Type} (f : BAA) (a0 : A) :=
  fix fold_right (l : nelist B) : A :=
  match l with
  | nebase bf b a0
  | necons b tf b (fold_right t)
  end.

Fixpoint combine {A B : Type} (l : nelist A) ( : nelist B) {struct l} : nelist (A × B) :=
  match l, with
    | nebase a, nebase bnebase (a, b)
    | nebase a, necons b lbnebase (a, b)
    | necons a la, nebase bnebase (a, b)
    | necons a la, necons b lbnecons (a, b) (combine la lb)
  end.

Definition alls {A B} c (l : nelist A) : nelist B := map (fun _c) l.

Predicates


Fixpoint In {A : Type} (x : A) (l : nelist A): Prop :=
  match l with
    | nebase ex = e
    | necons e x = e In x
  end.

Inductive Forall {A : Type} (P : AProp) : nelist AProp :=
  | Forall_nil : x : A, P xForall P (nebase x)
  | Forall_cons : (x : A) (l : nelist A), P xForall P lForall P (necons x l).

Inductive Forall2 {A B : Type} (R : ABProp) : nelist Anelist BProp :=
  | Forall2_nil : x y, R x yForall2 R (nebase x) (nebase y)
  | Forall2_cons : x y l , R x yForall2 R l Forall2 R (necons x l) (necons y ).

Inductive Exists {A : Type} (P : AProp) : nelist AProp :=
  | Exists_base : x, P xExists P (nebase x)
  | Exists_cons_hd : x l, P xExists P (necons x l)
  | Exists_cons_tl : x l, Exists P lExists P (necons x l).

Inductive NoDup {A : Type} : nelist AProp :=
    NoDup_base : x, NoDup (nebase x)
  | NoDup_cons : x l, ¬In x lNoDup lNoDup (necons x l).

Properties


Ltac inv H := inversion H; subst; clear H.

Lemma in_necons_spec:
   T xs x (y: T),
    In x (necons y xs) y = x In x xs.
Proof.
  induction xs; [firstorder|].
  split; intro H.
  - change (x = y In x (necons e xs)) in H.
    destruct H; [symmetry in H; now auto|].
    apply IHxs in H.
    destruct H; [subst e|]; simpl; now auto.
  - destruct H.
    subst y; simpl; now auto.
    apply IHxs in H.
    destruct H; [symmetry in H|]; simpl; now auto.
Qed.

About length


Lemma diff_length_nebase_necons : {A B} (a : A) (b : B) l, length (nebase a) length (necons b l).
Proof. intros A B a b [? | ? ?]; simpl; discriminate. Qed.

About nelist2list


Lemma nelist2list_non_empty : A (l : nelist A), nelist2list l nil.
Proof. intros A [e | e l]; simpl; discriminate. Qed.

Lemma nelist2list_In : {A} (x : A) l, List.In x (nelist2list l) In x l.
Proof. intros A x l. induction l; simpl; try rewrite IHl; intuition. Qed.

About map


Lemma map_compat {A B : Type} : Proper ((eq ==> eq) ==> eq ==> eq) (@map A B).
Proof. intros f g Hfg l Hl. subst . induction l; simpl; f_equal; auto. Qed.

Lemma map_compose:
   A B C (f: AB)(g: BC) xs,
    map g (map f xs) = map (fun xg (f x)) xs.
Proof.
  intros until xs.
  induction xs as [|x xs IH]; simpl; congruence.
Qed.

Definition injective {A B} (f : AB) := x y, f x = f yx = y.

Lemma map_In : {A B : Type} (f : AB), injective f
   l x, In (f x) (map f l) In x l.
Proof.
intros A B f Hf l x. induction l as [e | e l]; simpl.
+ split; intro Hin; inversion Hin; trivial; now apply Hf.
+ rewrite IHl. firstorder. subst. tauto.
Qed.

Lemma map_eq_nebase : {A B : Type} (f : AB) l y, map f l = nebase y x, l = nebase x f x = y.
Proof.
intros A B f l y. destruct l; simpl; split; intro H; decompose [ex and] H || inversion_clear H; subst; eauto.
- inversion H1. now subst.
- discriminate.
Qed.

Lemma map_eq_necons : {A B : Type} (f : AB) l y ,
  map f l = necons y x l´´, l = necons x l´´ f x = y map f l´´ = .
Proof.
intros A B f l y . destruct l; simpl; split; intro H; decompose [ex and] H || inversion_clear H; subst; eauto.
- discriminate.
- inversion H0. now subst.
Qed.

Lemma map_length : {A B : Type} (f : AB) l, length (map f l) = length l.
Proof. intros A B f l. induction l; simpl; auto. Qed.

Lemma nelist2list_map : {A B : Type} (f : AB) l,
  nelist2list (map f l) = List.map f (nelist2list l).
Proof. intros A B f l. induction l; simpl; try rewrite IHl; reflexivity. Qed.

About Forall


Lemma Forall_forall : {A : Type} P l, Forall P l x : A, In x lP x.
Proof.
intros A P l. induction l; simpl.
+ split; intro Hin.
  - intros. subst. now inversion_clear Hin.
  - constructor. now apply Hin.
+ split; intro Hin.
  - intros. inversion_clear Hin. rewrite IHl in ×. destruct H; subst; auto.
  - constructor.
    × apply Hin. now left.
    × rewrite IHl. intros. apply Hin. auto.
Qed.

Lemma nelist2list_Forall : {A} P (l : nelist A), List.Forall P (nelist2list l) Forall P l.
Proof. intros A P l. rewrite Forall_forall, List.Forall_forall. now setoid_rewrite nelist2list_In. Qed.

Lemma Forall_map : {A B} (f : AB) P l, Forall P (map f l) Forall (fun xP (f x)) l.
Proof. intros A B f P l. induction l; split; intro Hl; inv Hl; constructor; try rewrite IHl in *; auto. Qed.

Lemma Forall2_length: {A B : Type} (R : ABProp) l1 l2,
  Forall2 R l1 l2length l1 = length l2.
Proof. intros A B R l1. induction l1; intros [|] Hall; inversion_clear Hall; simpl; auto. Qed.

Lemma Forall2_det : {A B : Type} (R : ABProp),
  ( x y1 y2, R x y1R x y2y1 = y2) →
   xs ys1 ys2, Forall2 R xs ys1Forall2 R xs ys2ys1 = ys2.
Proof.
intros A B R HR xs. induction xs as [x | x xs]; intros ys1 ys2 Hall1 Hall2.
- inv Hall1. inv Hall2. f_equal. eauto.
- inv Hall1. inv Hall2. f_equal; eauto.
Qed.

Lemma Forall2_map_l : {A B C} (f : AB) (R : BCProp) l1 l2,
  Forall2 R (map f l1) l2 Forall2 (fun x yR (f x) y) l1 l2.
Proof.
intros A B C f R l1. induction l1; intro l2; split; intro Hl; inv Hl; simpl;
now constructor; trivial; now apply IHl1.
Qed.

Lemma Forall2_map_r : {A B C} (f : AC) (R : BCProp) l1 l2,
  Forall2 R l1 (map f l2) Forall2 (fun x yR x (f y)) l1 l2.
Proof.
intros A B C f R l1 l2. revert l1. induction l2; intro l1; split; intro Hl; inv Hl; simpl;
now constructor; trivial; now apply IHl2.
Qed.

Corollary Forall2_map_lr : {A B C D} (f : AC) (g : BD) (R : CDProp) l1 l2,
  Forall2 R (map f l1) (map g l2) Forall2 (fun x yR (f x) (g y)) l1 l2.
Proof. intros. now rewrite Forall2_map_l, Forall2_map_r. Qed.

Lemma Forall2_eq : {A} l1 l2, Forall2 (@eq A) l1 l2 l1 = l2.
Proof.
intros A l1 l2. split; intro Heq; subst.
- revert l2 Heq. induction l1; intros l2 Heq; inv Heq; trivial; f_equal; auto.
- induction l2; constructor; auto.
Qed.

About Exists


Lemma Exists_exists : {A : Type} P l, Exists P l x : A, In x l P x.
Proof.
intros A P l. induction l; simpl.
× split; intro Hin.
  + e. inversion_clear Hin. tauto.
  + destruct Hin as [? [? ?]]. subst. now constructor.
× split; intro Hin.
  + inversion_clear Hin.
    - e. tauto.
    - rewrite IHl in H. destruct H as [x ?]. x. tauto.
  + destruct Hin as [x [[? | ?] ?]].
    - subst. now constructor 2.
    - constructor 3. rewrite IHl. x. tauto.
Qed.

Lemma nelist2list_Exists : {A} P (l : nelist A), List.Exists P (nelist2list l) Exists P l.
Proof. intros A P l. rewrite Exists_exists, List.Exists_exists. now setoid_rewrite nelist2list_In. Qed.

About NoDup


Lemma nelist2list_NoDup : {A} (l : nelist A), List.NoDup (nelist2list l) NoDup l.
Proof.
intros A l. induction l; simpl.
- split; intro Hnodup; inv Hnodup; repeat constructor; intuition.
- split; intro Hnodup; inv Hnodup; repeat constructor;
  now rewrite ?IHl in *; try (rewrite <- nelist2list_In || rewrite nelist2list_In).
Qed.