Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: de-mathlib Nat.binaryRec #314

Closed
wants to merge 35 commits into from
Closed
Show file tree
Hide file tree
Changes from 10 commits
Commits
Show all changes
35 commits
Select commit Hold shift + click to select a range
dde0327
feat: de-mathlib `Nat.binaryRec`
FR-vdash-bot Oct 22, 2023
148c53f
specialize
FR-vdash-bot Oct 22, 2023
6b9474b
lint
FR-vdash-bot Oct 22, 2023
230166f
inline
FR-vdash-bot Oct 22, 2023
730782c
generalize & golf
FR-vdash-bot Oct 22, 2023
9da444d
mathlib dislike `binaryRec` with `elab_as_elim`
FR-vdash-bot Oct 22, 2023
6c6cd81
move & golf
FR-vdash-bot Oct 23, 2023
b33a41f
Merge remote-tracking branch 'upstream/main' into bitwise
FR-vdash-bot Nov 15, 2023
ff6786b
fix
FR-vdash-bot Nov 15, 2023
e638457
Merge remote-tracking branch 'upstream/main' into bitwise
FR-vdash-bot Dec 15, 2023
cff0bf7
fix merge
FR-vdash-bot Dec 15, 2023
615cfc9
fix merge & golf
FR-vdash-bot Dec 15, 2023
6982af0
feat: more `Bool` lemmas
FR-vdash-bot Dec 15, 2023
2aacb40
style
FR-vdash-bot Dec 15, 2023
a138ebb
Merge branch 'more_lemmas_for_bool' into bitwise
FR-vdash-bot Dec 15, 2023
3fed6e4
fix
FR-vdash-bot Dec 15, 2023
7607939
fix
FR-vdash-bot Dec 15, 2023
abd94e2
fix
FR-vdash-bot Dec 15, 2023
daf7634
fix
FR-vdash-bot Dec 15, 2023
198f289
Merge branch 'more_lemmas_for_bool' into bitwise
FR-vdash-bot Dec 15, 2023
3e9efc4
feat: more lemmas about equality
FR-vdash-bot Dec 15, 2023
fbf3b82
no `↦`
FR-vdash-bot Dec 15, 2023
03265b6
lint
FR-vdash-bot Dec 15, 2023
7738c3a
Merge branch 'more_lemmas_about_eq' into bitwise
FR-vdash-bot Dec 15, 2023
661debc
fix name
FR-vdash-bot Dec 18, 2023
ba2cdc5
it was not `elab_as_elim` in mathlib but eventually I think it should be
FR-vdash-bot Dec 19, 2023
f536348
implicit
FR-vdash-bot Dec 19, 2023
bd518d9
suggestion
FR-vdash-bot Dec 29, 2023
c5b411a
Merge remote-tracking branch 'upstream/main' into bitwise
FR-vdash-bot Jan 19, 2024
1736834
fix merge
FR-vdash-bot Jan 19, 2024
8901a71
Merge remote-tracking branch 'upstream/main' into more_lemmas_for_bool
FR-vdash-bot Jan 20, 2024
86627c6
Update Bool.lean
FR-vdash-bot Jan 20, 2024
70bbab4
Merge branch 'more_lemmas_for_bool' into bitwise
FR-vdash-bot Jan 20, 2024
67605c3
fix merge
FR-vdash-bot Jan 20, 2024
89a4e75
Merge remote-tracking branch 'upstream/main' into bitwise
FR-vdash-bot Jan 26, 2024
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
97 changes: 97 additions & 0 deletions Std/Data/Nat/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -135,3 +135,100 @@ of a number.

/-- `testBit m n` returns whether the `(n+1)` least significant bit is `1` or `0`-/
def testBit (m n : Nat) : Bool := (m >>> n) &&& 1 != 0

/-- `bodd n` returns `true` if `n` is odd-/
def bodd (n : Nat) : Bool :=
(1 &&& n) != 0

