@@ -50,6 +50,9 @@ by simp only [countp_eq_length_filter, filter_append, length_append]
50
50
lemma countp_pos {l} : 0 < countp p l ↔ ∃ a ∈ l, p a :=
51
51
by simp only [countp_eq_length_filter, length_pos_iff_exists_mem, mem_filter, exists_prop]
52
52
53
+ lemma countp_eq_length {l} : countp p l = l.length ↔ ∀ a ∈ l, p a :=
54
+ by rw [countp_eq_length_filter, filter_length_eq_length]
55
+
53
56
lemma length_filter_lt_length_iff_exists (l) : length (filter p l) < length l ↔ ∃ x ∈ l, ¬p x :=
54
57
by rw [length_eq_countp_add_countp p l, ← countp_pos, countp_eq_length_filter, lt_add_iff_pos_right]
55
58
@@ -122,6 +125,9 @@ decidable.by_contradiction $ λ h', h $ count_pos.1 (nat.pos_of_ne_zero h')
122
125
lemma not_mem_of_count_eq_zero {a : α} {l : list α} (h : count a l = 0 ) : a ∉ l :=
123
126
λ h', (count_pos.2 h').ne' h
124
127
128
+ lemma count_eq_length {a : α} {l} : count a l = l.length ↔ ∀ b ∈ l, a = b :=
129
+ by rw [count, countp_eq_length]
130
+
125
131
@[simp] lemma count_repeat (a : α) (n : ℕ) : count a (repeat a n) = n :=
126
132
by rw [count, countp_eq_length_filter, filter_eq_self.2 , length_repeat];
127
133
exact λ b m, (eq_of_mem_repeat m).symm
157
163
begin
158
164
induction l with y l IH generalizing x,
159
165
{ simp },
160
- { rw map_cons,
161
- by_cases h : x = y,
162
- { simpa [h] using IH _ },
163
- { simpa [h, hf.ne h] using IH _ } }
166
+ { simp [map_cons, count_cons', IH, hf.eq_iff] }
164
167
end
165
168
166
169
lemma count_le_count_map [decidable_eq β] (l : list α) (f : α → β) (x : α) :
0 commit comments