Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

Require Setoid.
Require Import PeanoNat Le Gt Minus Bool Lt.

Set Implicit Arguments.
(* Set Universe Polymorphism. *)

(******************************************************************)

Basics: definition of polymorphic lists and some operations

(******************************************************************)
The definition of list is now in Init/Datatypes, as well as the definitions of length and app
Open Scope list_scope.
Standard notations for lists. In a special module to avoid conflicts.
Module ListNotations.
Notation "[ ]" := nil (format "[ ]") : list_scope.
Notation "[ x ]" := (cons x nil) : list_scope.
Notation "[ x ; y ; .. ; z ]" :=  (cons x (cons y .. (cons z nil) ..)) : list_scope.
End ListNotations.

Import ListNotations.

Section Lists.

  Variable A : Type.
Head and tail
  Definition hd (default:A) (l:list A) :=
    match l with
      | [] => default
      | x :: _ => x
    end.

  Definition hd_error (l:list A) :=
    match l with
      | [] => None
      | x :: _ => Some x
    end.

  Definition tl (l:list A) :=
    match l with
      | [] => nil
      | a :: m => m
    end.
The In predicate
  Fixpoint In (a:A) (l:list A) : Prop :=
    match l with
      | [] => False
      | b :: m => b = a \/ In a m
    end.

End Lists.

Section Facts.

  Variable A : Type.

Generic facts

  
Discrimination
  
A:Type

forall (x : A) (l : list A), [] <> x :: l
A:Type

forall (x : A) (l : list A), [] <> x :: l
intros; discriminate. Qed.
Destruction
  
A:Type

forall l : list A, {x : A & {tl0 : list A | l = x :: tl0}} + {l = []}
A:Type

forall l : list A, {x : A & {tl0 : list A | l = x :: tl0}} + {l = []}
A:Type

{x : A & {tl0 : list A | [] = x :: tl0}} + {[] = []}
A:Type
a:A
tail:list A
IHtail:{x : A & {tl0 : list A | tail = x :: tl0}} + {tail = []}
{x : A & {tl0 : list A | a :: tail = x :: tl0}} + {a :: tail = []}
A:Type
a:A
tail:list A
IHtail:{x : A & {tl0 : list A | tail = x :: tl0}} + {tail = []}

{x : A & {tl0 : list A | a :: tail = x :: tl0}} + {a :: tail = []}
left; exists a, tail; reflexivity. Qed.
A:Type

forall (l : list A) (a : A) (r : list A), hd_error l = Some a /\ tl l = r <-> l = a :: r
A:Type

forall (l : list A) (a : A) (r : list A), hd_error l = Some a /\ tl l = r <-> l = a :: r
A:Type

forall (a : A) (r : list A), hd_error [] = Some a /\ tl [] = r <-> [] = a :: r
A:Type
x:A
xs:list A
forall (a : A) (r : list A), hd_error (x :: xs) = Some a /\ tl (x :: xs) = r <-> x :: xs = a :: r
A:Type

forall (a : A) (r : list A), hd_error [] = Some a /\ tl [] = r <-> [] = a :: r
A:Type
a:A
r:list A

None = Some a /\ [] = r <-> [] = a :: r
split; firstorder discriminate.
A:Type
x:A
xs:list A

forall (a : A) (r : list A), hd_error (x :: xs) = Some a /\ tl (x :: xs) = r <-> x :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A

hd_error (x :: xs) = Some a /\ tl (x :: xs) = r <-> x :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A

Some x = Some a /\ xs = r <-> x :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A

Some x = Some a /\ xs = r -> x :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A
x :: xs = a :: r -> Some x = Some a /\ xs = r
A:Type
x:A
xs:list A
a:A
r:list A

Some x = Some a /\ xs = r -> x :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A
H1:Some x = Some a
H2:xs = r

x :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A
H1:Some x = Some a
H2:xs = r
H0:x = a

a :: xs = a :: r
A:Type
x:A
xs:list A
a:A
r:list A
H1:Some x = Some a
H2:xs = r
H0:x = a

a :: r = a :: r
reflexivity.
A:Type
x:A
xs:list A
a:A
r:list A

x :: xs = a :: r -> Some x = Some a /\ xs = r
A:Type
x:A
xs:list A
a:A
r:list A
H:x :: xs = a :: r
H1:x = a
H2:xs = r

Some a = Some a /\ r = r
A:Type
a:A
r:list A
H:a :: r = a :: r

Some a = Some a /\ r = r
auto. Qed.
A:Type

forall (l : list A) (a : A), hd_error l = Some a -> l <> []
A:Type

forall (l : list A) (a : A), hd_error l = Some a -> l <> []
A:Type

forall (l : list A) (a : A), match l with | [] => None | x :: _ => Some x end = Some a -> l <> []
destruct l; now discriminate. Qed.
A:Type
l:list A

length l = 0 <-> l = []
A:Type
l:list A

length l = 0 <-> l = []
split; [now destruct l | now intros ->]. Qed.

Head and tail

  
A:Type

hd_error [] = None
A:Type

hd_error [] = None
simpl; reflexivity. Qed.
A:Type

forall (l : list A) (x : A), hd_error (x :: l) = Some x
A:Type

forall (l : list A) (x : A), hd_error (x :: l) = Some x
intros; simpl; reflexivity. Qed. (************************)

Facts about In

  (************************)
Characterization of In
  
A:Type

forall (a : A) (l : list A), In a (a :: l)
A:Type

forall (a : A) (l : list A), In a (a :: l)
simpl; auto. Qed.
A:Type

forall (a b : A) (l : list A), In b l -> In b (a :: l)
A:Type

forall (a b : A) (l : list A), In b l -> In b (a :: l)
simpl; auto. Qed.
A:Type
x, a:A
l:list A

~ In x (a :: l) <-> x <> a /\ ~ In x l
A:Type
x, a:A
l:list A

~ In x (a :: l) <-> x <> a /\ ~ In x l
A:Type
x, a:A
l:list A

~ (a = x \/ In x l) <-> x <> a /\ ~ In x l
intuition. Qed.
A:Type

forall a : A, ~ In a []
A:Type

forall a : A, ~ In a []
unfold not; intros a H; inversion_clear H. Qed.
A:Type

forall (x : A) (l : list A), In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2
A:Type

forall (x : A) (l : list A), In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2
A:Type
x, a:A
l:list A
IHl:In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2
H:a = x

exists l1 l2 : list A, a :: l = l1 ++ x :: l2
A:Type
x, a:A
l:list A
IHl:In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2
H:In x l
exists l1 l2 : list A, a :: l = l1 ++ x :: l2
A:Type
x:A
l:list A
IHl:In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2

exists l1 l2 : list A, x :: l = l1 ++ x :: l2
A:Type
x, a:A
l:list A
IHl:In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2
H:In x l
exists l1 l2 : list A, a :: l = l1 ++ x :: l2
A:Type
x, a:A
l:list A
IHl:In x l -> exists l1 l2 : list A, l = l1 ++ x :: l2
H:In x l

exists l1 l2 : list A, a :: l = l1 ++ x :: l2
A:Type
x, a:A
l:list A
IHl:In x l -> exists l0 l3 : list A, l = l0 ++ x :: l3
H:In x l
l1, l2:list A
H0:l = l1 ++ x :: l2

exists l0 l3 : list A, a :: l = l0 ++ x :: l3
A:Type
x, a:A
l:list A
IHl:In x l -> exists l0 l3 : list A, l = l0 ++ x :: l3
H:In x l
l1, l2:list A
H0:l = l1 ++ x :: l2

a :: l = a :: l1 ++ x :: l2
A:Type
x, a:A
l:list A
IHl:In x l -> exists l0 l3 : list A, l = l0 ++ x :: l3
H:In x l
l1, l2:list A
H0:l = l1 ++ x :: l2

l = l1 ++ x :: l2
auto. Qed.
Inversion
  
A:Type

forall (a b : A) (l : list A), In b (a :: l) -> a = b \/ In b l
A:Type

forall (a b : A) (l : list A), In b (a :: l) -> a = b \/ In b l
intros a b l H; inversion_clear H; auto. Qed.
Decidability of In
  
A:Type

(forall x y : A, {x = y} + {x <> y}) -> forall (a : A) (l : list A), {In a l} + {~ In a l}
A:Type

(forall x y : A, {x = y} + {x <> y}) -> forall (a : A) (l : list A), {In a l} + {~ In a l}
A:Type
H:forall x y : A, {x = y} + {x <> y}
a:A

{In a []} + {~ In a []}
A:Type
H:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:{In a l} + {~ In a l}
{In a (a0 :: l)} + {~ In a (a0 :: l)}
A:Type
H:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:{In a l} + {~ In a l}

{In a (a0 :: l)} + {~ In a (a0 :: l)}
A:Type
H:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:{In a l} + {~ In a l}
n:a0 <> a

{a0 = a \/ In a l} + {~ (a0 = a \/ In a l)}
A:Type
H:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
n0:~ In a l
n:a0 <> a

{a0 = a \/ In a l} + {~ (a0 = a \/ In a l)}
right; unfold not; intros [Hc1| Hc2]; auto. Defined. (**************************)

Facts about app

  (**************************)
Discrimination
  
A:Type

forall (x y : list A) (a : A), [] <> x ++ a :: y
A:Type

forall (x y : list A) (a : A), [] <> x ++ a :: y
A:Type

forall (x y : list A) (a : A), [] = x ++ a :: y -> False
A:Type
y:list A
a:A
H:[] = a :: y

False
A:Type
a:A
l, y:list A
a0:A
H:[] = a :: l ++ a0 :: y
False
A:Type
a:A
l, y:list A
a0:A
H:[] = a :: l ++ a0 :: y

False
discriminate H. Qed.
Concat with nil
  
A:Type

forall l : list A, [] ++ l = l
A:Type

forall l : list A, [] ++ l = l
reflexivity. Qed.
A:Type

forall l : list A, l ++ [] = l
A:Type

forall l : list A, l ++ [] = l
induction l; simpl; f_equal; auto. Qed. (* begin hide *) (* Deprecated *)
A:Type

forall l : list A, l = l ++ []
A:Type

forall l : list A, l = l ++ []
symmetry; apply app_nil_r. Qed. (* end hide *)
app is associative
  
A:Type

forall l m n : list A, l ++ m ++ n = (l ++ m) ++ n
A:Type

forall l m n : list A, l ++ m ++ n = (l ++ m) ++ n
intros l m n; induction l; simpl; f_equal; auto. Qed. (* begin hide *) (* Deprecated *)
A:Type

forall l m n : list A, (l ++ m) ++ n = l ++ m ++ n
A:Type

forall l m n : list A, (l ++ m) ++ n = l ++ m ++ n
auto using app_assoc. Qed. Hint Resolve app_assoc_reverse : core. (* end hide *)
app commutes with cons
  
A:Type

forall (x y : list A) (a : A), a :: x ++ y = (a :: x) ++ y
A:Type

forall (x y : list A) (a : A), a :: x ++ y = (a :: x) ++ y
auto. Qed.
Facts deduced from the result of a concatenation
  
A:Type

forall l l' : list A, l ++ l' = [] -> l = [] /\ l' = []
A:Type

forall l l' : list A, l ++ l' = [] -> l = [] /\ l' = []
A:Type
x:A
l:list A

x :: l ++ [] = [] -> x :: l = [] /\ [] = []
A:Type
x:A
l:list A
y:A
l':list A
x :: l ++ y :: l' = [] -> x :: l = [] /\ y :: l' = []
A:Type
x:A
l:list A
y:A
l':list A

x :: l ++ y :: l' = [] -> x :: l = [] /\ y :: l' = []
intros H; discriminate H. Qed.
A:Type

forall (x y : list A) (a : A), x ++ y = [a] -> x = [] /\ y = [a] \/ x = [a] /\ y = []
A:Type

forall (x y : list A) (a : A), x ++ y = [a] -> x = [] /\ y = [a] \/ x = [a] /\ y = []
A:Type

forall a : A, [] = [a] -> [] = [] /\ [] = [a] \/ [] = [a] /\ [] = []
A:Type
a:A
l:list A
forall a0 : A, a :: l = [a0] -> [] = [] /\ a :: l = [a0] \/ [] = [a0] /\ a :: l = []
A:Type
a:A
l:list A
forall a0 : A, a :: l ++ [] = [a0] -> a :: l = [] /\ [] = [a0] \/ a :: l = [a0] /\ [] = []
A:Type
a:A
l:list A
a0:A
l0:list A
forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A

forall a0 : A, a :: l = [a0] -> [] = [] /\ a :: l = [a0] \/ [] = [a0] /\ a :: l = []
A:Type
a:A
l:list A
forall a0 : A, a :: l ++ [] = [a0] -> a :: l = [] /\ [] = [a0] \/ a :: l = [a0] /\ [] = []
A:Type
a:A
l:list A
a0:A
l0:list A
forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A

forall a0 : A, a :: l ++ [] = [a0] -> a :: l = [] /\ [] = [a0] \/ a :: l = [a0] /\ [] = []
A:Type
a:A
l:list A
a0:A
l0:list A
forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
H:a :: l ++ [] = [a0]

a :: l = [a0]
A:Type
a:A
l:list A
a0:A
l0:list A
forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
H:a :: l ++ [] = [a0]

a :: l ++ [] = [a0] -> a :: l = [a0]
A:Type
a:A
l:list A
a0:A
l0:list A
forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
H:a :: l ++ [] = [a0]
E:l ++ [] = l

a :: l ++ [] = [a0] -> a :: l = [a0]
A:Type
a:A
l:list A
a0:A
l0:list A
forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
l0:list A

forall a1 : A, a :: l ++ a0 :: l0 = [a1] -> a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
l0:list A
a1:A
H:a :: l ++ a0 :: l0 = [a1]

a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
l0:list A
a1:A
H:a = a1
H0:l ++ a0 :: l0 = []

a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
A:Type
a:A
l:list A
a0:A
l0:list A
a1:A
H:a = a1
H0:l ++ a0 :: l0 = []
H1:[] = l ++ a0 :: l0

a :: l = [] /\ a0 :: l0 = [a1] \/ a :: l = [a1] /\ a0 :: l0 = []
apply app_cons_not_nil in H1 as []. Qed.
A:Type

forall (x y : list A) (a b : A), x ++ [a] = y ++ [b] -> x = y /\ a = b
A:Type

forall (x y : list A) (a b : A), x ++ [a] = y ++ [b] -> x = y /\ a = b
A:Type

forall a b : A, [a] = [b] -> [] = [] /\ a = b
A:Type
a:A
l:list A
forall a0 b : A, [a0] = a :: l ++ [b] -> [] = a :: l /\ a0 = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a b : A), l ++ [a] = y ++ [b] -> l = y /\ a = b
forall a b : A, x :: l ++ [a] = [b] -> x :: l = [] /\ a = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a0 b : A), l ++ [a0] = y ++ [b] -> l = y /\ a0 = b
a:A
l0:list A
forall a0 b : A, x :: l ++ [a0] = a :: l0 ++ [b] -> x :: l = a :: l0 /\ a0 = b
A:Type

forall a b : A, [a] = [b] -> [] = [] /\ a = b
A:Type
a, b:A
H:[a] = [b]

[] = [] /\ a = b
A:Type
a, b:A
H:[a] = [b]

a = b -> [] = [] /\ a = b
auto.
A:Type
a:A
l:list A

forall a0 b : A, [a0] = a :: l ++ [b] -> [] = a :: l /\ a0 = b
A:Type
a:A
l:list A
a0, b:A
H:[a0] = a :: l ++ [b]

[] = a :: l /\ a0 = b
A:Type
a:A
l:list A
a0, b:A
H1:a0 = a
H0:[] = l ++ [b]

[] = a :: l /\ a0 = b
apply app_cons_not_nil in H0 as [].
A:Type
x:A
l:list A
IHl:forall (y : list A) (a b : A), l ++ [a] = y ++ [b] -> l = y /\ a = b

forall a b : A, x :: l ++ [a] = [b] -> x :: l = [] /\ a = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a0 b0 : A), l ++ [a0] = y ++ [b0] -> l = y /\ a0 = b0
a, b:A
H:x :: l ++ [a] = [b]

x :: l = [] /\ a = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a0 b0 : A), l ++ [a0] = y ++ [b0] -> l = y /\ a0 = b0
a, b:A
H1:x = b
H0:l ++ [a] = []

x :: l = [] /\ a = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a0 b0 : A), l ++ [a0] = y ++ [b0] -> l = y /\ a0 = b0
a, b:A
H1:x = b
H0:l ++ [a] = []
H:[] = l ++ [a]

x :: l = [] /\ a = b
apply app_cons_not_nil in H as [].
A:Type
x:A
l:list A
IHl:forall (y : list A) (a0 b : A), l ++ [a0] = y ++ [b] -> l = y /\ a0 = b
a:A
l0:list A

forall a0 b : A, x :: l ++ [a0] = a :: l0 ++ [b] -> x :: l = a :: l0 /\ a0 = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a1 b0 : A), l ++ [a1] = y ++ [b0] -> l = y /\ a1 = b0
a:A
l0:list A
a0, b:A
H:x :: l ++ [a0] = a :: l0 ++ [b]

x :: l = a :: l0 /\ a0 = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a b0 : A), l ++ [a] = y ++ [b0] -> l = y /\ a = b0
l0:list A
a0, b:A
H0:l ++ [a0] = l0 ++ [b]

x :: l = x :: l0 /\ a0 = b
A:Type
x:A
l:list A
IHl:forall (y : list A) (a b : A), l ++ [a] = y ++ [b] -> l = y /\ a = b
a0:A
H0:l ++ [a0] = l ++ [a0]

x :: l = x :: l /\ a0 = a0
split; auto. Qed.
Compatibility with other operations
  
A:Type

forall l l' : list A, length (l ++ l') = length l + length l'
A:Type

forall l l' : list A, length (l ++ l') = length l + length l'
induction l; simpl; auto. Qed.
A:Type

forall (l m : list A) (a : A), In a (l ++ m) -> In a l \/ In a m
A:Type

forall (l m : list A) (a : A), In a (l ++ m) -> In a l \/ In a m
A:Type
l, m:list A
a:A

In a (l ++ m) -> In a l \/ In a m
A:Type
l, m:list A
a:A

forall (a0 : A) (l0 : list A), (In a (l0 ++ m) -> In a l0 \/ In a m) -> a0 = a \/ In a (l0 ++ m) -> (a0 = a \/ In a l0) \/ In a m
A:Type
l, m:list A
a, a0:A
y:list A
H:In a (y ++ m) -> In a y \/ In a m
H0:a0 = a \/ In a (y ++ m)

(a0 = a \/ In a y) \/ In a m
A:Type
l, m:list A
a, a0:A
y:list A
H:In a (y ++ m) -> In a y \/ In a m
H0:a0 = a \/ In a (y ++ m)

(a0 = a \/ In a y) \/ In a m
A:Type
l, m:list A
a, a0:A
y:list A
H:In a (y ++ m) -> In a y \/ In a m
H0:a0 = a \/ In a (y ++ m)

In a (y ++ m) -> (a0 = a \/ In a y) \/ In a m
A:Type
l, m:list A
a, a0:A
y:list A
H:In a (y ++ m) -> In a y \/ In a m
H0:a0 = a \/ In a (y ++ m)
H1:In a (y ++ m)

(a0 = a \/ In a y) \/ In a m
A:Type
l, m:list A
a, a0:A
y:list A
H:In a (y ++ m) -> In a y \/ In a m
H0:a0 = a \/ In a (y ++ m)
H1:In a (y ++ m)

(a0 = a \/ In a y) \/ In a m
elim (H H1); auto. Qed.
A:Type

forall (l m : list A) (a : A), In a l \/ In a m -> In a (l ++ m)
A:Type

forall (l m : list A) (a : A), In a l \/ In a m -> In a (l ++ m)
A:Type
l, m:list A
a:A

In a l \/ In a m -> In a (l ++ m)
A:Type
l, m:list A
a:A
H:False \/ In a m

In a m
A:Type
l, m:list A
a, H:A
forall l0 : list A, (In a l0 \/ In a m -> In a (l0 ++ m)) -> (H = a \/ In a l0) \/ In a m -> H = a \/ In a (l0 ++ m)
A:Type
l, m:list A
a:A
H:False \/ In a m

In a m
A:Type
l, m:list A
a, H:A
forall l0 : list A, (In a l0 \/ In a m -> In a (l0 ++ m)) -> (H = a \/ In a l0) \/ In a m -> H = a \/ In a (l0 ++ m)
A:Type
l, m:list A
a:A
H:False \/ In a m
H0:False

In a m
A:Type
l, m:list A
a, H:A
forall l0 : list A, (In a l0 \/ In a m -> In a (l0 ++ m)) -> (H = a \/ In a l0) \/ In a m -> H = a \/ In a (l0 ++ m)
A:Type
l, m:list A
a:A
H:False \/ In a m
H0:False

In a m
A:Type
l, m:list A
a, H:A
forall l0 : list A, (In a l0 \/ In a m -> In a (l0 ++ m)) -> (H = a \/ In a l0) \/ In a m -> H = a \/ In a (l0 ++ m)
A:Type
l, m:list A
a, H:A

forall l0 : list A, (In a l0 \/ In a m -> In a (l0 ++ m)) -> (H = a \/ In a l0) \/ In a m -> H = a \/ In a (l0 ++ m)
A:Type
l, m:list A
a, H:A
y:list A
H0:In a y \/ In a m -> In a (y ++ m)
H1:(H = a \/ In a y) \/ In a m

H = a \/ In a (y ++ m)
A:Type
l, m:list A
a, H:A
y:list A
H0:In a y \/ In a m -> In a (y ++ m)
H1:(H = a \/ In a y) \/ In a m

H = a \/ In a (y ++ m)
A:Type
l, m:list A
a, H:A
y:list A
H0:In a y \/ In a m -> In a (y ++ m)
H1:(H = a \/ In a y) \/ In a m

H = a \/ In a y -> H = a \/ In a (y ++ m)
A:Type
l, m:list A
a, H:A
y:list A
H0:In a y \/ In a m -> In a (y ++ m)
H1:(H = a \/ In a y) \/ In a m
H2:H = a \/ In a y

H = a \/ In a (y ++ m)
A:Type
l, m:list A
a, H:A
y:list A
H0:In a y \/ In a m -> In a (y ++ m)
H1:(H = a \/ In a y) \/ In a m
H2:H = a \/ In a y

H = a \/ In a (y ++ m)
elim H2; auto. Qed.
A:Type

forall (l l' : list A) (a : A), In a (l ++ l') <-> In a l \/ In a l'
A:Type

forall (l l' : list A) (a : A), In a (l ++ l') <-> In a l \/ In a l'
split; auto using in_app_or, in_or_app. Qed.
A:Type

forall l l1 l2 : list A, l ++ l1 = l ++ l2 -> l1 = l2
A:Type

forall l l1 l2 : list A, l ++ l1 = l ++ l2 -> l1 = l2
induction l; simpl; auto; injection 1; auto. Qed.
A:Type

forall l l1 l2 : list A, l1 ++ l = l2 ++ l -> l1 = l2
A:Type

forall l l1 l2 : list A, l1 ++ l = l2 ++ l -> l1 = l2
A:Type

forall l1 l2 l : list A, l1 ++ l = l2 ++ l -> l1 = l2
A:Type
x2:A
l2, l:list A
H:l = x2 :: l2 ++ l

[] = x2 :: l2
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l
x1 :: l1 = []
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l
x1 :: l1 = x2 :: l2
A:Type
x2:A
l2, l:list A
H:l = x2 :: l2 ++ l

~ length (x2 :: l2 ++ l) <= length l
A:Type
x2:A
l2, l:list A
H:l = x2 :: l2 ++ l
length (x2 :: l2 ++ l) <= length l
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l
x1 :: l1 = []
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l
x1 :: l1 = x2 :: l2
A:Type
x2:A
l2, l:list A
H:l = x2 :: l2 ++ l

length (x2 :: l2 ++ l) <= length l
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l
x1 :: l1 = []
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l
x1 :: l1 = x2 :: l2
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l

x1 :: l1 = []
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l
x1 :: l1 = x2 :: l2
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l

~ length (x1 :: l1 ++ l) <= length l
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l
length (x1 :: l1 ++ l) <= length l
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l
x1 :: l1 = x2 :: l2
A:Type
x1:A
l1:list A
IHl1:forall l2 l0 : list A, l1 ++ l0 = l2 ++ l0 -> l1 = l2
l:list A
H:x1 :: l1 ++ l = l

length (x1 :: l1 ++ l) <= length l
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l
x1 :: l1 = x2 :: l2
A:Type
x1:A
l1:list A
IHl1:forall l0 l3 : list A, l1 ++ l3 = l0 ++ l3 -> l1 = l0
x2:A
l2, l:list A
H:x1 :: l1 ++ l = x2 :: l2 ++ l

x1 :: l1 = x2 :: l2
injection H as H H0; f_equal; eauto. Qed. End Facts. Hint Resolve app_assoc app_assoc_reverse: datatypes. Hint Resolve app_comm_cons app_cons_not_nil: datatypes. Hint Immediate app_eq_nil: datatypes. Hint Resolve app_eq_unit app_inj_tail: datatypes. Hint Resolve in_eq in_cons in_inv in_nil in_app_or in_or_app: datatypes. (*******************************************)

Operations on the elements of a list

(*******************************************)

Section Elts.

  Variable A : Type.

  (*****************************)

Nth element of a list

  (*****************************)

  Fixpoint nth (n:nat) (l:list A) (default:A) {struct l} : A :=
    match n, l with
      | O, x :: l' => x
      | O, other => default
      | S m, [] => default
      | S m, x :: t => nth m t default
    end.

  Fixpoint nth_ok (n:nat) (l:list A) (default:A) {struct l} : bool :=
    match n, l with
      | O, x :: l' => true
      | O, other => false
      | S m, [] => false
      | S m, x :: t => nth_ok m t default
    end.

  
A:Type

forall (n : nat) (l : list A) (d : A), {In (nth n l d) l} + {nth n l d = d}
A:Type

forall (n : nat) (l : list A) (d : A), {In (nth n l d) l} + {nth n l d = d}
A:Type
d:A

forall n : nat, {In (nth n [] d) []} + {nth n [] d = d}
A:Type
a:A
l:list A
d:A
IHl:forall n : nat, {In (nth n l d) l} + {nth n l d = d}
forall n : nat, {In (nth n (a :: l) d) (a :: l)} + {nth n (a :: l) d = d}
A:Type
d:A

forall n : nat, {In (nth n [] d) []} + {nth n [] d = d}
right; destruct n; trivial.
A:Type
a:A
l:list A
d:A
IHl:forall n : nat, {In (nth n l d) l} + {nth n l d = d}

forall n : nat, {In (nth n (a :: l) d) (a :: l)} + {nth n (a :: l) d = d}
A:Type
a:A
l:list A
d:A
IHl:forall n : nat, {In (nth n l d) l} + {nth n l d = d}

{a = a \/ In a l} + {a = d}
A:Type
a:A
l:list A
d:A
IHl:forall n0 : nat, {In (nth n0 l d) l} + {nth n0 l d = d}
n:nat
{a = nth n l d \/ In (nth n l d) l} + {nth n l d = d}
A:Type
a:A
l:list A
d:A
IHl:forall n : nat, {In (nth n l d) l} + {nth n l d = d}

