Skip to content
This repository has been archived by the owner on Jul 24, 2024. It is now read-only.

[Merged by Bors] - refactor(data/list/basic): Remove many redundant hypotheses #12950

Closed
wants to merge 1 commit into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
17 changes: 9 additions & 8 deletions src/data/list/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -676,27 +676,28 @@ end
/-! ### last -/

@[simp] theorem last_cons {a : α} {l : list α} :
∀ (h : a :: l ≠ nil) (h₂ : l ≠ nil), last (a :: l) h₁ = last l h :=
∀ (h : l ≠ nil), last (a :: l) (cons_ne_nil a l) = last l h :=
by {induction l; intros, contradiction, reflexivity}

@[simp] theorem last_append {a : α} (l : list α) (h : l ++ [a] ≠ []) : last (l ++ [a]) h = a :=
@[simp] theorem last_append {a : α} (l : list α) :
last (l ++ [a]) (append_ne_nil_of_ne_nil_right l _ (cons_ne_nil a _)) = a :=
by induction l;
[refl, simp only [cons_append, last_cons _ (λ H, cons_ne_nil _ _ (append_eq_nil.1 H).2), *]]
[refl, simp only [cons_append, last_cons (λ H, cons_ne_nil _ _ (append_eq_nil.1 H).2), *]]

theorem last_concat {a : α} (l : list α) (h : concat l a ≠ []) : last (concat l a) h = a :=
theorem last_concat {a : α} (l : list α) : last (concat l a) (concat_ne_nil a l) = a :=
by simp only [concat_eq_append, last_append]

@[simp] theorem last_singleton (a : α) (h : [a] ≠ []) : last [a] h = a := rfl
@[simp] theorem last_singleton (a : α) : last [a] (cons_ne_nil a []) = a := rfl

@[simp] theorem last_cons_cons (a₁ a₂ : α) (l : list α) (h : a₁::a₂::l ≠ []) :
last (a₁::a₂::l) h = last (a₂::l) (cons_ne_nil a₂ l) := rfl
@[simp] theorem last_cons_cons (a₁ a₂ : α) (l : list α) :
last (a₁::a₂::l) (cons_ne_nil _ _) = last (a₂::l) (cons_ne_nil a₂ l) := rfl

theorem init_append_last : ∀ {l : list α} (h : l ≠ []), init l ++ [last l h] = l
| [] h := absurd rfl h
| [a] h := rfl
| (a::b::l) h :=
begin
rw [init, cons_append, last_cons (cons_ne_nil _ _) (cons_ne_nil _ _)],
rw [init, cons_append, last_cons (cons_ne_nil _ _)],
congr,
exact init_append_last (cons_ne_nil b l)
end
Expand Down
4 changes: 2 additions & 2 deletions src/data/nat/digits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -224,7 +224,7 @@ begin
{ simp, },
{ intros l m, apply w₁, exact list.mem_cons_of_mem _ m, },
{ intro h,
{ rw [list.last_cons _ h] at w₂,
{ rw [list.last_cons h] at w₂,
convert w₂, }}},
{ convert w₁ d (list.mem_cons_self _ _), simp, },
{ by_cases h' : L = [],
Expand All @@ -237,7 +237,7 @@ begin
apply nat.pos_of_ne_zero,
contrapose! w₂,
apply digits_zero_of_eq_zero _ w₂,
{ rw list.last_cons _ h',
{ rw list.last_cons h',
exact list.last_mem h', },
{ exact le_of_lt h, }, }, }, },
end
Expand Down