/-- `div2 n = ⌊n/2⌋` the greatest integer smaller than `n/2`-/
def div2 (n : Nat) : Nat :=
n >>> 1

/-- `bit b` appends the digit `b` to the binary representation of
its natural number input. -/
def bit (b : Bool) (n : Nat) : Nat :=
cond b (n + n + 1) (n + n)
joehendrix marked this conversation as resolved.
Show resolved Hide resolved

/-- `testBit m n` returns whether the `(n+1)ˢᵗ` least significant bit is `1` or `0`-/
def testBit (m n : Nat) : Bool :=
bodd (m >>> n)

theorem bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by
cases b <;> rw [Nat.mul_comm]
· exact congrArg (· + n) n.zero_add.symm
· exact congrArg (· + n + 1) n.zero_add.symm

theorem div2_val (n) : div2 n = n / 2 := rfl

theorem mod_two_eq_zero_or_one (n : Nat) : n % 2 = 0 ∨ n % 2 = 1 :=
match n % 2, @Nat.mod_lt n 2 (by simp) with
| 0, _ => .inl rfl
| 1, _ => .inr rfl

theorem mod_two_of_bodd (n : Nat) : n % 2 = cond (bodd n) 1 0 := by
dsimp [bodd, (· &&& ·), AndOp.and, land]
unfold bitwise
by_cases n0 : n = 0
· simp [n0]
· simp only [ite_false, decide_True, Bool.true_and, decide_eq_true_eq, n0,
show 1 / 2 = 0 by decide]
cases mod_two_eq_zero_or_one n with | _ h => simp [h]; rfl

theorem bodd_add_div2 (n : Nat) : 2 * div2 n + cond (bodd n) 1 0 = n := by
rw [← mod_two_of_bodd, div2_val, Nat.div_add_mod]

theorem bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n :=
(bit_val _ _).trans (bodd_add_div2 _)

theorem bit_eq_zero_iff {n : Nat} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by
cases n <;> cases b <;> simp [bit, Nat.succ_add]

/-- For a predicate `C : Nat → Sort u`, if instances can be
constructed for natural numbers of the form `bit b n`,
they can be constructed for any given natural number. -/
@[inline]
def bitCasesOn {C : Nat → Sort u} (n) (h : ∀ b n, C (bit b n)) : C n := bit_decomp n ▸ h _ _

theorem binaryRec_decreasing (h : n ≠ 0) : div2 n < n :=
div_lt_self (n.zero_lt_of_ne_zero h) (by decide)

/-- A recursion principle for `bit` representations of natural numbers.
For a predicate `C : Nat → Sort u`, if instances can be
constructed for natural numbers of the form `bit b n`,
they can be constructed for all natural numbers. -/
@[specialize]
def binaryRec {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) (n : Nat) : C n :=
if n0 : n = 0 then by rw [n0]; exact z
else by rw [← n.bit_decomp]; exact f n.bodd n.div2 (binaryRec z f n.div2)
decreasing_by exact binaryRec_decreasing n0

/-- The same as `binaryRec`, but the induction step can assume that if `n=0`,
the bit being appended is `true`-/
@[elab_as_elim, specialize]
def binaryRec' {C : Nat → Sort u} (z : C 0)
(f : ∀ b n, (n = 0 → b = true) → C n → C (bit b n)) : ∀ n, C n :=
binaryRec z fun b n ih =>
if h : n = 0 → b = true then f b n h ih
else by
have : bit b n = 0 := by
rw [bit_eq_zero_iff]
cases n <;> cases b <;> simp at h <;> simp [h]
exact this ▸ z

/-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/
@[elab_as_elim, specialize]
def binaryRecFromOne {C : Nat → Sort u} (z₀ : C 0) (z₁ : C 1)
(f : ∀ b n, n ≠ 0 → C n → C (bit b n)) : ∀ n, C n :=
binaryRec' z₀ fun b n h ih =>
if h' : n = 0 then by
rw [h', h h']
exact z₁
else f b n h' ih
/-!
### testBit
We define an operation for testing individual bits in the binary representation
of a number.
-/