{a = a \/ In a l} + {a = d}
left; auto.
A:Type
a:A
l:list A
d:A
IHl:forall n0 : nat, {In (nth n0 l d) l} + {nth n0 l d = d}
n:nat

{a = nth n l d \/ In (nth n l d) l} + {nth n l d = d}
destruct (IHl n); auto. Qed.
A:Type

forall (n : nat) (l : list A) (d a : A), In (nth n l d) l -> In (nth (S n) (a :: l) d) (a :: l)
A:Type

forall (n : nat) (l : list A) (d a : A), In (nth n l d) l -> In (nth (S n) (a :: l) d) (a :: l)
simpl; auto. Qed. Fixpoint nth_error (l:list A) (n:nat) {struct n} : option A := match n, l with | O, x :: _ => Some x | S n, _ :: l => nth_error l n | _, _ => None end. Definition nth_default (default:A) (l:list A) (n:nat) : A := match nth_error l n with | Some x => x | None => default end.
A:Type

forall (n : nat) (l : list A) (d : A), nth_default d l n = nth n l d
A:Type

forall (n : nat) (l : list A) (d : A), nth_default d l n = nth n l d
unfold nth_default; induction n; intros [ | ] ?; simpl; auto. Qed.
Results about nth
  
A:Type

forall (n : nat) (l : list A) (d : A), n < length l -> In (nth n l d) l
A:Type

forall (n : nat) (l : list A) (d : A), n < length l -> In (nth n l d) l
A:Type

forall (l : list A) (d : A), 1 <= length l -> In (nth 0 l d) l
A:Type
n:nat
hn:forall (l : list A) (d : A), S n <= length l -> In (nth n l d) l
forall (l : list A) (d : A), S (S n) <= length l -> In (nth (S n) l d) l
A:Type

forall (l : list A) (d : A), 1 <= length l -> In (nth 0 l d) l
destruct l; simpl; [ inversion 2 | auto ].
A:Type
n:nat
hn:forall (l : list A) (d : A), S n <= length l -> In (nth n l d) l

forall (l : list A) (d : A), S (S n) <= length l -> In (nth (S n) l d) l
A:Type
n:nat
hn:forall (l : list A) (d : A), S n <= length l -> In (nth n l d) l

A -> S (S n) <= 0 -> False
A:Type
n:nat
hn:forall (l0 : list A) (d : A), S n <= length l0 -> In (nth n l0 d) l0
a:A
l:list A
forall d : A, S (S n) <= S (length l) -> a = nth n l d \/ In (nth n l d) l
A:Type
n:nat
hn:forall (l : list A) (d : A), S n <= length l -> In (nth n l d) l

A -> S (S n) <= 0 -> False
inversion 2.
A:Type
n:nat
hn:forall (l0 : list A) (d : A), S n <= length l0 -> In (nth n l0 d) l0
a:A
l:list A

forall d : A, S (S n) <= S (length l) -> a = nth n l d \/ In (nth n l d) l
intros d ie; right; apply hn; auto with arith. Qed.
A:Type
l:list A
x, d:A

In x l -> exists n : nat, n < length l /\ nth n l d = x
A:Type
l:list A
x, d:A

In x l -> exists n : nat, n < length l /\ nth n l d = x
A:Type
x, d:A

In x [] -> exists n : nat, n < length [] /\ nth n [] d = x
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n : nat, n < length l /\ nth n l d = x
In x (a :: l) -> exists n : nat, n < length (a :: l) /\ nth n (a :: l) d = x
A:Type
x, d:A

In x [] -> exists n : nat, n < length [] /\ nth n [] d = x
easy.
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n : nat, n < length l /\ nth n l d = x

In x (a :: l) -> exists n : nat, n < length (a :: l) /\ nth n (a :: l) d = x
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n : nat, n < length l /\ nth n l d = x
H:a = x

exists n : nat, n < length (a :: l) /\ nth n (a :: l) d = x
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n : nat, n < length l /\ nth n l d = x
H:In x l
exists n : nat, n < length (a :: l) /\ nth n (a :: l) d = x
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n : nat, n < length l /\ nth n l d = x
H:a = x

exists n : nat, n < length (a :: l) /\ nth n (a :: l) d = x
subst; exists 0; simpl; auto with arith.
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n : nat, n < length l /\ nth n l d = x
H:In x l

exists n : nat, n < length (a :: l) /\ nth n (a :: l) d = x
A:Type
a:A
l:list A
x, d:A
IH:In x l -> exists n0 : nat, n0 < length l /\ nth n0 l d = x
H:In x l
n:nat
Hn:n < length l
Hn':nth n l d = x

exists n0 : nat, n0 < length (a :: l) /\ nth n0 (a :: l) d = x
exists (S n); simpl; auto with arith. Qed.
A:Type

forall (l : list A) (n : nat) (d : A), length l <= n -> nth n l d = d
A:Type

forall (l : list A) (n : nat) (d : A), length l <= n -> nth n l d = d
A:Type
a:A
l:list A
IHl:forall (n : nat) (d0 : A), length l <= n -> nth n l d0 = d0
d:A
H:S (length l) <= 0

a = d
A:Type
a:A
l:list A
IHl:forall (n0 : nat) (d0 : A), length l <= n0 -> nth n0 l d0 = d0
n:nat
d:A
H:S (length l) <= S n
nth n l d = d
A:Type
a:A
l:list A
IHl:forall (n : nat) (d0 : A), length l <= n -> nth n l d0 = d0
d:A
H:S (length l) <= 0

a = d
inversion H.
A:Type
a:A
l:list A
IHl:forall (n0 : nat) (d0 : A), length l <= n0 -> nth n0 l d0 = d0
n:nat
d:A
H:S (length l) <= S n

nth n l d = d
apply IHl; auto with arith. Qed.
A:Type

