Skip to content
Merged
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
7 changes: 7 additions & 0 deletions src/Init/ByCases.lean
Original file line number Diff line number Diff line change
Expand Up @@ -44,3 +44,10 @@ theorem apply_ite (f : α → β) (P : Prop) [Decidable P] (x y : α) :
/-- A `dite` whose results do not actually depend on the condition may be reduced to an `ite`. -/
@[simp] theorem dite_eq_ite [Decidable P] :
(dite P (fun _ => a) (fun _ => b)) = ite P a b := rfl

-- Remark: dite and ite are "defally equal" when we ignore the proofs.
@[deprecated dite_eq_ite (since := "2025-10-29")]
theorem dif_eq_if (c : Prop) {h : Decidable c} {α : Sort u} (t : α) (e : α) : dite c (fun _ => t) (fun _ => e) = ite c t e :=
match h with
| isTrue _ => rfl
| isFalse _ => rfl
6 changes: 0 additions & 6 deletions src/Init/Core.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1196,12 +1196,6 @@ theorem dif_neg {c : Prop} {h : Decidable c} (hnc : ¬c) {α : Sort u} {t : c
| isTrue hc => absurd hc hnc
| isFalse _ => rfl

-- Remark: dite and ite are "defally equal" when we ignore the proofs.
theorem dif_eq_if (c : Prop) {h : Decidable c} {α : Sort u} (t : α) (e : α) : dite c (fun _ => t) (fun _ => e) = ite c t e :=
match h with
| isTrue _ => rfl
| isFalse _ => rfl

@[macro_inline]
instance {c t e : Prop} [dC : Decidable c] [dT : Decidable t] [dE : Decidable e] : Decidable (if c then t else e) :=
match dC with
Expand Down
4 changes: 2 additions & 2 deletions src/Init/Data/Array/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2252,8 +2252,6 @@ theorem flatMap_def {xs : Array α} {f : α → Array β} : xs.flatMap f = flatt
rcases xs with ⟨l⟩
simp [flatten_toArray, Function.comp_def, List.flatMap_def]

@[simp, grind =] theorem flatMap_empty {β} {f : α → Array β} : (#[] : Array α).flatMap f = #[] := rfl

theorem flatMap_toList {xs : Array α} {f : α → List β} :
xs.toList.flatMap f = (xs.flatMap (fun a => (f a).toArray)).toList := by
rcases xs with ⟨l⟩
Expand All @@ -2264,6 +2262,7 @@ theorem flatMap_toList {xs : Array α} {f : α → List β} :
rcases xs with ⟨l⟩
simp

@[deprecated List.flatMap_toArray_cons (since := "2025-10-29")]
theorem flatMap_toArray_cons {β} {f : α → Array β} {a : α} {as : List α} :
(a :: as).toArray.flatMap f = f a ++ as.toArray.flatMap f := by
simp [flatMap]
Expand All @@ -2274,6 +2273,7 @@ theorem flatMap_toArray_cons {β} {f : α → Array β} {a : α} {as : List α}
intro cs
induction as generalizing cs <;> simp_all

@[deprecated List.flatMap_toArray (since := "2025-10-29")]
theorem flatMap_toArray {β} {f : α → Array β} {as : List α} :
as.toArray.flatMap f = (as.flatMap (fun a => (f a).toList)).toArray := by
simp
Expand Down
5 changes: 2 additions & 3 deletions src/Init/Data/BitVec/Bitblast.lean
Original file line number Diff line number Diff line change
Expand Up @@ -635,12 +635,11 @@ theorem mulRec_eq_mul_signExtend_setWidth (x y : BitVec w) (s : Nat) :
simp only [mulRec_zero_eq, ofNat_eq_ofNat, Nat.reduceAdd]
by_cases y.getLsbD 0
case pos hy =>
simp only [hy, ↓reduceIte, setWidth_one_eq_ofBool_getLsb_zero,
ofBool_true, ofNat_eq_ofNat]
simp only [hy, ↓reduceIte, setWidth_one, ofBool_true, ofNat_eq_ofNat]
rw [setWidth_ofNat_one_eq_ofNat_one_of_lt (by omega)]
simp
case neg hy =>
simp [hy, setWidth_one_eq_ofBool_getLsb_zero]
simp [hy, setWidth_one]
case succ s' hs =>
rw [mulRec_succ_eq, hs]
have heq :
Expand Down
35 changes: 21 additions & 14 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -64,6 +64,7 @@ theorem getElem?_eq_none_iff {l : BitVec w} : l[n]? = none ↔ w ≤ n := by
theorem none_eq_getElem?_iff {l : BitVec w} : none = l[n]? ↔ w ≤ n := by
simp

@[simp]
theorem getElem?_eq_none {l : BitVec w} (h : w ≤ n) : l[n]? = none := getElem?_eq_none_iff.mpr h

theorem getElem?_eq (l : BitVec w) (i : Nat) :
Expand Down Expand Up @@ -158,7 +159,8 @@ theorem getLsbD_eq_getMsbD (x : BitVec w) (i : Nat) : x.getLsbD i = (decide (i <
apply getLsbD_of_ge
omega

@[simp] theorem getElem?_of_ge (x : BitVec w) (i : Nat) (ge : w ≤ i) : x[i]? = none := by
@[deprecated getElem?_eq_none (since := "2025-10-29")]
theorem getElem?_of_ge (x : BitVec w) (i : Nat) (ge : w ≤ i) : x[i]? = none := by
simp [ge]

@[simp] theorem getMsb?_of_ge (x : BitVec w) (i : Nat) (ge : w ≤ i) : getMsb? x i = none := by
Expand Down Expand Up @@ -415,12 +417,18 @@ theorem getElem?_zero_ofNat_one : (BitVec.ofNat (w+1) 1)[0]? = some true := by

-- This does not need to be a `@[simp]` theorem as it is already handled by `getElem?_eq_getElem`.
theorem getElem?_zero_ofBool (b : Bool) : (ofBool b)[0]? = some b := by
simp only [ofBool, ofNat_eq_ofNat, cond_eq_if]
simp only [ofBool, ofNat_eq_ofNat, cond_eq_ite]
split <;> simp_all

@[simp, grind =] theorem getElem_zero_ofBool (b : Bool) : (ofBool b)[0] = b := by
@[simp, grind =]
theorem getElem_ofBool_zero {b : Bool} : (ofBool b)[0] = b := by
rw [getElem_eq_iff, getElem?_zero_ofBool]


@[deprecated getElem_ofBool_zero (since := "2025-10-29")]
theorem getElem_zero_ofBool (b : Bool) : (ofBool b)[0] = b := by
simp

theorem getElem?_succ_ofBool (b : Bool) (i : Nat) : (ofBool b)[i + 1]? = none := by
simp

Expand All @@ -431,8 +439,6 @@ theorem getLsbD_ofBool (b : Bool) (i : Nat) : (ofBool b).getLsbD i = ((i = 0) &&
· simp only [ofBool, ofNat_eq_ofNat, cond_true, getLsbD_ofNat, Bool.and_true]
by_cases hi : i = 0 <;> simp [hi] <;> omega

theorem getElem_ofBool_zero {b : Bool} : (ofBool b)[0] = b := by simp

@[simp]
theorem getElem_ofBool {b : Bool} {h : i < 1}: (ofBool b)[i] = b := by
simp [← getLsbD_eq_getElem]
Expand Down Expand Up @@ -580,7 +586,7 @@ theorem toInt_eq_toNat_bmod (x : BitVec n) : x.toInt = Int.bmod x.toNat (2^n) :=
simp only [toInt_eq_toNat_cond]
split
next g =>
rw [Int.bmod_pos] <;> simp only [←Int.natCast_emod, toNat_mod_cancel]
rw [Int.bmod_eq_emod_of_lt] <;> simp only [←Int.natCast_emod, toNat_mod_cancel]
omega
next g =>
rw [Int.bmod_neg] <;> simp only [←Int.natCast_emod, toNat_mod_cancel]
Expand Down Expand Up @@ -988,7 +994,14 @@ theorem msb_setWidth' (x : BitVec w) (h : w ≤ v) : (x.setWidth' h).msb = (deci
theorem msb_setWidth'' (x : BitVec w) : (x.setWidth (k + 1)).msb = x.getLsbD k := by
simp [BitVec.msb, getMsbD]

/-- Truncating to width 1 produces a bitvector equal to the least significant bit. -/
theorem setWidth_one {x : BitVec w} :
x.setWidth 1 = ofBool (x.getLsbD 0) := by
ext i
simp [show i = 0 by omega]

/-- zero extending a bitvector to width 1 equals the boolean of the lsb. -/
@[deprecated setWidth_one (since := "2025-10-29")]
theorem setWidth_one_eq_ofBool_getLsb_zero (x : BitVec w) :
x.setWidth 1 = BitVec.ofBool (x.getLsbD 0) := by
ext i h
Expand All @@ -1004,12 +1017,6 @@ theorem setWidth_ofNat_one_eq_ofNat_one_of_lt {v w : Nat} (hv : 0 < v) :
have hv := (@Nat.testBit_one_eq_true_iff_self_eq_zero i)
by_cases h : Nat.testBit 1 i = true <;> simp_all

/-- Truncating to width 1 produces a bitvector equal to the least significant bit. -/
theorem setWidth_one {x : BitVec w} :
x.setWidth 1 = ofBool (x.getLsbD 0) := by
ext i
simp [show i = 0 by omega]

@[simp, grind =] theorem setWidth_ofNat_of_le (h : v ≤ w) (x : Nat) : setWidth v (BitVec.ofNat w x) = BitVec.ofNat v x := by
apply BitVec.eq_of_toNat_eq
simp only [toNat_setWidth, toNat_ofNat]
Expand Down Expand Up @@ -1639,11 +1646,11 @@ theorem not_def {x : BitVec v} : ~~~x = allOnes v ^^^ x := rfl

@[simp] theorem ofInt_negSucc_eq_not_ofNat {w n : Nat} :
BitVec.ofInt w (Int.negSucc n) = ~~~.ofNat w n := by
simp only [BitVec.ofInt, Int.toNat, Int.ofNat_eq_coe, toNat_eq, toNat_ofNatLT, toNat_not,
simp only [BitVec.ofInt, Int.toNat, Int.ofNat_eq_natCast, toNat_eq, toNat_ofNatLT, toNat_not,
toNat_ofNat]
cases h : Int.negSucc n % ((2 ^ w : Nat) : Int)
case ofNat =>
rw [Int.ofNat_eq_coe, Int.negSucc_emod] at h
rw [Int.ofNat_eq_natCast, Int.negSucc_emod] at h
· dsimp only
omega
· omega
Expand Down
3 changes: 2 additions & 1 deletion src/Init/Data/Bool.lean
Original file line number Diff line number Diff line change
Expand Up @@ -514,6 +514,7 @@ theorem exists_bool {p : Bool → Prop} : (∃ b, p b) ↔ p false ∨ p true :=
theorem cond_eq_ite {α} (b : Bool) (t e : α) : cond b t e = if b then t else e := by
cases b <;> simp

@[deprecated cond_eq_ite (since := "2025-10-29")]
theorem cond_eq_if : (bif b then x else y) = (if b then x else y) := cond_eq_ite b x y

@[simp] theorem cond_not (b : Bool) (t e : α) : cond (!b) t e = cond b e t := by
Expand Down Expand Up @@ -612,7 +613,7 @@ theorem decide_beq_decide (p q : Prop) [dpq : Decidable (p ↔ q)] [dp : Decidab

end Bool

export Bool (cond_eq_if xor and or not)
export Bool (cond_eq_if cond_eq_ite xor and or not)

/-! ### decide -/

Expand Down
6 changes: 3 additions & 3 deletions src/Init/Data/Dyadic/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -287,7 +287,7 @@ theorem toRat_add (x y : Dyadic) : toRat (x + y) = toRat x + toRat y := by
· rename_i h
cases Int.sub_eq_iff_eq_add.mp h
rw [toRat_ofOdd_eq_mkRat, Rat.mkRat_eq_iff (NeZero.ne _) (NeZero.ne _)]
simp only [succ_eq_add_one, Int.ofNat_eq_coe, Int.add_shiftLeft, ← Int.shiftLeft_add,
simp only [succ_eq_add_one, Int.ofNat_eq_natCast, Int.add_shiftLeft, ← Int.shiftLeft_add,
Int.natCast_mul, Int.natCast_shiftLeft, Int.shiftLeft_mul_shiftLeft, Int.add_mul]
congr 2 <;> omega
· rename_i h
Expand Down Expand Up @@ -438,7 +438,7 @@ theorem toDyadic_mkRat (a : Int) (b : Nat) (prec : Int) :
rcases h : mkRat a b with ⟨n, d, hnz, hr⟩
obtain ⟨m, hm, rfl, rfl⟩ := Rat.mkRat_num_den hb h
cases prec
· simp only [Rat.toDyadic, Int.ofNat_eq_coe, Int.toNat_natCast, Int.toNat_neg_natCast,
· simp only [Rat.toDyadic, Int.ofNat_eq_natCast, Int.toNat_natCast, Int.toNat_neg_natCast,
shiftLeft_zero, Int.natCast_mul]
rw [Int.mul_comm d, ← Int.ediv_ediv (by simp), ← Int.shiftLeft_mul,
Int.mul_ediv_cancel _ (by simpa using hm)]
Expand All @@ -463,7 +463,7 @@ theorem toRat_toDyadic (x : Rat) (prec : Int) :
rw [Rat.floor_def, Int.shiftLeft_eq, Nat.shiftLeft_eq]
match prec with
| .ofNat prec =>
simp only [Int.ofNat_eq_coe, Int.toNat_natCast, Int.toNat_neg_natCast, Nat.pow_zero,
simp only [Int.ofNat_eq_natCast, Int.toNat_natCast, Int.toNat_neg_natCast, Nat.pow_zero,
Nat.mul_one]
have : (2 ^ prec : Rat) = ((2 ^ prec : Nat) : Rat) := by simp
rw [Rat.zpow_natCast, this, Rat.mul_def']
Expand Down
2 changes: 1 addition & 1 deletion src/Init/Data/Dyadic/Round.lean
Original file line number Diff line number Diff line change
Expand Up @@ -36,7 +36,7 @@ theorem roundDown_le {x : Dyadic} {prec : Int} : roundDown x prec ≤ x :=
refine Lean.Grind.OrderedRing.mul_le_mul_of_nonneg_right ?_ (Rat.zpow_nonneg (by decide))
rw [Int.shiftRight_eq_div_pow]
rw [← Lean.Grind.Field.IsOrdered.mul_le_mul_iff_of_pos_right (c := 2^(Int.ofNat l)) (Rat.zpow_pos (by decide))]
simp only [Int.natCast_pow, Int.cast_ofNat_Int, Int.ofNat_eq_coe]
simp only [Int.natCast_pow, Int.cast_ofNat_Int, Int.ofNat_eq_natCast]
rw [Rat.mul_assoc, ← Rat.zpow_add (by decide), Int.add_left_neg, Rat.zpow_zero, Rat.mul_one]
have : (2 : Rat) ^ (l : Int) = (2 ^ l : Int) := by
rw [Rat.zpow_natCast, Rat.intCast_pow, Rat.intCast_ofNat]
Expand Down
12 changes: 7 additions & 5 deletions src/Init/Data/Fin/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -547,7 +547,7 @@ theorem succ_cast_eq {n' : Nat} (i : Fin n) (h : n = n') :
@[simp] theorem cast_castSucc {n' : Nat} {h : n + 1 = n' + 1} {i : Fin n} :
i.castSucc.cast h = (i.cast (Nat.succ.inj h)).castSucc := rfl

theorem castSucc_lt_succ (i : Fin n) : i.castSucc < i.succ :=
theorem castSucc_lt_succ {i : Fin n} : i.castSucc < i.succ :=
lt_def.2 <| by simp only [coe_castSucc, val_succ, Nat.lt_succ_self]

theorem le_castSucc_iff {i : Fin (n + 1)} {j : Fin n} : i ≤ j.castSucc ↔ i < j.succ := by
Expand Down Expand Up @@ -587,15 +587,20 @@ theorem castSucc_pos [NeZero n] {i : Fin n} (h : 0 < i) : 0 < i.castSucc := by
theorem castSucc_ne_zero_iff [NeZero n] {a : Fin n} : a.castSucc ≠ 0 ↔ a ≠ 0 :=
not_congr <| castSucc_eq_zero_iff

@[simp, grind _=_]
theorem castSucc_succ (i : Fin n) : i.succ.castSucc = i.castSucc.succ := rfl

@[deprecated castSucc_succ (since := "2025-10-29")]
theorem castSucc_fin_succ (n : Nat) (j : Fin n) :
j.succ.castSucc = (j.castSucc).succ := by simp [Fin.ext_iff]
j.succ.castSucc = (j.castSucc).succ := by simp

@[simp]
theorem coeSucc_eq_succ {a : Fin n} : a.castSucc + 1 = a.succ := by
cases n
· exact a.elim0
· simp [Fin.ext_iff, add_def, Nat.mod_eq_of_lt (Nat.succ_lt_succ a.is_lt)]

@[deprecated castSucc_lt_succ (since := "2025-10-29")]
theorem lt_succ {a : Fin n} : a.castSucc < a.succ := by
rw [castSucc, lt_def, coe_castAdd, val_succ]; exact Nat.lt_succ_self a.val

Expand Down Expand Up @@ -699,9 +704,6 @@ theorem rev_castSucc (k : Fin n) : rev (castSucc k) = succ (rev k) := k.rev_cast

theorem rev_succ (k : Fin n) : rev (succ k) = castSucc (rev k) := k.rev_addNat 1

@[simp, grind _=_]
theorem castSucc_succ (i : Fin n) : i.succ.castSucc = i.castSucc.succ := rfl

@[simp, grind =]
theorem castLE_refl (h : n ≤ n) (i : Fin n) : i.castLE h = i := rfl

Expand Down
5 changes: 4 additions & 1 deletion src/Init/Data/Int/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -80,7 +80,10 @@ protected theorem zero_ne_one : (0 : Int) ≠ 1 := nofun

/-! ## Coercions -/

@[simp] theorem ofNat_eq_coe : Int.ofNat n = Nat.cast n := rfl
@[simp] theorem ofNat_eq_natCast (n : Nat) : Int.ofNat n = n := rfl

@[deprecated ofNat_eq_natCast (since := "2025-10-29")]
theorem ofNat_eq_coe : Int.ofNat n = Nat.cast n := rfl

@[simp] theorem ofNat_zero : ((0 : Nat) : Int) = 0 := rfl

Expand Down
12 changes: 6 additions & 6 deletions src/Init/Data/Int/Bitwise/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -47,10 +47,10 @@ theorem shiftRight_zero (n : Int) : n >>> 0 = n := by
simp [Int.shiftRight_eq_div_pow]

theorem le_shiftRight_of_nonpos {n : Int} {s : Nat} (h : n ≤ 0) : n ≤ n >>> s := by
simp only [Int.shiftRight_eq, Int.shiftRight, Int.ofNat_eq_coe]
simp only [Int.shiftRight_eq, Int.shiftRight, Int.ofNat_eq_natCast]
split
case _ _ _ m =>
simp only [ofNat_eq_coe] at h
simp only [ofNat_eq_natCast] at h
by_cases hm : m = 0
· simp [hm]
· omega
Expand All @@ -61,14 +61,14 @@ theorem le_shiftRight_of_nonpos {n : Int} {s : Nat} (h : n ≤ 0) : n ≤ n >>>
omega

theorem shiftRight_le_of_nonneg {n : Int} {s : Nat} (h : 0 ≤ n) : n >>> s ≤ n := by
simp only [Int.shiftRight_eq, Int.shiftRight, Int.ofNat_eq_coe]
simp only [Int.shiftRight_eq, Int.shiftRight, Int.ofNat_eq_natCast]
split
case _ _ _ m =>
simp only [Int.ofNat_eq_coe] at h
simp only [Int.ofNat_eq_natCast] at h
by_cases hm : m = 0
· simp [hm]
· have := Nat.shiftRight_le m s
rw [ofNat_eq_coe]
rw [ofNat_eq_natCast]
omega
case _ _ _ m =>
omega
Expand Down Expand Up @@ -108,7 +108,7 @@ theorem shiftLeft_succ (m : Int) (n : Nat) : m <<< (n + 1) = (m <<< n) * 2 := by
change Int.shiftLeft _ _ = Int.shiftLeft _ _ * 2
match m with
| (m : Nat) =>
dsimp only [Int.shiftLeft, Int.ofNat_eq_coe]
dsimp only [Int.shiftLeft, Int.ofNat_eq_natCast]
rw [Nat.shiftLeft_succ, Nat.mul_comm, natCast_mul, ofNat_two]
| Int.negSucc m =>
dsimp only [Int.shiftLeft]
Expand Down
Loading
Loading