/-- `testBit m n` returns whether the `(n+1)` least significant bit is `1` or `0`-/
def testBit (m n : Nat) : Bool := (m >>> n) &&& 1 != 0
148 changes: 146 additions & 2 deletions Std/Data/Nat/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,8 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro
-/
import Std.Logic
import Std.Tactic.Basic
import Std.Tactic.Alias
import Std.Tactic.RCases
import Std.Data.Bool
import Std.Data.Nat.Init.Lemmas
import Std.Data.Nat.Basic
import Std.Data.Ord
Expand Down Expand Up @@ -1046,6 +1047,27 @@ theorem mul_mod (a b n : Nat) : a * b % n = (a % n) * (b % n) % n := by
theorem add_mod (a b n : Nat) : (a + b) % n = ((a % n) + (b % n)) % n := by
rw [add_mod_mod, mod_add_mod]

@[simp]
theorem mod_two_ne_one (n : Nat) : ¬n % 2 = 1 ↔ n % 2 = 0 := by
cases mod_two_eq_zero_or_one n with | _ h => simp [h]

@[simp]
theorem mod_two_ne_zero (n : Nat) : ¬n % 2 = 0 ↔ n % 2 = 1 := by
cases mod_two_eq_zero_or_one n with | _ h => simp [h]

@[simp]
theorem mod_two_add_succ_mod_two (n : Nat) : n % 2 + (n + 1) % 2 = 1 := by
rw [add_mod]
cases mod_two_eq_zero_or_one n with | _ h => simp [h]

@[simp]
theorem succ_mod_two_add_mod_two (n : Nat) : (n + 1) % 2 + n % 2 = 1 := by
rw [Nat.add_comm, mod_two_add_succ_mod_two]

theorem succ_mod_two_eq_one_sub_mod_two (n : Nat) : (n + 1) % 2 = 1 - n % 2 := by
rw [add_mod]
cases mod_two_eq_zero_or_one n with | _ h => simp [h]

/-! ### pow -/

theorem pow_succ' {m n : Nat} : m ^ n.succ = m * m ^ n := by
Expand Down Expand Up @@ -1288,7 +1310,7 @@ theorem shiftRight_eq_div_pow (m : Nat) : ∀ n, m >>> n = m / 2 ^ n
| 0 => (Nat.div_one _).symm
| k + 1 => by
rw [shiftRight_add, shiftRight_eq_div_pow m k]
simp [Nat.div_div_eq_div_mul, ← Nat.pow_succ, shiftRight_succ]
simp [Nat.div_div_eq_div_mul, ← Nat.pow_succ]

theorem mul_add_div {m : Nat} (m_pos : m > 0) (x y : Nat) : (m * x + y) / m = x + y / m := by
FR-vdash-bot marked this conversation as resolved.
Show resolved Hide resolved
match x with
Expand All @@ -1310,3 +1332,125 @@ theorem mul_add_mod (m x y : Nat) : (m * x + y) % m = y % m := by
cases n
· exact (m % 0).div_zero
· case succ n => exact Nat.div_eq_of_lt (m.mod_lt n.succ_pos)

/-! ### shiftLeft -/

theorem shiftLeft_zero (m) : m <<< 0 = m := rfl

theorem shiftLeft_succ (m n) : m <<< (n + 1) = 2 * (m <<< n) := by
simp only [shiftLeft_eq, Nat.pow_add, Nat.pow_one, ← Nat.mul_assoc, Nat.mul_comm]

/-! ### bitwise -/

@[simp]
theorem bitwise_zero_left (m : Nat) : bitwise f 0 m = if f false true then m else 0 :=
rfl

@[simp]
theorem bitwise_zero_right (n : Nat) : bitwise f n 0 = if f true false then n else 0 := by
unfold bitwise
simp only [ite_self, decide_False, Nat.zero_div, ite_true]
cases n <;> simp