forall (l : list A) (n : nat) (d d' : A), n < length l -> nth n l d = nth n l d'
A:Type

forall (l : list A) (n : nat) (d d' : A), n < length l -> nth n l d = nth n l d'
A:Type

forall (n : nat) (d d' : A), n < length [] -> nth n [] d = nth n [] d'
A:Type
a:A
l:list A
IHl:forall (n : nat) (d d' : A), n < length l -> nth n l d = nth n l d'
forall (n : nat) (d d' : A), n < length (a :: l) -> nth n (a :: l) d = nth n (a :: l) d'
A:Type

forall (n : nat) (d d' : A), n < length [] -> nth n [] d = nth n [] d'
inversion 1.
A:Type
a:A
l:list A
IHl:forall (n : nat) (d d' : A), n < length l -> nth n l d = nth n l d'

forall (n : nat) (d d' : A), n < length (a :: l) -> nth n (a :: l) d = nth n (a :: l) d'
intros [|n] d d'; simpl; auto with arith. Qed.
A:Type

forall (l l' : list A) (d : A) (n : nat), n < length l -> nth n (l ++ l') d = nth n l d
A:Type

forall (l l' : list A) (d : A) (n : nat), n < length l -> nth n (l ++ l') d = nth n l d
A:Type

forall (l' : list A) (d : A) (n : nat), n < length [] -> nth n ([] ++ l') d = nth n [] d
A:Type
a:A
l:list A
IHl:forall (l' : list A) (d : A) (n : nat), n < length l -> nth n (l ++ l') d = nth n l d
forall (l' : list A) (d : A) (n : nat), n < length (a :: l) -> nth n ((a :: l) ++ l') d = nth n (a :: l) d
A:Type

forall (l' : list A) (d : A) (n : nat), n < length [] -> nth n ([] ++ l') d = nth n [] d
inversion 1.
A:Type
a:A
l:list A
IHl:forall (l' : list A) (d : A) (n : nat), n < length l -> nth n (l ++ l') d = nth n l d

forall (l' : list A) (d : A) (n : nat), n < length (a :: l) -> nth n ((a :: l) ++ l') d = nth n (a :: l) d
intros l' d [|n]; simpl; auto with arith. Qed.
A:Type

forall (l l' : list A) (d : A) (n : nat), n >= length l -> nth n (l ++ l') d = nth (n - length l) l' d
A:Type

forall (l l' : list A) (d : A) (n : nat), n >= length l -> nth n (l ++ l') d = nth (n - length l) l' d
A:Type
a:A
l:list A
IHl:forall (l'0 : list A) (d0 : A) (n : nat), n >= length l -> nth n (l ++ l'0) d0 = nth (n - length l) l'0 d0
l':list A
d:A

0 >= length (a :: l) -> nth 0 ((a :: l) ++ l') d = nth (0 - length (a :: l)) l' d
A:Type
a:A
l:list A
IHl:forall (l'0 : list A) (d0 : A) (n0 : nat), n0 >= length l -> nth n0 (l ++ l'0) d0 = nth (n0 - length l) l'0 d0
l':list A
d:A
n:nat
S n >= length (a :: l) -> nth (S n) ((a :: l) ++ l') d = nth (S n - length (a :: l)) l' d
A:Type
a:A
l:list A
IHl:forall (l'0 : list A) (d0 : A) (n : nat), n >= length l -> nth n (l ++ l'0) d0 = nth (n - length l) l'0 d0
l':list A
d:A

0 >= length (a :: l) -> nth 0 ((a :: l) ++ l') d = nth (0 - length (a :: l)) l' d
inversion 1.
A:Type
a:A
l:list A
IHl:forall (l'0 : list A) (d0 : A) (n0 : nat), n0 >= length l -> nth n0 (l ++ l'0) d0 = nth (n0 - length l) l'0 d0
l':list A
d:A
n:nat

S n >= length (a :: l) -> nth (S n) ((a :: l) ++ l') d = nth (S n - length (a :: l)) l' d
intros; simpl; rewrite IHl; auto with arith. Qed.
A:Type
n:nat
l:list A
d:A

n < length l -> exists l1 l2 : list A, l = l1 ++ nth n l d :: l2 /\ length l1 = n
A:Type
n:nat
l:list A
d:A

n < length l -> exists l1 l2 : list A, l = l1 ++ nth n l d :: l2 /\ length l1 = n
A:Type
n:nat
d:A

forall l : list A, n < length l -> exists l1 l2 : list A, l = l1 ++ nth n l d :: l2 /\ length l1 = n
A:Type
d, a:A
l:list A
H:0 < length (a :: l)

exists l1 l2 : list A, a :: l = l1 ++ nth 0 (a :: l) d :: l2 /\ length l1 = 0
A:Type
n:nat
d:A
IH:forall l0 : list A, n < length l0 -> exists l1 l2 : list A, l0 = l1 ++ nth n l0 d :: l2 /\ length l1 = n
a:A
l:list A
H:S n < length (a :: l)
exists l1 l2 : list A, a :: l = l1 ++ nth (S n) (a :: l) d :: l2 /\ length l1 = S n
A:Type
d, a:A
l:list A
H:0 < length (a :: l)

exists l1 l2 : list A, a :: l = l1 ++ nth 0 (a :: l) d :: l2 /\ length l1 = 0
exists nil; exists l; now simpl.
A:Type
n:nat
d:A
IH:forall l0 : list A, n < length l0 -> exists l1 l2 : list A, l0 = l1 ++ nth n l0 d :: l2 /\ length l1 = n
a:A
l:list A
H:S n < length (a :: l)

exists l1 l2 : list A, a :: l = l1 ++ nth (S n) (a :: l) d :: l2 /\ length l1 = S n
A:Type
n:nat
d:A
IH:forall l0 : list A, n < length l0 -> exists l3 l4 : list A, l0 = l3 ++ nth n l0 d :: l4 /\ length l3 = n
a:A
l:list A
H:S n < length (a :: l)
l1, l2:list A
Hl:l = l1 ++ nth n l d :: l2
Hl1:length l1 = n

exists l0 l3 : list A, a :: l = l0 ++ nth (S n) (a :: l) d :: l3 /\ length l0 = S n
exists (a::l1); exists l2; simpl; split; now f_equal. Qed.
Results about nth_error
  
A:Type
l:list A
n:nat
x:A

nth_error l n = Some x -> In x l
A:Type
l:list A
n:nat
x:A

nth_error l n = Some x -> In x l
A:Type
l:list A
x:A

forall n : nat, nth_error l n = Some x -> In x l
A:Type
a:A
l:list A
x:A
IH:forall n : nat, nth_error l n = Some x -> In x l

Some a = Some x -> a = x \/ In x l
A:Type
a:A
l:list A
x:A
IH:forall n0 : nat, nth_error l n0 = Some x -> In x l
n:nat
nth_error l n = Some x -> a = x \/ In x l
A:Type
a:A
l:list A
x:A
IH:forall n : nat, nth_error l n = Some x -> In x l

Some a = Some x -> a = x \/ In x l
injection 1; auto.
A:Type
a:A
l:list A
x:A
IH:forall n0 : nat, nth_error l n0 = Some x -> In x l
n:nat

nth_error l n = Some x -> a = x \/ In x l
eauto. Qed.
A:Type
l:list A
x:A

In x l -> exists n : nat, nth_error l n = Some x
A:Type
l:list A
x:A

In x l -> exists n : nat, nth_error l n = Some x
A:Type
x:A

In x [] -> exists n : nat, nth_error [] n = Some x
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n : nat, nth_error l n = Some x
In x (a :: l) -> exists n : nat, nth_error (a :: l) n = Some x
A:Type
x:A

In x [] -> exists n : nat, nth_error [] n = Some x
easy.
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n : nat, nth_error l n = Some x

In x (a :: l) -> exists n : nat, nth_error (a :: l) n = Some x
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n : nat, nth_error l n = Some x
H:a = x

exists n : nat, nth_error (a :: l) n = Some x
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n : nat, nth_error l n = Some x
H:In x l
exists n : nat, nth_error (a :: l) n = Some x
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n : nat, nth_error l n = Some x
H:a = x

exists n : nat, nth_error (a :: l) n = Some x
subst; exists 0; simpl; auto with arith.
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n : nat, nth_error l n = Some x
H:In x l

exists n : nat, nth_error (a :: l) n = Some x
A:Type
a:A
l:list A
x:A
IH:In x l -> exists n0 : nat, nth_error l n0 = Some x
H:In x l
n:nat
Hn:nth_error l n = Some x

exists n0 : nat, nth_error (a :: l) n0 = Some x
exists (S n); simpl; auto with arith. Qed.
A:Type
l:list A
n:nat

nth_error l n = None <-> length l <= n
A:Type
l:list A
n:nat

nth_error l n = None <-> length l <= n
A:Type
l:list A

forall n : nat, nth_error l n = None <-> length l <= n
A:Type

None = None <-> 0 <= 0
A:Type
n:nat
None = None <-> 0 <= S n
A:Type
a:A
l:list A
IHl:forall n : nat, nth_error l n = None <-> length l <= n
Some a = None <-> S (length l) <= 0
A:Type
a:A
l:list A
IHl:forall n0 : nat, nth_error l n0 = None <-> length l <= n0
n:nat
nth_error l n = None <-> S (length l) <= S n
A:Type

None = None <-> 0 <= 0
split; auto.
A:Type
n:nat

None = None <-> 0 <= S n
split; auto with arith.
A:Type
a:A
l:list A
IHl:forall n : nat, nth_error l n = None <-> length l <= n

Some a = None <-> S (length l) <= 0
split; now auto with arith.
A:Type
a:A
l:list A
IHl:forall n0 : nat, nth_error l n0 = None <-> length l <= n0
n:nat

nth_error l n = None <-> S (length l) <= S n
rewrite IHl; split; auto with arith. Qed.
A:Type
l:list A
n:nat

nth_error l n <> None <-> n < length l
A:Type
l:list A
n:nat

nth_error l n <> None <-> n < length l
A:Type
l:list A

forall n : nat, nth_error l n <> None <-> n < length l
A:Type

None <> None <-> 0 < 0
A:Type
n:nat
None <> None <-> S n < 0
A:Type
a:A
l:list A
IHl:forall n : nat, nth_error l n <> None <-> n < length l
Some a <> None <-> 0 < S (length l)
A:Type
a:A
l:list A
IHl:forall n0 : nat, nth_error l n0 <> None <-> n0 < length l
n:nat
nth_error l n <> None <-> S n < S (length l)
A:Type

None <> None <-> 0 < 0
split; [now destruct 1 | inversion 1].
A:Type
n:nat

None <> None <-> S n < 0
split; [now destruct 1 | inversion 1].
A:Type
a:A
l:list A
IHl:forall n : nat, nth_error l n <> None <-> n < length l

Some a <> None <-> 0 < S (length l)
split; now auto with arith.
A:Type
a:A
l:list A
IHl:forall n0 : nat, nth_error l n0 <> None <-> n0 < length l
n:nat

nth_error l n <> None <-> S n < S (length l)
rewrite IHl; split; auto with arith. Qed.
A:Type
l:list A
n:nat
a:A

nth_error l n = Some a -> exists l1 l2 : list A, l = l1 ++ a :: l2 /\ length l1 = n
A:Type
l:list A
n:nat
a:A

nth_error l n = Some a -> exists l1 l2 : list A, l = l1 ++ a :: l2 /\ length l1 = n
A:Type
n:nat
a:A

forall l : list A, nth_error l n = Some a -> exists l1 l2 : list A, l = l1 ++ a :: l2 /\ length l1 = n
A:Type
a, x:A
l:list A
H:Some x = Some a

exists l1 l2 : list A, x :: l = l1 ++ a :: l2 /\ length l1 = 0
A:Type
n:nat
a:A
IH:forall l0 : list A, nth_error l0 n = Some a -> exists l1 l2 : list A, l0 = l1 ++ a :: l2 /\ length l1 = n
x:A
l:list A
H:nth_error l n = Some a
exists l1 l2 : list A, x :: l = l1 ++ a :: l2 /\ length l1 = S n
A:Type
a, x:A
l:list A
H:Some x = Some a

exists l1 l2 : list A, x :: l = l1 ++ a :: l2 /\ length l1 = 0
A:Type
a, x:A
l:list A
H:Some x = Some a

x :: l = [] ++ a :: l /\ length [] = 0
now injection H as ->.
A:Type
n:nat
a:A
IH:forall l0 : list A, nth_error l0 n = Some a -> exists l1 l2 : list A, l0 = l1 ++ a :: l2 /\ length l1 = n
x:A
l:list A
H:nth_error l n = Some a

exists l1 l2 : list A, x :: l = l1 ++ a :: l2 /\ length l1 = S n
A:Type
n:nat
a:A
IH:forall l0 : list A, nth_error l0 n = Some a -> exists l3 l4 : list A, l0 = l3 ++ a :: l4 /\ length l3 = n
x:A
l:list A
H:nth_error l n = Some a
l1, l2:list A
H1:l = l1 ++ a :: l2
H2:length l1 = n

exists l0 l3 : list A, x :: l = l0 ++ a :: l3 /\ length l0 = S n
exists (x::l1); exists l2; simpl; split; now f_equal. Qed.
A:Type
l, l':list A
n:nat

n < length l -> nth_error (l ++ l') n = nth_error l n
A:Type
l, l':list A
n:nat

n < length l -> nth_error (l ++ l') n = nth_error l n
A:Type
l':list A
n:nat

forall l : list A, n < length l -> nth_error (l ++ l') n = nth_error l n
A:Type
l':list A
n:nat
IHn:forall l0 : list A, n < length l0 -> nth_error (l0 ++ l') n = nth_error l0 n
a:A
l:list A
H:S n < length (a :: l)

nth_error ((a :: l) ++ l') (S n) = nth_error (a :: l) (S n)
A:Type
l':list A
n:nat
IHn:forall l0 : list A, n < length l0 -> nth_error (l0 ++ l') n = nth_error l0 n
a:A
l:list A
H:S n < S (length l)

nth_error (l ++ l') n = nth_error l n
A:Type
l':list A
n:nat
IHn:forall l0 : list A, n < length l0 -> nth_error (l0 ++ l') n = nth_error l0 n
a:A
l:list A
H:S n < S (length l)

n < length l
auto with arith. Qed.
A:Type
l, l':list A
n:nat

length l <= n -> nth_error (l ++ l') n = nth_error l' (n - length l)
A:Type
l, l':list A
n:nat

length l <= n -> nth_error (l ++ l') n = nth_error l' (n - length l)
A:Type
l':list A
n:nat

forall l : list A, length l <= n -> nth_error (l ++ l') n = nth_error l' (n - length l)
A:Type
l':list A
n:nat
IHn:forall l0 : list A, length l0 <= n -> nth_error (l0 ++ l') n = nth_error l' (n - length l0)
a:A
l:list A
H:length (a :: l) <= S n

nth_error ((a :: l) ++ l') (S n) = nth_error l' (S n - length (a :: l))
A:Type
l':list A
n:nat
IHn:forall l0 : list A, length l0 <= n -> nth_error (l0 ++ l') n = nth_error l' (n - length l0)
a:A
l:list A
H:S (length l) <= S n

nth_error (l ++ l') n = nth_error l' (n - length l)
A:Type
l':list A
n:nat
IHn:forall l0 : list A, length l0 <= n -> nth_error (l0 ++ l') n = nth_error l' (n - length l0)
a:A
l:list A
H:S (length l) <= S n

length l <= n
auto with arith. Qed. (*****************)

Remove

  (*****************)

  Hypothesis eq_dec : forall x y : A, {x = y}+{x <> y}.

  Fixpoint remove (x : A) (l : list A) : list A :=
    match l with
      | [] => []
      | y::tl => if (eq_dec x y) then remove x tl else y::(remove x tl)
    end.

  
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall (l : list A) (x : A), ~ In x (remove x l)
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall (l : list A) (x : A), ~ In x (remove x l)
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A
l:list A
IHl:forall x0 : A, ~ In x0 (remove x0 l)

forall x0 : A, ~ In x0 (remove x0 (x :: l))
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
x:A
l:list A
IHl:forall x0 : A, ~ In x0 (remove x0 l)
y:A
yeqx:y = x

~ In y (remove y l)
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
x:A
l:list A
IHl:forall x0 : A, ~ In x0 (remove x0 l)
y:A
yneqx:y <> x
~ In y (x :: remove y l)
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
x:A
l:list A
IHl:forall x0 : A, ~ In x0 (remove x0 l)
y:A
yneqx:y <> x

~ In y (x :: remove y l)
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
x:A
l:list A
IHl:forall x0 : A, ~ In x0 (remove x0 l)
y:A
yneqx:y <> x
H:In y (remove y l)

False
apply (IHl y); assumption. Qed. (******************************)

Last element of a list

(******************************)
last l d returns the last element of the list l, or the default value d if l is empty.
  Fixpoint last (l:list A) (d:A) : A :=
  match l with
    | [] => d
    | [a] => a
    | a :: l => last l d
  end.
removelast l remove the last element of l
  Fixpoint removelast (l:list A) : list A :=
    match l with
      | [] =>  []
      | [a] => []
      | a :: l => a :: removelast l
    end.

  
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall (l : list A) (d : A), l <> [] -> l = removelast l ++ [last l d]
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall (l : list A) (d : A), l <> [] -> l = removelast l ++ [last l d]
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall d : A, [] <> [] -> [] = removelast [] ++ [last [] d]
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall d : A, l <> [] -> l = removelast l ++ [last l d]
forall d : A, a :: l <> [] -> a :: l = removelast (a :: l) ++ [last (a :: l) d]
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall d : A, l <> [] -> l = removelast l ++ [last l d]

forall d : A, a :: l <> [] -> a :: l = removelast (a :: l) ++ [last (a :: l) d]
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall d0 : A, l <> [] -> l = removelast l ++ [last l d0]
d:A

a :: l = removelast (a :: l) ++ [last (a :: l) d]
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:forall d0 : A, a0 :: l <> [] -> a0 :: l = removelast (a0 :: l) ++ [last (a0 :: l) d0]
d:A

a :: a0 :: l = removelast (a :: a0 :: l) ++ [last (a :: a0 :: l) d]
pattern (a0::l) at 1; rewrite IHl with d; auto; discriminate. Qed.
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l : list A, l <> [] -> {l' : list A & {a : A | l = l' ++ [a]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l : list A, l <> [] -> {l' : list A & {a : A | l = l' ++ [a]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

[] <> [] -> {l' : list A & {a : A | [] = l' ++ [a]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:l <> [] -> {l' : list A & {a0 : A | l = l' ++ [a0]}}
a :: l <> [] -> {l' : list A & {a0 : A | a :: l = l' ++ [a0]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:l <> [] -> {l' : list A & {a0 : A | l = l' ++ [a0]}}

a :: l <> [] -> {l' : list A & {a0 : A | a :: l = l' ++ [a0]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:l <> [] -> {l' : list A & {a0 : A | l = l' ++ [a0]}}

{l' : list A & {a0 : A | a :: l = l' ++ [a0]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
IHl:[] <> [] -> {l' : list A & {a0 : A | [] = l' ++ [a0]}}

{l' : list A & {a0 : A | [a] = l' ++ [a0]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:a0 :: l <> [] -> {l' : list A & {a1 : A | a0 :: l = l' ++ [a1]}}
{l' : list A & {a1 : A | a :: a0 :: l = l' ++ [a1]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:a0 :: l <> [] -> {l' : list A & {a1 : A | a0 :: l = l' ++ [a1]}}

{l' : list A & {a1 : A | a :: a0 :: l = l' ++ [a1]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l, l':list A
a':A
H:a0 :: l = l' ++ [a']

{l'0 : list A & {a1 : A | a :: a0 :: l = l'0 ++ [a1]}}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l, l':list A
a':A
H:a0 :: l = l' ++ [a']

{l'0 : list A & {a1 : A | a :: l' ++ [a'] = l'0 ++ [a1]}}
exists (a::l'), a'; auto. Qed.
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l l' : list A, l' <> [] -> removelast (l ++ l') = l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l l' : list A, l' <> [] -> removelast (l ++ l') = l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l' : list A, l' <> [] -> removelast ([] ++ l') = [] ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l' : list A, l' <> [] -> removelast (l ++ l') = l ++ removelast l'
forall l' : list A, l' <> [] -> removelast ((a :: l) ++ l') = (a :: l) ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l' : list A, l' <> [] -> removelast (l ++ l') = l ++ removelast l'

forall l' : list A, l' <> [] -> removelast ((a :: l) ++ l') = (a :: l) ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast (l ++ l'0) = l ++ removelast l'0
l':list A
H:l' <> []

match l ++ l' with | [] => [] | _ :: _ => a :: removelast (l ++ l') end = a :: l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast (l ++ l'0) = l ++ removelast l'0
l':list A
H:l' <> []

l ++ l' <> []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast (l ++ l'0) = l ++ removelast l'0
l':list A
H:l' <> []
H0:l ++ l' <> []
match l ++ l' with | [] => [] | _ :: _ => a :: removelast (l ++ l') end = a :: l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
IHl:forall l'0 : list A, l'0 <> [] -> removelast ([] ++ l'0) = [] ++ removelast l'0
l':list A
H:l' <> []

[] ++ l' <> []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast ((a0 :: l) ++ l'0) = (a0 :: l) ++ removelast l'0
l':list A
H:l' <> []
(a0 :: l) ++ l' <> []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast (l ++ l'0) = l ++ removelast l'0
l':list A
H:l' <> []
H0:l ++ l' <> []
match l ++ l' with | [] => [] | _ :: _ => a :: removelast (l ++ l') end = a :: l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a, a0:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast ((a0 :: l) ++ l'0) = (a0 :: l) ++ removelast l'0
l':list A
H:l' <> []

(a0 :: l) ++ l' <> []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast (l ++ l'0) = l ++ removelast l'0
l':list A
H:l' <> []
H0:l ++ l' <> []
match l ++ l' with | [] => [] | _ :: _ => a :: removelast (l ++ l') end = a :: l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:forall l'0 : list A, l'0 <> [] -> removelast (l ++ l'0) = l ++ removelast l'0
l':list A
H:l' <> []
H0:l ++ l' <> []

match l ++ l' with | [] => [] | _ :: _ => a :: removelast (l ++ l') end = a :: l ++ removelast l'
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
a:A
l, l':list A
IHl:removelast (l ++ l') = l ++ removelast l'
H:l' <> []
H0:l ++ l' <> []

match l ++ l' with | [] => [] | _ :: _ => a :: removelast (l ++ l') end = a :: l ++ removelast l'
destruct (l++l'); [elim H0; auto|f_equal; auto]. Qed. (******************************************)

Counting occurrences of an element

  (******************************************)

  Fixpoint count_occ (l : list A) (x : A) : nat :=
    match l with
      | [] => 0
      | y :: tl =>
        let n := count_occ tl x in
        if eq_dec y x then S n else n
    end.
Compatibility of count_occ with operations on list
  
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

In x l <-> count_occ l x > 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

In x l <-> count_occ l x > 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A

False <-> 0 > 0
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
y:A
l:list A
x:A
IHl:In x l <-> count_occ l x > 0
y = x \/ In x l <-> (if eq_dec y x then S (count_occ l x) else count_occ l x) > 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A

False <-> 0 > 0
split; [destruct 1 | apply gt_irrefl].
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
y:A
l:list A
x:A
IHl:In x l <-> count_occ l x > 0

y = x \/ In x l <-> (if eq_dec y x then S (count_occ l x) else count_occ l x) > 0
destruct eq_dec as [->|Hneq]; rewrite IHl; intuition. Qed.
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

~ In x l <-> count_occ l x = 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

~ In x l <-> count_occ l x = 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

~ count_occ l x > 0 <-> count_occ l x = 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

~ 0 < count_occ l x <-> count_occ l x = 0
now rewrite Nat.nlt_ge, Nat.le_0_r. Qed.
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A

count_occ [] x = 0
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A

count_occ [] x = 0
reflexivity. Qed.
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ l x = 0) <-> l = []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ l x = 0) <-> l = []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ l x = 0) -> l = []
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
l:list A
l = [] -> forall x : A, count_occ l x = 0
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ l x = 0) -> l = []
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A
l:list A
IHl:(forall x0 : A, count_occ l x0 = 0) -> l = []

(forall x0 : A, count_occ (x :: l) x0 = 0) -> x :: l = []
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A
l:list A
IHl:(forall x0 : A, count_occ l x0 = 0) -> l = []
H:forall x0 : A, count_occ (x :: l) x0 = 0

x :: l = []
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A
l:list A
IHl:(forall x0 : A, count_occ l x0 = 0) -> l = []
H:count_occ (x :: l) x = 0

x :: l = []
A:Type
eq_dec:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A
l:list A
IHl:(forall x0 : A, count_occ l x0 = 0) -> l = []
H:(if eq_dec x x then S (count_occ l x) else count_occ l x) = 0

x :: l = []
destruct eq_dec as [_|NEQ]; [discriminate|now elim NEQ].
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}
l:list A

l = [] -> forall x : A, count_occ l x = 0
now intros ->. Qed.
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A

x = y -> count_occ (x :: l) y = S (count_occ l y)
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A

x = y -> count_occ (x :: l) y = S (count_occ l y)
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A
H:x = y

count_occ (x :: l) y = S (count_occ l y)
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A
H:x = y

(if eq_dec x y then S (count_occ l y) else count_occ l y) = S (count_occ l y)
now destruct (eq_dec x y). Qed.
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A

x <> y -> count_occ (x :: l) y = count_occ l y
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A

x <> y -> count_occ (x :: l) y = count_occ l y
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A
H:x <> y

count_occ (x :: l) y = count_occ l y
A:Type
eq_dec:forall x0 y0 : A, {x0 = y0} + {x0 <> y0}
l:list A
x, y:A
H:x <> y

(if eq_dec x y then S (count_occ l y) else count_occ l y) = count_occ l y
now destruct (eq_dec x y). Qed. End Elts. (*******************************)

Manipulating whole lists

(*******************************)

Section ListOps.

  Variable A : Type.

  (*************************)

Reverse

  (*************************)

  Fixpoint rev (l:list A) : list A :=
    match l with
      | [] => []
      | x :: l' => rev l' ++ [x]
    end.

  
A:Type

forall x y : list A, rev (x ++ y) = rev y ++ rev x
A:Type

forall x y : list A, rev (x ++ y) = rev y ++ rev x
A:Type

forall y : list A, rev ([] ++ y) = rev y ++ rev []
A:Type
a:A
l:list A
IHl:forall y : list A, rev (l ++ y) = rev y ++ rev l
forall y : list A, rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type

rev ([] ++ []) = rev [] ++ rev []
A:Type
a:A
l:list A
rev ([] ++ a :: l) = rev (a :: l) ++ rev []
A:Type
a:A
l:list A
IHl:forall y : list A, rev (l ++ y) = rev y ++ rev l
forall y : list A, rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type

[] = []
A:Type
a:A
l:list A
rev ([] ++ a :: l) = rev (a :: l) ++ rev []
A:Type
a:A
l:list A
IHl:forall y : list A, rev (l ++ y) = rev y ++ rev l
forall y : list A, rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type
a:A
l:list A

rev ([] ++ a :: l) = rev (a :: l) ++ rev []
A:Type
a:A
l:list A
IHl:forall y : list A, rev (l ++ y) = rev y ++ rev l
forall y : list A, rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type
a:A
l:list A

rev l ++ [a] = (rev l ++ [a]) ++ []
A:Type
a:A
l:list A
IHl:forall y : list A, rev (l ++ y) = rev y ++ rev l
forall y : list A, rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type
a:A
l:list A
IHl:forall y : list A, rev (l ++ y) = rev y ++ rev l

forall y : list A, rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type
a:A
l:list A
IHl:forall y0 : list A, rev (l ++ y0) = rev y0 ++ rev l
y:list A

rev ((a :: l) ++ y) = rev y ++ rev (a :: l)
A:Type
a:A
l:list A
IHl:forall y0 : list A, rev (l ++ y0) = rev y0 ++ rev l
y:list A

rev (l ++ y) ++ [a] = rev y ++ rev l ++ [a]
A:Type
a:A
l:list A
IHl:forall y0 : list A, rev (l ++ y0) = rev y0 ++ rev l
y:list A

(rev y ++ rev l) ++ [a] = rev y ++ rev l ++ [a]
rewrite app_assoc; trivial. Qed.
A:Type

forall (l : list A) (a : A), rev (l ++ [a]) = a :: rev l
A:Type

forall (l : list A) (a : A), rev (l ++ [a]) = a :: rev l
A:Type
l:list A
a:A

rev (l ++ [a]) = a :: rev l
apply (rev_app_distr l [a]); simpl; auto. Qed.
A:Type

forall l : list A, rev (rev l) = l
A:Type

forall l : list A, rev (rev l) = l
A:Type

rev (rev []) = []
A:Type
a:A
l:list A
IHl:rev (rev l) = l
rev (rev (a :: l)) = a :: l
A:Type
a:A
l:list A
IHl:rev (rev l) = l

rev (rev (a :: l)) = a :: l
A:Type
a:A
l:list A
IHl:rev (rev l) = l

rev (rev l ++ [a]) = a :: l
A:Type
a:A
l:list A
IHl:rev (rev l) = l

a :: rev (rev l) = a :: l
rewrite IHl; auto. Qed.
Compatibility with other operations
  
A:Type

forall (l : list A) (x : A), In x l <-> In x (rev l)
A:Type

forall (l : list A) (x : A), In x l <-> In x (rev l)
A:Type

forall x : A, In x [] <-> In x (rev [])
A:Type
a:A
l:list A
IHl:forall x : A, In x l <-> In x (rev l)
forall x : A, In x (a :: l) <-> In x (rev (a :: l))
A:Type
a:A
l:list A
IHl:forall x : A, In x l <-> In x (rev l)

forall x : A, In x (a :: l) <-> In x (rev (a :: l))
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A

In x (a :: l) <-> In x (rev (a :: l))
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A

a = x \/ In x l <-> In x (rev l ++ [a])
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H0:a = x

In x (rev l ++ [a])
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H0:In x l
In x (rev l ++ [a])
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H:In x (rev l ++ [a])
a = x \/ In x l
A:Type
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A

In x (rev l ++ [x])
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H0:In x l
In x (rev l ++ [a])
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H:In x (rev l ++ [a])
a = x \/ In x l
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H0:In x l

In x (rev l ++ [a])
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H:In x (rev l ++ [a])
a = x \/ In x l
A:Type
a:A
l:list A
IHl:forall x0 : A, In x0 l <-> In x0 (rev l)
x:A
H:In x (rev l ++ [a])

a = x \/ In x l
destruct (in_app_or _ _ _ H); firstorder. Qed.
A:Type

forall l : list A, length (rev l) = length l
A:Type

forall l : list A, length (rev l) = length l
A:Type
a:A
l:list A
IHl:length (rev l) = length l

length (rev l ++ [a]) = S (length l)
A:Type
a:A
l:list A
IHl:length (rev l) = length l

length (rev l) + length [a] = S (length l)
A:Type
a:A
l:list A
IHl:length (rev l) = length l

length l + length [a] = S (length l)
A:Type
a:A
l:list A
IHl:length (rev l) = length l

length l + 1 = S (length l)
elim (length l); simpl; auto. Qed.
A:Type

forall (l : list A) (d : A) (n : nat), n < length l -> nth n (rev l) d = nth (length l - S n) l d
A:Type

forall (l : list A) (d : A) (n : nat), n < length l -> nth n (rev l) d = nth (length l - S n) l d
A:Type

forall (d : A) (n : nat), n < length [] -> nth n (rev []) d = nth (length [] - S n) [] d
A:Type
a:A
l:list A
IHl:forall (d : A) (n : nat), n < length l -> nth n (rev l) d = nth (length l - S n) l d
forall (d : A) (n : nat), n < length (a :: l) -> nth n (rev (a :: l)) d = nth (length (a :: l) - S n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d : A) (n : nat), n < length l -> nth n (rev l) d = nth (length l - S n) l d

forall (d : A) (n : nat), n < length (a :: l) -> nth n (rev (a :: l)) d = nth (length (a :: l) - S n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < length (a :: l)

nth n (rev (a :: l)) d = nth (length (a :: l) - S n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)

nth n (rev (a :: l)) d = nth (length (a :: l) - S n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)

nth n (rev l ++ [a]) d = nth (length (a :: l) - S n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)

nth n (rev l ++ [a]) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
H1:n = length l

nth (length l) (rev l ++ [a]) d = nth (length l - length l) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
nth n (rev l ++ [a]) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
H1:n = length l

nth (length l) (rev l ++ [a]) d = a
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
nth n (rev l ++ [a]) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
H1:n = length l

nth (length (rev l)) (rev l ++ [a]) d = a
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
nth n (rev l ++ [a]) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
H1:n = length l

nth (length (rev l) - length (rev l)) [a] d = a
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
nth n (rev l ++ [a]) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l

nth n (rev l ++ [a]) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l

nth n (rev l) d = nth (length l - n) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
n < length (rev l)
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l

nth n (rev l) d = nth (1 + length l - (1 + n)) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
n < length (rev l)
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l

nth n (rev l) d = nth (S (length l) - (1 + n)) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
n < length (rev l)
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l

nth n (rev l) d = nth (S (length l - (1 + n))) (a :: l) d
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l
n < length (rev l)
A:Type
a:A
l:list A
IHl:forall (d0 : A) (n0 : nat), n0 < length l -> nth n0 (rev l) d0 = nth (length l - S n0) l d0
d:A
n:nat
H:n < S (length l)
m:nat
H1:S n <= length l
H0:m = length l

n < length (rev l)
rewrite rev_length; auto. Qed.
An alternative tail-recursive definition for reverse
  Fixpoint rev_append (l l': list A) : list A :=
    match l with
      | [] => l'
      | a::l => rev_append l (a::l')
    end.

  Definition rev' l : list A := rev_append l [].

  
A:Type

forall l l' : list A, rev_append l l' = rev l ++ l'
A:Type

forall l l' : list A, rev_append l l' = rev l ++ l'
A:Type
a:A
l:list A
IHl:forall l'0 : list A, rev_append l l'0 = rev l ++ l'0
l':list A

rev_append l (a :: l') = (rev l ++ [a]) ++ l'
rewrite <- app_assoc; firstorder. Qed.
A:Type

forall l : list A, rev l = rev_append l []
A:Type

forall l : list A, rev l = rev_append l []
A:Type
l:list A

rev l = rev l ++ []
rewrite app_nil_r; trivial. Qed. (*********************************************)
Reverse Induction Principle on Lists
(*********************************************)

  Section Reverse_Induction.

    
A:Type

forall P : list A -> Prop, P [] -> (forall (a : A) (l : list A), P (rev l) -> P (rev (a :: l))) -> forall l : list A, P (rev l)
A:Type

forall P : list A -> Prop, P [] -> (forall (a : A) (l : list A), P (rev l) -> P (rev (a :: l))) -> forall l : list A, P (rev l)
induction l; auto. Qed.
A:Type

forall P : list A -> Prop, P [] -> (forall (x : A) (l : list A), P l -> P (l ++ [x])) -> forall l : list A, P l
A:Type

forall P : list A -> Prop, P [] -> (forall (x : A) (l : list A), P l -> P (l ++ [x])) -> forall l : list A, P l
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A

P l
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A

rev (rev l) = l -> P l
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A
E:rev (rev l) = l

P (rev (rev l))
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A
E:rev (rev l) = l

P []
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A
E:rev (rev l) = l
forall (a : A) (l0 : list A), P (rev l0) -> P (rev (a :: l0))
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A
E:rev (rev l) = l

forall (a : A) (l0 : list A), P (rev l0) -> P (rev (a :: l0))
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l0 : list A), P l0 -> P (l0 ++ [x])
l:list A
E:rev (rev l) = l

forall (a : A) (l0 : list A), P (rev l0) -> P (rev l0 ++ [a])
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l1 : list A), P l1 -> P (l1 ++ [x])
l:list A
E:rev (rev l) = l
a:A
l0:list A
H1:P (rev l0)

P (rev l0 ++ [a])
A:Type
P:list A -> Prop
H:P []
H0:forall (x : A) (l1 : list A), P l1 -> P (l1 ++ [x])
l:list A
E:rev (rev l) = l
a:A
l0:list A
H1:P (rev l0)

P (rev l0)
auto. Qed. End Reverse_Induction. (*************************)

Concatenation

  (*************************)

  Fixpoint concat (l : list (list A)) : list A :=
  match l with
  | nil => nil
  | cons x l => x ++ concat l
  end.

  
A:Type

concat [] = []
A:Type

concat [] = []
reflexivity. Qed.
A:Type

forall (x : list A) (l : list (list A)), concat (x :: l) = x ++ concat l
A:Type

forall (x : list A) (l : list (list A)), concat (x :: l) = x ++ concat l
reflexivity. Qed.
A:Type

forall l1 l2 : list (list A), concat (l1 ++ l2) = concat l1 ++ concat l2
A:Type

forall l1 l2 : list (list A), concat (l1 ++ l2) = concat l1 ++ concat l2
A:Type
l2:list (list A)

concat l2 = concat l2
A:Type
x:list A
l1:list (list A)
IH:forall l0 : list (list A), concat (l1 ++ l0) = concat l1 ++ concat l0
l2:list (list A)
x ++ concat (l1 ++ l2) = (x ++ concat l1) ++ concat l2
A:Type
l2:list (list A)

concat l2 = concat l2
reflexivity.
A:Type
x:list A
l1:list (list A)
IH:forall l0 : list (list A), concat (l1 ++ l0) = concat l1 ++ concat l0
l2:list (list A)

x ++ concat (l1 ++ l2) = (x ++ concat l1) ++ concat l2
rewrite IH; apply app_assoc. Qed. (***********************************)

Decidable equality on lists

  (***********************************)

  Hypothesis eq_dec : forall (x y : A), {x = y}+{x <> y}.

  
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l l' : list A, {l = l'} + {l <> l'}
A:Type
eq_dec:forall x y : A, {x = y} + {x <> y}

forall l l' : list A, {l = l'} + {l <> l'}
decide equality. Defined. End ListOps. (***************************************************)

Applying functions to the elements of a list

(***************************************************)

(************)

Map

(************)

Section Map.
  Variables (A : Type) (B : Type).
  Variable f : A -> B.

  Fixpoint map (l:list A) : list B :=
    match l with
      | [] => []
      | a :: t => (f a) :: (map t)
    end.

  
A, B:Type
f:A -> B
x:A
l:list A

map (x :: l) = f x :: map l
A, B:Type
f:A -> B
x:A
l:list A

map (x :: l) = f x :: map l
reflexivity. Qed.
A, B:Type
f:A -> B

forall (l : list A) (x : A), In x l -> In (f x) (map l)
A, B:Type
f:A -> B

forall (l : list A) (x : A), In x l -> In (f x) (map l)
induction l; firstorder (subst; auto). Qed.
A, B:Type
f:A -> B

forall (l : list A) (y : B), In y (map l) <-> (exists x : A, f x = y /\ In x l)
A, B:Type
f:A -> B

forall (l : list A) (y : B), In y (map l) <-> (exists x : A, f x = y /\ In x l)
induction l; firstorder (subst; auto). Qed.
A, B:Type
f:A -> B

forall l : list A, length (map l) = length l
A, B:Type
f:A -> B

forall l : list A, length (map l) = length l
induction l; simpl; auto. Qed.
A, B:Type
f:A -> B

forall (l : list A) (d : A) (n : nat), nth n (map l) (f d) = f (nth n l d)
A, B:Type
f:A -> B

forall (l : list A) (d : A) (n : nat), nth n (map l) (f d) = f (nth n l d)
induction l; simpl map; destruct n; firstorder. Qed.
A, B:Type
f:A -> B

forall (n : nat) (l : list A) (d : A), nth_error l n = Some d -> nth_error (map l) n = Some (f d)
A, B:Type
f:A -> B

forall (n : nat) (l : list A) (d : A), nth_error l n = Some d -> nth_error (map l) n = Some (f d)
induction n; intros [ | ] ? Heq; simpl in *; inversion Heq; auto. Qed.
A, B:Type
f:A -> B

forall l l' : list A, map (l ++ l') = map l ++ map l'
A, B:Type
f:A -> B

forall l l' : list A, map (l ++ l') = map l ++ map l'
A, B:Type
f:A -> B
a:A
l:list A
IHl:forall l' : list A, map (l ++ l') = map l ++ map l'

forall l' : list A, f a :: map (l ++ l') = f a :: map l ++ map l'
intros; rewrite IHl; auto. Qed.
A, B:Type
f:A -> B

forall l : list A, map (rev l) = rev (map l)
A, B:Type
f:A -> B

forall l : list A, map (rev l) = rev (map l)
A, B:Type
f:A -> B
a:A
l:list A
IHl:map (rev l) = rev (map l)

map (rev l ++ [a]) = rev (map l) ++ [f a]
A, B:Type
f:A -> B
a:A
l:list A
IHl:map (rev l) = rev (map l)

map (rev l) ++ map [a] = rev (map l) ++ [f a]
rewrite IHl; auto. Qed.
A, B:Type
f:A -> B

forall l : list A, map l = [] -> l = []
A, B:Type
f:A -> B

forall l : list A, map l = [] -> l = []
destruct l; simpl; reflexivity || discriminate. Qed.
map and count of occurrences
  Hypothesis decA: forall x1 x2 : A, {x1 = x2} + {x1 <> x2}.
  Hypothesis decB: forall y1 y2 : B, {y1 = y2} + {y1 <> y2}.
  Hypothesis Hfinjective: forall x1 x2: A, (f x1) = (f x2) -> x1 = x2.

  
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
x:A
l:list A

count_occ decA l x = count_occ decB (map l) (f x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
x:A
l:list A

count_occ decA l x = count_occ decB (map l) (f x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
l:list A

forall x : A, count_occ decA l x = count_occ decB (map l) (f x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
x:A

0 = 0
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
Hrec:forall x0 : A, count_occ decA l' x0 = count_occ decB (map l') (f x0)
x:A
(if decA a x then S (count_occ decA l' x) else count_occ decA l' x) = (if decB (f a) (f x) then S (count_occ decB (map l') (f x)) else count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
x:A

0 = 0
reflexivity.
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
Hrec:forall x0 : A, count_occ decA l' x0 = count_occ decB (map l') (f x0)
x:A

(if decA a x then S (count_occ decA l' x) else count_occ decA l' x) = (if decB (f a) (f x) then S (count_occ decB (map l') (f x)) else count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)

(if decA a x then S (count_occ decA l' x) else count_occ decA l' x) = (if decB (f a) (f x) then S (count_occ decB (map l') (f x)) else count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a = f x

S (count_occ decA l' x) = S (count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a <> f x
S (count_occ decA l' x) = count_occ decB (map l') (f x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a <> x
H2:f a = f x
count_occ decA l' x = S (count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a <> x
H2:f a <> f x
count_occ decA l' x = count_occ decB (map l') (f x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a = f x

S (count_occ decA l' x) = S (count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a = f x

S (count_occ decB (map l') (f x)) = S (count_occ decB (map l') (f x))
reflexivity.
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a <> f x

S (count_occ decA l' x) = count_occ decB (map l') (f x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a <> f x

f a = f x
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a = x
H2:f a <> f x

f x = f x
reflexivity.
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a <> x
H2:f a = f x

count_occ decA l' x = S (count_occ decB (map l') (f x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
a, x:A
Hfinjective:a = x
l':list A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a <> x
H2:f a = f x

count_occ decA l' x = S (count_occ decB (map l') (f x))
contradiction H1.
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2
a:A
l':list A
x:A
Hrec:count_occ decA l' x = count_occ decB (map l') (f x)
H1:a <> x
H2:f a <> f x

count_occ decA l' x = count_occ decB (map l') (f x)
assumption. Qed.
flat_map
  Definition flat_map (f:A -> list B) :=
    fix flat_map (l:list A) : list B :=
    match l with
      | nil => nil
      | cons x t => (f x)++(flat_map t)
    end.

  
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2

forall (f0 : A -> list B) (l : list A) (y : B), In y (flat_map f0 l) <-> (exists x : A, In x l /\ In y (f0 x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
Hfinjective:forall x1 x2 : A, f x1 = f x2 -> x1 = x2

forall (f0 : A -> list B) (l : list A) (y : B), In y (flat_map f0 l) <-> (exists x : A, In x l /\ In y (f0 x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}

forall (f0 : A -> list B) (l : list A) (y : B), In y (flat_map f0 l) <-> (exists x : A, In x l /\ In y (f0 x))
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
y:B
H:False

exists x : A, False /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
y:B
H:exists x : A, False /\ In y (f0 x)
False
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:In y (f0 a ++ flat_map f0 l)
exists x : A, (a = x \/ In x l) /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)
In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
y:B
H:exists x : A, False /\ In y (f0 x)

False
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:In y (f0 a ++ flat_map f0 l)
exists x : A, (a = x \/ In x l) /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)
In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:In y (f0 a ++ flat_map f0 l)

exists x : A, (a = x \/ In x l) /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)
In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:In y (f0 a ++ flat_map f0 l)
H0:In y (f0 a)

exists x : A, (a = x \/ In x l) /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:In y (f0 a ++ flat_map f0 l)
H0:In y (flat_map f0 l)
exists x : A, (a = x \/ In x l) /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)
In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:In y (f0 a ++ flat_map f0 l)
H0:In y (flat_map f0 l)

exists x : A, (a = x \/ In x l) /\ In y (f0 x)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)
In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x0 : A, In x0 l /\ In y0 (f0 x0))
y:B
H:In y (f0 a ++ flat_map f0 l)
H0:In y (flat_map f0 l)
H1:In y (flat_map f0 l) -> exists x0 : A, In x0 l /\ In y (f0 x0)
x:A
H2:In x l
H3:In y (f0 x)

exists x0 : A, (a = x0 \/ In x0 l) /\ In y (f0 x0)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)
In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)

In y (f0 a ++ flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x : A, In x l /\ In y0 (f0 x))
y:B
H:exists x : A, (a = x \/ In x l) /\ In y (f0 x)

In y (f0 a) \/ In y (flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x0 : A, In x0 l /\ In y0 (f0 x0))
y:B
x:A
H:a = x
H1:In y (f0 x)

In y (f0 a) \/ In y (flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x0 : A, In x0 l /\ In y0 (f0 x0))
y:B
x:A
H:In x l
H1:In y (f0 x)
In y (f0 a) \/ In y (flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x0 : A, In x0 l /\ In y0 (f0 x0))
y:B
x:A
H:In x l
H1:In y (f0 x)

In y (f0 a) \/ In y (flat_map f0 l)
A, B:Type
f:A -> B
decA:forall x1 x2 : A, {x1 = x2} + {x1 <> x2}
decB:forall y1 y2 : B, {y1 = y2} + {y1 <> y2}
f0:A -> list B
a:A
l:list A
IHl:forall y0 : B, In y0 (flat_map f0 l) <-> (exists x0 : A, In x0 l /\ In y0 (f0 x0))
y:B
x:A
H:In x l
H1:In y (f0 x)
H2:(exists x0 : A, In x0 l /\ In y (f0 x0)) -> In y (flat_map f0 l)

exists x0 : A, In x0 l /\ In y (f0 x0)
exists x; auto. Qed. End Map.

forall (A B : Type) (f : A -> list B) (l : list A), flat_map f l = concat (map f l)

forall (A B : Type) (f : A -> list B) (l : list A), flat_map f l = concat (map f l)
A, B:Type
f:A -> list B

[] = []
A, B:Type
f:A -> list B
x:A
l:list A
IH:flat_map f l = concat (map f l)
f x ++ flat_map f l = f x ++ concat (map f l)
A, B:Type
f:A -> list B

[] = []
reflexivity.
A, B:Type
f:A -> list B
x:A
l:list A
IH:flat_map f l = concat (map f l)

f x ++ flat_map f l = f x ++ concat (map f l)
rewrite IH; reflexivity. Qed.

forall (A B : Type) (f : A -> B) (l : list (list A)), map f (concat l) = concat (map (map f) l)

forall (A B : Type) (f : A -> B) (l : list (list A)), map f (concat l) = concat (map (map f) l)
A, B:Type
f:A -> B

[] = []
A, B:Type
f:A -> B
x:list A
l:list (list A)
IH:map f (concat l) = concat (map (map f) l)
map f (x ++ concat l) = map f x ++ concat (map (map f) l)
A, B:Type
f:A -> B

[] = []
reflexivity.
A, B:Type
f:A -> B
x:list A
l:list (list A)
IH:map f (concat l) = concat (map (map f) l)

map f (x ++ concat l) = map f x ++ concat (map (map f) l)
rewrite map_app, IH; reflexivity. Qed.

forall (A : Type) (l : list A), map (fun x : A => x) l = l

forall (A : Type) (l : list A), map (fun x : A => x) l = l
induction l; simpl; auto; rewrite IHl; auto. Qed.

forall (A B C : Type) (f : A -> B) (g : B -> C) (l : list A), map g (map f l) = map (fun x : A => g (f x)) l

forall (A B C : Type) (f : A -> B) (g : B -> C) (l : list A), map g (map f l) = map (fun x : A => g (f x)) l
A, B, C:Type
f:A -> B
g:B -> C
a:A
l:list A
IHl:map g (map f l) = map (fun x : A => g (f x)) l

g (f a) :: map g (map f l) = g (f a) :: map (fun x : A => g (f x)) l
rewrite IHl; auto. Qed.

forall (A B : Type) (f g : A -> B) (l : list A), (forall a : A, In a l -> f a = g a) -> map f l = map g l

forall (A B : Type) (f g : A -> B) (l : list A), (forall a : A, In a l -> f a = g a) -> map f l = map g l
A, B:Type
f, g:A -> B
a:A
l:list A
IHl:(forall a0 : A, In a0 l -> f a0 = g a0) -> map f l = map g l

(forall a0 : A, a = a0 \/ In a0 l -> f a0 = g a0) -> f a :: map f l = g a :: map g l
intros; rewrite H by intuition; rewrite IHl; auto. Qed.

forall (A B : Type) (f g : A -> B) (l : list A), map f l = map g l -> forall a : A, In a l -> f a = g a

forall (A B : Type) (f g : A -> B) (l : list A), map f l = map g l -> forall a : A, In a l -> f a = g a
induction l; intros [=] ? []; subst; auto. Qed. Arguments ext_in_map [A B f g l].

forall (A B : Type) (f g : A -> B) (l : list A), map f l = map g l <-> (forall a : A, In a l -> f a = g a)

forall (A B : Type) (f g : A -> B) (l : list A), map f l = map g l <-> (forall a : A, In a l -> f a = g a)
split; [apply ext_in_map | apply map_ext_in]. Qed. Arguments map_ext_in_iff {A B f g l}.

forall (A B : Type) (f g : A -> B), (forall a : A, f a = g a) -> forall l : list A, map f l = map g l

forall (A B : Type) (f g : A -> B), (forall a : A, f a = g a) -> forall l : list A, map f l = map g l
intros; apply map_ext_in; auto. Qed. (************************************)
Left-to-right iterator on lists
(************************************)

Section Fold_Left_Recursor.
  Variables (A : Type) (B : Type).
  Variable f : A -> B -> A.

  Fixpoint fold_left (l:list B) (a0:A) : A :=
    match l with
      | nil => a0
      | cons b t => fold_left t (f a0 b)
    end.

  
A, B:Type
f:A -> B -> A

forall (l l' : list B) (i : A), fold_left (l ++ l') i = fold_left l' (fold_left l i)
A, B:Type
f:A -> B -> A

forall (l l' : list B) (i : A), fold_left (l ++ l') i = fold_left l' (fold_left l i)
A, B:Type
f:A -> B -> A

forall (l' : list B) (i : A), fold_left ([] ++ l') i = fold_left l' (fold_left [] i)
A, B:Type
f:A -> B -> A
a:B
l:list B
IHl:forall (l' : list B) (i : A), fold_left (l ++ l') i = fold_left l' (fold_left l i)
forall (l' : list B) (i : A), fold_left ((a :: l) ++ l') i = fold_left l' (fold_left (a :: l) i)
A, B:Type
f:A -> B -> A
a:B
l:list B
IHl:forall (l' : list B) (i : A), fold_left (l ++ l') i = fold_left l' (fold_left l i)

forall (l' : list B) (i : A), fold_left ((a :: l) ++ l') i = fold_left l' (fold_left (a :: l) i)
A, B:Type
f:A -> B -> A
a:B
l:list B
IHl:forall (l'0 : list B) (i0 : A), fold_left (l ++ l'0) i0 = fold_left l'0 (fold_left l i0)
l':list B
i:A

fold_left ((a :: l) ++ l') i = fold_left l' (fold_left (a :: l) i)
A, B:Type
f:A -> B -> A
a:B
l:list B
IHl:forall (l'0 : list B) (i0 : A), fold_left (l ++ l'0) i0 = fold_left l'0 (fold_left l i0)
l':list B
i:A

fold_left (l ++ l') (f i a) = fold_left l' (fold_left l (f i a))
auto. Qed. End Fold_Left_Recursor.

forall (A : Type) (l : list A), fold_left (fun (x : nat) (_ : A) => S x) l 0 = length l

forall (A : Type) (l : list A), fold_left (fun (x : nat) (_ : A) => S x) l 0 = length l
A:Type
l:list A

fold_left (fun (x : nat) (_ : A) => S x) l 0 = length l
A:Type
l:list A

forall n : nat, fold_left (fun (x : nat) (_ : A) => S x) l n = n + length l
A:Type
a:A
l:list A
IHl:forall n : nat, fold_left (fun (x : nat) (_ : A) => S x) l n = n + length l

forall n : nat, fold_left (fun (x : nat) (_ : A) => S x) l (S n) = n + S (length l)
A:Type
a:A
l:list A
IHl:forall n0 : nat, fold_left (fun (x : nat) (_ : A) => S x) l n0 = n0 + length l
n:nat

S n + length l = n + S (length l)
simpl; auto with arith. Qed. (************************************)
Right-to-left iterator on lists
(************************************)

Section Fold_Right_Recursor.
  Variables (A : Type) (B : Type).
  Variable f : B -> A -> A.
  Variable a0 : A.

  Fixpoint fold_right (l:list B) : A :=
    match l with
      | nil => a0
      | cons b t => f b (fold_right t)
    end.

End Fold_Right_Recursor.

  

forall (A B : Type) (f : A -> B -> B) (l l' : list A) (i : B), fold_right f i (l ++ l') = fold_right f (fold_right f i l') l

forall (A B : Type) (f : A -> B -> B) (l l' : list A) (i : B), fold_right f i (l ++ l') = fold_right f (fold_right f i l') l
A, B:Type
f:A -> B -> B

forall (l' : list A) (i : B), fold_right f i ([] ++ l') = fold_right f (fold_right f i l') []
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall (l' : list A) (i : B), fold_right f i (l ++ l') = fold_right f (fold_right f i l') l
forall (l' : list A) (i : B), fold_right f i ((a :: l) ++ l') = fold_right f (fold_right f i l') (a :: l)
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall (l' : list A) (i : B), fold_right f i (l ++ l') = fold_right f (fold_right f i l') l

forall (l' : list A) (i : B), fold_right f i ((a :: l) ++ l') = fold_right f (fold_right f i l') (a :: l)
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall (l'0 : list A) (i0 : B), fold_right f i0 (l ++ l'0) = fold_right f (fold_right f i0 l'0) l
l':list A
i:B

f a (fold_right f i (l ++ l')) = f a (fold_right f (fold_right f i l') l)
f_equal; auto. Qed.

forall (A B : Type) (f : A -> B -> B) (l : list A) (i : B), fold_right f i (rev l) = fold_left (fun (x : B) (y : A) => f y x) l i

forall (A B : Type) (f : A -> B -> B) (l : list A) (i : B), fold_right f i (rev l) = fold_left (fun (x : B) (y : A) => f y x) l i
A, B:Type
f:A -> B -> B

forall i : B, fold_right f i (rev []) = fold_left (fun (x : B) (y : A) => f y x) [] i
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall i : B, fold_right f i (rev l) = fold_left (fun (x : B) (y : A) => f y x) l i
forall i : B, fold_right f i (rev (a :: l)) = fold_left (fun (x : B) (y : A) => f y x) (a :: l) i
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall i : B, fold_right f i (rev l) = fold_left (fun (x : B) (y : A) => f y x) l i

forall i : B, fold_right f i (rev (a :: l)) = fold_left (fun (x : B) (y : A) => f y x) (a :: l) i
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall i0 : B, fold_right f i0 (rev l) = fold_left (fun (x : B) (y : A) => f y x) l i0
i:B

fold_right f i (rev (a :: l)) = fold_left (fun (x : B) (y : A) => f y x) (a :: l) i
A, B:Type
f:A -> B -> B
a:A
l:list A
IHl:forall i0 : B, fold_right f i0 (rev l) = fold_left (fun (x : B) (y : A) => f y x) l i0
i:B

fold_right f i (rev l ++ [a]) = fold_left (fun (x : B) (y : A) => f y x) l (f a i)
rewrite fold_right_app; simpl; auto. Qed.

forall (A : Type) (f : A -> A -> A), (forall x y z : A, f x (f y z) = f (f x y) z) -> forall a0 : A, (forall y : A, f a0 y = f y a0) -> forall l : list A, fold_left f l a0 = fold_right f a0 l

forall (A : Type) (f : A -> A -> A), (forall x y z : A, f x (f y z) = f (f x y) z) -> forall a0 : A, (forall y : A, f a0 y = f y a0) -> forall l : list A, fold_left f l a0 = fold_right f a0 l
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
l:list A

fold_left f l a0 = fold_right f a0 l
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a1:A
l:list A
IHl:fold_left f l a0 = fold_right f a0 l

fold_left f (a1 :: l) a0 = fold_right f a0 (a1 :: l)
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a1:A
l:list A
IHl:fold_left f l a0 = fold_right f a0 l

fold_left f l (f a0 a1) = f a1 (fold_right f a0 l)
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a1:A
l:list A
IHl:fold_left f l a0 = fold_right f a0 l

fold_left f l (f a0 a1) = f a1 (fold_left f l a0)
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a1:A
l:list A

fold_left f l (f a0 a1) = f a1 (fold_left f l a0)
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
l:list A

forall a1 : A, fold_left f l (f a0 a1) = f a1 (fold_left f l a0)
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a:A
l:list A
IHl:forall a1 : A, fold_left f l (f a0 a1) = f a1 (fold_left f l a0)

forall a1 : A, fold_left f (a :: l) (f a0 a1) = f a1 (fold_left f (a :: l) a0)
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a:A
l:list A
IHl:forall a1 : A, fold_left f l (f a0 a1) = f a1 (fold_left f l a0)

forall a1 : A, fold_left f l (f (f a0 a1) a) = f a1 (fold_left f l (f a0 a))
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a:A
l:list A
IHl:forall a2 : A, fold_left f l (f a0 a2) = f a2 (fold_left f l a0)
a1:A

fold_left f l (f (f a0 a1) a) = f a1 (fold_left f l (f a0 a))
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a:A
l:list A
IHl:forall a2 : A, fold_left f l (f a0 a2) = f a2 (fold_left f l a0)
a1:A

fold_left f l (f a0 (f a1 a)) = f a1 (fold_left f l (f a0 a))
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a:A
l:list A
IHl:forall a2 : A, fold_left f l (f a0 a2) = f a2 (fold_left f l a0)
a1:A

f (f a1 a) (fold_left f l a0) = f a1 (fold_left f l (f a0 a))
A:Type
f:A -> A -> A
assoc:forall x y z : A, f x (f y z) = f (f x y) z
a0:A
comma0:forall y : A, f a0 y = f y a0
a:A
l:list A
IHl:forall a2 : A, fold_left f l (f a0 a2) = f a2 (fold_left f l a0)
a1:A

f (f a1 a) (fold_left f l a0) = f a1 (f a (fold_left f l a0))
auto. Qed.
(list_power x y) is y^x, or the set of sequences of elts of y indexed by elts of x, sorted in lexicographic order.
  Fixpoint list_power (A B:Type)(l:list A) (l':list B) :
    list (list (A * B)) :=
    match l with
      | nil => cons nil nil
      | cons x t =>
	flat_map (fun f:list (A * B) => map (fun y:B => cons (x, y) f) l')
        (list_power t l')
    end.


  (*************************************)

Boolean operations over lists

  (*************************************)

  Section Bool.
    Variable A : Type.
    Variable f : A -> bool.
find whether a boolean function can be satisfied by an elements of the list.
    Fixpoint existsb (l:list A) : bool :=
      match l with
	| nil => false
	| a::l => f a || existsb l
      end.

    
A:Type
f:A -> bool

forall l : list A, existsb l = true <-> (exists x : A, In x l /\ f x = true)
A:Type
f:A -> bool

forall l : list A, existsb l = true <-> (exists x : A, In x l /\ f x = true)
A:Type
f:A -> bool
H:false = true

exists x : A, False /\ f x = true
A:Type
f:A -> bool
H:exists x : A, False /\ f x = true
false = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:f a || existsb l = true
exists x : A, (a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:exists x : A, (a = x \/ In x l) /\ f x = true
f a || existsb l = true
A:Type
f:A -> bool
H:exists x : A, False /\ f x = true

false = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:f a || existsb l = true
exists x : A, (a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:exists x : A, (a = x \/ In x l) /\ f x = true
f a || existsb l = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:f a || existsb l = true

exists x : A, (a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:exists x : A, (a = x \/ In x l) /\ f x = true
f a || existsb l = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x : A, In x l /\ f x = true
H0:(exists x : A, In x l /\ f x = true) -> existsb l = true
H1:exists x : A, (a = x \/ In x l) /\ f x = true

f a || existsb l = true
A:Type
f:A -> bool
a:A
l:list A
H:existsb l = true -> exists x0 : A, In x0 l /\ f x0 = true
x:A
H2:f x = true
H0:forall x0 : A, (fun x1 : A => In x1 l /\ f x1 = true) x0 -> existsb l = true
H1:a = x

f a || existsb l = true
A:Type
f:A -> bool
l:list A
H:existsb l = true -> exists x0 : A, In x0 l /\ f x0 = true
x:A
H2:f x = true
H0:forall x0 : A, (fun x1 : A => In x1 l /\ f x1 = true) x0 -> existsb l = true

f x || existsb l = true
rewrite H2; auto. Qed.
A:Type
f:A -> bool

forall (l : list A) (n : nat) (d : A), n < length l -> existsb l = false -> f (nth n l d) = false
A:Type
f:A -> bool

forall (l : list A) (n : nat) (d : A), n < length l -> existsb l = false -> f (nth n l d) = false
A:Type
f:A -> bool

forall (n : nat) (d : A), n < length [] -> existsb [] = false -> f (nth n [] d) = false
A:Type
f:A -> bool
a:A
l:list A
IHl:forall (n : nat) (d : A), n < length l -> existsb l = false -> f (nth n l d) = false
forall (n : nat) (d : A), n < length (a :: l) -> existsb (a :: l) = false -> f (nth n (a :: l) d) = false
A:Type
f:A -> bool
a:A
l:list A
IHl:forall (n : nat) (d : A), n < length l -> existsb l = false -> f (nth n l d) = false

forall (n : nat) (d : A), n < length (a :: l) -> existsb (a :: l) = false -> f (nth n (a :: l) d) = false
A:Type
f:A -> bool
a:A
l:list A
IHl:forall (n0 : nat) (d0 : A), n0 < length l -> existsb l = false -> f (nth n0 l d0) = false
n:nat
d:A
H:n < S (length l)
H0:f a || existsb l = false

f match n with | 0 => a | S m => nth m l d end = false
A:Type
f:A -> bool
a:A
l:list A
IHl:forall (n0 : nat) (d0 : A), n0 < length l -> existsb l = false -> f (nth n0 l d0) = false
n:nat
d:A
H:n < S (length l)
H1:f a = false
H2:existsb l = false

f match n with | 0 => a | S m => nth m l d end = false
A:Type
f:A -> bool
a:A
l:list A
IHl:forall (n0 : nat) (d0 : A), n0 < length l -> existsb l = false -> f (nth n0 l d0) = false
n:nat
d:A
H:S n < S (length l)
H1:f a = false
H2:existsb l = false

f (nth n l d) = false
rewrite IHl; auto with arith. Qed.
A:Type
f:A -> bool

forall l1 l2 : list A, existsb (l1 ++ l2) = existsb l1 || existsb l2
A:Type
f:A -> bool

forall l1 l2 : list A, existsb (l1 ++ l2) = existsb l1 || existsb l2
A:Type
f:A -> bool
l2:list A

existsb l2 = existsb l2
A:Type
f:A -> bool
a:A
l1:list A
IHl1:forall l0 : list A, existsb (l1 ++ l0) = existsb l1 || existsb l0
l2:list A
f a || existsb (l1 ++ l2) = f a || existsb l1 || existsb l2
A:Type
f:A -> bool
a:A
l1:list A
IHl1:forall l0 : list A, existsb (l1 ++ l0) = existsb l1 || existsb l0
l2:list A

f a || existsb (l1 ++ l2) = f a || existsb l1 || existsb l2
case (f a); simpl; solve[auto]. Qed.
find whether a boolean function is satisfied by all the elements of a list.
    Fixpoint forallb (l:list A) : bool :=
      match l with
	| nil => true
	| a::l => f a && forallb l
      end.

    
A:Type
f:A -> bool

forall l : list A, forallb l = true <-> (forall x : A, In x l -> f x = true)
A:Type
f:A -> bool

forall l : list A, forallb l = true <-> (forall x : A, In x l -> f x = true)
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x0 : A, In x0 l -> f x0 = true
H0:(forall x0 : A, In x0 l -> f x0 = true) -> forallb l = true
H1:f a && forallb l = true
x:A
H3:a = x

f x = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x0 : A, In x0 l -> f x0 = true
H0:(forall x0 : A, In x0 l -> f x0 = true) -> forallb l = true
H1:f a && forallb l = true
x:A
H3:In x l
f x = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true
f a && forallb l = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x0 : A, In x0 l -> f x0 = true
H0:(forall x0 : A, In x0 l -> f x0 = true) -> forallb l = true
H1:f a && forallb l = true
x:A
H3:a = x
H2:f a = true
H4:forallb l = true

f x = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x0 : A, In x0 l -> f x0 = true
H0:(forall x0 : A, In x0 l -> f x0 = true) -> forallb l = true
H1:f a && forallb l = true
x:A
H3:In x l
f x = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true
f a && forallb l = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x0 : A, In x0 l -> f x0 = true
H0:(forall x0 : A, In x0 l -> f x0 = true) -> forallb l = true
H1:f a && forallb l = true
x:A
H3:In x l

f x = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true
f a && forallb l = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true

f a && forallb l = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true

forallb l = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true
H2:forallb l = true
f a && forallb l = true
A:Type
f:A -> bool
a:A
l:list A
H:forallb l = true -> forall x : A, In x l -> f x = true
H0:(forall x : A, In x l -> f x = true) -> forallb l = true
H1:forall x : A, a = x \/ In x l -> f x = true
H2:forallb l = true

f a && forallb l = true
rewrite H1; auto. Qed.
A:Type
f:A -> bool

forall l1 l2 : list A, forallb (l1 ++ l2) = forallb l1 && forallb l2
A:Type
f:A -> bool

forall l1 l2 : list A, forallb (l1 ++ l2) = forallb l1 && forallb l2
A:Type
f:A -> bool

forall l2 : list A, forallb l2 = forallb l2
A:Type
f:A -> bool
a:A
l1:list A
IHl1:forall l2 : list A, forallb (l1 ++ l2) = forallb l1 && forallb l2
forall l2 : list A, f a && forallb (l1 ++ l2) = f a && forallb l1 && forallb l2
A:Type
f:A -> bool
a:A
l1:list A
IHl1:forall l2 : list A, forallb (l1 ++ l2) = forallb l1 && forallb l2

forall l2 : list A, f a && forallb (l1 ++ l2) = f a && forallb l1 && forallb l2
case (f a); simpl; solve[auto]. Qed.
filter
    Fixpoint filter (l:list A) : list A :=
      match l with
	| nil => nil
	| x :: l => if f x then x::(filter l) else filter l
      end.

    
A:Type
f:A -> bool

forall (x : A) (l : list A), In x (filter l) <-> In x l /\ f x = true
A:Type
f:A -> bool

forall (x : A) (l : list A), In x (filter l) <-> In x l /\ f x = true
A:Type
f:A -> bool
x:A

False <-> False /\ f x = true
A:Type
f:A -> bool
x, a:A
l:list A
IHl:In x (filter l) <-> In x l /\ f x = true
In x (if f a then a :: filter l else filter l) <-> (a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
x, a:A
l:list A
IHl:In x (filter l) <-> In x l /\ f x = true

In x (if f a then a :: filter l else filter l) <-> (a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
x, a:A
l:list A
IHl:In x (filter l) <-> In x l /\ f x = true

In x (if f a then a :: filter l else filter l) <-> (a = x \/ In x l) /\ f x = true
case_eq (f a); intros; simpl; intuition congruence. Qed.
find
    Fixpoint find (l:list A) : option A :=
      match l with
	| nil => None
	| x :: tl => if f x then Some x else find tl
      end.

    
A:Type
f:A -> bool
l:list A
x:A

find l = Some x -> In x l /\ f x = true
A:Type
f:A -> bool
l:list A
x:A

find l = Some x -> In x l /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
x:A
IH:find l = Some x -> In x l /\ f x = true

(if f a then Some a else find l) = Some x -> (a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
x:A
IH:find l = Some x -> In x l /\ f x = true
Ha:f a = true
Eq:Some a = Some x

(a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
x:A
IH:find l = Some x -> In x l /\ f x = true
Ha:f a = false
Eq:find l = Some x
(a = x \/ In x l) /\ f x = true
A:Type
f:A -> bool
a:A
l:list A
x:A
IH:find l = Some x -> In x l /\ f x = true
Ha:f a = true
Eq:Some a = Some x

(a = x \/ In x l) /\ f x = true
injection Eq as ->; auto.
A:Type
f:A -> bool
a:A
l:list A
x:A
IH:find l = Some x -> In x l /\ f x = true
Ha:f a = false
Eq:find l = Some x

(a = x \/ In x l) /\ f x = true
destruct (IH Eq); auto. Qed.
A:Type
f:A -> bool
l:list A

find l = None -> forall x : A, In x l -> f x = false
A:Type
f:A -> bool
l:list A

find l = None -> forall x : A, In x l -> f x = false
A:Type
f:A -> bool
a:A
l:list A
IH:find l = None -> forall x : A, In x l -> f x = false

(if f a then Some a else find l) = None -> forall x : A, a = x \/ In x l -> f x = false
A:Type
f:A -> bool
a:A
l:list A
IH:find l = None -> forall x0 : A, In x0 l -> f x0 = false
Ha:f a = false
Eq:find l = None
x:A
IN:a = x \/ In x l

f x = false
destruct IN as [<-|IN]; auto. Qed.
partition
    Fixpoint partition (l:list A) : list A * list A :=
      match l with
	| nil => (nil, nil)
	| x :: tl => let (g,d) := partition tl in
	  if f x then (x::g,d) else (g,x::d)
      end.

  
A:Type
f:A -> bool
a:A
l, l1, l2:list A

partition l = (l1, l2) -> f a = true -> partition (a :: l) = (a :: l1, l2)
A:Type
f:A -> bool
a:A
l, l1, l2:list A

partition l = (l1, l2) -> f a = true -> partition (a :: l) = (a :: l1, l2)
A:Type
f:A -> bool
a:A
l, l1, l2:list A

partition l = (l1, l2) -> f a = true -> (let (g, d) := partition l in if f a then (a :: g, d) else (g, a :: d)) = (a :: l1, l2)
now intros -> ->. Qed.
A:Type
f:A -> bool
a:A
l, l1, l2:list A

partition l = (l1, l2) -> f a = false -> partition (a :: l) = (l1, a :: l2)
A:Type
f:A -> bool
a:A
l, l1, l2:list A

partition l = (l1, l2) -> f a = false -> partition (a :: l) = (l1, a :: l2)
A:Type
f:A -> bool
a:A
l, l1, l2:list A

partition l = (l1, l2) -> f a = false -> (let (g, d) := partition l in if f a then (a :: g, d) else (g, a :: d)) = (l1, a :: l2)
now intros -> ->. Qed.
A:Type
f:A -> bool
l, l1, l2:list A

partition l = (l1, l2) -> length l = length l1 + length l2
A:Type
f:A -> bool
l, l1, l2:list A

partition l = (l1, l2) -> length l = length l1 + length l2
A:Type
f:A -> bool
l:list A

forall l1 l2 : list A, partition l = (l1, l2) -> length l = length l1 + length l2
A:Type
f:A -> bool
l1, l2:list A

partition [] = (l1, l2) -> length [] = length l1 + length l2
A:Type
f:A -> bool
a:A
l':list A
Hrec:forall l0 l3 : list A, partition l' = (l0, l3) -> length l' = length l0 + length l3
l1, l2:list A
partition (a :: l') = (l1, l2) -> length (a :: l') = length l1 + length l2
A:Type
f:A -> bool
l1, l2:list A

partition [] = (l1, l2) -> length [] = length l1 + length l2
now intros [= <- <- ].
A:Type
f:A -> bool
a:A
l':list A
Hrec:forall l0 l3 : list A, partition l' = (l0, l3) -> length l' = length l0 + length l3
l1, l2:list A

partition (a :: l') = (l1, l2) -> length (a :: l') = length l1 + length l2
A:Type
f:A -> bool
a:A
l':list A
Hrec:forall l0 l3 : list A, partition l' = (l0, l3) -> length l' = length l0 + length l3
l1, l2:list A

(let (g, d) := partition l' in if f a then (a :: g, d) else (g, a :: d)) = (l1, l2) -> S (length l') = length l1 + length l2
destruct (f a), (partition l') as (left, right); intros [= <- <- ]; simpl; rewrite (Hrec left right); auto. Qed.
A:Type
f:A -> bool
l:list A

partition l = ([], []) <-> l = []
A:Type
f:A -> bool
l:list A

partition l = ([], []) <-> l = []
A:Type
f:A -> bool
l:list A

partition l = ([], []) -> l = []
A:Type
f:A -> bool
l:list A
l = [] -> partition l = ([], [])
A:Type
f:A -> bool
l:list A

partition l = ([], []) -> l = []
A:Type
f:A -> bool

partition [] = ([], []) -> [] = []
A:Type
f:A -> bool
a:A
l':list A
partition (a :: l') = ([], []) -> a :: l' = []
A:Type
f:A -> bool

partition [] = ([], []) -> [] = []
intuition.
A:Type
f:A -> bool
a:A
l':list A

partition (a :: l') = ([], []) -> a :: l' = []
A:Type
f:A -> bool
a:A
l':list A

(let (g, d) := partition l' in if f a then (a :: g, d) else (g, a :: d)) = ([], []) -> a :: l' = []
destruct (f a), (partition l'); now intros [= -> ->].
A:Type
f:A -> bool
l:list A

l = [] -> partition l = ([], [])
now intros ->. Qed.
A:Type
f:A -> bool
l, l1, l2:list A

partition l = (l1, l2) -> forall x : A, In x l <-> In x l1 \/ In x l2
A:Type
f:A -> bool
l, l1, l2:list A

partition l = (l1, l2) -> forall x : A, In x l <-> In x l1 \/ In x l2
A:Type
f:A -> bool
l:list A

forall l1 l2 : list A, partition l = (l1, l2) -> forall x : A, In x l <-> In x l1 \/ In x l2
A:Type
f:A -> bool
l1, l2:list A
Eq:([], []) = (l1, l2)
x:A

False <-> In x l1 \/ In x l2
A:Type
f:A -> bool
a:A
l':list A
Hrec:forall l0 l3 : list A, partition l' = (l0, l3) -> forall x0 : A, In x0 l' <-> In x0 l0 \/ In x0 l3
l1, l2:list A
Eq:(let (g, d) := partition l' in if f a then (a :: g, d) else (g, a :: d)) = (l1, l2)
x:A
a = x \/ In x l' <-> In x l1 \/ In x l2
A:Type
f:A -> bool
l1, l2:list A
Eq:([], []) = (l1, l2)
x:A

False <-> In x l1 \/ In x l2
A:Type
f:A -> bool
x:A

False <-> In x [] \/ In x []
tauto.
A:Type
f:A -> bool
a:A
l':list A
Hrec:forall l0 l3 : list A, partition l' = (l0, l3) -> forall x0 : A, In x0 l' <-> In x0 l0 \/ In x0 l3
l1, l2:list A
Eq:(let (g, d) := partition l' in if f a then (a :: g, d) else (g, a :: d)) = (l1, l2)
x:A

a = x \/ In x l' <-> In x l1 \/ In x l2
A:Type
f:A -> bool
a:A
l', left, right:list A
Hrec:forall l0 l3 : list A, (left, right) = (l0, l3) -> forall x0 : A, In x0 l' <-> In x0 l0 \/ In x0 l3
l1, l2:list A
Eq:(if f a then (a :: left, right) else (left, a :: right)) = (l1, l2)
x:A

a = x \/ In x l' <-> In x l1 \/ In x l2
A:Type
f:A -> bool
a:A
l', left, right:list A
x:A
Hrec:In x l' <-> In x left \/ In x right
l1, l2:list A
Eq:(if f a then (a :: left, right) else (left, a :: right)) = (l1, l2)

a = x \/ In x l' <-> In x l1 \/ In x l2
destruct (f a); injection Eq as <- <-; simpl; tauto. Qed. End Bool. (******************************************************)

Operations on lists of pairs or lists of lists

  (******************************************************)

  Section ListPairs.
    Variables (A : Type) (B : Type).
split derives two lists from a list of pairs
    Fixpoint split (l:list (A*B)) : list A * list B :=
      match l with
	| [] => ([], [])
	| (x,y) :: tl => let (left,right) := split tl in (x::left, y::right)
      end.

    
A, B:Type

forall (l : list (A * B)) (p : A * B), In p l -> In (fst p) (fst (split l))
A, B:Type

forall (l : list (A * B)) (p : A * B), In p l -> In (fst p) (fst (split l))
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall p0 : A * B, In p0 l -> In (fst p0) (fst (split l))
p:(A * B)%type
H:a = p \/ In p l

In (fst p) (fst (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right)))
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (fst p) l0
a0:A
b:B
H:(a, b0) = (a0, b) \/ In (a0, b) l

a = a0 \/ In a0 l0
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (fst p) l0
a0:A
b:B
H:(a, b0) = (a0, b)

a = a0 \/ In a0 l0
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (fst p) l0
a0:A
b:B
H:In (a0, b) l
a = a0 \/ In a0 l0
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (fst p) l0
a0:A
b:B
H:In (a0, b) l

a = a0 \/ In a0 l0
right; apply (IHl (a0,b) H). Qed.
A, B:Type

forall (l : list (A * B)) (p : A * B), In p l -> In (snd p) (snd (split l))
A, B:Type

forall (l : list (A * B)) (p : A * B), In p l -> In (snd p) (snd (split l))
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall p0 : A * B, In p0 l -> In (snd p0) (snd (split l))
p:(A * B)%type
H:a = p \/ In p l

In (snd p) (snd (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right)))
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (snd p) l1
a0:A
b:B
H:(a, b0) = (a0, b) \/ In (a0, b) l

b0 = b \/ In b l1
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (snd p) l1
a0:A
b:B
H:(a, b0) = (a0, b)

b0 = b \/ In b l1
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (snd p) l1
a0:A
b:B
H:In (a0, b) l
b0 = b \/ In b l1
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall p : A * B, In p l -> In (snd p) l1
a0:A
b:B
H:In (a0, b) l

b0 = b \/ In b l1
right; apply (IHl (a0,b) H). Qed.
A, B:Type

forall (l : list (A * B)) (n : nat) (d : A * B), nth n l d = (nth n (fst (split l)) (fst d), nth n (snd (split l)) (snd d))
A, B:Type

forall (l : list (A * B)) (n : nat) (d : A * B), nth n l d = (nth n (fst (split l)) (fst d), nth n (snd (split l)) (snd d))
A, B:Type

forall (n : nat) (d : A * B), nth n [] d = (nth n (fst (split [])) (fst d), nth n (snd (split [])) (snd d))
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall (n : nat) (d : A * B), nth n l d = (nth n (fst (split l)) (fst d), nth n (snd (split l)) (snd d))
forall (n : nat) (d : A * B), nth n (a :: l) d = (nth n (fst (split (a :: l))) (fst d), nth n (snd (split (a :: l))) (snd d))
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall (n : nat) (d : A * B), nth n l d = (nth n (fst (split l)) (fst d), nth n (snd (split l)) (snd d))

forall (n : nat) (d : A * B), nth n (a :: l) d = (nth n (fst (split (a :: l))) (fst d), nth n (snd (split (a :: l))) (snd d))
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall (n : nat) (d : A * B), nth n l d = (nth n (fst (split l)) (fst d), nth n (snd (split l)) (snd d))
a0:A
b:B

a = (nth 0 (fst (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) a0, nth 0 (snd (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) b)
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall (n0 : nat) (d : A * B), nth n0 l d = (nth n0 (fst (split l)) (fst d), nth n0 (snd (split l)) (snd d))
n:nat
a0:A
b:B
nth n l (a0, b) = (nth (S n) (fst (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) a0, nth (S n) (snd (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) b)
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:forall (n0 : nat) (d : A * B), nth n0 l d = (nth n0 (fst (split l)) (fst d), nth n0 (snd (split l)) (snd d))
n:nat
a0:A
b:B

nth n l (a0, b) = (nth (S n) (fst (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) a0, nth (S n) (snd (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) b)
A, B:Type
a:A
b0:B
l:list (A * B)
l0:list A
l1:list B
IHl:forall (n0 : nat) (d : A * B), nth n0 l d = (nth n0 l0 (fst d), nth n0 l1 (snd d))
n:nat
a0:A
b:B

nth n l (a0, b) = (nth n l0 a0, nth n l1 b)
apply IHl. Qed.
A, B:Type

forall l : list (A * B), length (fst (split l)) = length l
A, B:Type

forall l : list (A * B), length (fst (split l)) = length l
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:length (fst (split l)) = length l

length (fst (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) = S (length l)
destruct a; destruct (split l); simpl; auto. Qed.
A, B:Type

forall l : list (A * B), length (snd (split l)) = length l
A, B:Type

forall l : list (A * B), length (snd (split l)) = length l
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:length (snd (split l)) = length l

length (snd (let (x, y) := a in let (left, right) := split l in (x :: left, y :: right))) = S (length l)
destruct a; destruct (split l); simpl; auto. Qed.
combine is the opposite of split. Lists given to combine are meant to be of same length. If not, combine stops on the shorter list
    Fixpoint combine (l : list A) (l' : list B) : list (A*B) :=
      match l,l' with
	| x::tl, y::tl' => (x,y)::(combine tl tl')
	| _, _ => nil
      end.

    
A, B:Type

forall l : list (A * B), let (l1, l2) := split l in combine l1 l2 = l
A, B:Type

forall l : list (A * B), let (l1, l2) := split l in combine l1 l2 = l
A, B:Type

let (l1, l2) := split [] in combine l1 l2 = []
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:let (l1, l2) := split l in combine l1 l2 = l
let (l1, l2) := split (a :: l) in combine l1 l2 = a :: l
A, B:Type
a:(A * B)%type
l:list (A * B)
IHl:let (l1, l2) := split l in combine l1 l2 = l

let (l1, l2) := split (a :: l) in combine l1 l2 = a :: l
A, B:Type
a:A
b:B
l:list (A * B)
IHl:let (l1, l2) := split l in combine l1 l2 = l

let (l1, l2) := let (left, right) := split l in (a :: left, b :: right) in combine l1 l2 = (a, b) :: l
A, B:Type
a:A
b:B
l:list (A * B)
l0:list A
l1:list B
IHl:combine l0 l1 = l

(a, b) :: combine l0 l1 = (a, b) :: l
f_equal; auto. Qed.
A, B:Type

forall (l : list A) (l' : list B), length l = length l' -> split (combine l l') = (l, l')
A, B:Type

forall (l : list A) (l' : list B), length l = length l' -> split (combine l l') = (l, l')
A, B:Type
a:A
l:list A
IHl:forall l'0 : list B, length l = length l'0 -> split (combine l l'0) = (l, l'0)
b:B
l':list B

S (length l) = S (length l') -> (let (left, right) := split (combine l l') in (a :: left, b :: right)) = (a :: l, b :: l')
now intros [= ->%IHl]. Qed.
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In x l
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In x l
A, B:Type

forall (l' : list B) (x : A) (y : B), In (x, y) (combine [] l') -> In x []
A, B:Type
a:A
l:list A
IHl:forall (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In x l
forall (l' : list B) (x : A) (y : B), In (x, y) (combine (a :: l) l') -> In x (a :: l)
A, B:Type
a:A
l:list A
IHl:forall (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In x l

forall (l' : list B) (x : A) (y : B), In (x, y) (combine (a :: l) l') -> In x (a :: l)
A, B:Type
a:A
l:list A
IHl:forall (l' : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l') -> In x0 l
x:A
y:B
H:False

a = x \/ In x l
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In x0 l
b:B
l':list B
x:A
y:B
H:(a, b) = (x, y) \/ In (x, y) (combine l l')
a = x \/ In x l
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In x0 l
b:B
l':list B
x:A
y:B
H:(a, b) = (x, y) \/ In (x, y) (combine l l')

a = x \/ In x l
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In x0 l
b:B
l':list B
x:A
y:B
H:(a, b) = (x, y)

a = x \/ In x l
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In x0 l
b:B
l':list B
x:A
y:B
H:In (x, y) (combine l l')
a = x \/ In x l
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In x0 l
b:B
l':list B
x:A
y:B
H:In (x, y) (combine l l')

a = x \/ In x l
right; apply IHl with l' y; auto. Qed.
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In y l'
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In y l'
A, B:Type

forall (l' : list B) (x : A) (y : B), In (x, y) (combine [] l') -> In y l'
A, B:Type
a:A
l:list A
IHl:forall (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In y l'
forall (l' : list B) (x : A) (y : B), In (x, y) (combine (a :: l) l') -> In y l'
A, B:Type
a:A
l:list A
IHl:forall (l' : list B) (x : A) (y : B), In (x, y) (combine l l') -> In y l'

forall (l' : list B) (x : A) (y : B), In (x, y) (combine (a :: l) l') -> In y l'
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In y0 l'0
b:B
l':list B
x:A
y:B
H:(a, b) = (x, y) \/ In (x, y) (combine l l')

b = y \/ In y l'
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In y0 l'0
b:B
l':list B
x:A
y:B
H:(a, b) = (x, y)

b = y \/ In y l'
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In y0 l'0
b:B
l':list B
x:A
y:B
H:In (x, y) (combine l l')
b = y \/ In y l'
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (x0 : A) (y0 : B), In (x0, y0) (combine l l'0) -> In y0 l'0
b:B
l':list B
x:A
y:B
H:In (x, y) (combine l l')

b = y \/ In y l'
right; apply IHl with x; auto. Qed.
A, B:Type

forall (l : list A) (l' : list B), length (combine l l') = Init.Nat.min (length l) (length l')
A, B:Type

forall (l : list A) (l' : list B), length (combine l l') = Init.Nat.min (length l) (length l')
A, B:Type

forall l' : list B, length (combine [] l') = Init.Nat.min (length []) (length l')
A, B:Type
a:A
l:list A
IHl:forall l' : list B, length (combine l l') = Init.Nat.min (length l) (length l')
forall l' : list B, length (combine (a :: l) l') = Init.Nat.min (length (a :: l)) (length l')
A, B:Type
a:A
l:list A
IHl:forall l' : list B, length (combine l l') = Init.Nat.min (length l) (length l')

forall l' : list B, length (combine (a :: l) l') = Init.Nat.min (length (a :: l)) (length l')
destruct l'; simpl; auto. Qed.
A, B:Type

forall (l : list A) (l' : list B) (n : nat) (x : A) (y : B), length l = length l' -> nth n (combine l l') (x, y) = (nth n l x, nth n l' y)
A, B:Type

forall (l : list A) (l' : list B) (n : nat) (x : A) (y : B), length l = length l' -> nth n (combine l l') (x, y) = (nth n l x, nth n l' y)
A, B:Type
n:nat
x:A
y:B
H:length [] = length []

nth n (combine [] []) (x, y) = (nth n [] x, nth n [] y)
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (n0 : nat) (x0 : A) (y0 : B), length l = length l'0 -> nth n0 (combine l l'0) (x0, y0) = (nth n0 l x0, nth n0 l'0 y0)
b:B
l':list B
n:nat
x:A
y:B
H:length (a :: l) = length (b :: l')
nth n (combine (a :: l) (b :: l')) (x, y) = (nth n (a :: l) x, nth n (b :: l') y)
A, B:Type
a:A
l:list A
IHl:forall (l'0 : list B) (n0 : nat) (x0 : A) (y0 : B), length l = length l'0 -> nth n0 (combine l l'0) (x0, y0) = (nth n0 l x0, nth n0 l'0 y0)
b:B
l':list B
n:nat
x:A
y:B
H:length (a :: l) = length (b :: l')

nth n (combine (a :: l) (b :: l')) (x, y) = (nth n (a :: l) x, nth n (b :: l') y)
destruct n; simpl in *; auto. Qed.
list_prod has the same signature as combine, but unlike combine, it adds every possible pairs, not only those at the same position.
    Fixpoint list_prod (l:list A) (l':list B) :
      list (A * B) :=
      match l with
	| nil => nil
	| cons x t => (map (fun y:B => (x, y)) l')++(list_prod t l')
      end.

    
A, B:Type

forall (x : A) (y : B) (l : list B), In y l -> In (x, y) (map (fun y0 : B => (x, y0)) l)
A, B:Type

forall (x : A) (y : B) (l : list B), In y l -> In (x, y) (map (fun y0 : B => (x, y0)) l)
induction l; [ simpl; auto | simpl; destruct 1 as [H1| ]; [ left; rewrite H1; trivial | right; auto ] ]. Qed.
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In x l -> In y l' -> In (x, y) (list_prod l l')
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In x l -> In y l' -> In (x, y) (list_prod l l')
induction l; [ simpl; tauto | simpl; intros; apply in_or_app; destruct H; [ left; rewrite H; apply in_prod_aux; assumption | right; auto ] ]. Qed.
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In (x, y) (list_prod l l') <-> In x l /\ In y l'
A, B:Type

forall (l : list A) (l' : list B) (x : A) (y : B), In (x, y) (list_prod l l') <-> In x l /\ In y l'
A, B:Type
l:list A
l':list B
x:A
y:B

In (x, y) (list_prod l l') -> In x l /\ In y l'
A, B:Type
l':list B
x:A
y:B
H:False

False /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H:In (x, y) (map (fun y0 : B => (a, y0)) l' ++ list_prod l l')
(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H:In (x, y) (map (fun y0 : B => (a, y0)) l' ++ list_prod l l')

(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H0:In (x, y) (map (fun y0 : B => (a, y0)) l')

(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H0:In (x, y) (list_prod l l')
(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H0:In (x, y) (map (fun y0 : B => (a, y0)) l')
H1:In (x, y) (map (fun y0 : B => (a, y0)) l') -> exists x0 : B, (a, x0) = (x, y) /\ In x0 l'

(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H0:In (x, y) (list_prod l l')
(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
z:B
H2:(a, z) = (x, y)
H3:In z l'

(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H0:In (x, y) (list_prod l l')
(a = x \/ In x l) /\ In y l'
A, B:Type
a:A
l:list A
l':list B
x:A
y:B
IHl:In (x, y) (list_prod l l') -> In x l /\ In y l'
H0:In (x, y) (list_prod l l')

(a = x \/ In x l) /\ In y l'
intuition. Qed.
A, B:Type

forall (l : list A) (l' : list B), length (list_prod l l') = length l * length l'
A, B:Type

forall (l : list A) (l' : list B), length (list_prod l l') = length l * length l'
A, B:Type
a:A
l:list A
IHl:forall l' : list B, length (list_prod l l') = length l * length l'

forall l' : list B, length (map (fun y : B => (a, y)) l' ++ list_prod l l') = length l' + length l * length l'
A, B:Type
a:A
l:list A
IHl:forall l'0 : list B, length (list_prod l l'0) = length l * length l'0
l':list B

length (map (fun y : B => (a, y)) l' ++ list_prod l l') = length l' + length l * length l'
A, B:Type
a:A
l:list A
IHl:forall l'0 : list B, length (list_prod l l'0) = length l * length l'0
l':list B

length (map (fun y : B => (a, y)) l') + length (list_prod l l') = length l' + length l * length l'
A, B:Type
a:A
l:list A
IHl:forall l'0 : list B, length (list_prod l l'0) = length l * length l'0
l':list B

length l' + length (list_prod l l') = length l' + length l * length l'
auto. Qed. End ListPairs. (*****************************************)

Miscellaneous operations on lists

(*****************************************)



(******************************)

Length order of lists

(******************************)

Section length_order.
  Variable A : Type.

  Definition lel (l m:list A) := length l <= length m.

  Variables a b : A.
  Variables l m n : list A.

  
A:Type
a, b:A
l, m, n:list A

lel l l
A:Type
a, b:A
l, m, n:list A

lel l l
unfold lel; auto with arith. Qed.
A:Type
a, b:A
l, m, n:list A

lel l m -> lel m n -> lel l n
A:Type
a, b:A
l, m, n:list A

lel l m -> lel m n -> lel l n
A:Type
a, b:A
l, m, n:list A
H:length l <= length m
H0:length m <= length n

length l <= length n
A:Type
a, b:A
l, m, n:list A
H:length l <= length m
H0:length m <= length n

length l <= length n
apply le_trans with (length m); auto with arith. Qed.
A:Type
a, b:A
l, m, n:list A

lel l m -> lel (a :: l) (b :: m)
A:Type
a, b:A
l, m, n:list A

lel l m -> lel (a :: l) (b :: m)
unfold lel; simpl; auto with arith. Qed.
A:Type
a, b:A
l, m, n:list A

lel l m -> lel l (b :: m)
A:Type
a, b:A
l, m, n:list A

lel l m -> lel l (b :: m)
unfold lel; simpl; auto with arith. Qed.
A:Type
a, b:A
l, m, n:list A

lel (a :: l) (b :: m) -> lel l m
A:Type
a, b:A
l, m, n:list A

lel (a :: l) (b :: m) -> lel l m
unfold lel; simpl; auto with arith. Qed.
A:Type
a, b:A
l, m, n:list A

forall l' : list A, lel l' [] -> [] = l'
A:Type
a, b:A
l, m, n:list A

forall l' : list A, lel l' [] -> [] = l'
A:Type
a, b:A
l, m, n, l':list A

forall (a0 : A) (l0 : list A), (lel l0 [] -> [] = l0) -> lel (a0 :: l0) [] -> [] = a0 :: l0
A:Type
a, b:A
l, m, n, l':list A
a':A
y:list A
H:lel y [] -> [] = y
H0:lel (a' :: y) []

[] = a' :: y
A:Type
a, b:A
l, m, n, l':list A
a':A
y:list A
H:lel y [] -> [] = y
H0:lel (a' :: y) []

[] = a' :: y
absurd (S (length y) <= 0); auto with arith. Qed. End length_order. Hint Resolve lel_refl lel_cons_cons lel_cons lel_nil lel_nil nil_cons: datatypes. (******************************)

Set inclusion on list

(******************************)

Section SetIncl.

  Variable A : Type.

  Definition incl (l m:list A) := forall a:A, In a l -> In a m.
  Hint Unfold incl : core.

  
A:Type

forall l : list A, incl l l
A:Type

forall l : list A, incl l l
auto. Qed. Hint Resolve incl_refl : core.
A:Type

forall (a : A) (l m : list A), incl l m -> incl l (a :: m)
A:Type

forall (a : A) (l m : list A), incl l m -> incl l (a :: m)
auto with datatypes. Qed. Hint Immediate incl_tl : core.
A:Type

forall l m n : list A, incl l m -> incl m n -> incl l n
A:Type

forall l m n : list A, incl l m -> incl m n -> incl l n
auto. Qed.
A:Type

forall l m n : list A, incl l n -> incl l (n ++ m)
A:Type

forall l m n : list A, incl l n -> incl l (n ++ m)
auto with datatypes. Qed. Hint Immediate incl_appl : core.
A:Type

forall l m n : list A, incl l n -> incl l (m ++ n)
A:Type

forall l m n : list A, incl l n -> incl l (m ++ n)
auto with datatypes. Qed. Hint Immediate incl_appr : core.
A:Type

forall (a : A) (l m : list A), In a m -> incl l m -> incl (a :: l) m
A:Type

forall (a : A) (l m : list A), In a m -> incl l m -> incl (a :: l) m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l

In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l

In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l

a = a0 -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l
In a0 l -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l

a = a0 -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l
In a0 l -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l
H2:a = a0

a = a0 -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l
In a0 l -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l
H2:a = a0

a = a0 -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l
In a0 l -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l

In a0 l -> In a0 m
A:Type
a:A
l, m:list A
H:In a m
H0:forall a1 : A, In a1 l -> In a1 m
a0:A
H1:a = a0 \/ In a0 l

In a0 l -> In a0 m
auto. Qed. Hint Resolve incl_cons : core.
A:Type

forall l m n : list A, incl l n -> incl m n -> incl (l ++ m) n
A:Type

forall l m n : list A, incl l n -> incl m n -> incl (l ++ m) n
A:Type
l, m, n:list A
H:forall a0 : A, In a0 l -> In a0 n
H0:forall a0 : A, In a0 m -> In a0 n
a:A
H1:In a (l ++ m)

In a n
A:Type
l, m, n:list A
H:forall a0 : A, In a0 l -> In a0 n
H0:forall a0 : A, In a0 m -> In a0 n
a:A
H1:In a (l ++ m)

In a n
elim (in_app_or _ _ _ H1); auto. Qed. Hint Resolve incl_app : core. End SetIncl. Hint Resolve incl_refl incl_tl incl_tran incl_appl incl_appr incl_cons incl_app: datatypes. (**************************************)

Cutting a list at some position

(**************************************)

Section Cutting.

  Variable A : Type.

  Fixpoint firstn (n:nat)(l:list A) : list A :=
    match n with
      | 0 => nil
      | S n => match l with
		 | nil => nil
		 | a::l => a::(firstn n l)
	       end
    end.

  
A:Type
n:nat

firstn n [] = []
A:Type
n:nat

firstn n [] = []
induction n; now simpl. Qed.
A:Type
n:nat
a:A
l:list A

firstn (S n) (a :: l) = a :: firstn n l
A:Type
n:nat
a:A
l:list A

firstn (S n) (a :: l) = a :: firstn n l
now simpl. Qed.
A:Type
l:list A

firstn (length l) l = l
A:Type
l:list A

firstn (length l) l = l
induction l as [| ? ? H]; simpl; [reflexivity | now rewrite H]. Qed.
A:Type
n:nat

forall l : list A, length l <= n -> firstn n l = l
A:Type
n:nat

forall l : list A, length l <= n -> firstn n l = l
A:Type

forall l : list A, length l <= 0 -> firstn 0 l = l
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
forall l : list A, length l <= S k -> firstn (S k) l = l
A:Type

forall l : list A, length l <= 0 -> firstn 0 l = l
A:Type
l:list A

length l <= 0 -> firstn 0 l = l
A:Type
l:list A
H:length l <= 0
H1:length l = 0

firstn (length l) l = l
A:Type
l:list A
H:length l <= 0
H1:l = []

firstn (length l) l = l
A:Type
H:length [] <= 0

firstn (length []) [] = []
now simpl.
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l

forall l : list A, length l <= S k -> firstn (S k) l = l
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l

0 <= S k -> [] = []
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
x:A
xs:list A
S (length xs) <= S k -> x :: firstn k xs = x :: xs
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l

0 <= S k -> [] = []
now reflexivity.
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
x:A
xs:list A

S (length xs) <= S k -> x :: firstn k xs = x :: xs
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
x:A
xs:list A

S (length xs) <= S k -> x :: firstn k xs = x :: xs
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
x:A
xs:list A
H:S (length xs) <= S k

x :: firstn k xs = x :: xs
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
x:A
xs:list A
H:length xs <= k

x :: firstn k xs = x :: xs
A:Type
k:nat
iHk:forall l : list A, length l <= k -> firstn k l = l
x:A
xs:list A
H:length xs <= k

firstn k xs = xs
apply iHk, H. Qed.
A:Type
l:list A

firstn 0 l = []
A:Type
l:list A

firstn 0 l = []
now simpl. Qed.
A:Type
n:nat

forall l : list A, length (firstn n l) <= n
A:Type
n:nat

forall l : list A, length (firstn n l) <= n
A:Type
k:nat
iHk:forall l : list A, length (firstn k l) <= k

0 <= S k
A:Type
k:nat
iHk:forall l : list A, length (firstn k l) <= k
x:A
xs:list A
S (length (firstn k xs)) <= S k
A:Type
k:nat
iHk:forall l : list A, length (firstn k l) <= k

0 <= S k
auto with arith.
A:Type
k:nat
iHk:forall l : list A, length (firstn k l) <= k
x:A
xs:list A

S (length (firstn k xs)) <= S k
apply Peano.le_n_S, iHk. Qed.
A:Type

forall (l : list A) (n : nat), n <= length l -> length (firstn n l) = n
A:Type

forall (l : list A) (n : nat), n <= length l -> length (firstn n l) = n
A:Type

forall n : nat, n <= length [] -> length (firstn n []) = n
A:Type
x:A
xs:list A
Hrec:forall n : nat, n <= length xs -> length (firstn n xs) = n
forall n : nat, n <= length (x :: xs) -> length (firstn n (x :: xs)) = n
A:Type

forall n : nat, n <= length [] -> length (firstn n []) = n
A:Type

forall n : nat, n <= 0 -> length (firstn n []) = n
A:Type
n:nat
H:n <= 0

length (firstn n []) = n
A:Type
n:nat
H:0 = n

length (firstn n []) = n
A:Type
n:nat
H:0 = n

length (firstn 0 []) = 0
now simpl.
A:Type
x:A
xs:list A
Hrec:forall n : nat, n <= length xs -> length (firstn n xs) = n

forall n : nat, n <= length (x :: xs) -> length (firstn n (x :: xs)) = n
A:Type
x:A
xs:list A
Hrec:forall n : nat, n <= length xs -> length (firstn n xs) = n

0 <= length (x :: xs) -> length (firstn 0 (x :: xs)) = 0
A:Type
x:A
xs:list A
Hrec:forall n0 : nat, n0 <= length xs -> length (firstn n0 xs) = n0
n:nat
S n <= length (x :: xs) -> length (firstn (S n) (x :: xs)) = S n
A:Type
x:A
xs:list A
Hrec:forall n : nat, n <= length xs -> length (firstn n xs) = n

0 <= length (x :: xs) -> length (firstn 0 (x :: xs)) = 0
now simpl.
A:Type
x:A
xs:list A
Hrec:forall n0 : nat, n0 <= length xs -> length (firstn n0 xs) = n0
n:nat

S n <= length (x :: xs) -> length (firstn (S n) (x :: xs)) = S n
A:Type
x:A
xs:list A
Hrec:forall n0 : nat, n0 <= length xs -> length (firstn n0 xs) = n0
n:nat

S n <= S (length xs) -> S (length (firstn n xs)) = S n
A:Type
x:A
xs:list A
Hrec:forall n0 : nat, n0 <= length xs -> length (firstn n0 xs) = n0
n:nat
H:S n <= S (length xs)

S (length (firstn n xs)) = S n
A:Type
x:A
xs:list A
Hrec:forall n0 : nat, n0 <= length xs -> length (firstn n0 xs) = n0
n:nat
H:n <= length xs

S (length (firstn n xs)) = S n
now rewrite (Hrec n H). Qed.
A:Type
n:nat

forall l1 l2 : list A, firstn n (l1 ++ l2) = firstn n l1 ++ firstn (n - length l1) l2
A:Type
n:nat

forall l1 l2 : list A, firstn n (l1 ++ l2) = firstn n l1 ++ firstn (n - length l1) l2
A:Type
l1, l2:list A

firstn 0 (l1 ++ l2) = firstn 0 l1 ++ firstn (0 - length l1) l2
A:Type
k:nat
iHk:forall l0 l3 : list A, firstn k (l0 ++ l3) = firstn k l0 ++ firstn (k - length l0) l3
l1, l2:list A
firstn (S k) (l1 ++ l2) = firstn (S k) l1 ++ firstn (S k - length l1) l2
A:Type
l1, l2:list A

firstn 0 (l1 ++ l2) = firstn 0 l1 ++ firstn (0 - length l1) l2
now simpl.
A:Type
k:nat
iHk:forall l0 l3 : list A, firstn k (l0 ++ l3) = firstn k l0 ++ firstn (k - length l0) l3
l1, l2:list A

firstn (S k) (l1 ++ l2) = firstn (S k) l1 ++ firstn (S k - length l1) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
l2:list A

firstn (S k) ([] ++ l2) = firstn (S k) [] ++ firstn (S k - length []) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
x:A
xs, l2:list A
firstn (S k) ((x :: xs) ++ l2) = firstn (S k) (x :: xs) ++ firstn (S k - length (x :: xs)) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
l2:list A

firstn (S k) ([] ++ l2) = firstn (S k) [] ++ firstn (S k - length []) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
l2:list A

firstn (S k) ([] ++ l2) = [] ++ firstn (S k - 0) l2
now rewrite 2!app_nil_l, <- minus_n_O.
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
x:A
xs, l2:list A

firstn (S k) ((x :: xs) ++ l2) = firstn (S k) (x :: xs) ++ firstn (S k - length (x :: xs)) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
x:A
xs, l2:list A

firstn (S k) (x :: xs ++ l2) = firstn (S k) (x :: xs) ++ firstn (S k - length (x :: xs)) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
x:A
xs, l2:list A

x :: firstn k (xs ++ l2) = x :: firstn k xs ++ firstn (k - length xs) l2
A:Type
k:nat
iHk:forall l1 l0 : list A, firstn k (l1 ++ l0) = firstn k l1 ++ firstn (k - length l1) l0
x:A
xs, l2:list A

firstn k (xs ++ l2) = firstn k xs ++ firstn (k - length xs) l2
apply iHk. Qed.
A:Type
n:nat

forall l1 l2 : list A, firstn (length l1 + n) (l1 ++ l2) = l1 ++ firstn n l2
A:Type
n:nat

forall l1 l2 : list A, firstn (length l1 + n) (l1 ++ l2) = l1 ++ firstn n l2
A:Type
l1, l2:list A

firstn (length l1 + 0) (l1 ++ l2) = l1 ++ firstn 0 l2
A:Type
k:nat
iHk:forall l0 l3 : list A, firstn (length l0 + k) (l0 ++ l3) = l0 ++ firstn k l3
l1, l2:list A
firstn (length l1 + S k) (l1 ++ l2) = l1 ++ firstn (S k) l2
A:Type
l1, l2:list A

firstn (length l1 + 0) (l1 ++ l2) = l1 ++ firstn 0 l2
A:Type
l1, l2:list A

firstn (length l1 + 0) (l1 ++ l2) = l1 ++ []
A:Type
l1, l2:list A

firstn (length l1) (l1 ++ l2) = l1
A:Type
l1, l2:list A

firstn (length l1) l1 ++ firstn (length l1 - length l1) l2 = l1
A:Type
l1, l2:list A

firstn (length l1) l1 ++ firstn 0 l2 = l1
A:Type
l1, l2:list A

firstn (length l1) l1 ++ [] = l1
A:Type
l1, l2:list A

firstn (length l1) l1 = l1
apply firstn_all.
A:Type
k:nat
iHk:forall l0 l3 : list A, firstn (length l0 + k) (l0 ++ l3) = l0 ++ firstn k l3
l1, l2:list A

firstn (length l1 + S k) (l1 ++ l2) = l1 ++ firstn (S k) l2
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A

firstn (length l1 + S k) (l1 ++ []) = l1 ++ firstn (S k) []
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A
x:A
xs:list A
firstn (length l1 + S k) (l1 ++ x :: xs) = l1 ++ firstn (S k) (x :: xs)
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A

firstn (length l1 + S k) (l1 ++ []) = l1 ++ firstn (S k) []
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A

firstn (length l1 + S k) (l1 ++ []) = l1 ++ []
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A

firstn (length l1 + S k) l1 = l1
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A

length l1 <= length l1 + S k
auto with arith.
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A
x:A
xs:list A

firstn (length l1 + S k) (l1 ++ x :: xs) = l1 ++ firstn (S k) (x :: xs)
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A
x:A
xs:list A

firstn (length l1 + S k) l1 ++ firstn (length l1 + S k - length l1) (x :: xs) = l1 ++ firstn (S k) (x :: xs)
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A
x:A
xs:list A

length l1 + S k - length l1 = S k
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A
x:A
xs:list A
H0:length l1 + S k - length l1 = S k
firstn (length l1 + S k) l1 ++ firstn (length l1 + S k - length l1) (x :: xs) = l1 ++ firstn (S k) (x :: xs)
A:Type
k:nat
iHk:forall l0 l2 : list A, firstn (length l0 + k) (l0 ++ l2) = l0 ++ firstn k l2
l1:list A
x:A
xs:list A
H0:length l1 + S k - length l1 = S k

firstn (length l1 + S k) l1 ++ firstn (length l1 + S k - length l1) (x :: xs) = l1 ++ firstn (S k) (x :: xs)
rewrite H0, firstn_all2; [reflexivity | auto with arith]. Qed.
A:Type

forall (l : list A) (i j : nat), firstn i (firstn j l) = firstn (Init.Nat.min i j) l
A:Type

forall (l : list A) (i j : nat), firstn i (firstn j l) = firstn (Init.Nat.min i j) l
A:Type

forall i j : nat, firstn i (firstn j []) = firstn (Init.Nat.min i j) []
A:Type
x:A
xs:list A
Hl:forall i j : nat, firstn i (firstn j xs) = firstn (Init.Nat.min i j) xs
forall i j : nat, firstn i (firstn j (x :: xs)) = firstn (Init.Nat.min i j) (x :: xs)
A:Type

forall i j : nat, firstn i (firstn j []) = firstn (Init.Nat.min i j) []
A:Type
i, j:nat

firstn i (firstn j []) = firstn (Init.Nat.min i j) []
A:Type
i, j:nat

firstn i (firstn j []) = firstn (Init.Nat.min i j) []
now rewrite ?firstn_nil.
A:Type
x:A
xs:list A
Hl:forall i j : nat, firstn i (firstn j xs) = firstn (Init.Nat.min i j) xs

forall i j : nat, firstn i (firstn j (x :: xs)) = firstn (Init.Nat.min i j) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i j : nat, firstn i (firstn j xs) = firstn (Init.Nat.min i j) xs

forall j : nat, firstn 0 (firstn j (x :: xs)) = firstn (Init.Nat.min 0 j) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i0 j : nat, firstn i0 (firstn j xs) = firstn (Init.Nat.min i0 j) xs
i:nat
forall j : nat, firstn (S i) (firstn j (x :: xs)) = firstn (Init.Nat.min (S i) j) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i j : nat, firstn i (firstn j xs) = firstn (Init.Nat.min i j) xs

forall j : nat, firstn 0 (firstn j (x :: xs)) = firstn (Init.Nat.min 0 j) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i j0 : nat, firstn i (firstn j0 xs) = firstn (Init.Nat.min i j0) xs
j:nat

firstn 0 (firstn j (x :: xs)) = firstn (Init.Nat.min 0 j) (x :: xs)
now simpl.
A:Type
x:A
xs:list A
Hl:forall i0 j : nat, firstn i0 (firstn j xs) = firstn (Init.Nat.min i0 j) xs
i:nat

forall j : nat, firstn (S i) (firstn j (x :: xs)) = firstn (Init.Nat.min (S i) j) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i0 j : nat, firstn i0 (firstn j xs) = firstn (Init.Nat.min i0 j) xs
i:nat

firstn (S i) (firstn 0 (x :: xs)) = firstn (Init.Nat.min (S i) 0) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i0 j0 : nat, firstn i0 (firstn j0 xs) = firstn (Init.Nat.min i0 j0) xs
i, j:nat
firstn (S i) (firstn (S j) (x :: xs)) = firstn (Init.Nat.min (S i) (S j)) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i0 j : nat, firstn i0 (firstn j xs) = firstn (Init.Nat.min i0 j) xs
i:nat

firstn (S i) (firstn 0 (x :: xs)) = firstn (Init.Nat.min (S i) 0) (x :: xs)
now simpl.
A:Type
x:A
xs:list A
Hl:forall i0 j0 : nat, firstn i0 (firstn j0 xs) = firstn (Init.Nat.min i0 j0) xs
i, j:nat

firstn (S i) (firstn (S j) (x :: xs)) = firstn (Init.Nat.min (S i) (S j)) (x :: xs)
A:Type
x:A
xs:list A
Hl:forall i0 j0 : nat, firstn i0 (firstn j0 xs) = firstn (Init.Nat.min i0 j0) xs
i, j:nat

x :: firstn i (firstn j xs) = x :: firstn (Init.Nat.min i j) xs
A:Type
x:A
xs:list A
Hl:forall i0 j0 : nat, firstn i0 (firstn j0 xs) = firstn (Init.Nat.min i0 j0) xs
i, j:nat

firstn i (firstn j xs) = firstn (Init.Nat.min i j) xs
apply Hl. Qed. Fixpoint skipn (n:nat)(l:list A) : list A := match n with | 0 => l | S n => match l with | nil => nil | a::l => skipn n l end end.
A:Type

forall (m n : nat) (l : list A), firstn m (skipn n l) = skipn n (firstn (n + m) l)
A:Type

forall (m n : nat) (l : list A), firstn m (skipn n l) = skipn n (firstn (n + m) l)
now intros m; induction n; intros []; simpl; destruct m. Qed.
A:Type

forall (m n : nat) (l : list A), skipn m (firstn n l) = firstn (n - m) (skipn m l)
A:Type

forall (m n : nat) (l : list A), skipn m (firstn n l) = firstn (n - m) (skipn m l)
now induction m; intros [] []; simpl; rewrite ?firstn_nil. Qed.
A:Type

forall l : list A, skipn 0 l = l
A:Type

forall l : list A, skipn 0 l = l
reflexivity. Qed.
A:Type

forall n : nat, skipn n ([] : list A) = []
A:Type

forall n : nat, skipn n ([] : list A) = []
now intros []. Qed.
A:Type
n:nat
a:A
l:list A

skipn (S n) (a :: l) = skipn n l
A:Type
n:nat
a:A
l:list A

skipn (S n) (a :: l) = skipn n l
reflexivity. Qed.
A:Type

forall l : list A, skipn (length l) l = []
A:Type

forall l : list A, skipn (length l) l = []
now induction l. Qed.
A:Type
n:nat

forall l : list A, length l <= n -> skipn n l = []
A:Type
n:nat

forall l : list A, length l <= n -> skipn n l = []
A:Type
n:nat
l:list A
L:length l - n = 0

skipn n (firstn (length l) l) = []
now rewrite skipn_firstn_comm, L. Qed.
A:Type

forall (n : nat) (l : list A), firstn n l ++ skipn n l = l
A:Type

forall (n : nat) (l : list A), firstn n l ++ skipn n l = l
A:Type

forall l : list A, firstn 0 l ++ skipn 0 l = l
A:Type
n:nat
IHn:forall l : list A, firstn n l ++ skipn n l = l
forall l : list A, firstn (S n) l ++ skipn (S n) l = l
A:Type
n:nat
IHn:forall l : list A, firstn n l ++ skipn n l = l

forall l : list A, firstn (S n) l ++ skipn (S n) l = l
A:Type
n:nat
IHn:forall l0 : list A, firstn n l0 ++ skipn n l0 = l0
a:A
l:list A

a :: firstn n l ++ skipn n l = a :: l
f_equal; auto. Qed.
A:Type

forall (n : nat) (l : list A), length (firstn n l) = Init.Nat.min n (length l)
A:Type

forall (n : nat) (l : list A), length (firstn n l) = Init.Nat.min n (length l)
induction n; destruct l; simpl; auto. Qed.
A:Type
n:nat

forall l : list A, length (skipn n l) = length l - n
A:Type
n:nat

forall l : list A, length (skipn n l) = length l - n
A:Type

forall l : list A, length (skipn 0 l) = length l - 0
A:Type
n:nat
IHn:forall l : list A, length (skipn n l) = length l - n
forall l : list A, length (skipn (S n) l) = length l - S n
A:Type

forall l : list A, length (skipn 0 l) = length l - 0
intros l; simpl; rewrite Nat.sub_0_r; reflexivity.
A:Type
n:nat
IHn:forall l : list A, length (skipn n l) = length l - n

forall l : list A, length (skipn (S n) l) = length l - S n
destruct l; simpl; auto. Qed.
A:Type
l:list A

skipn (length l) l = []
A:Type
l:list A

skipn (length l) l = []
now induction l. Qed.
A:Type
n:nat

forall l1 l2 : list A, skipn n (l1 ++ l2) = skipn n l1 ++ skipn (n - length l1) l2
A:Type
n:nat

forall l1 l2 : list A, skipn n (l1 ++ l2) = skipn n l1 ++ skipn (n - length l1) l2
induction n; auto; intros [|]; simpl; auto. Qed.
A:Type

forall (x : nat) (l : list A), firstn x l = rev (skipn (length l - x) (rev l))
A:Type

forall (x : nat) (l : list A), firstn x l = rev (skipn (length l - x) (rev l))
A:Type
x:nat
l:list A

firstn x l = rev (skipn (length l - x) (rev (firstn x l ++ skipn x l)))
A:Type
x:nat
l:list A

firstn x l = firstn x l ++ rev (skipn (length l - x) (rev (skipn x l)))
A:Type
x:nat
l:list A

length (rev (skipn (length l - x) (rev (skipn x l)))) = 0
repeat rewrite rev_length, skipn_length; apply Nat.sub_diag. Qed.
A:Type

forall (x : nat) (l : list A), firstn x (rev l) = rev (skipn (length l - x) l)
A:Type

forall (x : nat) (l : list A), firstn x (rev l) = rev (skipn (length l - x) l)
now intros x l; rewrite firstn_skipn_rev, rev_involutive, rev_length. Qed.
A:Type

forall (x : nat) (l : list A), skipn x (rev l) = rev (firstn (length l - x) l)
A:Type

forall (x : nat) (l : list A), skipn x (rev l) = rev (firstn (length l - x) l)
A:Type
x:nat
l:list A

skipn x (rev l) = skipn (length (rev l) - (length (rev l) - x)) (rev l)
A:Type
x:nat
l:list A
L:length (rev l) <= x

skipn x (rev l) = skipn (length (rev l) - (length (rev l) - x)) (rev l)
A:Type
x:nat
l:list A
L:x <= length (rev l)
skipn x (rev l) = skipn (length (rev l) - (length (rev l) - x)) (rev l)
A:Type
x:nat
l:list A
L:length (rev l) <= x

skipn x (rev l) = skipn (length (rev l) - (length (rev l) - x)) (rev l)
A:Type
x:nat
l:list A
L:length (rev l) - x = 0

[] = skipn (length (rev l) - (length (rev l) - x)) (rev l)
now rewrite L, Nat.sub_0_r, skipn_none.
A:Type
x:nat
l:list A
L:x <= length (rev l)

skipn x (rev l) = skipn (length (rev l) - (length (rev l) - x)) (rev l)
A:Type
x:nat
l:list A
L:x <= length (rev l)

skipn x (rev l) = skipn (length (rev l) + x - length (rev l)) (rev l)
A:Type
x:nat
l:list A
L:x <= length (rev l)
length (rev l) + x - length (rev l) = length (rev l) - (length (rev l) - x)
A:Type
x:nat
l:list A
L:x <= length (rev l)

skipn x (rev l) = skipn x (rev l)
A:Type
x:nat
l:list A
L:x <= length (rev l)
length (rev l) + x - length (rev l) = length (rev l) - (length (rev l) - x)
A:Type
x:nat
l:list A
L:x <= length (rev l)

length (rev l) + x - length (rev l) = length (rev l) - (length (rev l) - x)
A:Type
x:nat
l:list A
L:x <= length (rev l)

length (rev l) + x - (length (rev l) - x + x) = length (rev l) - (length (rev l) - x)
now rewrite <-!(Nat.add_comm x), <-minus_plus_simpl_l_reverse. Qed.
A:Type

forall (n : nat) (l : list A), n < length l -> removelast (firstn (S n) l) = firstn n l
A:Type

forall (n : nat) (l : list A), n < length l -> removelast (firstn (S n) l) = firstn n l
A:Type

0 < length [] -> removelast (firstn 1 []) = firstn 0 []
A:Type
a:A
l:list A
0 < length (a :: l) -> removelast (firstn 1 (a :: l)) = firstn 0 (a :: l)
A:Type
n:nat
IHn:forall l : list A, n < length l -> removelast (firstn (S n) l) = firstn n l
S n < length [] -> removelast (firstn (S (S n)) []) = firstn (S n) []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
S n < length (a :: l) -> removelast (firstn (S (S n)) (a :: l)) = firstn (S n) (a :: l)
A:Type
a:A
l:list A

0 < length (a :: l) -> removelast (firstn 1 (a :: l)) = firstn 0 (a :: l)
A:Type
n:nat
IHn:forall l : list A, n < length l -> removelast (firstn (S n) l) = firstn n l
S n < length [] -> removelast (firstn (S (S n)) []) = firstn (S n) []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
S n < length (a :: l) -> removelast (firstn (S (S n)) (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l : list A, n < length l -> removelast (firstn (S n) l) = firstn n l

S n < length [] -> removelast (firstn (S (S n)) []) = firstn (S n) []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
S n < length (a :: l) -> removelast (firstn (S (S n)) (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A

S n < length (a :: l) -> removelast (firstn (S (S n)) (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < length (a :: l)

removelast (firstn (S (S n)) (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

removelast (firstn (S (S n)) (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

removelast ([a] ++ firstn (S n) l) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

removelast ([a] ++ firstn (S n) l) = a :: firstn n l
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

[a] ++ removelast (firstn (S n) l) = a :: firstn n l
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)
firstn (S n) l <> []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> removelast (firstn (S n) l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

firstn (S n) l <> []
A:Type
n:nat
a:A
H:S n < 1

[] <> []
A:Type
n:nat
a:A
H0:S (S n) <= 0

[] <> []
inversion_clear H0. Qed.
A:Type

forall (n : nat) (l : list A), n < length l -> firstn n (removelast l) = firstn n l
A:Type

forall (n : nat) (l : list A), n < length l -> firstn n (removelast l) = firstn n l
A:Type

0 < length [] -> firstn 0 (removelast []) = firstn 0 []
A:Type
a:A
l:list A
0 < length (a :: l) -> firstn 0 (removelast (a :: l)) = firstn 0 (a :: l)
A:Type
n:nat
IHn:forall l : list A, n < length l -> firstn n (removelast l) = firstn n l
S n < length [] -> firstn (S n) (removelast []) = firstn (S n) []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
S n < length (a :: l) -> firstn (S n) (removelast (a :: l)) = firstn (S n) (a :: l)
A:Type
a:A
l:list A

0 < length (a :: l) -> firstn 0 (removelast (a :: l)) = firstn 0 (a :: l)
A:Type
n:nat
IHn:forall l : list A, n < length l -> firstn n (removelast l) = firstn n l
S n < length [] -> firstn (S n) (removelast []) = firstn (S n) []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
S n < length (a :: l) -> firstn (S n) (removelast (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l : list A, n < length l -> firstn n (removelast l) = firstn n l

S n < length [] -> firstn (S n) (removelast []) = firstn (S n) []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
S n < length (a :: l) -> firstn (S n) (removelast (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A

S n < length (a :: l) -> firstn (S n) (removelast (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
H:S n < length (a :: l)

firstn (S n) (removelast (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

firstn (S n) (removelast (a :: l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

firstn (S n) (removelast ([a] ++ l)) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

firstn (S n) ([a] ++ removelast l) = firstn (S n) (a :: l)
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)
l <> []
A:Type
n:nat
IHn:forall l0 : list A, n < length l0 -> firstn n (removelast l0) = firstn n l0
a:A
l:list A
H:S n < S (length l)

l <> []
intro H0; rewrite H0 in H; inversion_clear H; inversion_clear H1. Qed. End Cutting. (**********************************************************************)

Predicate for List addition/removal (no need for decidability)

(**********************************************************************)

Section Add.

  Variable A : Type.

  (* [Add a l l'] means that [l'] is exactly [l], with [a] added
     once somewhere *)
  Inductive Add (a:A) : list A -> list A -> Prop :=
    | Add_head l : Add a l (a::l)
    | Add_cons x l l' : Add a l l' -> Add a (x::l) (x::l').

  
A:Type
a:A
l1, l2:list A

Add a (l1 ++ l2) (l1 ++ a :: l2)
A:Type
a:A
l1, l2:list A

Add a (l1 ++ l2) (l1 ++ a :: l2)
induction l1; simpl; now constructor. Qed.
A:Type
a:A
l, l':list A

Add a l l' -> exists l1 l2 : list A, l = l1 ++ l2 /\ l' = l1 ++ a :: l2
A:Type
a:A
l, l':list A

Add a l l' -> exists l1 l2 : list A, l = l1 ++ l2 /\ l' = l1 ++ a :: l2
A:Type
a:A
l:list A

exists l1 l2 : list A, l = l1 ++ l2 /\ a :: l = l1 ++ a :: l2
A:Type
a, x:A
l, l':list A
H:Add a l l'
IHAdd:exists l1 l2 : list A, l = l1 ++ l2 /\ l' = l1 ++ a :: l2
exists l1 l2 : list A, x :: l = l1 ++ l2 /\ x :: l' = l1 ++ a :: l2
A:Type
a:A
l:list A

exists l1 l2 : list A, l = l1 ++ l2 /\ a :: l = l1 ++ a :: l2
exists nil; exists l; split; trivial.
A:Type
a, x:A
l, l':list A
H:Add a l l'
IHAdd:exists l1 l2 : list A, l = l1 ++ l2 /\ l' = l1 ++ a :: l2

exists l1 l2 : list A, x :: l = l1 ++ l2 /\ x :: l' = l1 ++ a :: l2
A:Type
a, x:A
l, l':list A
H:Add a l l'
l1, l2:list A
Hl:l = l1 ++ l2
Hl':l' = l1 ++ a :: l2

exists l0 l3 : list A, x :: l = l0 ++ l3 /\ x :: l' = l0 ++ a :: l3
exists (x::l1); exists l2; split; simpl; f_equal; trivial. Qed.
A:Type
a:A
l, l':list A

Add a l l' -> forall x : A, In x l' <-> In x (a :: l)
A:Type
a:A
l, l':list A

Add a l l' -> forall x : A, In x l' <-> In x (a :: l)
induction 1; intros; simpl in *; rewrite ?IHAdd; tauto. Qed.
A:Type
a:A
l, l':list A

Add a l l' -> length l' = S (length l)
A:Type
a:A
l, l':list A

Add a l l' -> length l' = S (length l)
induction 1; simpl; auto with arith. Qed.
A:Type
a:A
l:list A

In a l -> exists l' : list A, Add a l' l
A:Type
a:A
l:list A

In a l -> exists l' : list A, Add a l' l
A:Type
a:A
l:list A
Ha:In a l

exists l' : list A, Add a l' l
A:Type
a:A
l1, l2:list A
Ha:In a (l1 ++ a :: l2)

exists l' : list A, Add a l' (l1 ++ a :: l2)
A:Type
a:A
l1, l2:list A
Ha:In a (l1 ++ a :: l2)

Add a (l1 ++ l2) (l1 ++ a :: l2)
apply Add_app. Qed.
A:Type
a:A
l, u, v:list A

~ In a l -> incl (a :: l) v -> Add a u v -> incl l u
A:Type
a:A
l, u, v:list A

~ In a l -> incl (a :: l) v -> Add a u v -> incl l u
A:Type
a:A
l, u, v:list A
Ha:~ In a l
H:incl (a :: l) v
AD:Add a u v
y:A
Hy:In y l

In y u
A:Type
a:A
l, u, v:list A
Ha:~ In a l
H:incl (a :: l) v
AD:Add a u v
y:A
Hy:In y l

In y (a :: u)
A:Type
a:A
l, u, v:list A
Ha:~ In a l
H:incl (a :: l) v
AD:Add a u v
y:A
Hy:In y l
Hy':In y (a :: u)
In y u
A:Type
a:A
l, u, v:list A
Ha:~ In a l
H:incl (a :: l) v
AD:Add a u v
y:A
Hy:In y l

In y (a :: u)
A:Type
a:A
l, u, v:list A
Ha:~ In a l
H:incl (a :: l) v
AD:Add a u v
y:A
Hy:In y l

In y v
apply H; simpl; auto.
A:Type
a:A
l, u, v:list A
Ha:~ In a l
H:incl (a :: l) v
AD:Add a u v
y:A
Hy:In y l
Hy':In y (a :: u)

In y u
destruct Hy'; [ subst; now elim Ha | trivial ]. Qed. End Add. (********************************)

Lists without redundancy

(********************************)

Section ReDun.

  Variable A : Type.

  Inductive NoDup : list A -> Prop :=
    | NoDup_nil : NoDup nil
    | NoDup_cons : forall x l, ~ In x l -> NoDup l -> NoDup (x::l).

  
A:Type
a:A
l, l':list A

Add a l l' -> NoDup l' <-> NoDup l /\ ~ In a l
A:Type
a:A
l, l':list A

Add a l l' -> NoDup l' <-> NoDup l /\ ~ In a l
A:Type
a:A
l:list A

NoDup (a :: l) <-> NoDup l /\ ~ In a l
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
NoDup (x :: l') <-> NoDup (x :: l) /\ ~ In a (x :: l)
A:Type
a:A
l:list A

NoDup (a :: l) <-> NoDup l /\ ~ In a l
split; [ inversion_clear 1; now split | now constructor ].
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l

NoDup (x :: l') <-> NoDup (x :: l) /\ ~ In a (x :: l)
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l

NoDup (x :: l') -> NoDup (x :: l) /\ ~ In a (x :: l)
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
NoDup (x :: l) /\ ~ In a (x :: l) -> NoDup (x :: l')
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l

NoDup (x :: l') -> NoDup (x :: l) /\ ~ In a (x :: l)
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
H0:~ In x l'
H1:NoDup l'

NoDup (x :: l) /\ ~ In a (x :: l)
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
H0:~ In x l'
H1:NoDup l /\ ~ In a l

NoDup (x :: l) /\ ~ In a (x :: l)
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
H0:~ In x (a :: l)
H1:NoDup l /\ ~ In a l

NoDup (x :: l) /\ ~ In a (x :: l)
simpl in *; split; try constructor; intuition.
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l

NoDup (x :: l) /\ ~ In a (x :: l) -> NoDup (x :: l')
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
N:NoDup (x :: l)
IN:~ In a (x :: l)

NoDup (x :: l')
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l

NoDup (x :: l')
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l

~ In x l'
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l
NoDup l'
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l

~ In x l'
rewrite (Add_in AD); simpl in *; intuition.
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l

NoDup l'
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l

NoDup l /\ ~ In a l
A:Type
a, x:A
l, l':list A
AD:Add a l l'
IH:NoDup l' <-> NoDup l /\ ~ In a l
IN:~ In a (x :: l)
H:~ In x l
H0:NoDup l

~ In a l
simpl in *; intuition. Qed.
A:Type
l, l':list A
a:A

NoDup (l ++ a :: l') -> NoDup (l ++ l') /\ ~ In a (l ++ l')
A:Type
l, l':list A
a:A

NoDup (l ++ a :: l') -> NoDup (l ++ l') /\ ~ In a (l ++ l')
A:Type
l, l':list A
a:A

Add a (l ++ l') (l ++ a :: l')
apply Add_app. Qed.
A:Type
l, l':list A
a:A

NoDup (l ++ a :: l') -> NoDup (l ++ l')
A:Type
l, l':list A
a:A

NoDup (l ++ a :: l') -> NoDup (l ++ l')
A:Type
l, l':list A
a:A
H:NoDup (l ++ a :: l')

NoDup (l ++ l')
now apply NoDup_remove with a. Qed.
A:Type
l, l':list A
a:A

NoDup (l ++ a :: l') -> ~ In a (l ++ l')
A:Type
l, l':list A
a:A

NoDup (l ++ a :: l') -> ~ In a (l ++ l')
A:Type
l, l':list A
a:A
H:NoDup (l ++ a :: l')

~ In a (l ++ l')
now apply NoDup_remove. Qed.
A:Type
a:A
l:list A

NoDup (a :: l) <-> ~ In a l /\ NoDup l
A:Type
a:A
l:list A

NoDup (a :: l) <-> ~ In a l /\ NoDup l
A:Type
a:A
l:list A

NoDup (a :: l) -> ~ In a l /\ NoDup l
A:Type
a:A
l:list A
~ In a l /\ NoDup l -> NoDup (a :: l)
A:Type
a:A
l:list A

NoDup (a :: l) -> ~ In a l /\ NoDup l
A:Type
a:A
l:list A
H0:~ In a l
H1:NoDup l

~ In a l /\ NoDup l
now split.
A:Type
a:A
l:list A

~ In a l /\ NoDup l -> NoDup (a :: l)
now constructor. Qed.
Effective computation of a list without duplicates
  Hypothesis decA: forall x y : A, {x = y} + {x <> y}.

  Fixpoint nodup (l : list A) : list A :=
    match l with
      | [] => []
      | x::xs => if in_dec decA x xs then nodup xs else x::(nodup xs)
    end.

  
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

In x (nodup l) <-> In x l
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A

In x (nodup l) <-> In x l
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A

False <-> False
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
Hrec:In x (nodup l') <-> In x l'
In x (if in_dec decA a l' then nodup l' else a :: nodup l') <-> a = x \/ In x l'
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
x:A

False <-> False
reflexivity.
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
Hrec:In x (nodup l') <-> In x l'

In x (if in_dec decA a l' then nodup l' else a :: nodup l') <-> a = x \/ In x l'
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
Hrec:In x (nodup l') <-> In x l'
i:In a l'

In x l' <-> a = x \/ In x l'
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
Hrec:In x (nodup l') <-> In x l'
n:~ In a l'
a = x \/ In x l' <-> a = x \/ In x l'
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
Hrec:In x (nodup l') <-> In x l'
i:In a l'

In x l' <-> a = x \/ In x l'
intuition; now subst.
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
Hrec:In x (nodup l') <-> In x l'
n:~ In a l'

a = x \/ In x l' <-> a = x \/ In x l'
reflexivity. Qed.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup (nodup l)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup (nodup l)
A:Type
decA:forall x y : A, {x = y} + {x <> y}

NoDup []
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup (nodup l')
NoDup (if in_dec decA a l' then nodup l' else a :: nodup l')
A:Type
decA:forall x y : A, {x = y} + {x <> y}

NoDup []
constructor.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup (nodup l')

NoDup (if in_dec decA a l' then nodup l' else a :: nodup l')
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup (nodup l')
i:In a l'

NoDup (nodup l')
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup (nodup l')
n:~ In a l'
NoDup (a :: nodup l')
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup (nodup l')
i:In a l'

NoDup (nodup l')
assumption.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup (nodup l')
n:~ In a l'

NoDup (a :: nodup l')
constructor; [ now rewrite nodup_In | assumption]. Qed.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A

nodup k = a :: l -> ~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A

nodup k = a :: l -> ~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A
H:nodup k = a :: l

~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A
H:nodup k = a :: l

NoDup (a :: l)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A
H:nodup k = a :: l
H':NoDup (a :: l)
~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A
H:nodup k = a :: l

NoDup (a :: l)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A
H:nodup k = a :: l

NoDup (nodup k)
apply NoDup_nodup.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
k, l:list A
a:A
H:nodup k = a :: l
H':NoDup (a :: l)

~ In a l
now inversion_clear H'. Qed.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l <-> (forall x : A, count_occ decA l x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l <-> (forall x : A, count_occ decA l x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}

NoDup [] <-> (forall x : A, count_occ decA [] x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup l' <-> (forall x : A, count_occ decA l' x <= 1)
NoDup (a :: l') <-> (forall x : A, count_occ decA (a :: l') x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}

NoDup [] <-> (forall x : A, count_occ decA [] x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}

(A -> 0 <= 1) -> NoDup []
constructor.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup l' <-> (forall x : A, count_occ decA l' x <= 1)

NoDup (a :: l') <-> (forall x : A, count_occ decA (a :: l') x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
Hrec:NoDup l' <-> (forall x : A, count_occ decA l' x <= 1)

count_occ decA l' a = 0 /\ (forall x : A, count_occ decA l' x <= 1) <-> (forall x : A, count_occ decA (a :: l') x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A

count_occ decA l' a = 0 /\ (forall x : A, count_occ decA l' x <= 1) <-> (forall x : A, count_occ decA (a :: l') x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A

count_occ decA l' a = 0 /\ (forall x : A, count_occ decA l' x <= 1) -> forall x : A, count_occ decA (a :: l') x <= 1
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
(forall x : A, count_occ decA (a :: l') x <= 1) -> count_occ decA l' a = 0 /\ (forall x : A, count_occ decA l' x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A

count_occ decA l' a = 0 /\ (forall x : A, count_occ decA l' x <= 1) -> forall x : A, count_occ decA (a :: l') x <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
Ha:count_occ decA l' a = 0
H:forall x0 : A, count_occ decA l' x0 <= 1
x:A

count_occ decA (a :: l') x <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
Ha:count_occ decA l' a = 0
H:forall x0 : A, count_occ decA l' x0 <= 1
x:A

(if decA a x then S (count_occ decA l' x) else count_occ decA l' x) <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
Ha:count_occ decA l' a = 0
H:forall x0 : A, count_occ decA l' x0 <= 1
x:A
e:a = x

S (count_occ decA l' x) <= 1
subst; now rewrite Ha.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A

(forall x : A, count_occ decA (a :: l') x <= 1) -> count_occ decA l' a = 0 /\ (forall x : A, count_occ decA l' x <= 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
H:forall x : A, count_occ decA (a :: l') x <= 1

count_occ decA l' a = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
H:forall x : A, count_occ decA (a :: l') x <= 1
forall x : A, count_occ decA l' x <= 1
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
H:forall x : A, count_occ decA (a :: l') x <= 1

count_occ decA l' a = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
H:count_occ decA (a :: l') a <= 1

count_occ decA l' a = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
H:S (count_occ decA l' a) <= 1

count_occ decA l' a = 0
now inversion H.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l':list A
H:forall x : A, count_occ decA (a :: l') x <= 1

forall x : A, count_occ decA l' x <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
H:forall x0 : A, count_occ decA (a :: l') x0 <= 1
x:A

count_occ decA l' x <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
H:count_occ decA (a :: l') x <= 1

count_occ decA l' x <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
H:(if decA a x then S (count_occ decA l' x) else count_occ decA l' x) <= 1

count_occ decA l' x <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l':list A
x:A
e:a = x
H:S (count_occ decA l' x) <= 1

count_occ decA l' x <= 1
now apply Nat.lt_le_incl. Qed.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l <-> (forall x : A, In x l -> count_occ decA l x = 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l <-> (forall x : A, In x l -> count_occ decA l x = 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ decA l x <= 1) <-> (forall x : A, In x l -> count_occ decA l x = 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ decA l x <= 1) <-> (forall x : A, count_occ decA l x > 0 -> count_occ decA l x = 1)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

(forall x : A, count_occ decA l x <= 1) <-> (forall x : A, 1 <= count_occ decA l x -> count_occ decA l x = 1)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:n <= 1

1 <= n -> n = 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:1 <= n -> n = 1
n <= 1
(* the rest would be solved by omega if we had it here... *)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:n <= 1

1 <= n -> n = 1
now apply Nat.le_antisymm.
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:1 <= n -> n = 1

n <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:1 <= n -> n = 1
H0:1 <= n

n <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:1 <= n -> n = 1
H0:n < 1
n <= 1
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:1 <= n -> n = 1
H0:1 <= n

n <= 1
rewrite H; trivial.
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
l:list A
x:A
n:nat
H:1 <= n -> n = 1
H0:n < 1

n <= 1
now apply Nat.lt_le_incl. Qed.
Alternative characterisations of being without duplicates, thanks to nth_error and nth
  
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l <-> (forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l <-> (forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l -> forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

NoDup l -> forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
i, j:nat
Hi:i < length []
E:nth_error [] i = nth_error [] j

i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> nth_error l i0 = nth_error l j0 -> i0 = j0
i, j:nat
Hi:i < length (a :: l)
E:nth_error (a :: l) i = nth_error (a :: l) j
i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
i, j:nat
Hi:i < length []
E:nth_error [] i = nth_error [] j

i = j
inversion Hi.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> nth_error l i0 = nth_error l j0 -> i0 = j0
i, j:nat
Hi:i < length (a :: l)
E:nth_error (a :: l) i = nth_error (a :: l) j

i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> nth_error l i = nth_error l j0 -> i = j0
j:nat
Hi:0 < S (length l)
E:Some a = nth_error l j

0 = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> nth_error l i0 = nth_error l j -> i0 = j
i:nat
Hi:S i < S (length l)
E:nth_error l i = Some a
S i = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> nth_error l i0 = nth_error l j0 -> i0 = j0
i, j:nat
Hi:S i < S (length l)
E:nth_error l i = nth_error l j
S i = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> nth_error l i = nth_error l j0 -> i = j0
j:nat
Hi:0 < S (length l)
E:Some a = nth_error l j

0 = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> nth_error l i = nth_error l j0 -> i = j0
j:nat
Hi:0 < S (length l)
E:Some a = nth_error l j

In a l
eapply nth_error_In; eauto.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> nth_error l i0 = nth_error l j -> i0 = j
i:nat
Hi:S i < S (length l)
E:nth_error l i = Some a

S i = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> nth_error l i0 = nth_error l j -> i0 = j
i:nat
Hi:S i < S (length l)
E:nth_error l i = Some a

In a l
eapply nth_error_In; eauto.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> nth_error l i0 = nth_error l j0 -> i0 = j0
i, j:nat
Hi:S i < S (length l)
E:nth_error l i = nth_error l j

S i = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> nth_error l i0 = nth_error l j0 -> i0 = j0
i, j:nat
Hi:S i < S (length l)
E:nth_error l i = nth_error l j

i = j
apply IH; auto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A

(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j

~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j
NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j

~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j
Ha:In a l

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j
Ha:exists n : nat, nth_error l n = Some a

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j
n:nat
Hn:nth_error l n = Some a

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j
n:nat
Hn:nth_error l n = Some a
H0:n < length l

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
n:nat
H:0 < length (a :: l) -> nth_error (a :: l) 0 = nth_error (a :: l) (S n) -> 0 = S n
Hn:nth_error l n = Some a
H0:n < length l

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
n:nat
H:0 < S (length l) -> Some a = nth_error l n -> 0 = S n
Hn:nth_error l n = Some a
H0:n < length l

False
discriminate H; auto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j

NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> nth_error (a :: l) i = nth_error (a :: l) j -> i = j

forall i j : nat, i < length l -> nth_error l i = nth_error l j -> i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
IHl:(forall i0 j0 : nat, i0 < length l -> nth_error l i0 = nth_error l j0 -> i0 = j0) -> NoDup l
H:forall i0 j0 : nat, i0 < length (a :: l) -> nth_error (a :: l) i0 = nth_error (a :: l) j0 -> i0 = j0
i, j:nat
Hi:i < length l
E:nth_error l i = nth_error l j

i = j
apply eq_add_S, H; simpl; auto with arith. } Qed.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A

NoDup l <-> (forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A

NoDup l <-> (forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j)
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A

NoDup l -> forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A
(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A

NoDup l -> forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d:A
i, j:nat
Hi:i < length []
Hj:j < length []
E:nth i [] d = nth j [] d

i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> j0 < length l -> nth i0 l d = nth j0 l d -> i0 = j0
i, j:nat
Hi:i < length (a :: l)
Hj:j < length (a :: l)
E:nth i (a :: l) d = nth j (a :: l) d
i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d:A
i, j:nat
Hi:i < length []
Hj:j < length []
E:nth i [] d = nth j [] d

i = j
inversion Hi.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> j0 < length l -> nth i0 l d = nth j0 l d -> i0 = j0
i, j:nat
Hi:i < length (a :: l)
Hj:j < length (a :: l)
E:nth i (a :: l) d = nth j (a :: l) d

i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> j0 < length l -> nth i l d = nth j0 l d -> i = j0
j:nat
Hi:0 < S (length l)
Hj:S j < S (length l)
E:a = nth j l d

0 = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> j < length l -> nth i0 l d = nth j l d -> i0 = j
i:nat
Hi:S i < S (length l)
Hj:0 < S (length l)
E:nth i l d = a
S i = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> j0 < length l -> nth i0 l d = nth j0 l d -> i0 = j0
i, j:nat
Hi:S i < S (length l)
Hj:S j < S (length l)
E:nth i l d = nth j l d
S i = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> j0 < length l -> nth i l d = nth j0 l d -> i = j0
j:nat
Hi:0 < S (length l)
Hj:S j < S (length l)
E:a = nth j l d

0 = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> j0 < length l -> nth i l d = nth j0 l d -> i = j0
j:nat
Hi:0 < S (length l)
Hj:S j < S (length l)
E:a = nth j l d

In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d:A
l:list A
j:nat
Hal:~ In (nth j l d) l
Hl:NoDup l
IH:forall i j0 : nat, i < length l -> j0 < length l -> nth i l d = nth j0 l d -> i = j0
Hi:0 < S (length l)
Hj:S j < S (length l)

In (nth j l d) l
apply nth_In; auto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> j < length l -> nth i0 l d = nth j l d -> i0 = j
i:nat
Hi:S i < S (length l)
Hj:0 < S (length l)
E:nth i l d = a

S i = 0
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> j < length l -> nth i0 l d = nth j l d -> i0 = j
i:nat
Hi:S i < S (length l)
Hj:0 < S (length l)
E:nth i l d = a

In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d:A
l:list A
i:nat
Hal:~ In (nth i l d) l
Hl:NoDup l
IH:forall i0 j : nat, i0 < length l -> j < length l -> nth i0 l d = nth j l d -> i0 = j
Hi:S i < S (length l)
Hj:0 < S (length l)

In (nth i l d) l
apply nth_In; auto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> j0 < length l -> nth i0 l d = nth j0 l d -> i0 = j0
i, j:nat
Hi:S i < S (length l)
Hj:S j < S (length l)
E:nth i l d = nth j l d

S i = S j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
d, a:A
l:list A
Hal:~ In a l
Hl:NoDup l
IH:forall i0 j0 : nat, i0 < length l -> j0 < length l -> nth i0 l d = nth j0 l d -> i0 = j0
i, j:nat
Hi:S i < S (length l)
Hj:S j < S (length l)
E:nth i l d = nth j l d

i = j
apply IH; auto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A

(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
d:A

(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j

~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j
NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j

~ In a l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j
Ha:In a l

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j
Ha:exists n : nat, n < length l /\ nth n l ?d = a

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j
n:nat
Hn:n < length l
Hn':nth n l ?d = a

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
n:nat
H:0 < length (a :: l) -> S n < length (a :: l) -> nth 0 (a :: l) d = nth (S n) (a :: l) d -> 0 = S n
Hn:n < length l
Hn':nth n l ?d = a

False
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
n:nat
H:0 < S (length l) -> S n < S (length l) -> a = nth n l d -> 0 = S n
Hn:n < length l
Hn':nth n l ?d = a

False
discriminate H; eauto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j

NoDup l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j) -> NoDup l
H:forall i j : nat, i < length (a :: l) -> j < length (a :: l) -> nth i (a :: l) d = nth j (a :: l) d -> i = j

forall i j : nat, i < length l -> j < length l -> nth i l d = nth j l d -> i = j
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
d:A
IHl:(forall i0 j0 : nat, i0 < length l -> j0 < length l -> nth i0 l d = nth j0 l d -> i0 = j0) -> NoDup l
H:forall i0 j0 : nat, i0 < length (a :: l) -> j0 < length (a :: l) -> nth i0 (a :: l) d = nth j0 (a :: l) d -> i0 = j0
i, j:nat
Hi:i < length l
Hj:j < length l
E:nth i l d = nth j l d

i = j
apply eq_add_S, H; simpl; auto with arith. } Qed.
Having NoDup hypotheses bring more precise facts about incl.
  
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l, l':list A

NoDup l -> incl l l' -> length l <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l, l':list A

NoDup l -> incl l l' -> length l <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l, l':list A
N:NoDup l

incl l l' -> length l <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
N:NoDup l

forall l' : list A, incl l l' -> length l <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}

forall l' : list A, incl [] l' -> 0 <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l' : list A, incl l l' -> length l <= length l'
forall l' : list A, incl (a :: l) l' -> S (length l) <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}

forall l' : list A, incl [] l' -> 0 <= length l'
auto with arith.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l' : list A, incl l l' -> length l <= length l'

forall l' : list A, incl (a :: l) l' -> S (length l) <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'

S (length l) <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'

In a l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'
S (length l) <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'

In a l'
apply H; simpl; auto.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

S (length l) <= length l'
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

S (length l) <= S (length l'')
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

length l <= length l''
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, incl l l'0 -> length l <= length l'0
l':list A
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

incl l l''
now apply incl_Add_inv with a l'. Qed.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l, l':list A

NoDup l -> length l' <= length l -> incl l l' -> incl l' l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l, l':list A

NoDup l -> length l' <= length l -> incl l l' -> incl l' l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l, l':list A
N:NoDup l

length l' <= length l -> incl l l' -> incl l' l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
l:list A
N:NoDup l

forall l' : list A, length l' <= length l -> incl l l' -> incl l' l
A:Type
decA:forall x y : A, {x = y} + {x <> y}

forall l' : list A, length l' <= length [] -> incl [] l' -> incl l' []
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l' : list A, length l' <= length l -> incl l l' -> incl l' l
forall l' : list A, length l' <= length (a :: l) -> incl (a :: l) l' -> incl l' (a :: l)
A:Type
decA:forall x y : A, {x = y} + {x <> y}

forall l' : list A, length l' <= length [] -> incl [] l' -> incl l' []
destruct l'; easy.
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l' : list A, length l' <= length l -> incl l l' -> incl l' l

forall l' : list A, length l' <= length (a :: l) -> incl (a :: l) l' -> incl l' (a :: l)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
Hx:In x l'

In x (a :: l)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
Hx:In x l'

In a l'
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
Hx:In x l'
l'':list A
AD:Add a l'' l'
In x (a :: l)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
Hx:In x l'

In a l'
apply H; simpl; auto.
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
Hx:In x l'
l'':list A
AD:Add a l'' l'

In x (a :: l)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
l'':list A
Hx:In x (a :: l'')
AD:Add a l'' l'

In x (a :: l)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
l'':list A
Hx:a = x \/ In x l''
AD:Add a l'' l'

In x (a :: l)
A:Type
decA:forall x0 y : A, {x0 = y} + {x0 <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
x:A
l'':list A
Hx:In x l''
AD:Add a l'' l'

In x l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

forall x : A, In x l'' -> In x l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

length l'' <= length l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'
incl l l''
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

length l'' <= length l
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

S (length l'') <= S (length l)
now rewrite <- (Add_length AD).
A:Type
decA:forall x y : A, {x = y} + {x <> y}
a:A
l:list A
Hal:~ In a l
N:NoDup l
IH:forall l'0 : list A, length l'0 <= length l -> incl l l'0 -> incl l'0 l
l':list A
E:length l' <= length (a :: l)
H:incl (a :: l) l'
l'':list A
AD:Add a l'' l'

incl l l''
now apply incl_Add_inv with a l'. Qed. End ReDun.
NoDup and map
NB: the reciprocal result holds only for injective functions, see FinFun.v
A, B:Type
f:A -> B
l:list A

NoDup (map f l) -> NoDup l
A, B:Type
f:A -> B
l:list A

NoDup (map f l) -> NoDup l
A, B:Type
f:A -> B
a:A
l:list A
IHl:NoDup (map f l) -> NoDup l
H0:~ In (f a) (map f l)
H1:NoDup (map f l)

~ In a l
A, B:Type
f:A -> B
a:A
l:list A
IHl:NoDup (map f l) -> NoDup l
H0:~ In (f a) (map f l)
H1:NoDup (map f l)
H:In a l

False
now apply (in_map f) in H. Qed. (***********************************)

Sequence of natural numbers

(***********************************)

Section NatSeq.
seq computes the sequence of len contiguous integers that starts at start. For instance, seq 2 3 is 2::3::4::nil.
  Fixpoint seq (start len:nat) : list nat :=
    match len with
      | 0 => nil
      | S len => start :: seq (S start) len
    end.

  

forall len start : nat, length (seq start len) = len

forall len start : nat, length (seq start len) = len
induction len; simpl; auto. Qed.

forall len start n d : nat, n < len -> nth n (seq start len) d = start + n

forall len start n d : nat, n < len -> nth n (seq start len) d = start + n
start, n, d:nat
H:n < 0

nth n (seq start 0) d = start + n
len:nat
IHlen:forall start0 n0 d0 : nat, n0 < len -> nth n0 (seq start0 len) d0 = start0 + n0
start, n, d:nat
H:n < S len
nth n (seq start (S len)) d = start + n
len:nat
IHlen:forall start0 n0 d0 : nat, n0 < len -> nth n0 (seq start0 len) d0 = start0 + n0
start, n, d:nat
H:n < S len

nth n (seq start (S len)) d = start + n
len:nat
IHlen:forall start0 n0 d0 : nat, n0 < len -> nth n0 (seq start0 len) d0 = start0 + n0
start, n, d:nat
H:n < S len

nth n (start :: seq (S start) len) d = start + n
len:nat
IHlen:forall start0 n d0 : nat, n < len -> nth n (seq start0 len) d0 = start0 + n
start, d:nat
H:0 < S len

start = start + 0
len:nat
IHlen:forall start0 n0 d0 : nat, n0 < len -> nth n0 (seq start0 len) d0 = start0 + n0
start, n, d:nat
H:S n < S len
nth n (seq (S start) len) d = start + S n
len:nat
IHlen:forall start0 n0 d0 : nat, n0 < len -> nth n0 (seq start0 len) d0 = start0 + n0
start, n, d:nat
H:S n < S len

nth n (seq (S start) len) d = start + S n
rewrite IHlen;simpl; auto with arith. Qed.

forall len start : nat, map S (seq start len) = seq (S start) len

forall len start : nat, map S (seq start len) = seq (S start) len
len:nat
IHlen:forall start : nat, map S (seq start len) = seq (S start) len

forall start : nat, S start :: map S (seq (S start) len) = S start :: seq (S (S start)) len
len:nat
IHlen:forall start0 : nat, map S (seq start0 len) = seq (S start0) len
start:nat

S start :: map S (seq (S start) len) = S start :: seq (S (S start)) len
len:nat
IHlen:forall start0 : nat, map S (seq start0 len) = seq (S start0) len
start:nat

S start :: seq (S (S start)) len = S start :: seq (S (S start)) len
auto with arith. Qed.
len, start, n:nat

In n (seq start len) <-> start <= n < start + len
len, start, n:nat

In n (seq start len) <-> start <= n < start + len
len, n:nat

forall start : nat, In n (seq start len) <-> start <= n < start + len
n, start:nat

False <-> start <= n < start + 0
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat
start = n \/ In n (seq (S start) len) <-> start <= n < start + S len
n, start:nat

False <-> start <= n < start + 0
n, start:nat

False <-> start <= n < start
n, start:nat

start <= n < start -> False
n, start:nat
H:start <= n
H':n < start

False
apply (Lt.lt_irrefl _ (Lt.le_lt_trans _ _ _ H H')).
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat

start = n \/ In n (seq (S start) len) <-> start <= n < start + S len
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat

start = n \/ S start <= n < S (start + len) -> start <= n < S (start + len)
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat
start <= n < S (start + len) -> start = n \/ S start <= n < S (start + len)
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat

start = n \/ S start <= n < S (start + len) -> start <= n < S (start + len)
intros [H|H]; subst; intuition auto with arith.
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat

start <= n < S (start + len) -> start = n \/ S start <= n < S (start + len)
len, n:nat
IHlen:forall start0 : nat, In n (seq start0 len) <-> start0 <= n < start0 + len
start:nat
H:start <= n
H':n < S (start + len)

start = n \/ S start <= n < S (start + len)
destruct (Lt.le_lt_or_eq _ _ H); intuition. Qed.
len, start:nat

NoDup (seq start len)
len, start:nat

NoDup (seq start len)
len:nat
IHlen:forall start0 : nat, NoDup (seq start0 len)
start:nat

~ In start (seq (S start) len)
len:nat
IHlen:forall start0 : nat, NoDup (seq start0 len)
start:nat

~ S start <= start < S start + len
len:nat
IHlen:forall start0 : nat, NoDup (seq start0 len)
start:nat
H:S start <= start

False
apply (Lt.lt_irrefl _ H). Qed.

forall len1 len2 start : nat, seq start (len1 + len2) = seq start len1 ++ seq (start + len1) len2

forall len1 len2 start : nat, seq start (len1 + len2) = seq start len1 ++ seq (start + len1) len2
len2, start:nat

seq start len2 = seq (start + 0) len2
len1':nat
IHlen:forall len0 start0 : nat, seq start0 (len1' + len0) = seq start0 len1' ++ seq (start0 + len1') len0
len2, start:nat
start :: seq (S start) (len1' + len2) = start :: seq (S start) len1' ++ seq (start + S len1') len2
len2, start:nat

seq start len2 = seq (start + 0) len2
now rewrite Nat.add_0_r.
len1':nat
IHlen:forall len0 start0 : nat, seq start0 (len1' + len0) = seq start0 len1' ++ seq (start0 + len1') len0
len2, start:nat

start :: seq (S start) (len1' + len2) = start :: seq (S start) len1' ++ seq (start + S len1') len2
now rewrite Nat.add_succ_r, IHlen. Qed. End NatSeq. Section Exists_Forall.

Existential and universal predicates over lists

  Variable A:Type.

  Section One_predicate.

    Variable P:A->Prop.

    Inductive Exists : list A -> Prop :=
      | Exists_cons_hd : forall x l, P x -> Exists (x::l)
      | Exists_cons_tl : forall x l, Exists l -> Exists (x::l).

    Hint Constructors Exists : core.

    
A:Type
P:A -> Prop
l:list A

Exists l <-> (exists x : A, In x l /\ P x)
A:Type
P:A -> Prop
l:list A

Exists l <-> (exists x : A, In x l /\ P x)
A:Type
P:A -> Prop
l:list A

Exists l -> exists x : A, In x l /\ P x
A:Type
P:A -> Prop
l:list A
(exists x : A, In x l /\ P x) -> Exists l
A:Type
P:A -> Prop
l:list A

Exists l -> exists x : A, In x l /\ P x
induction 1; firstorder.
A:Type
P:A -> Prop
l:list A

(exists x : A, In x l /\ P x) -> Exists l
induction l; firstorder; subst; auto. Qed.
A:Type
P:A -> Prop

Exists [] <-> False
A:Type
P:A -> Prop

Exists [] <-> False
split; inversion 1. Qed.
A:Type
P:A -> Prop
x:A
l:list A

Exists (x :: l) <-> P x \/ Exists l
A:Type
P:A -> Prop
x:A
l:list A

Exists (x :: l) <-> P x \/ Exists l
split; inversion 1; auto. Qed.
A:Type
P:A -> Prop
l:list A

(forall x : A, {P x} + {~ P x}) -> {Exists l} + {~ Exists l}
A:Type
P:A -> Prop
l:list A

(forall x : A, {P x} + {~ P x}) -> {Exists l} + {~ Exists l}
A:Type
P:A -> Prop
l:list A
Pdec:forall x : A, {P x} + {~ P x}

{Exists l} + {~ Exists l}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

{Exists []} + {~ Exists []}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hrec:{Exists l'} + {~ Exists l'}
{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

{Exists []} + {~ Exists []}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

~ Exists []
abstract now rewrite Exists_nil.
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hrec:{Exists l'} + {~ Exists l'}

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':Exists l'

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':Exists l'

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':Exists l'

Exists (a :: l')
now apply Exists_cons_tl.
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
Ha:P a

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
Ha:~ P a
{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
Ha:P a

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
Ha:P a

Exists (a :: l')
now apply Exists_cons_hd.
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
Ha:~ P a

{Exists (a :: l')} + {~ Exists (a :: l')}
A:Type
P:A -> Prop
a:A
l':list A
Pdec:forall x : A, {P x} + {~ P x}
Hl':~ Exists l'
Ha:~ P a

~ Exists (a :: l')
abstract now inversion 1. Defined. Inductive Forall : list A -> Prop := | Forall_nil : Forall nil | Forall_cons : forall x l, P x -> Forall l -> Forall (x::l). Hint Constructors Forall : core.
A:Type
P:A -> Prop
l:list A

Forall l <-> (forall x : A, In x l -> P x)
A:Type
P:A -> Prop
l:list A

Forall l <-> (forall x : A, In x l -> P x)
A:Type
P:A -> Prop
l:list A

Forall l -> forall x : A, In x l -> P x
A:Type
P:A -> Prop
l:list A
(forall x : A, In x l -> P x) -> Forall l
A:Type
P:A -> Prop
l:list A

Forall l -> forall x : A, In x l -> P x
induction 1; firstorder; subst; auto.
A:Type
P:A -> Prop
l:list A

(forall x : A, In x l -> P x) -> Forall l
induction l; firstorder. Qed.
A:Type
P:A -> Prop

forall (a : A) (l : list A), Forall (a :: l) -> P a
A:Type
P:A -> Prop

forall (a : A) (l : list A), Forall (a :: l) -> P a
intros; inversion H; trivial. Qed.
A:Type
P:A -> Prop

forall Q : list A -> Type, Q [] -> (forall (b : A) (l : list A), P b -> Q (b :: l)) -> forall l : list A, Forall l -> Q l
A:Type
P:A -> Prop

forall Q : list A -> Type, Q [] -> (forall (b : A) (l : list A), P b -> Q (b :: l)) -> forall l : list A, Forall l -> Q l
intros Q H H'; induction l; intro; [|eapply H', Forall_inv]; eassumption. Qed.
A:Type
P:A -> Prop

(forall x : A, {P x} + {~ P x}) -> forall l : list A, {Forall l} + {~ Forall l}
A:Type
P:A -> Prop

(forall x : A, {P x} + {~ P x}) -> forall l : list A, {Forall l} + {~ Forall l}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

forall l : list A, {Forall l} + {~ Forall l}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

{Forall []} + {~ Forall []}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hrec:{Forall l'} + {~ Forall l'}
{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

{Forall []} + {~ Forall []}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}

Forall []
apply Forall_nil.
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hrec:{Forall l'} + {~ Forall l'}

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':~ Forall l'
{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'
Ha:P a

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'
Ha:~ P a
{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'
Ha:P a

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'
Ha:P a

Forall (a :: l')
now apply Forall_cons.
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'
Ha:~ P a

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':Forall l'
Ha:~ P a

~ Forall (a :: l')
abstract now inversion 1.
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':~ Forall l'

{Forall (a :: l')} + {~ Forall (a :: l')}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
a:A
l':list A
Hl':~ Forall l'

~ Forall (a :: l')
abstract now inversion 1. Defined. End One_predicate.
A:Type

forall (P : A -> Prop) (x0 : A) (xs : list A), Forall P (x0 :: xs) -> Forall P xs
A:Type

forall (P : A -> Prop) (x0 : A) (xs : list A), Forall P (x0 :: xs) -> Forall P xs
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)

Forall P xs
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)

forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)

forall x : A, In x (x0 :: xs) -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)

Forall P (x0 :: xs)
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x

forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x

forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
H1:forall x : A, In x xs -> P x
forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x1 : A, In x1 (x0 :: xs) -> P x1
x:A
H2:In x xs

P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
H1:forall x : A, In x xs -> P x
forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x1 : A, In x1 (x0 :: xs) -> P x1
x:A
H2:In x xs

In x (x0 :: xs)
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
H1:forall x : A, In x xs -> P x
forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x1 : A, In x1 (x0 :: xs) -> P x1
x:A
H2:In x xs

In x xs
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
H1:forall x : A, In x xs -> P x
forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x : A, In x (x0 :: xs) -> P x
H1:forall x : A, In x xs -> P x

forall x : A, In x xs -> P x
A:Type
P:A -> Prop
x0:A
xs:list A
H:Forall P (x0 :: xs)
H0:forall x1 : A, In x1 (x0 :: xs) -> P x1
H1:forall x1 : A, In x1 xs -> P x1
x:A
H2:In x xs

P x
apply (H1 x H2). Qed.
A:Type

forall P Q : A -> Prop, (forall x : A, P x -> Q x) -> forall xs : list A, Exists P xs -> Exists Q xs
A:Type

forall P Q : A -> Prop, (forall x : A, P x -> Q x) -> forall xs : list A, Exists P xs -> Exists Q xs
A:Type
P, Q:A -> Prop
H:forall x : A, P x -> Q x
xs:list A
H0:Exists P xs

Exists Q xs
A:Type
P, Q:A -> Prop
H:forall x0 : A, P x0 -> Q x0
x:A
l:list A
H0:P x

Exists Q (x :: l)
A:Type
P, Q:A -> Prop
H:forall x0 : A, P x0 -> Q x0
x:A
l:list A
H0:Exists P l
IHExists:Exists Q l
Exists Q (x :: l)
A:Type
P, Q:A -> Prop
H:forall x0 : A, P x0 -> Q x0
x:A
l:list A
H0:Exists P l
IHExists:Exists Q l

Exists Q (x :: l)
apply (Exists_cons_tl x IHExists). Qed.
A:Type
P:A -> Prop
l:list A

Forall (fun x : A => ~ P x) l <-> ~ Exists P l
A:Type
P:A -> Prop
l:list A

Forall (fun x : A => ~ P x) l <-> ~ Exists P l
A:Type
P:A -> Prop
l:list A

(forall x : A, In x l -> ~ P x) <-> ~ (exists x : A, In x l /\ P x)
firstorder. Qed.
A:Type
P:A -> Prop
l:list A

(forall x : A, P x \/ ~ P x) -> Exists (fun x : A => ~ P x) l <-> ~ Forall P l
A:Type
P:A -> Prop
l:list A

(forall x : A, P x \/ ~ P x) -> Exists (fun x : A => ~ P x) l <-> ~ Forall P l
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, P x \/ ~ P x

Exists (fun x : A => ~ P x) l <-> ~ Forall P l
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, P x \/ ~ P x

Exists (fun x : A => ~ P x) l -> ~ Forall P l
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, P x \/ ~ P x
~ Forall P l -> Exists (fun x : A => ~ P x) l
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, P x \/ ~ P x

Exists (fun x : A => ~ P x) l -> ~ Forall P l
rewrite Forall_forall, Exists_exists; firstorder.
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, P x \/ ~ P x

~ Forall P l -> Exists (fun x : A => ~ P x) l
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P l

Exists (fun x : A => ~ P x) l
A:Type
P:A -> Prop
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P []

Exists (fun x : A => ~ P x) []
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Exists (fun x : A => ~ P x) (a :: l)
A:Type
P:A -> Prop
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P []

Exists (fun x : A => ~ P x) []
A:Type
P:A -> Prop
Dec:forall x : A, P x \/ ~ P x

Forall P []
constructor.
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l

Exists (fun x : A => ~ P x) (a :: l)
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Ha:P a

Exists (fun x : A => ~ P x) (a :: l)
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Ha:~ P a
Exists (fun x : A => ~ P x) (a :: l)
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Ha:P a

Exists (fun x : A => ~ P x) (a :: l)
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Ha:P a

~ Forall P l
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Ha:P a
NF:Forall P l

Forall P (a :: l)
now constructor.
A:Type
P:A -> Prop
a:A
l:list A
Dec:forall x : A, P x \/ ~ P x
NF:~ Forall P (a :: l)
IH:~ Forall P l -> Exists (fun x : A => ~ P x) l
Ha:~ P a

Exists (fun x : A => ~ P x) (a :: l)
now apply Exists_cons_hd. Qed.
A:Type
P:A -> Prop
l:list A

(forall x : A, {P x} + {~ P x}) -> ~ Forall P l -> Exists (fun x : A => ~ P x) l
A:Type
P:A -> Prop
l:list A

(forall x : A, {P x} + {~ P x}) -> ~ Forall P l -> Exists (fun x : A => ~ P x) l
A:Type
P:A -> Prop
l:list A
Dec:forall x : A, {P x} + {~ P x}

~ Forall P l -> Exists (fun x : A => ~ P x) l
A:Type
P:A -> Prop
l:list A
Dec:forall x0 : A, {P x0} + {~ P x0}
x:A

P x \/ ~ P x
destruct (Dec x); auto. Qed.
A:Type
P:A -> Prop

(forall x : A, {P x} + {~ P x}) -> forall l : list A, {Forall P l} + {Exists (fun x : A => ~ P x) l}
A:Type
P:A -> Prop

(forall x : A, {P x} + {~ P x}) -> forall l : list A, {Forall P l} + {Exists (fun x : A => ~ P x) l}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
l:list A

{Forall P l} + {Exists (fun x : A => ~ P x) l}
A:Type
P:A -> Prop
Pdec:forall x : A, {P x} + {~ P x}
l:list A
n:~ Forall P l

Exists (fun x : A => ~ P x) l
now apply neg_Forall_Exists_neg. Defined.
A:Type

forall P Q : A -> Prop, (forall a : A, P a -> Q a) -> forall l : list A, Forall P l -> Forall Q l
A:Type

forall P Q : A -> Prop, (forall a : A, P a -> Q a) -> forall l : list A, Forall P l -> Forall Q l
A:Type
P, Q:A -> Prop
H:forall a : A, P a -> Q a
l:list A

Forall P l -> Forall Q l
A:Type
P, Q:A -> Prop
H:forall a : A, P a -> Q a
l:list A

(forall x : A, In x l -> P x) -> forall x : A, In x l -> Q x
firstorder. Qed. End Exists_Forall. Hint Constructors Exists : core. Hint Constructors Forall : core. Section Forall2.
Forall2: stating that elements of two lists are pairwise related.
  Variables A B : Type.
  Variable R : A -> B -> Prop.

  Inductive Forall2 : list A -> list B -> Prop :=
    | Forall2_nil : Forall2 [] []
    | Forall2_cons : forall x y l l',
      R x y -> Forall2 l l' -> Forall2 (x::l) (y::l').

  Hint Constructors Forall2 : core.

  
A, B:Type
R:A -> B -> Prop

Forall2 [] []
A, B:Type
R:A -> B -> Prop

Forall2 [] []
intros; apply Forall2_nil. Qed.
A, B:Type
R:A -> B -> Prop

forall (l1 l2 : list A) (l' : list B), Forall2 (l1 ++ l2) l' -> exists l1' l2' : list B, Forall2 l1 l1' /\ Forall2 l2 l2' /\ l' = l1' ++ l2'
A, B:Type
R:A -> B -> Prop

forall (l1 l2 : list A) (l' : list B), Forall2 (l1 ++ l2) l' -> exists l1' l2' : list B, Forall2 l1 l1' /\ Forall2 l2 l2' /\ l' = l1' ++ l2'
A, B:Type
R:A -> B -> Prop
l2:list A
l':list B
H:Forall2 ([] ++ l2) l'

exists l1' l2' : list B, Forall2 [] l1' /\ Forall2 l2 l2' /\ l' = l1' ++ l2'
A, B:Type
R:A -> B -> Prop
a:A
l1:list A
IHl1:forall (l0 : list A) (l'0 : list B), Forall2 (l1 ++ l0) l'0 -> exists l1' l2' : list B, Forall2 l1 l1' /\ Forall2 l0 l2' /\ l'0 = l1' ++ l2'
l2:list A
l':list B
H:Forall2 ((a :: l1) ++ l2) l'
exists l1' l2' : list B, Forall2 (a :: l1) l1' /\ Forall2 l2 l2' /\ l' = l1' ++ l2'
A, B:Type
R:A -> B -> Prop
a:A
l1:list A
IHl1:forall (l0 : list A) (l'0 : list B), Forall2 (l1 ++ l0) l'0 -> exists l1' l2' : list B, Forall2 l1 l1' /\ Forall2 l0 l2' /\ l'0 = l1' ++ l2'
l2:list A
l':list B
H:Forall2 ((a :: l1) ++ l2) l'

exists l1' l2' : list B, Forall2 (a :: l1) l1' /\ Forall2 l2 l2' /\ l' = l1' ++ l2'
A, B:Type
R:A -> B -> Prop
a:A
l1:list A
IHl1:forall (l0 : list A) (l' : list B), Forall2 (l1 ++ l0) l' -> exists l1' l2' : list B, Forall2 l1 l1' /\ Forall2 l0 l2' /\ l' = l1' ++ l2'
l2:list A
y:B
l'0:list B
H2:R a y
H4:Forall2 (l1 ++ l2) l'0

exists l1' l2' : list B, Forall2 (a :: l1) l1' /\ Forall2 l2 l2' /\ y :: l'0 = l1' ++ l2'
A, B:Type
R:A -> B -> Prop
a:A
l1:list A
IHl1:forall (l0 : list A) (l' : list B), Forall2 (l1 ++ l0) l' -> exists l1'0 l2'0 : list B, Forall2 l1 l1'0 /\ Forall2 l0 l2'0 /\ l' = l1'0 ++ l2'0
l2:list A
y:B
H2:R a y
l1', l2':list B
Hl1:Forall2 l1 l1'
Hl2:Forall2 l2 l2'

exists l1'0 l2'0 : list B, Forall2 (a :: l1) l1'0 /\ Forall2 l2 l2'0 /\ y :: l1' ++ l2' = l1'0 ++ l2'0
exists (y::l1'), l2'; simpl; auto. Qed.
A, B:Type
R:A -> B -> Prop

forall (l1' l2' : list B) (l : list A), Forall2 l (l1' ++ l2') -> exists l1 l2 : list A, Forall2 l1 l1' /\ Forall2 l2 l2' /\ l = l1 ++ l2
A, B:Type
R:A -> B -> Prop

forall (l1' l2' : list B) (l : list A), Forall2 l (l1' ++ l2') -> exists l1 l2 : list A, Forall2 l1 l1' /\ Forall2 l2 l2' /\ l = l1 ++ l2
A, B:Type
R:A -> B -> Prop
l2':list B
l:list A
H:Forall2 l ([] ++ l2')

exists l1 l2 : list A, Forall2 l1 [] /\ Forall2 l2 l2' /\ l = l1 ++ l2
A, B:Type
R:A -> B -> Prop
a:B
l1':list B
IHl1':forall (l2'0 : list B) (l0 : list A), Forall2 l0 (l1' ++ l2'0) -> exists l1 l2 : list A, Forall2 l1 l1' /\ Forall2 l2 l2'0 /\ l0 = l1 ++ l2
l2':list B
l:list A
H:Forall2 l ((a :: l1') ++ l2')
exists l1 l2 : list A, Forall2 l1 (a :: l1') /\ Forall2 l2 l2' /\ l = l1 ++ l2
A, B:Type
R:A -> B -> Prop
a:B
l1':list B
IHl1':forall (l2'0 : list B) (l0 : list A), Forall2 l0 (l1' ++ l2'0) -> exists l1 l2 : list A, Forall2 l1 l1' /\ Forall2 l2 l2'0 /\ l0 = l1 ++ l2
l2':list B
l:list A
H:Forall2 l ((a :: l1') ++ l2')

exists l1 l2 : list A, Forall2 l1 (a :: l1') /\ Forall2 l2 l2' /\ l = l1 ++ l2
A, B:Type
R:A -> B -> Prop
a:B
l1':list B
IHl1':forall (l2'0 : list B) (l : list A), Forall2 l (l1' ++ l2'0) -> exists l1 l2 : list A, Forall2 l1 l1' /\ Forall2 l2 l2'0 /\ l = l1 ++ l2
l2':list B
x:A
l0:list A
H3:R x a
H4:Forall2 l0 (l1' ++ l2')

exists l1 l2 : list A, Forall2 l1 (a :: l1') /\ Forall2 l2 l2' /\ x :: l0 = l1 ++ l2
A, B:Type
R:A -> B -> Prop
a:B
l1':list B
IHl1':forall (l2'0 : list B) (l : list A), Forall2 l (l1' ++ l2'0) -> exists l0 l3 : list A, Forall2 l0 l1' /\ Forall2 l3 l2'0 /\ l = l0 ++ l3
l2':list B
x:A
H3:R x a
l1, l2:list A
Hl1:Forall2 l1 l1'
Hl2:Forall2 l2 l2'

exists l0 l3 : list A, Forall2 l0 (a :: l1') /\ Forall2 l3 l2' /\ x :: l1 ++ l2 = l0 ++ l3
exists (x::l1), l2; simpl; auto. Qed.
A, B:Type
R:A -> B -> Prop

forall (l1 l2 : list A) (l1' l2' : list B), Forall2 l1 l1' -> Forall2 l2 l2' -> Forall2 (l1 ++ l2) (l1' ++ l2')
A, B:Type
R:A -> B -> Prop

forall (l1 l2 : list A) (l1' l2' : list B), Forall2 l1 l1' -> Forall2 l2 l2' -> Forall2 (l1 ++ l2) (l1' ++ l2')
A, B:Type
R:A -> B -> Prop
l1, l2:list A
l1', l2':list B
H:Forall2 l1 l1'
H0:Forall2 l2 l2'

Forall2 (l1 ++ l2) (l1' ++ l2')
induction l1 in l1', H, H0 |- *; inversion H; subst; simpl; auto. Qed. End Forall2. Hint Constructors Forall2 : core. Section ForallPairs.
ForallPairs : specifies that a certain relation should always hold when inspecting all possible pairs of elements of a list.
  Variable A : Type.
  Variable R : A -> A -> Prop.

  Definition ForallPairs l :=
    forall a b, In a l -> In b l -> R a b.
ForallOrdPairs : we still check a relation over all pairs of elements of a list, but now the order of elements matters.
  Inductive ForallOrdPairs : list A -> Prop :=
    | FOP_nil : ForallOrdPairs nil
    | FOP_cons : forall a l,
      Forall (R a) l -> ForallOrdPairs l -> ForallOrdPairs (a::l).

  Hint Constructors ForallOrdPairs : core.

  
A:Type
R:A -> A -> Prop

forall l : list A, ForallOrdPairs l -> forall x y : A, In x l -> In y l -> x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop

forall l : list A, ForallOrdPairs l -> forall x y : A, In x l -> In y l -> x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop

forall x y : A, In x [] -> In y [] -> x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
a:A
l:list A
H:Forall (R a) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x y : A, In x l -> In y l -> x = y \/ R x y \/ R y x
forall x y : A, In x (a :: l) -> In y (a :: l) -> x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
a:A
l:list A
H:Forall (R a) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x y : A, In x l -> In y l -> x = y \/ R x y \/ R y x

forall x y : A, In x (a :: l) -> In y (a :: l) -> x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
l:list A
x:A
H:Forall (R x) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x0 y0 : A, In x0 l -> In y0 l -> x0 = y0 \/ R x0 y0 \/ R y0 x0
y:A
H2:In y l

x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
l:list A
y:A
H:Forall (R y) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x0 y0 : A, In x0 l -> In y0 l -> x0 = y0 \/ R x0 y0 \/ R y0 x0
x:A
H1:In x l
x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
l:list A
x:A
H:Forall (R x) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x0 y0 : A, In x0 l -> In y0 l -> x0 = y0 \/ R x0 y0 \/ R y0 x0
y:A
H2:In y l

R x y
A:Type
R:A -> A -> Prop
l:list A
y:A
H:Forall (R y) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x0 y0 : A, In x0 l -> In y0 l -> x0 = y0 \/ R x0 y0 \/ R y0 x0
x:A
H1:In x l
x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
l:list A
y:A
H:Forall (R y) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x0 y0 : A, In x0 l -> In y0 l -> x0 = y0 \/ R x0 y0 \/ R y0 x0
x:A
H1:In x l

x = y \/ R x y \/ R y x
A:Type
R:A -> A -> Prop
l:list A
y:A
H:Forall (R y) l
H0:ForallOrdPairs l
IHForallOrdPairs:forall x0 y0 : A, In x0 l -> In y0 l -> x0 = y0 \/ R x0 y0 \/ R y0 x0
x:A
H1:In x l

R y x
apply -> Forall_forall; eauto. Qed.
ForallPairs implies ForallOrdPairs. The reverse implication is true only when R is symmetric and reflexive.
  
A:Type
R:A -> A -> Prop
l:list A

ForallPairs l -> ForallOrdPairs l
A:Type
R:A -> A -> Prop
l:list A

ForallPairs l -> ForallOrdPairs l
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l

ForallPairs (a :: l) -> ForallOrdPairs (a :: l)
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)

ForallOrdPairs (a :: l)
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)

Forall (R a) l
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)
ForallOrdPairs l
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)

forall x : A, In x l -> R a x
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)
ForallOrdPairs l
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)

ForallOrdPairs l
A:Type
R:A -> A -> Prop
a:A
l:list A
IHl:ForallPairs l -> ForallOrdPairs l
H:ForallPairs (a :: l)

ForallPairs l
red; intros; apply H; simpl; auto. Qed.
A:Type
R:A -> A -> Prop

(forall x : A, R x x) -> (forall x y : A, R x y -> R y x) -> forall l : list A, ForallOrdPairs l -> ForallPairs l
A:Type
R:A -> A -> Prop

(forall x : A, R x x) -> (forall x y : A, R x y -> R y x) -> forall l : list A, ForallOrdPairs l -> ForallPairs l
A:Type
R:A -> A -> Prop
Refl:forall x0 : A, R x0 x0
Sym:forall x0 y0 : A, R x0 y0 -> R y0 x0
l:list A
Hl:ForallOrdPairs l
x, y:A
Hx:In x l
Hy:In y l

R x y
destruct (ForallOrdPairs_In Hl _ _ Hx Hy); subst; intuition. Qed. End ForallPairs.

Inversion of predicates over lists based on head symbol

Ltac is_list_constr c :=
 match c with
  | nil => idtac
  | (_::_) => idtac
  | _ => fail
 end.

Ltac invlist f :=
 match goal with
  | H:f ?l |- _ => is_list_constr l; inversion_clear H; invlist f
  | H:f _ ?l |- _ => is_list_constr l; inversion_clear H; invlist f
  | H:f _ _ ?l |- _ => is_list_constr l; inversion_clear H; invlist f
  | H:f _ _ _ ?l |- _ => is_list_constr l; inversion_clear H; invlist f
  | H:f _ _ _ _ ?l |- _ => is_list_constr l; inversion_clear H; invlist f
  | _ => idtac
 end.

Exporting hints and tactics

Hint Rewrite
  rev_involutive (* rev (rev l) = l *)
  rev_unit (* rev (l ++ a :: nil) = a :: rev l *)
  map_nth (* nth n (map f l) (f d) = f (nth n l d) *)
  map_length (* length (map f l) = length l *)
  seq_length (* length (seq start len) = len *)
  app_length (* length (l ++ l') = length l + length l' *)
  rev_length (* length (rev l) = length l *)
  app_nil_r (* l ++ nil = l *)
  : list.

Ltac simpl_list := autorewrite with list.
Ltac ssimpl_list := autorewrite with list using simpl.

(* begin hide *)
(* Compatibility notations after the migration of [list] to [Datatypes] *)
Notation list := list (only parsing).
Notation list_rect := list_rect (only parsing).
Notation list_rec := list_rec (only parsing).
Notation list_ind := list_ind (only parsing).
Notation nil := nil (only parsing).
Notation cons := cons (only parsing).
Notation length := length (only parsing).
Notation app := app (only parsing).
(* Compatibility Names *)
Notation tail := tl (only parsing).
Notation head := hd_error (only parsing).
Notation head_nil := hd_error_nil (only parsing).
Notation head_cons := hd_error_cons (only parsing).
Notation ass_app := app_assoc (only parsing).
Notation app_ass := app_assoc_reverse (only parsing).
Notation In_split := in_split (only parsing).
Notation In_rev := in_rev (only parsing).
Notation In_dec := in_dec (only parsing).
Notation distr_rev := rev_app_distr (only parsing).
Notation rev_acc := rev_append (only parsing).
Notation rev_acc_rev := rev_append_rev (only parsing).
Notation AllS := Forall (only parsing). (* was formerly in TheoryList *)

Hint Resolve app_nil_end : datatypes.
(* end hide *)

Section Repeat.

  Variable A : Type.
  Fixpoint repeat (x : A) (n: nat ) :=
    match n with
      | O => []
      | S k => x::(repeat x k)
    end.

  
A:Type
x:A
n:nat

Datatypes.length (repeat x n) = n
A:Type
x:A
n:nat

Datatypes.length (repeat x n) = n
induction n as [| k Hrec]; simpl; rewrite ?Hrec; reflexivity. Qed.
A:Type
n:nat
x, y:A

In y (repeat x n) -> y = x
A:Type
n:nat
x, y:A

In y (repeat x n) -> y = x
induction n as [|k Hrec]; simpl; destruct 1; auto. Qed. End Repeat. (* Unset Universe Polymorphism. *)