@@ -375,11 +375,6 @@ begin
375
375
exists_and_distrib_left] } }
376
376
end
377
377
378
- @[simp] theorem split_at_eq_take_drop : ∀ (n : ℕ) (l : list α), split_at n l = (take n l, drop n l)
379
- | 0 a := rfl
380
- | (succ n) [] := rfl
381
- | (succ n) (x :: xs) := by simp only [split_at, split_at_eq_take_drop n xs, take, drop]
382
-
383
378
@[simp] theorem take_append_drop : ∀ (n : ℕ) (l : list α), take n l ++ drop n l = l
384
379
| 0 a := rfl
385
380
| (succ n) [] := rfl
@@ -2732,6 +2727,8 @@ end
2732
2727
2733
2728
attribute [simp] join
2734
2729
2730
+ @[simp] lemma join_nil {α : Type u} : [([] : list α)].join = [] := rfl
2731
+
2735
2732
@[simp] theorem join_eq_nil : ∀ {L : list (list α)}, join L = [] ↔ ∀ l ∈ L, l = []
2736
2733
| [] := iff_of_true rfl (forall_mem_nil _)
2737
2734
| (l::L) := by simp only [join, append_eq_nil, join_eq_nil, forall_mem_cons]
@@ -2837,6 +2834,60 @@ begin
2837
2834
rw [← drop_take_succ_join_eq_nth_le, ← drop_take_succ_join_eq_nth_le, join_eq, length_eq] }
2838
2835
end
2839
2836
2837
+ /-! ### intersperse -/
2838
+ @[simp] lemma intersperse_nil {α : Type u} (a : α) : intersperse a [] = [] := rfl
2839
+
2840
+ @[simp] lemma intersperse_singleton {α : Type u} (a b : α) : intersperse a [b] = [b] := rfl
2841
+
2842
+ @[simp] lemma intersperse_cons_cons {α : Type u} (a b c : α) (tl : list α) :
2843
+ intersperse a (b :: c :: tl) = b :: a :: intersperse a (c :: tl) := rfl
2844
+
2845
+ /-! ### split_at and split_on -/
2846
+
2847
+ @[simp] theorem split_at_eq_take_drop : ∀ (n : ℕ) (l : list α), split_at n l = (take n l, drop n l)
2848
+ | 0 a := rfl
2849
+ | (succ n) [] := rfl
2850
+ | (succ n) (x :: xs) := by simp only [split_at, split_at_eq_take_drop n xs, take, drop]
2851
+
2852
+ @[simp] lemma split_on_nil {α : Type u} [decidable_eq α] (a : α) : [].split_on a = [[]] := rfl
2853
+
2854
+ /-- An auxiliary definition for proving a specification lemma for `split_on_p`.
2855
+
2856
+ `split_on_p_aux' P xs ys` splits the list `ys ++ xs` at every element satisfying `P`,
2857
+ where `ys` is an accumulating parameter for the initial segment of elements not satisfying `P`.
2858
+ -/
2859
+ def split_on_p_aux' {α : Type u} (P : α → Prop ) [decidable_pred P] : list α → list α → list (list α)
2860
+ | [] xs := [xs]
2861
+ | (h :: t) xs :=
2862
+ if P h then xs :: split_on_p_aux' t []
2863
+ else split_on_p_aux' t (xs ++ [h])
2864
+
2865
+ lemma split_on_p_aux_eq {α : Type u} (P : α → Prop ) [decidable_pred P] (xs ys : list α) :
2866
+ split_on_p_aux' P xs ys = split_on_p_aux P xs ((++) ys) :=
2867
+ begin
2868
+ induction xs with a t ih generalizing ys; simp! only [append_nil, eq_self_iff_true, and_self],
2869
+ split_ifs; rw ih,
2870
+ { refine ⟨rfl, rfl⟩ },
2871
+ { congr, ext, simp }
2872
+ end
2873
+
2874
+ lemma split_on_p_aux_nil {α : Type u} (P : α → Prop ) [decidable_pred P] (xs : list α) :
2875
+ split_on_p_aux P xs id = split_on_p_aux' P xs [] :=
2876
+ by { rw split_on_p_aux_eq, refl }
2877
+
2878
+ /-- The original list `L` can be recovered by joining the lists produced by `split_on_p p L`,
2879
+ interspersed with the elements `L.filter p`. -/
2880
+ lemma split_on_p_spec {α : Type u} (p : α → Prop ) [decidable_pred p] (as : list α) :
2881
+ join (zip_with (++) (split_on_p p as) ((as.filter p).map (λ x, [x]) ++ [[]])) = as :=
2882
+ begin
2883
+ rw [split_on_p, split_on_p_aux_nil],
2884
+ suffices : ∀ xs,
2885
+ join (zip_with (++) (split_on_p_aux' p as xs) ((as.filter p).map(λ x, [x]) ++ [[]])) = xs ++ as,
2886
+ { rw this , refl },
2887
+ induction as; intro; simp! only [split_on_p_aux', append_nil],
2888
+ split_ifs; simp [zip_with, join, *],
2889
+ end
2890
+
2840
2891
/-! ### lexicographic ordering -/
2841
2892
2842
2893
/-- Given a strict order `<` on `α`, the lexicographic strict order on `list α`, for which
0 commit comments