Built with Alectryon, running Coq+SerAPI v8.18.0+0.18.1. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use instead of Ctrl.
[Loading ML file ring_plugin.cmxs (using legacy method) ... done]
[Loading ML file zify_plugin.cmxs (using legacy method) ... done]
[Loading ML file micromega_plugin.cmxs (using legacy method) ... done]
[Loading ML file btauto_plugin.cmxs (using legacy method) ... done]
[Loading ML file coq-itauto.plugin ... done]
From stdpp Require Import prelude. From VLSM.Lib Require Import Preamble ListExtras.

Utility: Std++ General Results

A: Type
l: list A
n: nat
x: A

x ∈ take n l → x ∈ l
A: Type
l: list A
n: nat
x: A

x ∈ take n l → x ∈ l
A: Type
l: list A
x: A

n : nat, x ∈ take n l → x ∈ l
A: Type
x: A
n: nat
H: x ∈ take n []

x ∈ []
A: Type
a: A
l: list A
x: A
IHl: n : nat, x ∈ take n l → x ∈ l
n: nat
H: x ∈ take n (a :: l)
x ∈ a :: l
A: Type
x: A
n: nat
H: x ∈ take n []

x ∈ []
by simpl in H; destruct n; simpl in H; inversion H.
A: Type
a: A
l: list A
x: A
IHl: n : nat, x ∈ take n l → x ∈ l
n: nat
H: x ∈ take n (a :: l)

x ∈ a :: l
A: Type
a: A
l: list A
x: A
IHl: n : nat, x ∈ take n l → x ∈ l
n: nat
H: x ∈ take (S n) (a :: l)

x ∈ a :: l
A: Type
a: A
l: list A
x: A
IHl: n : nat, x ∈ take n l → x ∈ l
n: nat
H: x ∈ a :: take n l

x ∈ a :: l
A: Type
a: A
l: list A
IHl: n : nat, a ∈ take n l → a ∈ l
n: nat

a ∈ a :: l
A: Type
a: A
l: list A
x: A
IHl: n : nat, x ∈ take n l → x ∈ l
n: nat
H: x ∈ take n l
x ∈ a :: l
A: Type
a: A
l: list A
IHl: n : nat, a ∈ take n l → a ∈ l
n: nat

a ∈ a :: l
by left.
A: Type
a: A
l: list A
x: A
IHl: n : nat, x ∈ take n l → x ∈ l
n: nat
H: x ∈ take n l

x ∈ a :: l
by right; eapply IHl. Qed.
A, B: Type
f: A → B
l: list A

map f (tail l) = tail (map f l)
A, B: Type
f: A → B
l: list A

map f (tail l) = tail (map f l)
by destruct l. Qed.
A: Type
l: list A

nth_error l (length l - 1) = last l
A: Type
l: list A

nth_error l (length l - 1) = last l
A: Type
a: A
l: list A
IHl: nth_error l (length l - 1) = last l

nth_error (a :: l) (length (a :: l) - 1) = last (a :: l)
A: Type
a, a0: A
l: list A
IHl: nth_error (a0 :: l) (length l - 0) = last (a0 :: l)

nth_error (a0 :: l) (length l) = last (a0 :: l)
by rewrite <- IHl, Nat.sub_0_r. Qed.
A: Type
l: list A

last_error l = last l
A: Type
l: list A

last_error l = last l
A: Type
a: A
l: list A
IHl: last_error l = last l

last_error (a :: l) = last (a :: l)
A: Type
a: A
l: list A

last_error (a :: l) = match last_error l with | Some y => Some y | None => Some a end
A: Type
a, a0: A
l: list A

match l with | [] => a0 | _ :: _ => List.last l a end = List.last l a0
A: Type
a, a0, a1: A
l: list A
IHl: match l with | [] => a0 | _ :: _ => List.last l a end = List.last l a0

match l with | [] => a1 | _ :: _ => List.last l a end = match l with | [] => a1 | _ :: _ => List.last l a0 end
A: Type
a, a0, a1: A
l: list A
IHl: match l with | [] => a0 | _ :: _ => List.last l a end = List.last l a0

match l with | [] => a1 | _ :: _ => List.last l a end = match l with | [] => a1 | _ :: _ => match l with | [] => a0 | _ :: _ => List.last l a end end
by destruct l. Qed.
A: Type
f: A → bool

l : list A, existsb f l = true ↔ Exists (λ x : A, f x = true) l
A: Type
f: A → bool

l : list A, existsb f l = true ↔ Exists (λ x : A, f x = true) l
A: Type
f: A → bool
l: list A

existsb f l = true ↔ Exists (λ x : A, f x = true) l
A: Type
f: A → bool
l: list A

( x : A, In x l ∧ f x = true) ↔ ( x : A, x ∈ l ∧ f x = true)
A: Type
f: A → bool
l: list A
x: A

In x l ∧ f x = true ↔ x ∈ l ∧ f x = true
by rewrite elem_of_list_In. Qed.
A: Type
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P l

(prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P l

(prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
p: P x

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
n: ¬ P x
(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
p: P x

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
p: P x

P x ∧ l ++ [x] = l ++ [x] ++ [] ∧ ¬ Exists P []
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
p: P x

P x ∧ l ++ [x] = l ++ [x] ++ [] ∧ ¬ False
by itauto.
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P (l ++ [x])
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
n: ¬ P x

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
Hsomething: Exists P l ∨ Exists P [x]
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
n: ¬ P x

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
H: Exists P l
IHl: Exists P l → (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
n: ¬ P x

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
l: list A
P: A → Prop
Pdec: a : A, Decision (P a)
IHl: (prefix suffix : list A) (last : A), P last ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
n: ¬ P x

(prefix suffix : list A) (last : A), P last ∧ l ++ [x] = prefix ++ [last] ++ suffix ∧ ¬ Exists P suffix
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

(prefix0 suffix0 : list A) (last0 : A), P last0 ∧ (prefix ++ [last] ++ suffix) ++ [x] = prefix0 ++ [last0] ++ suffix0 ∧ ¬ Exists P suffix0
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ (prefix ++ [last] ++ suffix) ++ [x] = prefix ++ [last] ++ suffix ++ [x] ∧ ¬ Exists P (suffix ++ [x])
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ (prefix ++ last :: suffix) ++ [x] = prefix ++ last :: suffix ++ [x] ∧ ¬ Exists P (suffix ++ [x])
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ prefix ++ (last :: suffix) ++ [x] = prefix ++ last :: suffix ++ [x] ∧ ¬ Exists P (suffix ++ [x])
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ prefix ++ last :: suffix ++ [x] = prefix ++ last :: suffix ++ [x] ∧ ¬ Exists P (suffix ++ [x])
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ prefix ++ last :: suffix ++ [x] = prefix ++ last :: suffix ++ [x] ∧ ¬ (Exists P suffix ∨ Exists P [x])
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ prefix ++ last :: suffix ++ [x] = prefix ++ last :: suffix ++ [x] ∧ ¬ (Exists P suffix ∨ P x ∨ Exists P [])
A: Type
x: A
P: A → Prop
Pdec: a : A, Decision (P a)
prefix, suffix: list A
last: A
Hf: P last
Hnone_after: ¬ Exists P suffix
n: ¬ P x

P last ∧ prefix ++ last :: suffix ++ [x] = prefix ++ last :: suffix ++ [x] ∧ ¬ (Exists P suffix ∨ P x ∨ False)
by itauto. Qed.
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (last : A), f last = true ∧ l = prefix ++ [last] ++ suffix ∧ existsb f suffix = false
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (last : A), f last = true ∧ l = prefix ++ [last] ++ suffix ∧ existsb f suffix = false
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (last : A), f last = true ∧ l = prefix ++ [last] ++ suffix ∧ existsb f suffix ≠ true
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (last : A), f last = true ∧ l = prefix ++ [last] ++ suffix ∧ ¬ Exists (λ x : A, f x = true) suffix
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

a : A, Decision (f a = true)
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true
Exists (λ x : A, f x = true) l
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

a : A, Decision (f a = true)
by typeclasses eauto.
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

Exists (λ x : A, f x = true) l
by apply existsb_Exists. Qed.
A: Type
f: A → bool

l : list A, existsb f l = false ↔ ( x : A, x ∈ l → f x = false)
A: Type
f: A → bool

l : list A, existsb f l = false ↔ ( x : A, x ∈ l → f x = false)
A: Type
f: A → bool
l: list A

existsb f l = false ↔ ( x : A, x ∈ l → f x = false)
A: Type
f: A → bool
l: list A

existsb f l ≠ true ↔ ( x : A, x ∈ l → f x ≠ true)
by rewrite existsb_Exists, <- Forall_Exists_neg, Forall_forall. Qed.
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (first : A), f first = true ∧ l = prefix ++ [first] ++ suffix ∧ existsb f prefix = false
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (first : A), f first = true ∧ l = prefix ++ [first] ++ suffix ∧ existsb f prefix = false
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (first : A), f first = true ∧ l = prefix ++ [first] ++ suffix ∧ existsb f prefix ≠ true
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

(prefix suffix : list A) (first : A), f first = true ∧ l = prefix ++ [first] ++ suffix ∧ ¬ Exists (λ x : A, f x = true) prefix
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

a : A, Decision (f a = true)
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true
Exists (λ x : A, f x = true) l
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

a : A, Decision (f a = true)
by typeclasses eauto.
A: Type
l: list A
f: A → bool
Hsomething: existsb f l = true

Exists (λ x : A, f x = true) l
by apply existsb_Exists. Qed. (* Returns all elements <<X>> of <<l>> such that <<X>> does not compare less than any other element w.r.t to the precedes relation. *) Definition maximal_elements_list {A} (precedes : relation A) `{!RelDecision precedes} (l : list A) : list A := filter (fun a => Forall (fun b => ~ precedes a b) l) l.

maximal_elements_list Nat.lt [1; 4; 2; 4] = [4; 4]

maximal_elements_list Nat.lt [1; 4; 2; 4] = [4; 4]
by itauto. Qed.

maximal_elements_list Nat.le [1; 4; 2; 4] = []

maximal_elements_list Nat.le [1; 4; 2; 4] = []
by itauto. Qed.
Returns all elements x of a set S such that x does not compare less than any other element in S w.r.t to a given precedes relation.
Definition maximal_elements_set
  `{HfinSetMessage : FinSet A SetA}
   (precedes : relation A) `{!RelDecision precedes} (s : SetA)
   : SetA :=
    filter (fun a => set_Forall (fun b => ~ precedes a b) s) s.

A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A

( a : A, a ∈ l → P a ↔ Q a) → filter P l = filter Q l
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A

( a : A, a ∈ l → P a ↔ Q a) → filter P l = filter Q l
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, a ∈ [] → P a ↔ Q a

filter P [] = filter Q []
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
l: list A
IHl: ( a : A, a ∈ l → P a ↔ Q a) → filter P l = filter Q l
H1: a0 : A, a0 ∈ a :: l → P a0 ↔ Q a0
filter P (a :: l) = filter Q (a :: l)
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, a ∈ [] → P a ↔ Q a

filter P [] = filter Q []
by rewrite 2 filter_nil.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
l: list A
IHl: ( a : A, a ∈ l → P a ↔ Q a) → filter P l = filter Q l
H1: a0 : A, a0 ∈ a :: l → P a0 ↔ Q a0

filter P (a :: l) = filter Q (a :: l)
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
l: list A
IHl: ( a : A, a ∈ l → P a ↔ Q a) → filter P l = filter Q l
H1: a0 : A, a0 ∈ a :: l → P a0 ↔ Q a0

(if decide (P a) then a :: filter P l else filter P l) = (if decide (Q a) then a :: filter Q l else filter Q l)
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
l: list A
IHl: ( a : A, a ∈ l → P a ↔ Q a) → filter P l = filter Q l
H1: a0 : A, a0 = a ∨ a0 ∈ l → P a0 ↔ Q a0

(if decide (P a) then a :: filter P l else filter P l) = (if decide (Q a) then a :: filter Q l else filter Q l)
by destruct (decide (P a)), (decide (Q a)); [rewrite IHl | ..]; firstorder. Qed.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A

filter P l = filter Q l → a : A, a ∈ l → P a ↔ Q a
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A

filter P l = filter Q l → a : A, a ∈ l → P a ↔ Q a
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A
H1: filter P l = filter Q l
a: A
H2: a ∈ l

P a ↔ Q a
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A
H1: filter P l = filter Q l
a: A
H2: a ∈ l
H3: P a

Q a
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A
H1: filter P l = filter Q l
a: A
H2: a ∈ l
H3: Q a
P a
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A
H1: filter P l = filter Q l
a: A
H2: a ∈ l
H3: P a

Q a
by eapply elem_of_list_filter; rewrite <- H1; apply elem_of_list_filter.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
l: list A
H1: filter P l = filter Q l
a: A
H2: a ∈ l
H3: Q a

P a
by eapply elem_of_list_filter; rewrite H1; apply elem_of_list_filter. Qed.
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X

filter P l = filter Q l ↔ filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X

filter P l = filter Q l ↔ filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter P l = filter Q l

filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
filter P l = filter Q l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter P l = filter Q l

filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter P l = filter Q l
Hext: a : X, a ∈ l → P a ↔ Q a

filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter P l = filter Q l
Hext: a : X, a ∈ l → P a ↔ Q a

a : X, a ∈ l → ¬ P a ↔ ¬ Q a
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter P l = filter Q l
Hext: a : X, a ∈ l → P a ↔ Q a
a: X
H2: a ∈ l

¬ P a ↔ ¬ Q a
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter P l = filter Q l
a: X
Hext: P a ↔ Q a
H2: a ∈ l

¬ P a ↔ ¬ Q a
by rewrite Hext; itauto.
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l

filter P l = filter Q l
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l

a : X, a ∈ l → P a ↔ Q a
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
a: X
H2: a ∈ l

P a ↔ Q a
X: Type
P, Q: X → Prop
H: x : X, Decision (P x)
H0: x : X, Decision (Q x)
l: list X
H1: filter (λ x : X, ¬ P x) l = filter (λ x : X, ¬ Q x) l
a: X
H2: a ∈ l
Hext: ¬ P a ↔ ¬ Q a

P a ↔ Q a
by destruct (decide (P a)), (decide (Q a)); itauto. Qed.
A: Type
l, l': list A
a: A

NoDup (l ++ a :: l') → NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A

NoDup (l ++ a :: l') → NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnda: NoDup (l ++ a :: l')

NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnda: NoDup l ∧ ( x : A, x ∈ l → x ∉ a :: l') ∧ NoDup (a :: l')

NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → x ∉ a :: l'
Hnda: NoDup (a :: l')

NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → x ∉ a :: l'
Hnda: (a ∉ l') ∧ NoDup l'

NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → ¬ (x = a ∨ x ∈ l')
Hnda: (a ∉ l') ∧ NoDup l'

NoDup (l ++ l') ∧ a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → ¬ (x = a ∨ x ∈ l')
Ha': a ∉ l'
Hnd': NoDup l'

NoDup (l ++ l')
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → ¬ (x = a ∨ x ∈ l')
Ha': a ∉ l'
Hnd': NoDup l'
a ∉ l ++ l'
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → ¬ (x = a ∨ x ∈ l')
Ha': a ∉ l'
Hnd': NoDup l'

NoDup (l ++ l')
by apply NoDup_app; firstorder.
A: Type
l, l': list A
a: A
Hnd: NoDup l
Ha: x : A, x ∈ l → ¬ (x = a ∨ x ∈ l')
Ha': a ∉ l'
Hnd': NoDup l'

a ∉ l ++ l'
by rewrite elem_of_app; firstorder. Qed.
A: Type
is: list A

i : nat, is_Some (is !! i) → j : nat, j < i → is_Some (is !! j)
A: Type
is: list A

i : nat, is_Some (is !! i) → j : nat, j < i → is_Some (is !! j)
A: Type
is: list A
i: nat
H: is_Some (is !! i)
j: nat
H0: j < i

j < length is
by etransitivity; [| apply lookup_lt_is_Some]. Qed.
A: Type
EqDecision0: EqDecision A

(l : list A) (a : A), a ∉ l → list_difference l [a] = l
A: Type
EqDecision0: EqDecision A

(l : list A) (a : A), a ∉ l → list_difference l [a] = l
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∉ l → list_difference l [a] = l

a ∉ a0 :: l → list_difference (a0 :: l) [a] = a0 :: l
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∉ l → list_difference l [a] = l
Hna0: a ≠ a0
Hnal: a ∉ l

(if decide_rel elem_of a0 [a] then list_difference l [a] else a0 :: list_difference l [a]) = a0 :: l
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∉ l → list_difference l [a] = l
Hna0: a ≠ a0
Hnal: a ∉ l
Ha0: a0 ∉ [a]

a0 :: list_difference l [a] = a0 :: l
by rewrite IHl. Qed.
A: Type
EqDecision0: EqDecision A

(l : list A) (a : A), a ∈ l → length (list_difference l [a]) < length l
A: Type
EqDecision0: EqDecision A

(l : list A) (a : A), a ∈ l → length (list_difference l [a]) < length l
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l

a ∈ a0 :: l → length (if decide_rel elem_of a0 [a] then list_difference l [a] else a0 :: list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
Ha0: a0 = a

a ∈ a0 :: l → length (list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
Ha0: a0 ≠ a
a ∈ a0 :: l → length (a0 :: list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
Ha0: a0 = a

a ∈ a0 :: l → length (list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l

length (list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
e: a ∈ l

length (list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
n: a ∉ l
length (list_difference l [a]) < S (length l)
A: Type
EqDecision0: EqDecision A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
e: a ∈ l

length (list_difference l [a]) < S (length l)
by etransitivity; [apply IHl | lia].
A: Type
EqDecision0: EqDecision A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
n: a ∉ l

length (list_difference l [a]) < S (length l)
by rewrite list_difference_singleton_not_in; [lia |].
A: Type
EqDecision0: EqDecision A
a0: A
l: list A
a: A
IHl: a ∈ l → length (list_difference l [a]) < length l
Ha0: a0 ≠ a

a ∈ a0 :: l → length (a0 :: list_difference l [a]) < S (length l)
by inversion 1; subst; [done |]; cbn; spec IHl; [| lia]. Qed.
A: Type
EqDecision0: EqDecision A

l1 l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
A: Type
EqDecision0: EqDecision A

l1 l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a

l2 : list A, a :: l1 ⊆ l2 → length (a :: l1) > length l2 → (i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2

(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
e: a ∈ l1

(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
e: a ∈ l1

(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
e: a ∈ l1

(i2 : nat) (a0 : A), 0 ≠ i2 ∧ (a :: l1) !! 0 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
i2: nat
Hi2: l1 !! i2 = Some a

(i2 : nat) (a0 : A), 0 ≠ i2 ∧ (a :: l1) !! 0 = Some a0 ∧ (a :: l1) !! i2 = Some a0
by exists (S i2), a.
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1

(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
i1, i2: nat
a': A
Hi12: i1 ≠ i2
Hli1: l1 !! i1 = Some a'
Hli2: l1 !! i2 = Some a'

(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
l1 ⊆ list_difference l2 [a]
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
length l1 > length (list_difference l2 [a])
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
i1, i2: nat
a': A
Hi12: i1 ≠ i2
Hli1: l1 !! i1 = Some a'
Hli2: l1 !! i2 = Some a'

(i1 i2 : nat) (a0 : A), i1 ≠ i2 ∧ (a :: l1) !! i1 = Some a0 ∧ (a :: l1) !! i2 = Some a0
by exists (S i1), (S i2), a'; cbn; itauto.
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1

l1 ⊆ list_difference l2 [a]
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
x: A
Hx: x ∈ l1

x ∈ list_difference l2 [a]
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1
x: A
Hx: x ∈ l1

x ∈ l2 ∧ x ≠ a
by split; [apply Hl12; right | by contradict n; subst].
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: length (a :: l1) > length l2
n: a ∉ l1

length l1 > length (list_difference l2 [a])
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: S (length l1) > length l2
n: a ∉ l1

length l1 > length (list_difference l2 [a])
A: Type
EqDecision0: EqDecision A
a: A
l1: list A
IHl1: l2 : list A, l1 ⊆ l2 → length l1 > length l2 → (i1 i2 : nat) (a : A), i1 ≠ i2 ∧ l1 !! i1 = Some a ∧ l1 !! i2 = Some a
l2: list A
Hl12: a :: l1 ⊆ l2
Hlen12: S (length l1) > length l2
n: a ∉ l1
Ha: a ∈ l2

length l1 > length (list_difference l2 [a])
by specialize (list_difference_singleton_length_in _ _ Ha) as Hlen'; lia. Qed.
A: Type
R: A → A → Prop
l: list A

ForAllSuffix2 R l ↔ ( (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b)
A: Type
R: A → A → Prop
l: list A

ForAllSuffix2 R l ↔ ( (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b)
A: Type
R: A → A → Prop
l: list A

ForAllSuffix2 R l → (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
A: Type
R: A → A → Prop
l: list A
( (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b) → ForAllSuffix2 R l
A: Type
R: A → A → Prop
l: list A

ForAllSuffix2 R l → (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
A: Type
R: A → A → Prop
a: A
l: list A
H: match l with | [] => True | b :: _ => R a b end
H0: ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) l
IHForAllSuffix: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b

(n : nat) (a0 b : A), (a :: l) !! n = Some a0 → l !! n = Some b → R a0 b
A: Type
R: A → A → Prop
a: A
l: list A
H: match l with | [] => True | b :: _ => R a b end
H0: ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) l
IHForAllSuffix: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b

a0 b : A, Some a = Some a0 → l !! 0 = Some b → R a0 b
A: Type
R: A → A → Prop
a: A
l: list A
H: match l with | [] => True | b :: _ => R a b end
H0: ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) l
IHForAllSuffix: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
n': nat
a b : A, l !! n' = Some a → l !! S n' = Some b → R a b
A: Type
R: A → A → Prop
a: A
l: list A
H: match l with | [] => True | b :: _ => R a b end
H0: ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) l
IHForAllSuffix: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b

a0 b : A, Some a = Some a0 → l !! 0 = Some b → R a0 b
by destruct l; do 2 inversion 1; subst.
A: Type
R: A → A → Prop
a: A
l: list A
H: match l with | [] => True | b :: _ => R a b end
H0: ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) l
IHForAllSuffix: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
n': nat

a b : A, l !! n' = Some a → l !! S n' = Some b → R a b
by intros; eapply IHForAllSuffix.
A: Type
R: A → A → Prop
l: list A

( (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b) → ForAllSuffix2 R l
A: Type
R: A → A → Prop

( (n : nat) (a b : A), [] !! n = Some a → None = Some b → R a b) → ForAllSuffix2 R []
A: Type
R: A → A → Prop
a: A
IHl: ( (n : nat) (a b : A), [] !! n = Some a → [] !! S n = Some b → R a b) → ForAllSuffix2 R []
( (n : nat) (a0 b : A), [a] !! n = Some a0 → [] !! n = Some b → R a0 b) → ForAllSuffix2 R [a]
A: Type
R: A → A → Prop
a, b: A
l': list A
IHl: ( (n : nat) (a b0 : A), (b :: l') !! n = Some a → (b :: l') !! S n = Some b0 → R a b0) → ForAllSuffix2 R (b :: l')
( (n : nat) (a0 b0 : A), (a :: b :: l') !! n = Some a0 → (b :: l') !! n = Some b0 → R a0 b0) → ForAllSuffix2 R (a :: b :: l')
A: Type
R: A → A → Prop

( (n : nat) (a b : A), [] !! n = Some a → None = Some b → R a b) → ForAllSuffix2 R []
by constructor.
A: Type
R: A → A → Prop
a: A
IHl: ( (n : nat) (a b : A), [] !! n = Some a → [] !! S n = Some b → R a b) → ForAllSuffix2 R []

( (n : nat) (a0 b : A), [a] !! n = Some a0 → [] !! n = Some b → R a0 b) → ForAllSuffix2 R [a]
by repeat constructor.
A: Type
R: A → A → Prop
a, b: A
l': list A
IHl: ( (n : nat) (a b0 : A), (b :: l') !! n = Some a → (b :: l') !! S n = Some b0 → R a b0) → ForAllSuffix2 R (b :: l')

( (n : nat) (a0 b0 : A), (a :: b :: l') !! n = Some a0 → (b :: l') !! n = Some b0 → R a0 b0) → ForAllSuffix2 R (a :: b :: l')
A: Type
R: A → A → Prop
a, b: A
l': list A
IHl: ( (n : nat) (a b0 : A), (b :: l') !! n = Some a → (b :: l') !! S n = Some b0 → R a b0) → ForAllSuffix2 R (b :: l')
H: (n : nat) (a0 b0 : A), (a :: b :: l') !! n = Some a0 → (b :: l') !! n = Some b0 → R a0 b0

R a b
A: Type
R: A → A → Prop
a, b: A
l': list A
IHl: ( (n : nat) (a b0 : A), (b :: l') !! n = Some a → (b :: l') !! S n = Some b0 → R a b0) → ForAllSuffix2 R (b :: l')
H: (n : nat) (a0 b0 : A), (a :: b :: l') !! n = Some a0 → (b :: l') !! n = Some b0 → R a0 b0
ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) (b :: l')
A: Type
R: A → A → Prop
a, b: A
l': list A
IHl: ( (n : nat) (a b0 : A), (b :: l') !! n = Some a → (b :: l') !! S n = Some b0 → R a b0) → ForAllSuffix2 R (b :: l')
H: (n : nat) (a0 b0 : A), (a :: b :: l') !! n = Some a0 → (b :: l') !! n = Some b0 → R a0 b0

R a b
by apply (H 0).
A: Type
R: A → A → Prop
a, b: A
l': list A
IHl: ( (n : nat) (a b0 : A), (b :: l') !! n = Some a → (b :: l') !! S n = Some b0 → R a b0) → ForAllSuffix2 R (b :: l')
H: (n : nat) (a0 b0 : A), (a :: b :: l') !! n = Some a0 → (b :: l') !! n = Some b0 → R a0 b0

ForAllSuffix (λ l : list A, match l with | [] => True | [a] => True | a :: b :: _ => R a b end) (b :: l')
by apply IHl; intro n; apply (H (S n)). Qed.
x, y: nat

x ≤ y ↔ ( z : nat, y = x + z)
x, y: nat

x ≤ y ↔ ( z : nat, y = x + z)
x, y: nat

x ≤ y → z : nat, y = x + z
x, y: nat
( z : nat, y = x + z) → x ≤ y
x, y: nat

x ≤ y → z : nat, y = x + z
by exists (y - x); lia.
x, y: nat

( z : nat, y = x + z) → x ≤ y
by intros [z ->]; lia. Qed.
A: Type
R: A → A → Prop
HT: Transitive R
l: list A

ForAllSuffix2 R l ↔ ( (m n : nat) (a b : A), m < n → l !! m = Some a → l !! n = Some b → R a b)
A: Type
R: A → A → Prop
HT: Transitive R
l: list A

ForAllSuffix2 R l ↔ ( (m n : nat) (a b : A), m < n → l !! m = Some a → l !! n = Some b → R a b)
A: Type
R: A → A → Prop
HT: Transitive R
l: list A

( (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b) ↔ ( (m n : nat) (a b : A), m < n → l !! m = Some a → l !! n = Some b → R a b)
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b

(m n : nat) (a b : A), m < n → l !! m = Some a → l !! n = Some b → R a b
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
m, n: nat
a, b: A
Hlt: m < n

l !! m = Some a → l !! n = Some b → R a b
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
m: nat
a, b: A
k: nat

l !! m = Some a → l !! (k + S m) = Some b → R a b
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
m, k: nat
IHk: a b : A, l !! m = Some a → l !! (k + S m) = Some b → R a b

a b : A, l !! m = Some a → l !! S (k + S m) = Some b → R a b
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
m, k: nat
IHk: a b : A, l !! m = Some a → l !! (k + S m) = Some b → R a b
a, b: A
Ha: l !! m = Some a
Hb: l !! S (k + S m) = Some b

R a b
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
m, k: nat
IHk: a b : A, l !! m = Some a → l !! (k + S m) = Some b → R a b
a, b: A
Ha: l !! m = Some a
Hb: l !! S (k + S m) = Some b
Hlt: k + S m < length l

R a b
A: Type
R: A → A → Prop
HT: Transitive R
l: list A
Hall: (n : nat) (a b : A), l !! n = Some a → l !! S n = Some b → R a b
m, k: nat
IHk: a b : A, l !! m = Some a → l !! (k + S m) = Some b → R a b
a, b: A
Ha: l !! m = Some a
Hb: l !! S (k + S m) = Some b
c: A
Hc: l !! (k + S m) = Some c

R a b
by transitivity c; [apply IHk | eapply Hall]. Qed.
If the n-th element of l is x, then we can decompose long enough suffixes of l into x and a suffix shorter by 1.

(A : Type) (n : nat) (l : list A) (x : A), l !! n = Some x → lastn (length l - n) l = x :: lastn (length l - S n) l

(A : Type) (n : nat) (l : list A) (x : A), l !! n = Some x → lastn (length l - n) l = x :: lastn (length l - S n) l
A: Type
n: nat
l: list A
x: A
H: l !! n = Some x

lastn (length l - n) l = x :: lastn (length l - S n) l
A: Type
n: nat
l: list A
x: A
H: l !! n = Some x

rev (take (length l - n) (rev l)) = x :: rev (take (length l - S n) (rev l))
A: Type
n: nat
l: list A
x: A
H: l !! n = Some x

drop n l = x :: drop (S n) l
by apply drop_S. Qed.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)

( a : A, P a → Q a) → s : list A, filter P s ⊆ filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)

( a : A, P a → Q a) → s : list A, filter P s ⊆ filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
Hin: x ∈ (if decide (P a) then a :: filter P s else filter P s)

x ∈ (if decide (Q a) then a :: filter Q s else filter Q s)
by destruct (decide (P a)), (decide (Q a)); cbn in *; rewrite ?elem_of_cons in *; itauto. Qed.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
s: list A
Hfg: Forall (λ a : A, P a → Q a) s

length (filter P s) ≤ length (filter Q s)
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
s: list A
Hfg: Forall (λ a : A, P a → Q a) s

length (filter P s) ≤ length (filter Q s)
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
s: list A
Hfg: Forall (λ a : A, P a → Q a) (a :: s)
IHs: Forall (λ a : A, P a → Q a) s → length (filter P s) ≤ length (filter Q s)

length (filter P (a :: s)) ≤ length (filter Q (a :: s))
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
s: list A
Hfg: Forall (λ a : A, P a → Q a) (a :: s)
IHs: Forall (λ a : A, P a → Q a) s → length (filter P s) ≤ length (filter Q s)
H3: P a → Q a
H4: Forall (λ a : A, P a → Q a) s

length (filter P (a :: s)) ≤ length (filter Q (a :: s))
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
s: list A
Hfg: Forall (λ a : A, P a → Q a) (a :: s)
IHs: length (filter P s) ≤ length (filter Q s)
H3: P a → Q a
H4: Forall (λ a : A, P a → Q a) s

length (filter P (a :: s)) ≤ length (filter Q (a :: s))
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
a: A
s: list A
Hfg: Forall (λ a : A, P a → Q a) (a :: s)
IHs: length (filter P s) ≤ length (filter Q s)
H3: P a → Q a
H4: Forall (λ a : A, P a → Q a) s

length (if decide (P a) then a :: filter P s else filter P s) ≤ length (if decide (Q a) then a :: filter Q s else filter Q s)
by destruct (decide (P a)), (decide (Q a)); cbn; itauto lia. Qed.
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
n: nat
a: A
Hnth: nth_error (filter P l) n = Some a

nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
n: nat
a: A
Hnth: nth_error (filter P l) n = Some a

nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
n: nat

a : A, nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A

(n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
A: Type
P: A → Prop
H: x : A, Decision (P x)

(n : nat) (a : A), nth_error (filter P []) n = Some a → nth : nat, nth_error_filter_index P [] n = Some nth ∧ nth_error [] nth = Some a
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
(n : nat) (a0 : A), nth_error (filter P (a :: l)) n = Some a0 → nth : nat, nth_error_filter_index P (a :: l) n = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)

(n : nat) (a : A), nth_error (filter P []) n = Some a → nth : nat, nth_error_filter_index P [] n = Some nth ∧ nth_error [] nth = Some a
by intros []; cbn; inversion 1.
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a

(n : nat) (a0 : A), nth_error (filter P (a :: l)) n = Some a0 → nth : nat, nth_error_filter_index P (a :: l) n = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
Hnth: nth_error (filter P (a :: l)) n = Some a0

nth : nat, nth_error_filter_index P (a :: l) n = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
Hnth: nth_error (if decide (P a) then a :: filter P l else filter P l) n = Some a0

nth : nat, nth_error_filter_index P (a :: l) n = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
Hnth: nth_error (if decide (P a) then a :: filter P l else filter P l) n = Some a0

nth : nat, (if decide (P a) then match n with | 0 => Some 0 | S n' => option_map S (nth_error_filter_index P l n') end else option_map S (nth_error_filter_index P l n)) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
p: P a
Hnth: nth_error (a :: filter P l) n = Some a0

nth : nat, match n with | 0 => Some 0 | S n' => option_map S (nth_error_filter_index P l n') end = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
n0: ¬ P a
Hnth: nth_error (filter P l) n = Some a0
nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
p: P a
Hnth: nth_error (a :: filter P l) n = Some a0

nth : nat, match n with | 0 => Some 0 | S n' => option_map S (nth_error_filter_index P l n') end = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
a0: A
p: P a
Hnth: nth_error (a :: filter P l) 0 = Some a0

nth : nat, Some 0 = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
p: P a
Hnth: nth_error (a :: filter P l) (S n) = Some a0
nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
a0: A
p: P a
Hnth: nth_error (a :: filter P l) 0 = Some a0

nth : nat, Some 0 = Some nth ∧ nth_error (a :: l) nth = Some a0
by inversion Hnth; subst; exists 0.
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
p: P a
Hnth: nth_error (a :: filter P l) (S n) = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
p: P a
Hnth: nth_error (filter P l) n = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
n: nat
a0: A
IHl: nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a0
p: P a
Hnth: nth_error (filter P l) n = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
n: nat
a0: A
nth: nat
Hnth': nth_error_filter_index P l n = Some nth
Ha0: nth_error l nth = Some a0
p: P a
Hnth: nth_error (filter P l) n = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
n: nat
a0: A
nth: nat
Hnth': nth_error_filter_index P l n = Some nth
Ha0: nth_error l nth = Some a0
p: P a
Hnth: nth_error (filter P l) n = Some a0

option_map S (nth_error_filter_index P l n) = Some (S nth) ∧ nth_error (a :: l) (S nth) = Some a0
by rewrite Hnth'.
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
IHl: (n : nat) (a : A), nth_error (filter P l) n = Some a → nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a
n: nat
a0: A
n0: ¬ P a
Hnth: nth_error (filter P l) n = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
n: nat
a0: A
IHl: nth : nat, nth_error_filter_index P l n = Some nth ∧ nth_error l nth = Some a0
n0: ¬ P a
Hnth: nth_error (filter P l) n = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
n: nat
a0: A
nth: nat
Hnth': nth_error_filter_index P l n = Some nth
Ha0: nth_error l nth = Some a0
n0: ¬ P a
Hnth: nth_error (filter P l) n = Some a0

nth : nat, option_map S (nth_error_filter_index P l n) = Some nth ∧ nth_error (a :: l) nth = Some a0
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
n: nat
a0: A
nth: nat
Hnth': nth_error_filter_index P l n = Some nth
Ha0: nth_error l nth = Some a0
n0: ¬ P a
Hnth: nth_error (filter P l) n = Some a0

option_map S (nth_error_filter_index P l n) = Some (S nth) ∧ nth_error (a :: l) (S nth) = Some a0
by rewrite Hnth'. Qed.
A: Type
P: A → Prop
H: x : A, Decision (P x)
s1, s2: list A

s1 ⊆ s2 → filter P s1 ⊆ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
s1, s2: list A

s1 ⊆ s2 → filter P s1 ⊆ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
s2: list A
H0: [] ⊆ s2
x: A
H1: x ∈ filter P []

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
H1: x ∈ filter P (a :: s1)
x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
s2: list A
H0: [] ⊆ s2
x: A
H1: x ∈ filter P []

x ∈ filter P s2
by apply not_elem_of_nil in H1.
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
H1: x ∈ filter P (a :: s1)

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
H1: x ∈ (if decide (P a) then a :: filter P s1 else filter P s1)

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x ∈ a :: filter P s1

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
n: ¬ P a
H1: x ∈ filter P s1
x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x ∈ a :: filter P s1

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x = a ∨ x ∈ filter P s1

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x = a

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x ∈ filter P s1
x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x = a

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
p: P a

P a ∧ a ∈ s2
by split; [| apply H0; left].
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x ∈ filter P s1

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
p: P a
H1: x ∈ filter P s1

s1 ⊆ s2
by intros y Hel; apply H0; right.
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
n: ¬ P a
H1: x ∈ filter P s1

x ∈ filter P s2
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
s1, s2: list A
IHs1: s1 ⊆ s2 → filter P s1 ⊆ filter P s2
H0: a :: s1 ⊆ s2
x: A
n: ¬ P a
H1: x ∈ filter P s1

s1 ⊆ s2
by intros y Hel; apply H0; right. Qed.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)

( a : A, P a → Q a) → s : list A, filter P s ⊆ filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)

( a : A, P a → Q a) → s : list A, filter P s ⊆ filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
H2: x ∈ (if decide (P a) then a :: filter P s else filter P s)

x ∈ (if decide (Q a) then a :: filter Q s else filter Q s)
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
p: P a
H2: x = a ∨ x ∈ filter P s
q: Q a

x ∈ a :: filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
p: P a
H2: x = a ∨ x ∈ filter P s
n: ¬ Q a
x ∈ filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
n: ¬ P a
H2: x ∈ filter P s
q: Q a
x ∈ a :: filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
n: ¬ P a
H2: x ∈ filter P s
n0: ¬ Q a
x ∈ filter Q s
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
p: P a
H2: x = a ∨ x ∈ filter P s
q: Q a

x ∈ a :: filter Q s
by destruct H2 as [-> |]; [left | right; apply IHs].
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
p: P a
H2: x = a ∨ x ∈ filter P s
n: ¬ Q a

x ∈ filter Q s
by itauto.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
n: ¬ P a
H2: x ∈ filter P s
q: Q a

x ∈ a :: filter Q s
by right; apply IHs.
A: Type
P, Q: A → Prop
H: x : A, Decision (P x)
H0: x : A, Decision (Q x)
H1: a : A, P a → Q a
a: A
s: list A
IHs: filter P s ⊆ filter Q s
x: A
n: ¬ P a
H2: x ∈ filter P s
n0: ¬ Q a

x ∈ filter Q s
by itauto. Qed.
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A

Forall (λ a : A, ¬ P a) l ↔ filter P l = []
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A

Forall (λ a : A, ¬ P a) l ↔ filter P l = []
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A

( x : A, x ∈ l → ¬ P x) ↔ filter P l = []
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
Hnone: x : A, x ∈ l → ¬ P x

filter P l = []
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
Hnone: filter P l = []
x : A, x ∈ l → ¬ P x
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
Hnone: x : A, x ∈ l → ¬ P x

filter P l = []
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
Hnone: x : A, x ∈ a :: l → ¬ P x
IHl: ( x : A, x ∈ l → ¬ P x) → filter P l = []

(if decide (P a) then a :: filter P l else filter P l) = []
A: Type
P: A → Prop
H: x : A, Decision (P x)
a: A
l: list A
Hnone: x : A, x = a ∨ x ∈ l → ¬ P x
IHl: ( x : A, x ∈ l → ¬ P x) → filter P l = []

(if decide (P a) then a :: filter P l else filter P l) = []
by rewrite decide_False; auto.
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
Hnone: filter P l = []

x : A, x ∈ l → ¬ P x
A: Type
P: A → Prop
H: x : A, Decision (P x)
l: list A
Hnone: filter P l = []
x: A
Hel: x ∈ l
Px: P x

False
by eapply filter_nil_not_elem_of in Px. Qed.
A: Type
a, b: A
la1, la2, lb1, lb2: list A
Heq: la1 ++ a :: la2 = lb1 ++ b :: lb2
Ha: a ∉ b :: lb2

lab : list A, lb1 = la1 ++ a :: lab
A: Type
a, b: A
la1, la2, lb1, lb2: list A
Heq: la1 ++ a :: la2 = lb1 ++ b :: lb2
Ha: a ∉ b :: lb2

lab : list A, lb1 = la1 ++ a :: lab
A: Type
a, b: A
la1, la2, lb1: list A

lb2 : list A, la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
A: Type
a, b: A
la1, lb1: list A

la2 lb2 : list A, la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
A: Type
a: A
la1, lb1: list A

(b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
A: Type
a: A
la1: list A

(lb1 : list A) (b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
A: Type
la1: list A

(a : A) (lb1 : list A) (b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
A: Type
b: A
lb2: list A
Ha: b ∉ b :: lb2
Heq: b :: lb2 = b :: lb2

lab : list A, [] = b :: lab
A: Type
b0: A
lb1: list A
b: A
lb2: list A
Ha: b0 ∉ b :: lb2
Heq: b0 :: lb1 ++ b :: lb2 = b0 :: lb1 ++ b :: lb2
lab : list A, b0 :: lb1 = b0 :: lab
A: Type
la1: list A
IHla1: (a : A) (lb1 : list A) (b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
a0, b: A
la2: list A
Ha: a0 ∉ b :: la1 ++ a0 :: la2
Heq: b :: la1 ++ a0 :: la2 = b :: la1 ++ a0 :: la2
lab : list A, [] = b :: la1 ++ a0 :: lab
A: Type
la1: list A
IHla1: (a : A) (lb1 : list A) (b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
a0, b0: A
lb1: list A
b: A
la2, lb2: list A
Heq: b0 :: la1 ++ a0 :: la2 = b0 :: lb1 ++ b :: lb2
Ha: a0 ∉ b :: lb2
H1: la1 ++ a0 :: la2 = lb1 ++ b :: lb2
lab : list A, b0 :: lb1 = b0 :: la1 ++ a0 :: lab
A: Type
b: A
lb2: list A
Ha: b ∉ b :: lb2
Heq: b :: lb2 = b :: lb2

lab : list A, [] = b :: lab
by contradict Ha; left.
A: Type
b0: A
lb1: list A
b: A
lb2: list A
Ha: b0 ∉ b :: lb2
Heq: b0 :: lb1 ++ b :: lb2 = b0 :: lb1 ++ b :: lb2

lab : list A, b0 :: lb1 = b0 :: lab
by exists lb1.
A: Type
la1: list A
IHla1: (a : A) (lb1 : list A) (b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
a0, b: A
la2: list A
Ha: a0 ∉ b :: la1 ++ a0 :: la2
Heq: b :: la1 ++ a0 :: la2 = b :: la1 ++ a0 :: la2

lab : list A, [] = b :: la1 ++ a0 :: lab
by contradict Ha; rewrite elem_of_cons, elem_of_app, elem_of_cons; auto.
A: Type
la1: list A
IHla1: (a : A) (lb1 : list A) (b : A) (la2 lb2 : list A), la1 ++ a :: la2 = lb1 ++ b :: lb2 → a ∉ b :: lb2 → lab : list A, lb1 = la1 ++ a :: lab
a0, b0: A
lb1: list A
b: A
la2, lb2: list A
Heq: b0 :: la1 ++ a0 :: la2 = b0 :: lb1 ++ b :: lb2
Ha: a0 ∉ b :: lb2
H1: la1 ++ a0 :: la2 = lb1 ++ b :: lb2

lab : list A, b0 :: lb1 = b0 :: la1 ++ a0 :: lab
A: Type
la1: list A
a0: A
lb1: list A
IHla1: lab : list A, lb1 = la1 ++ a0 :: lab
b0, b: A
la2, lb2: list A
Heq: b0 :: la1 ++ a0 :: la2 = b0 :: lb1 ++ b :: lb2
Ha: a0 ∉ b :: lb2
H1: la1 ++ a0 :: la2 = lb1 ++ b :: lb2

lab : list A, b0 :: lb1 = b0 :: la1 ++ a0 :: lab
A: Type
la1: list A
a0: A
lb1, la0b: list A
Hla0b: lb1 = la1 ++ a0 :: la0b
b0, b: A
la2, lb2: list A
Heq: b0 :: la1 ++ a0 :: la2 = b0 :: lb1 ++ b :: lb2
Ha: a0 ∉ b :: lb2
H1: la1 ++ a0 :: la2 = lb1 ++ b :: lb2

lab : list A, b0 :: lb1 = b0 :: la1 ++ a0 :: lab
by exists la0b; subst. Qed.
l: list nat
nz: list_max l > 0

list_max l ∈ l
l: list nat
nz: list_max l > 0

list_max l ∈ l
a: nat
l: list nat
nz: a `max` list_max l > 0
IHl: list_max l > 0 → list_max l ∈ l

a `max` list_max l ∈ a :: l
a: nat
l: list nat
nz: a `max` list_max l > 0
IHl: list_max l > 0 → list_max l ∈ l

a `max` list_max l = a ∨ a `max` list_max l ∈ l
by destruct (Nat.max_spec_le a (list_max l)) as [[H ->] | [H ->]]; itauto lia. Qed.
A, B: Type
f: A → option B
l1, l2: list A
Hincl: l1 ⊆ l2

omap f l1 ⊆ omap f l2
A, B: Type
f: A → option B
l1, l2: list A
Hincl: l1 ⊆ l2

omap f l1 ⊆ omap f l2
by intros b; rewrite !elem_of_list_omap; firstorder. Qed.
A: Type
l: list (option A)
a: A

a ∈ cat_option l ↔ ( b : option A, b ∈ l ∧ b = Some a)
A: Type
l: list (option A)
a: A

a ∈ cat_option l ↔ ( b : option A, b ∈ l ∧ b = Some a)
by apply elem_of_list_omap. Qed.
l: list nat
Hne: l ≠ []

list_max l ∈ l
l: list nat
Hne: l ≠ []

list_max l ∈ l
l: list nat
Hne: l ≠ []
eq_max: list_max l = 0

0 ∈ l
l: list nat
Hne: l ≠ []
n: nat
eq_max: list_max l = S n
S n ∈ l
l: list nat
Hne: l ≠ []
eq_max: list_max l = 0

0 ∈ l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: list_max (n :: l) ≤ 0 ↔ Forall (λ k : nat, k ≤ 0) (n :: l)

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: list_max (n :: l) ≤ 0 → Forall (λ k : nat, k ≤ 0) (n :: l)

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: 00 → Forall (λ k : nat, k ≤ 0) (n :: l)

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: 00 → Forall (λ k : nat, k ≤ 0) (n :: l)

00
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: Forall (λ k : nat, k ≤ 0) (n :: l)
0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: Forall (λ k : nat, k ≤ 0) (n :: l)

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: x : nat, x ∈ n :: l → x ≤ 0

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: n ∈ n :: l → n ≤ 0

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: n ∈ n :: l → n ≤ 0

n ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: n ≤ 0
0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: n ≤ 0

0 ∈ n :: l
n: nat
l: list nat
Hne: n :: l ≠ []
eq_max: list_max (n :: l) = 0
Hle: n ≤ 0
Hn0: n = 0

0 ∈ n :: l
by rewrite Hn0; left.
l: list nat
Hne: l ≠ []
n: nat
eq_max: list_max l = S n

S n ∈ l
l: list nat
Hne: l ≠ []
n: nat
eq_max: list_max l = S n
Hmax: list_max l > 0 → list_max l ∈ l

S n ∈ l
by rewrite <- eq_max; itauto lia. Qed.
A: Type
EqDecision0: EqDecision A
l: list A
Hne: l ≠ []

mode l ≠ []
A: Type
EqDecision0: EqDecision A
l: list A
Hne: l ≠ []

mode l ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l: list A
Hne: a :: l ≠ []

mode (a :: l) ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []

mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'

mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'

list_max occurrences > 0
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'

list_max occurrences > 0
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'

(if decide (a = a) then S (count_occ decide_eq l a) else count_occ decide_eq l a) `max` foldr Init.Nat.max 0 (map (λ x : A, if decide (a = x) then S (count_occ decide_eq l x) else count_occ decide_eq l x) l) > 0
by rewrite decide_True; [lia |].
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0

mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0

a : A, count_occ decide_eq l' a = list_max occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences
mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0

a : A, count_occ decide_eq l' a = list_max occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0

list_max occurrences ∈ occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: list_max occurrences ∈ occurrences
a : A, count_occ decide_eq l' a = list_max occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0

list_max occurrences ∈ occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0

occurrences ≠ []
by destruct occurrences; [cbn in Hmaxp; lia |].
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: list_max occurrences ∈ occurrences

a : A, count_occ decide_eq l' a = list_max occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: y : A, list_max (map (count_occ decide_eq l') l') = count_occ decide_eq l' y ∧ y ∈ l'

a : A, count_occ decide_eq l' a = list_max occurrences
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
x: A
Heq: list_max (map (count_occ decide_eq l') l') = count_occ decide_eq l' x
Hin: x ∈ l'

a : A, count_occ decide_eq l' a = list_max occurrences
by rewrite Heqoccurrences; eauto.
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences

mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences

a : A, a ∈ mode l'
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences
H0: a : A, a ∈ mode l'
mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences

a : A, a ∈ mode l'
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
x: A
H: count_occ decide_eq l' x = list_max occurrences

a : A, a ∈ mode l'
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
x: A
H: count_occ decide_eq l' x = list_max occurrences

x ∈ mode l'
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
x: A
H: count_occ decide_eq l' x = list_max occurrences

x ∈ filter (λ a : A, count_occ decide_eq l' a = list_max (map (count_occ decide_eq l') l')) l'
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
x: A
H: count_occ decide_eq l' x = list_max occurrences

count_occ decide_eq l' x = list_max (map (count_occ decide_eq l') l') ∧ x ∈ l'
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
x: A
H: count_occ decide_eq l' x = list_max occurrences

x ∈ l' ↔ count_occ decide_eq l' x > 0 → count_occ decide_eq l' x = list_max (map (count_occ decide_eq l') l') ∧ x ∈ l'
by itauto congruence.
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences
H0: a : A, a ∈ mode l'

mode l' ≠ []
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences
H0: a : A, a ∈ []
contra: mode l' = []

False
A: Type
EqDecision0: EqDecision A
a: A
l, l': list A
Heql': l' = a :: l
Hne: l' ≠ []
occurrences: list nat
Heqoccurrences: occurrences = map (count_occ decide_eq l') l'
Hmaxp: list_max occurrences > 0
H: a : A, count_occ decide_eq l' a = list_max occurrences
x: A
H0: x ∈ []
contra: mode l' = []

False
by apply elem_of_nil in H0. Qed.
When a list contains two elements, either they are equal or we can split the list into three parts separated by the elements (and this can be done in two ways, depending on the order of the elements).

(A : Type) (l : list A) (x y : A), x ∈ l → y ∈ l → x = y ∨ ( l1 l2 l3 : list A, l = l1 ++ x :: l2 ++ y :: l3 ∨ l = l1 ++ y :: l2 ++ x :: l3)

(A : Type) (l : list A) (x y : A), x ∈ l → y ∈ l → x = y ∨ ( l1 l2 l3 : list A, l = l1 ++ x :: l2 ++ y :: l3 ∨ l = l1 ++ y :: l2 ++ x :: l3)
A: Type
x: list A
y, l: A
Hx: y ∈ x
Hy: l ∈ x

y = l ∨ ( l1 l2 l3 : list A, x = l1 ++ y :: l2 ++ l :: l3 ∨ x = l1 ++ l :: l2 ++ y :: l3)
A: Type
y, l: A
l1, l2: list A
Hy: l ∈ l1 ++ y :: l2

y = l ∨ ( l3 l4 l5 : list A, l1 ++ y :: l2 = l3 ++ y :: l4 ++ l :: l5 ∨ l1 ++ y :: l2 = l3 ++ l :: l4 ++ y :: l5)
A: Type
y, l: A
l1, l2: list A
Hy: l ∈ l1 ∨ l = y ∨ l ∈ l2

y = l ∨ ( l3 l4 l5 : list A, l1 ++ y :: l2 = l3 ++ y :: l4 ++ l :: l5 ∨ l1 ++ y :: l2 = l3 ++ l :: l4 ++ y :: l5)
A: Type
y, l: A
l1, l2: list A
Hy: l ∈ l1

y = l ∨ ( l3 l4 l5 : list A, l1 ++ y :: l2 = l3 ++ y :: l4 ++ l :: l5 ∨ l1 ++ y :: l2 = l3 ++ l :: l4 ++ y :: l5)
A: Type
y, l: A
l1, l2: list A
Hy: l ∈ l2
y = l ∨ ( l3 l4 l5 : list A, l1 ++ y :: l2 = l3 ++ y :: l4 ++ l :: l5 ∨ l1 ++ y :: l2 = l3 ++ l :: l4 ++ y :: l5)
A: Type
y, l: A
l1, l2: list A
Hy: l ∈ l1

y = l ∨ ( l3 l4 l5 : list A, l1 ++ y :: l2 = l3 ++ y :: l4 ++ l :: l5 ∨ l1 ++ y :: l2 = l3 ++ l :: l4 ++ y :: l5)
A: Type
y, l: A
l2, l11, l12: list A

y = l ∨ ( l1 l3 l4 : list A, (l11 ++ l :: l12) ++ y :: l2 = l1 ++ y :: l3 ++ l :: l4 ∨ (l11 ++ l :: l12) ++ y :: l2 = l1 ++ l :: l3 ++ y :: l4)
A: Type
y, l: A
l2, l11, l12: list A

(l11 ++ l :: l12) ++ y :: l2 = l11 ++ l :: l12 ++ y :: l2
by rewrite <- app_assoc.
A: Type
y, l: A
l1, l2: list A
Hy: l ∈ l2

y = l ∨ ( l3 l4 l5 : list A, l1 ++ y :: l2 = l3 ++ y :: l4 ++ l :: l5 ∨ l1 ++ y :: l2 = l3 ++ l :: l4 ++ y :: l5)
A: Type
y, l: A
l1, l21, l22: list A

y = l ∨ ( l2 l3 l4 : list A, l1 ++ y :: l21 ++ l :: l22 = l2 ++ y :: l3 ++ l :: l4 ∨ l1 ++ y :: l21 ++ l :: l22 = l2 ++ l :: l3 ++ y :: l4)
by right; exists l1, l21, l22; left; cbn. Qed.
A: Type
l1, l2: list (list A)

mjoin (l1 ++ l2) = mjoin l1 ++ mjoin l2
A: Type
l1, l2: list (list A)

mjoin (l1 ++ l2) = mjoin l1 ++ mjoin l2
A: Type
a: list A
l1, l2: list (list A)
IHl1: mjoin (l1 ++ l2) = mjoin l1 ++ mjoin l2

a ++ mjoin (l1 ++ l2) = (a ++ mjoin l1) ++ mjoin l2
A: Type
a: list A
l1, l2: list (list A)
IHl1: mjoin (l1 ++ l2) = mjoin l1 ++ mjoin l2

a ++ mjoin l1 ++ mjoin l2 = (a ++ mjoin l1) ++ mjoin l2
by rewrite app_assoc. Qed.
A, B: Type
f: A → list B
l1, l2: list A

(l1 ++ l2) ≫= f = (l1 ≫= f) ++ l2 ≫= f
A, B: Type
f: A → list B
l1, l2: list A

(l1 ++ l2) ≫= f = (l1 ≫= f) ++ l2 ≫= f
by induction l1; [| cbn; rewrite IHl1, app_assoc]. Qed.

(A B : Type) (f : A → list B) (l : list A), Forall (λ x : A, f x = []) l → l ≫= f = []

(A B : Type) (f : A → list B) (l : list A), Forall (λ x : A, f x = []) l → l ≫= f = []
A, B: Type
f: A → list B
x: A
l: list A
H: f x = []
H0: Forall (λ x : A, f x = []) l
IHForall: l ≫= f = []

f x ++ l ≫= f = []
by rewrite H, IHForall; cbn. Qed.

(A : Type) (l1 l2 l3 : list A), l1 ++ l2 ⊆ l3 → l1 ⊆ l3 ∧ l2 ⊆ l3

(A : Type) (l1 l2 l3 : list A), l1 ++ l2 ⊆ l3 → l1 ⊆ l3 ∧ l2 ⊆ l3

(A : Type) (l1 l2 l3 : list A), ( x : A, x ∈ l1 ++ l2 → x ∈ l3) → ( x : A, x ∈ l1 → x ∈ l3) ∧ ( x : A, x ∈ l2 → x ∈ l3)
A: Type
l1, l2, l3: list A
Hsub: x : A, x ∈ l1 ++ l2 → x ∈ l3

( x : A, x ∈ l1 → x ∈ l3) ∧ ( x : A, x ∈ l2 → x ∈ l3)
A: Type
l1, l2, l3: list A
Hsub: x : A, x ∈ l1 ++ l2 → x ∈ l3
x: A
Hin: x ∈ l1

x ∈ l3
A: Type
l1, l2, l3: list A
Hsub: x : A, x ∈ l1 ++ l2 → x ∈ l3
x: A
Hin: x ∈ l2
x ∈ l3
A: Type
l1, l2, l3: list A
Hsub: x : A, x ∈ l1 ++ l2 → x ∈ l3
x: A
Hin: x ∈ l1

x ∈ l3
by apply Hsub, elem_of_app; left.
A: Type
l1, l2, l3: list A
Hsub: x : A, x ∈ l1 ++ l2 → x ∈ l3
x: A
Hin: x ∈ l2

x ∈ l3
by apply Hsub, elem_of_app; right. Qed.
index: Type
f: index → nat

Proper (Permutation ==> eq) (sum_list_with f)
index: Type
f: index → nat

Proper (Permutation ==> eq) (sum_list_with f)
index: Type
f: index → nat
x: index
l, l': list index
H: l ≡ₚ l'
IHPermutation: sum_list_with f l = sum_list_with f l'

f x + sum_list_with f l = f x + sum_list_with f l'
index: Type
f: index → nat
x, y: index
l: list index
f y + (f x + sum_list_with f l) = f x + (f y + sum_list_with f l)
index: Type
f: index → nat
l, l', l'': list index
H: l ≡ₚ l'
H0: l' ≡ₚ l''
IHPermutation1: sum_list_with f l = sum_list_with f l'
IHPermutation2: sum_list_with f l' = sum_list_with f l''
sum_list_with f l = sum_list_with f l''
index: Type
f: index → nat
x: index
l, l': list index
H: l ≡ₚ l'
IHPermutation: sum_list_with f l = sum_list_with f l'

f x + sum_list_with f l = f x + sum_list_with f l'
by rewrite IHPermutation.
index: Type
f: index → nat
x, y: index
l: list index

f y + (f x + sum_list_with f l) = f x + (f y + sum_list_with f l)
by lia.
index: Type
f: index → nat
l, l', l'': list index
H: l ≡ₚ l'
H0: l' ≡ₚ l''
IHPermutation1: sum_list_with f l = sum_list_with f l'
IHPermutation2: sum_list_with f l' = sum_list_with f l''

sum_list_with f l = sum_list_with f l''
by congruence. Qed.
index: Type
f, g: index → nat
l: list index

( i : index, i ∈ l → f i = g i) → sum_list_with f l = sum_list_with g l
index: Type
f, g: index → nat
l: list index

( i : index, i ∈ l → f i = g i) → sum_list_with f l = sum_list_with g l
index: Type
f, g: index → nat
a: index
l: list index
IHl: ( i : index, i ∈ l → f i = g i) → sum_list_with f l = sum_list_with g l
Heq: i : index, i ∈ a :: l → f i = g i

f a + sum_list_with f l = g a + sum_list_with g l
index: Type
f, g: index → nat
a: index
l: list index
IHl: ( i : index, i ∈ l → f i = g i) → sum_list_with f l = sum_list_with g l
Heq: i : index, i ∈ a :: l → f i = g i

g a + sum_list_with f l = g a + sum_list_with g l
index: Type
f, g: index → nat
a: index
l: list index
IHl: ( i : index, i ∈ l → f i = g i) → sum_list_with f l = sum_list_with g l
Heq: i : index, i ∈ a :: l → f i = g i

i : index, i ∈ l → f i = g i
by intros; apply Heq; right. Qed.
index: Type
f: index → nat
l: list index

sum_list_with f l = 0 ↔ ( i : index, i ∈ l → f i = 0)
index: Type
f: index → nat
l: list index

sum_list_with f l = 0 ↔ ( i : index, i ∈ l → f i = 0)
index: Type
f: index → nat
l: list index

sum_list_with f l = 0 i : index, i ∈ l → f i = 0
index: Type
f: index → nat
l: list index
( i : index, i ∈ l → f i = 0) → sum_list_with f l = 0
index: Type
f: index → nat
l: list index

sum_list_with f l = 0 i : index, i ∈ l → f i = 0
index: Type
f: index → nat
l: list index
Hsum: sum_list_with f l = 0
i: index
Hi: i ∈ l

f i = 0
index: Type
f: index → nat
l: list index
Hsum: sum_list_with f l = 0
i: index
Hi: f i ≤ sum_list_with f l

f i = 0
by lia.
index: Type
f: index → nat
l: list index

( i : index, i ∈ l → f i = 0) → sum_list_with f l = 0
index: Type
f: index → nat
a: index
l: list index
IHl: ( i : index, i ∈ l → f i = 0) → sum_list_with f l = 0
Hall: i : index, i ∈ a :: l → f i = 0

f a + sum_list_with f l = 0
index: Type
f: index → nat
a: index
l: list index
IHl: ( i : index, i ∈ l → f i = 0) → sum_list_with f l = 0
Hall: i : index, i ∈ a :: l → f i = 0

0 + sum_list_with f l = 0
index: Type
f: index → nat
a: index
l: list index
IHl: ( i : index, i ∈ l → f i = 0) → sum_list_with f l = 0
Hall: i : index, i ∈ a :: l → f i = 0

i : index, i ∈ l → f i = 0
by intros; apply Hall; right. Qed.
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)

l : list (dsig P), NoDup l ↔ NoDup (map proj1_sig l)
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)

l : list (dsig P), NoDup l ↔ NoDup (map proj1_sig l)
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
l: list (dsig P)

NoDup l → NoDup (map proj1_sig l)
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
l: list (dsig P)
NoDup (map proj1_sig l) → NoDup l
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
l: list (dsig P)

NoDup l → NoDup (map proj1_sig l)
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
da: dsig P
dl: list (dsig P)
Hda: da ∉ dl
H: NoDup dl
IHNoDup: NoDup (map proj1_sig dl)

`da ∉ map proj1_sig dl
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
da: dsig P
dl: list (dsig P)
Hda: da ∉ dl
H: NoDup dl
IHNoDup: NoDup (map proj1_sig dl)

¬ ( y : {x : A | bool_decide (P x)}, `da = `y ∧ y ∈ dl)
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
da: dsig P
dl: list (dsig P)
Hda: da ∉ dl
H: NoDup dl
IHNoDup: NoDup (map proj1_sig dl)
_da: {x : A | bool_decide (P x)}
Heq: `da = `_da
H_da: _da ∈ dl

False
by apply dsig_eq in Heq as <-.
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
l: list (dsig P)

NoDup (map proj1_sig l) → NoDup l
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
a: dsig P
l: list (dsig P)
IHl: NoDup (map proj1_sig l) → NoDup l

NoDup (`a :: map proj1_sig l) → NoDup (a :: l)
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
a: dsig P
l: list (dsig P)
IHl: NoDup (map proj1_sig l) → NoDup l

(`a ∉ map proj1_sig l) ∧ NoDup (map proj1_sig l) → (a ∉ l) ∧ NoDup l
A: Type
P: A → Prop
Pdec: a : A, Decision (P a)
a: dsig P
l: list (dsig P)
IHl: NoDup (map proj1_sig l) → NoDup l
Ha: `a ∉ map proj1_sig l
H: NoDup (map proj1_sig l)

a ∉ l
by contradict Ha; apply elem_of_list_fmap; eexists. Qed.
A: Type
m: option A
Hsome: is_Some m
f: A
H: m = Some f

is_Some_proj Hsome = f
A: Type
m: option A
Hsome: is_Some m
f: A
H: m = Some f

is_Some_proj Hsome = f
by intros; subst. Qed. Definition set_Forall2 `{ElemOf A C} (R : relation A) (X : C) := forall x y : A, x ∈ X -> y ∈ X -> x <> y -> R x y. Definition set_Exists2 `{ElemOf A C} (R : relation A) (X : C) := exists x y : A, x ∈ X /\ y ∈ X /\ x <> y /\ R x y. Section sec_Forall2_Exists2_props. Context `{SemiSet A C} (R : relation A) .
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A

set_Forall2 R (∅ : C)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A

set_Forall2 R (∅ : C)
by unfold set_Forall2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x: A

set_Forall2 R ({[x]} : C)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x: A

set_Forall2 R ({[x]} : C)
by unfold set_Forall2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A

x ≠ y → set_Forall2 R ({[x; y]} : C) ↔ R x y ∧ R y x
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A

x ≠ y → set_Forall2 R ({[x; y]} : C) ↔ R x y ∧ R y x
by unfold set_Forall2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Forall2 R X → set_Forall2 R Y → ( x y : A, x ∈ X → y ∈ Y → x ≠ y → R x y ∧ R y x) → set_Forall2 R (X ∪ Y)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Forall2 R X → set_Forall2 R Y → ( x y : A, x ∈ X → y ∈ Y → x ≠ y → R x y ∧ R y x) → set_Forall2 R (X ∪ Y)
by unfold set_Forall2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Forall2 R (X ∪ Y) → set_Forall2 R X
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Forall2 R (X ∪ Y) → set_Forall2 R X
by unfold set_Forall2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Forall2 R (X ∪ Y) → set_Forall2 R Y
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Forall2 R (X ∪ Y) → set_Forall2 R Y
by unfold set_Forall2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A

¬ set_Exists2 R (∅ : C)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A

¬ set_Exists2 R (∅ : C)
by unfold set_Exists2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x: A

¬ set_Exists2 R ({[x]} : C)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x: A

¬ set_Exists2 R ({[x]} : C)
by unfold set_Exists2; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A

set_Exists2 R ({[x; y]} : C) ↔ x ≠ y ∧ (R x y ∨ R y x)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A

set_Exists2 R ({[x; y]} : C) ↔ x ≠ y ∧ (R x y ∨ R y x)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A

x ≠ y ∧ (R x y ∨ R y x) → set_Exists2 R {[x; y]}
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A
H4: x ≠ y
H5: R x y

set_Exists2 R {[x; y]}
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A
H4: x ≠ y
H5: R y x
set_Exists2 R {[x; y]}
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A
H4: x ≠ y
H5: R x y

set_Exists2 R {[x; y]}
by exists x, y; set_solver.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
x, y: A
H4: x ≠ y
H5: R y x

set_Exists2 R {[x; y]}
by exists y, x; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Exists2 R X → set_Exists2 R (X ∪ Y)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Exists2 R X → set_Exists2 R (X ∪ Y)
by intros [x [y Hxy]]; exists x, y; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Exists2 R Y → set_Exists2 R (X ∪ Y)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Exists2 R Y → set_Exists2 R (X ∪ Y)
by intros [x [y Hxy]]; exists x, y; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A

x ∈ X → y ∈ Y → x ≠ y → R x y → set_Exists2 R (X ∪ Y)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A

x ∈ X → y ∈ Y → x ≠ y → R x y → set_Exists2 R (X ∪ Y)
by exists x, y; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A

x ∈ X → y ∈ Y → x ≠ y → R y x → set_Exists2 R (X ∪ Y)
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A

x ∈ X → y ∈ Y → x ≠ y → R y x → set_Exists2 R (X ∪ Y)
by intros; exists y, x; set_solver. Qed.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Exists2 R (X ∪ Y) → set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C

set_Exists2 R (X ∪ Y) → set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
Hx: x ∈ X ∪ Y
Hy: y ∈ X ∪ Y
Hneq: x ≠ y
Hxy: R x y

set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ X
H5: y ∈ X
Hneq: x ≠ y
Hxy: R x y

set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ X
H5: y ∈ Y
Hneq: x ≠ y
Hxy: R x y
set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ Y
H5: y ∈ X
Hneq: x ≠ y
Hxy: R x y
set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ Y
H5: y ∈ Y
Hneq: x ≠ y
Hxy: R x y
set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ X
H5: y ∈ X
Hneq: x ≠ y
Hxy: R x y

set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
by left; exists x, y.
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ X
H5: y ∈ Y
Hneq: x ≠ y
Hxy: R x y

set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
by right; right; exists x, y; repeat split; [.. | left].
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ Y
H5: y ∈ X
Hneq: x ≠ y
Hxy: R x y

set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
by right; right; exists y, x; repeat split; [.. | right].
A, C: Type
H: ElemOf A C
H0: Empty C
H1: Singleton A C
H2: Union C
H3: SemiSet A C
R: relation A
X, Y: C
x, y: A
H4: x ∈ Y
H5: y ∈ Y
Hneq: x ≠ y
Hxy: R x y

set_Exists2 R X ∨ set_Exists2 R Y ∨ ( x y : A, x ∈ X ∧ y ∈ Y ∧ x ≠ y ∧ (R x y ∨ R y x))
by right; left; exists x, y. Qed. End sec_Forall2_Exists2_props.