Skip to content

Commit

Permalink
chore(algebra/group/basic): lemmas about bit0, bit1, and addition (
Browse files Browse the repository at this point in the history
  • Loading branch information
eric-wieser committed Jun 20, 2022
1 parent df50b88 commit ff40b2c
Show file tree
Hide file tree
Showing 2 changed files with 30 additions and 0 deletions.
22 changes: 22 additions & 0 deletions src/algebra/group/basic.lean
Expand Up @@ -116,6 +116,18 @@ by simp only [mul_left_comm, mul_comm]

end comm_semigroup

section add_comm_semigroup
variables {M : Type u} [add_comm_semigroup M]

lemma bit0_add (a b : M) : bit0 (a + b) = bit0 a + bit0 b :=
add_add_add_comm _ _ _ _
lemma bit1_add [has_one M] (a b : M) : bit1 (a + b) = bit0 a + bit1 b :=
(congr_arg (+ (1 : M)) $ bit0_add a b : _).trans (add_assoc _ _ _)
lemma bit1_add' [has_one M] (a b : M) : bit1 (a + b) = bit1 a + bit0 b :=
by rw [add_comm, bit1_add, add_comm]

end add_comm_semigroup

local attribute [simp] mul_assoc sub_eq_add_neg

section add_monoid
Expand Down Expand Up @@ -616,6 +628,16 @@ by rw [div_eq_iff_eq_mul, div_mul_eq_mul_div, div_eq_iff_eq_mul', mul_div_assoc]

end comm_group

section subtraction_comm_monoid
variables {M : Type u} [subtraction_comm_monoid M]

lemma bit0_sub (a b : M) : bit0 (a - b) = bit0 a - bit0 b :=
sub_add_sub_comm _ _ _ _
lemma bit1_sub [has_one M] (a b : M) : bit1 (a - b) = bit1 a - bit0 b :=
(congr_arg (+ (1 : M)) $ bit0_sub a b : _).trans $ sub_add_eq_add_sub _ _ _

end subtraction_comm_monoid

section commutator

/-- The commutator of two elements `g₁` and `g₂`. -/
Expand Down
8 changes: 8 additions & 0 deletions src/data/nat/basic.lean
Expand Up @@ -1489,6 +1489,14 @@ by unfold bodd div2; cases bodd_div2 n; refl
@[simp] lemma one_eq_bit1 {n : ℕ} : 1 = bit1 n ↔ n = 0 :=
⟨λ h, (@nat.bit1_inj 0 n h).symm, λ h, by subst h⟩

theorem bit_add : ∀ (b : bool) (n m : ℕ), bit b (n + m) = bit ff n + bit b m
| tt := bit1_add
| ff := bit0_add

theorem bit_add' : ∀ (b : bool) (n m : ℕ), bit b (n + m) = bit b n + bit ff m
| tt := bit1_add'
| ff := bit0_add

protected theorem bit0_le {n m : ℕ} (h : n ≤ m) : bit0 n ≤ bit0 m :=
add_le_add h h

Expand Down

0 comments on commit ff40b2c

Please sign in to comment.