Skip to content

Commit 97bc609

Browse files
authored
feat: add have forms of let_* simp lemmas (#8790)
This PR adds `have` forms of simp lemmas that will be used in a future `have` simplifier. This depends on #8751 and future elaboration changes, since these are meant to elaborate using `Expr.letE (nondep := true) ..` expressions; for now they are duplicates of the `letFun_*` lemmas.
1 parent cdc9231 commit 97bc609

File tree

1 file changed

+22
-0
lines changed

1 file changed

+22
-0
lines changed

src/Init/SimpLemmas.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -74,6 +74,28 @@ theorem let_body_congr {α : Sort u} {β : α → Sort v} {b b' : (a : α) →
7474
(a : α) (h : ∀ x, b x = b' x) : (let x := a; b x) = (let x := a; b' x) :=
7575
(funext h : b = b') ▸ rfl
7676

77+
theorem have_unused {α : Sort u} {β : Sort v} (a : α) {b b' : β}
78+
(h : b = b') : (have _ := a; b) = b' := h
79+
80+
theorem have_unused_dep {α : Sort u} {β : Sort v} (a : α) {b : α → β} {b' : β}
81+
(h : ∀ x, b x = b') : (have x := a; b x) = b' := h a
82+
83+
theorem have_congr {α : Sort u} {β : Sort v} {a a' : α} {f f' : α → β}
84+
(h₁ : a = a') (h₂ : ∀ x, f x = f' x) : (have x := a; f x) = (have x := a'; f' x) :=
85+
@congr α β f f' a a' (funext h₂) h₁
86+
87+
theorem have_val_congr {α : Sort u} {β : Sort v} {a a' : α} {f : α → β}
88+
(h : a = a') : (have x := a; f x) = (have x := a'; f x) :=
89+
@congrArg α β a a' f h
90+
91+
theorem have_body_congr_dep {α : Sort u} {β : α → Sort v} (a : α) {f f' : (x : α) → β x}
92+
(h : ∀ x, f x = f' x) : (have x := a; f x) = (have x := a; f' x) :=
93+
h a
94+
95+
theorem have_body_congr {α : Sort u} {β : Sort v} (a : α) {f f' : α → β}
96+
(h : ∀ x, f x = f' x) : (have x := a; f x) = (have x := a; f' x) :=
97+
h a
98+
7799
theorem letFun_unused {α : Sort u} {β : Sort v} (a : α) {b b' : β} (h : b = b') : @letFun α (fun _ => β) a (fun _ => b) = b' :=
78100
h
79101

0 commit comments

Comments
 (0)