@@ -2053,6 +2053,32 @@ theorem eq_msb_cons_setWidth (x : BitVec (w+1)) : x = (cons x.msb (x.setWidth w)
2053
2053
ext i
2054
2054
simp [cons]
2055
2055
2056
+
2057
+ theorem cons_append (x : BitVec w₁) (y : BitVec w₂) (a : Bool) :
2058
+ (cons a x) ++ y = (cons a (x ++ y)).cast (by omega) := by
2059
+ apply eq_of_toNat_eq
2060
+ simp only [toNat_append, toNat_cons, toNat_cast]
2061
+ rw [Nat.shiftLeft_add, Nat.shiftLeft_or_distrib, Nat.or_assoc]
2062
+
2063
+ theorem cons_append_append (x : BitVec w₁) (y : BitVec w₂) (z : BitVec w₃) (a : Bool) :
2064
+ (cons a x) ++ y ++ z = (cons a (x ++ y ++ z)).cast (by omega) := by
2065
+ ext i h
2066
+ simp only [cons, getLsbD_append, getLsbD_cast, getLsbD_ofBool, cast_cast]
2067
+ by_cases h₀ : i < w₁ + w₂ + w₃
2068
+ · simp only [h₀, ↓reduceIte]
2069
+ by_cases h₁ : i < w₃
2070
+ · simp [h₁]
2071
+ · simp only [h₁, ↓reduceIte]
2072
+ by_cases h₂ : i - w₃ < w₂
2073
+ · simp [h₂]
2074
+ · simp [h₂]
2075
+ omega
2076
+ · simp only [show ¬i - w₃ - w₂ < w₁ by omega, ↓reduceIte, show i - w₃ - w₂ - w₁ = 0 by omega,
2077
+ decide_true, Bool.true_and, h₀, show i - (w₁ + w₂ + w₃) = 0 by omega]
2078
+ by_cases h₂ : i < w₃
2079
+ · simp [h₂]; omega
2080
+ · simp [h₂]; omega
2081
+
2056
2082
/-! ### concat -/
2057
2083
2058
2084
@[simp] theorem toNat_concat (x : BitVec w) (b : Bool) :
@@ -3373,11 +3399,11 @@ theorem and_one_eq_setWidth_ofBool_getLsbD {x : BitVec w} :
3373
3399
ext (_ | i) h <;> simp [Bool.and_comm]
3374
3400
3375
3401
@[simp]
3376
- theorem replicate_zero_eq {x : BitVec w} : x.replicate 0 = 0 #0 := by
3402
+ theorem replicate_zero {x : BitVec w} : x.replicate 0 = 0 #0 := by
3377
3403
simp [replicate]
3378
3404
3379
3405
@[simp]
3380
- theorem replicate_succ_eq {x : BitVec w} :
3406
+ theorem replicate_succ {x : BitVec w} :
3381
3407
x.replicate (n + 1 ) =
3382
3408
(x ++ replicate n x).cast (by rw [Nat.mul_succ]; omega) := by
3383
3409
simp [replicate]
@@ -3389,7 +3415,7 @@ theorem getLsbD_replicate {n w : Nat} (x : BitVec w) :
3389
3415
induction n generalizing x
3390
3416
case zero => simp
3391
3417
case succ n ih =>
3392
- simp only [replicate_succ_eq , getLsbD_cast, getLsbD_append]
3418
+ simp only [replicate_succ , getLsbD_cast, getLsbD_append]
3393
3419
by_cases hi : i < w * (n + 1 )
3394
3420
· simp only [hi, decide_true, Bool.true_and]
3395
3421
by_cases hi' : i < w * n
@@ -3406,6 +3432,33 @@ theorem getElem_replicate {n w : Nat} (x : BitVec w) (h : i < w * n) :
3406
3432
simp only [← getLsbD_eq_getElem, getLsbD_replicate]
3407
3433
by_cases h' : w = 0 <;> simp [h'] <;> omega
3408
3434
3435
+ theorem append_assoc {x₁ : BitVec w₁} {x₂ : BitVec w₂} {x₃ : BitVec w₃} :
3436
+ (x₁ ++ x₂) ++ x₃ = (x₁ ++ (x₂ ++ x₃)).cast (by omega) := by
3437
+ induction w₁ generalizing x₂ x₃
3438
+ case zero => simp
3439
+ case succ n ih =>
3440
+ specialize @ih (setWidth n x₁)
3441
+ rw [← cons_msb_setWidth x₁, cons_append_append, ih, cons_append]
3442
+ ext j h
3443
+ simp [getLsbD_cons, show n + w₂ + w₃ = n + (w₂ + w₃) by omega]
3444
+
3445
+ theorem replicate_append_self {x : BitVec w} :
3446
+ x ++ x.replicate n = (x.replicate n ++ x).cast (by omega) := by
3447
+ induction n with
3448
+ | zero => simp
3449
+ | succ n ih =>
3450
+ rw [replicate_succ]
3451
+ conv => lhs; rw [ih]
3452
+ simp only [cast_cast, cast_eq]
3453
+ rw [← cast_append_left]
3454
+ · rw [append_assoc]; congr
3455
+ · rw [Nat.add_comm, Nat.mul_add, Nat.mul_one]; omega
3456
+
3457
+ theorem replicate_succ' {x : BitVec w} :
3458
+ x.replicate (n + 1 ) =
3459
+ (replicate n x ++ x).cast (by rw [Nat.mul_succ]) := by
3460
+ simp [replicate_append_self]
3461
+
3409
3462
/-! ### intMin -/
3410
3463
3411
3464
/-- The bitvector of width `w` that has the smallest value when interpreted as an integer. -/
@@ -3691,6 +3744,57 @@ theorem toInt_abs_eq_natAbs_of_ne_intMin {x : BitVec w} (hx : x ≠ intMin w) :
3691
3744
x.abs.toInt = x.toInt.natAbs := by
3692
3745
simp [toInt_abs_eq_natAbs, hx]
3693
3746
3747
+ /-! ### Reverse -/
3748
+
3749
+ theorem getLsbD_reverse {i : Nat} {x : BitVec w} :
3750
+ (x.reverse).getLsbD i = x.getMsbD i := by
3751
+ induction w generalizing i
3752
+ case zero => simp
3753
+ case succ n ih =>
3754
+ simp only [reverse, truncate_eq_setWidth, getLsbD_concat]
3755
+ rcases i with rfl | i
3756
+ · rfl
3757
+ · simp only [Nat.add_one_ne_zero, ↓reduceIte, Nat.add_one_sub_one, ih]
3758
+ rw [getMsbD_setWidth]
3759
+ simp only [show n - (n + 1 ) = 0 by omega, Nat.zero_le, decide_true, Bool.true_and]
3760
+ congr; omega
3761
+
3762
+ theorem getMsbD_reverse {i : Nat} {x : BitVec w} :
3763
+ (x.reverse).getMsbD i = x.getLsbD i := by
3764
+ simp only [getMsbD_eq_getLsbD, getLsbD_reverse]
3765
+ by_cases hi : i < w
3766
+ · simp only [hi, decide_true, show w - 1 - i < w by omega, Bool.true_and]
3767
+ congr; omega
3768
+ · simp [hi, show i ≥ w by omega]
3769
+
3770
+ theorem msb_reverse {x : BitVec w} :
3771
+ (x.reverse).msb = x.getLsbD 0 :=
3772
+ by rw [BitVec.msb, getMsbD_reverse]
3773
+
3774
+ theorem reverse_append {x : BitVec w} {y : BitVec v} :
3775
+ (x ++ y).reverse = (y.reverse ++ x.reverse).cast (by omega) := by
3776
+ ext i h
3777
+ simp only [getLsbD_append, getLsbD_reverse]
3778
+ by_cases hi : i < v
3779
+ · by_cases hw : w ≤ i
3780
+ · simp [getMsbD_append, getLsbD_cast, getLsbD_append, getLsbD_reverse, hw]
3781
+ · simp [getMsbD_append, getLsbD_cast, getLsbD_append, getLsbD_reverse, hw, show i < w by omega]
3782
+ · by_cases hw : w ≤ i
3783
+ · simp [getMsbD_append, getLsbD_cast, getLsbD_append, hw, show ¬ i < w by omega, getLsbD_reverse]
3784
+ · simp [getMsbD_append, getLsbD_cast, getLsbD_append, hw, show i < w by omega, getLsbD_reverse]
3785
+
3786
+ @[simp]
3787
+ theorem reverse_cast {w v : Nat} (h : w = v) (x : BitVec w) :
3788
+ (x.cast h).reverse = x.reverse.cast h := by
3789
+ subst h; simp
3790
+
3791
+ theorem reverse_replicate {n : Nat} {x : BitVec w} :
3792
+ (x.replicate n).reverse = (x.reverse).replicate n := by
3793
+ induction n with
3794
+ | zero => rfl
3795
+ | succ n ih =>
3796
+ conv => lhs; simp only [replicate_succ']
3797
+ simp [reverse_append, ih]
3694
3798
3695
3799
/-! ### Decidable quantifiers -/
3696
3800
@@ -3906,4 +4010,10 @@ abbrev shiftLeft_zero_eq := @shiftLeft_zero
3906
4010
@[deprecated ushiftRight_zero (since := "2024-10-27")]
3907
4011
abbrev ushiftRight_zero_eq := @ushiftRight_zero
3908
4012
4013
+ @[deprecated replicate_zero (since := "2025-01-08")]
4014
+ abbrev replicate_zero_eq := @replicate_zero
4015
+
4016
+ @[deprecated replicate_succ (since := "2025-01-08")]
4017
+ abbrev replicate_succ_eq := @replicate_succ
4018
+
3909
4019
end BitVec
0 commit comments