theorem bitwise_zero : bitwise f 0 0 = 0 := by
simp only [bitwise_zero_right, ite_self]

/-! ### bodd -/

@[simp]
theorem bodd_zero : bodd 0 = false :=
rfl

theorem bodd_one : bodd 1 = true :=
rfl

theorem bodd_two : bodd 2 = false :=
rfl

theorem bodd_eq_mod_two_bne_zero (n : Nat) : bodd n = (n % 2 != 0) := by
rw [mod_two_of_bodd]
cases bodd n with | false | true => rfl

@[simp]
theorem bodd_succ (n : Nat) : bodd (succ n) = not (bodd n) := by
simp only [bodd_eq_mod_two_bne_zero, succ_eq_add_one, succ_mod_two_eq_one_sub_mod_two]
cases mod_two_eq_zero_or_one n with | _ h => simp [h]

@[simp]
theorem bodd_add (m n : Nat) : bodd (m + n) = ((bodd m).xor (bodd n)) := by
induction n with
| zero => simp
| succ n ih => simp [add_succ, ih, Bool.xor_not]

@[simp]
theorem bodd_mul (m n : Nat) : bodd (m * n) = (bodd m && bodd n) := by
induction n with
| zero => simp
| succ n ih =>
simp [mul_succ, ih]
cases bodd m <;> cases bodd n <;> rfl

theorem bodd_bit (b n) : bodd (bit b n) = b := by
rw [bit_val, Nat.mul_comm, Nat.add_comm, bodd_add, bodd_mul]
cases b <;> cases bodd n <;> rfl

/-! ### div2 -/

@[simp]
theorem div2_zero : div2 0 = 0 :=
rfl

theorem div2_one : div2 1 = 0 :=
rfl

theorem div2_two : div2 2 = 1 :=
rfl

theorem div2_eq_div_two (n : Nat) : div2 n = n / 2 := rfl

@[simp]
theorem div2_succ (n : Nat) : div2 (succ n) = cond (bodd n) (succ (div2 n)) (div2 n) := by
apply Nat.eq_of_mul_eq_mul_left (by decide : 0 < 2)
apply Nat.add_right_cancel (m := cond (bodd (succ n)) 1 0)
rw (config := {occs := .pos [1]}) [bodd_add_div2, bodd_succ, ← bodd_add_div2 n]
cases bodd n <;> simp [succ_eq_add_one, Nat.add_comm 1, Nat.mul_add]

theorem div2_bit (b n) : div2 (bit b n) = n := by
rw [bit_val, div2_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add]
· cases b <;> decide
· decide

/-! ### binary rec/cases -/
FR-vdash-bot marked this conversation as resolved.
Show resolved Hide resolved

@[simp]
theorem binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) :
binaryRec z f 0 = z := by
rw [binaryRec]
rfl

theorem binaryRec_eq {C : Nat → Sort u} {z : C 0} {f : ∀ b n, C n → C (bit b n)} (b n)
(h : f false 0 z = z ∨ (n = 0 → b = true)) :
binaryRec z f (bit b n) = f b n (binaryRec z f n) := by
rw [binaryRec]
by_cases h : bit b n = 0
-- Note: this renames the original `h : f false 0 z = z` to `h'` and leaves `h : bit b n = 0`
case pos h' =>
obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h
simp only [or_false] at h'
exact h'.symm
case neg h' =>
simp only [dif_neg h]
generalize bit_decomp (bit b n) = e
revert e
rw [bodd_bit, div2_bit]
intros; rfl

theorem binaryRec_of_ne_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) {n}
(h : n ≠ 0) :
binaryRec z f n = bit_decomp n ▸ f n.bodd n.div2 (binaryRec z f n.div2) := by
rw [binaryRec, dif_neg h]
generalize bit_decomp n = e; revert e
apply bitCasesOn n
intro _ _
rw [bodd_bit, div2_bit]
intro; rfl
Loading