@@ -651,10 +651,6 @@ theorem length_dropLast : ∀ l : List α, length l.dropLast = length l - 1
651
651
simp
652
652
#align list.length_init List.length_dropLast
653
653
654
- -- Porting note: `rw [dropLast]` in Lean4 generates a goal `(b::l) ≠ []`
655
- -- so we use this lemma instead
656
- theorem dropLast_cons_cons (a b : α) (l : List α) : dropLast (a::b::l) = a::dropLast (b::l) := rfl
657
-
658
654
/-! ### getLast -/
659
655
660
656
@[simp]
@@ -699,7 +695,7 @@ theorem dropLast_append_getLast : ∀ {l : List α} (h : l ≠ []), dropLast l +
699
695
| [], h => absurd rfl h
700
696
| [a], h => rfl
701
697
| a :: b :: l, h => by
702
- rw [dropLast_cons_cons , cons_append, getLast_cons (cons_ne_nil _ _)]
698
+ rw [dropLast_cons₂ , cons_append, getLast_cons (cons_ne_nil _ _)]
703
699
congr
704
700
exact dropLast_append_getLast (cons_ne_nil b l)
705
701
#align list.init_append_last List.dropLast_append_getLast
@@ -782,7 +778,7 @@ theorem dropLast_append_getLast? : ∀ {l : List α}, ∀ a ∈ l.getLast?, drop
782
778
| [a], _, rfl => rfl
783
779
| a :: b :: l, c, hc => by
784
780
rw [getLast?_cons_cons] at hc
785
- rw [dropLast_cons_cons , cons_append, dropLast_append_getLast? _ hc]
781
+ rw [dropLast_cons₂ , cons_append, dropLast_append_getLast? _ hc]
786
782
#align list.init_append_last' List.dropLast_append_getLast?
787
783
788
784
theorem getLastI_eq_getLast? [Inhabited α] : ∀ l : List α, l.getLastI = l.getLast?.iget
@@ -1877,9 +1873,6 @@ theorem zipWith_flip (f : α → β → γ) : ∀ as bs, zipWith (flip f) bs as
1877
1873
1878
1874
/-! ### take, drop -/
1879
1875
1880
- @[simp]
1881
- theorem take_zero (l : List α) : take 0 l = [] :=
1882
- rfl
1883
1876
#align list.take_zero List.take_zero
1884
1877
1885
1878
#align list.take_nil List.take_nil
@@ -2043,7 +2036,7 @@ theorem dropLast_take {n : ℕ} {l : List α} (h : n < l.length) :
2043
2036
#align list.init_take List.dropLast_take
2044
2037
2045
2038
theorem dropLast_cons_of_ne_nil {α : Type *} {x : α}
2046
- {l : List α} (h : l ≠ []) : (x :: l).dropLast = x :: l.dropLast := by simp [h]
2039
+ {l : List α} (h : l ≠ []) : (x :: l).dropLast = x :: l.dropLast := by simp [h, dropLast ]
2047
2040
#align list.init_cons_of_ne_nil List.dropLast_cons_of_ne_nil
2048
2041
2049
2042
@[simp]
@@ -2343,26 +2336,9 @@ theorem foldr_ext (f g : α → β → β) (b : β) {l : List α} (H : ∀ a ∈
2343
2336
simp only [foldr, ih H.2 , H.1 ]
2344
2337
#align list.foldr_ext List.foldr_ext
2345
2338
2346
- @[simp]
2347
- theorem foldl_nil (f : α → β → α) (a : α) : foldl f a [] = a :=
2348
- rfl
2349
2339
#align list.foldl_nil List.foldl_nil
2350
-
2351
- @[simp]
2352
- theorem foldl_cons (f : α → β → α) (a : α) (b : β) (l : List β) :
2353
- foldl f a (b :: l) = foldl f (f a b) l :=
2354
- rfl
2355
2340
#align list.foldl_cons List.foldl_cons
2356
-
2357
- @[simp]
2358
- theorem foldr_nil (f : α → β → β) (b : β) : foldr f b [] = b :=
2359
- rfl
2360
2341
#align list.foldr_nil List.foldr_nil
2361
-
2362
- @[simp]
2363
- theorem foldr_cons (f : α → β → β) (b : β) (a : α) (l : List α) :
2364
- foldr f b (a :: l) = f a (foldr f b l) :=
2365
- rfl
2366
2342
#align list.foldr_cons List.foldr_cons
2367
2343
2368
2344
#align list.foldl_append List.foldl_append
@@ -2748,18 +2724,9 @@ section FoldlMFoldrM
2748
2724
2749
2725
variable {m : Type v → Type w} [Monad m]
2750
2726
2751
- @[simp]
2752
- theorem foldlM_nil (f : β → α → m β) {b} : List.foldlM f b [] = pure b :=
2753
- rfl
2754
2727
#align list.mfoldl_nil List.foldlM_nil
2755
-
2756
2728
-- Porting note: now in std
2757
2729
#align list.mfoldr_nil List.foldrM_nil
2758
-
2759
- @[simp]
2760
- theorem foldlM_cons {f : β → α → m β} {b a l} :
2761
- List.foldlM f b (a :: l) = f b a >>= fun b' => List.foldlM f b' l :=
2762
- rfl
2763
2730
#align list.mfoldl_cons List.foldlM_cons
2764
2731
2765
2732
/- Porting note: now in std; now assumes an instance of `LawfulMonad m`, so we make everything
@@ -2795,9 +2762,6 @@ end FoldlMFoldrM
2795
2762
2796
2763
/-! ### intersperse -/
2797
2764
2798
- @[simp]
2799
- theorem intersperse_nil {α : Type u} (a : α) : intersperse a [] = [] :=
2800
- rfl
2801
2765
#align list.intersperse_nil List.intersperse_nil
2802
2766
2803
2767
@[simp]
@@ -3195,18 +3159,13 @@ section find?
3195
3159
3196
3160
variable {p : α → Bool} {l : List α} {a : α}
3197
3161
3198
- @[simp]
3199
- theorem find?_nil (p : α → Bool) : find? p [] = none :=
3200
- rfl
3201
3162
#align list.find_nil List.find?_nil
3202
3163
3203
- -- Porting note: List.find? is given @[ simp ] in Std.Data.List.Init.Lemmas
3204
3164
-- @[ simp ]
3205
3165
-- Later porting note (at time of this lemma moving to Std): removing attribute `nolint simpNF`
3206
3166
attribute [simp 1100] find?_cons_of_pos
3207
3167
#align list.find_cons_of_pos List.find?_cons_of_pos
3208
3168
3209
- -- Porting note: List.find? is given @[ simp ] in Std.Data.List.Init.Lemmas
3210
3169
-- @[ simp ]
3211
3170
-- Later porting note (at time of this lemma moving to Std): removing attribute `nolint simpNF`
3212
3171
attribute [simp 1100] find?_cons_of_neg
@@ -3818,15 +3777,8 @@ theorem enum_nil : enum ([] : List α) = [] :=
3818
3777
rfl
3819
3778
#align list.enum_nil List.enum_nil
3820
3779
3821
- @[simp]
3822
- theorem enumFrom_nil (n : ℕ) : enumFrom n ([] : List α) = [] :=
3823
- rfl
3824
3780
#align list.enum_from_nil List.enumFrom_nil
3825
3781
3826
- @[simp]
3827
- theorem enumFrom_cons (x : α) (xs : List α) (n : ℕ) :
3828
- enumFrom n (x :: xs) = (n, x) :: enumFrom (n + 1 ) xs :=
3829
- rfl
3830
3782
#align list.enum_from_cons List.enumFrom_cons
3831
3783
3832
3784
@[simp]
0 commit comments