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

[Merged by Bors] - feat(data/fin): add some lemmas about coercions #2522

Closed
wants to merge 12 commits into from
55 changes: 55 additions & 0 deletions src/data/fin.lean
Original file line number Diff line number Diff line change
Expand Up @@ -98,6 +98,61 @@ attribute [simp] val_zero
@[simp] lemma coe_one {n : ℕ} : ((1 : fin (n+2)) : ℕ) = 1 := rfl
@[simp] lemma coe_two {n : ℕ} : ((2 : fin (n+3)) : ℕ) = 2 := rfl

/-- `a < b` as natural numbers if and only if `a < b` in `fin n`. -/
@[norm_cast, simp] lemma coe_fin_lt {n : ℕ} {a b : fin n} : (a : ℕ) < (b : ℕ) ↔ a < b :=
iff.rfl

/-- `a ≤ b` as natural numbers if and only if `a ≤ b` in `fin n`. -/
@[norm_cast, simp] lemma coe_fin_le {n : ℕ} {a b : fin n} : (a : ℕ) ≤ (b : ℕ) ↔ a ≤ b :=
iff.rfl

lemma val_add {n : ℕ} : ∀ a b : fin n, (a + b).val = (a.val + b.val) % n
| ⟨_, _⟩ ⟨_, _⟩ := rfl

lemma val_mul {n : ℕ} : ∀ a b : fin n, (a * b).val = (a.val * b.val) % n
| ⟨_, _⟩ ⟨_, _⟩ := rfl

lemma one_val {n : ℕ} : (1 : fin (n+1)).val = 1 % (n+1) := rfl

@[simp] lemma zero_val (n : ℕ) : (0 : fin (n+1)).val = 0 := rfl

@[simp]
lemma of_nat_eq_coe (n : ℕ) (a : ℕ) : (of_nat a : fin (n+1)) = a :=
begin
induction a with a ih, { refl },
ext, show (a+1) % (n+1) = fin.val (a+1 : fin (n+1)),
{ rw [val_add, ← ih, of_nat],
exact add_mod _ _ _ }
end

/-- Converting an in-range number to `fin (n + 1)` produces a result
whose value is the original number. -/
lemma coe_val_of_lt {n : ℕ} {a : ℕ} (h : a < n + 1) :
(a : fin (n + 1)).val = a :=
begin
rw ←of_nat_eq_coe,
exact nat.mod_eq_of_lt h
end

/-- Converting the value of a `fin (n + 1)` to `fin (n + 1)` results
in the same value. -/
@[simp] lemma coe_val_eq_self {n : ℕ} (a : fin (n + 1)) : (a.val : fin (n + 1)) = a :=
begin
rw fin.eq_iff_veq,
exact coe_val_of_lt a.is_lt
end

/-- Coercing an in-range number to `fin (n + 1)`, and converting back
to `ℕ`, results in that number. -/
lemma coe_coe_of_lt {n : ℕ} {a : ℕ} (h : a < n + 1) :
((a : fin (n + 1)) : ℕ) = a :=
coe_val_of_lt h

/-- Converting a `fin (n + 1)` to `ℕ` and back results in the same
value. -/
@[simp] lemma coe_coe_eq_self {n : ℕ} (a : fin (n + 1)) : ((a : ℕ) : fin (n + 1)) = a :=
coe_val_eq_self a

/-- Assume `k = l`. If two functions defined on `fin k` and `fin l` are equal on each element,
then they coincide (in the heq sense). -/
protected lemma heq_fun_iff {α : Type*} {k l : ℕ} (h : k = l) {f : fin k → α} {g : fin l → α} :
Expand Down
15 changes: 15 additions & 0 deletions src/data/nat/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -606,6 +606,21 @@ by rw [←nat.mod_add_div a c, ←nat.mod_add_div b c, ←h, ←nat.sub_sub, nat
lemma dvd_sub_mod (k : ℕ) : n ∣ (k - (k % n)) :=
⟨k / n, nat.sub_eq_of_eq_add (nat.mod_add_div k n).symm⟩

lemma add_mod (a b n : ℕ) : (a + b) % n = ((a % n) + (b % n)) % n :=
begin
conv_lhs {
rw [←mod_add_div a n, ←mod_add_div b n, ←add_assoc, add_mul_mod_self_left,
add_assoc, add_comm _ (b % n), ←add_assoc, add_mul_mod_self_left] }
end

lemma mul_mod (a b n : ℕ) : (a * b) % n = ((a % n) * (b % n)) % n :=
begin
conv_lhs {
rw [←mod_add_div a n, ←mod_add_div b n, right_distrib, left_distrib, left_distrib,
mul_assoc, mul_assoc, ←left_distrib n _ _, add_mul_mod_self_left,
mul_comm _ (n * (b / n)), mul_assoc, add_mul_mod_self_left] }
end

theorem add_pos_left {m : ℕ} (h : 0 < m) (n : ℕ) : 0 < m + n :=
calc
m + n > 0 + n : nat.add_lt_add_right h n
Expand Down
6 changes: 0 additions & 6 deletions src/data/nat/modeq.lean
Original file line number Diff line number Diff line change
Expand Up @@ -124,12 +124,6 @@ end)

end modeq

lemma add_mod (a b n : ℕ) : (a + b) % n = ((a % n) + (b % n)) % n :=
(nat.modeq.modeq_add (nat.modeq.mod_modeq a n) (nat.modeq.mod_modeq b n)).symm

jsm28 marked this conversation as resolved.
Show resolved Hide resolved
lemma mul_mod (a b n : ℕ) : (a * b) % n = ((a % n) * (b % n)) % n :=
(nat.modeq.modeq_mul (nat.modeq.mod_modeq a n) (nat.modeq.mod_modeq b n)).symm

@[simp] lemma mod_mul_right_mod (a b c : ℕ) : a % (b * c) % b = a % b :=
modeq.modeq_of_modeq_mul_right _ (modeq.mod_modeq _ _)

Expand Down
21 changes: 0 additions & 21 deletions src/data/zmod/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -73,16 +73,6 @@ def comm_semigroup (n : ℕ) : comm_semigroup (fin (n+1)) :=

local attribute [instance] fin.add_comm_semigroup fin.comm_semigroup

lemma val_add {n : ℕ} : ∀ a b : fin n, (a + b).val = (a.val + b.val) % n
| ⟨_, _⟩ ⟨_, _⟩ := rfl

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Should some of the other fin-related lemmas be moved to data.fin as well?

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I've moved the trivial lemmas val_mul, one_val and zero_val that seem reasonably to go along with val_add. The rest look more specific to how fin is being used to set up zmod n.

lemma val_mul {n : ℕ} : ∀ a b : fin n, (a * b).val = (a.val * b.val) % n
| ⟨_, _⟩ ⟨_, _⟩ := rfl

lemma one_val {n : ℕ} : (1 : fin (n+1)).val = 1 % (n+1) := rfl

@[simp] lemma zero_val (n : ℕ) : (0 : fin (n+1)).val = 0 := rfl

private lemma one_mul_aux (n : ℕ) (a : fin (n+1)) : (1 : fin (n+1)) * a = a :=
begin
cases n with n,
Expand Down Expand Up @@ -124,17 +114,6 @@ def comm_ring (n : ℕ) : comm_ring (fin (n+1)) :=
..fin.add_comm_semigroup n,
..fin.comm_semigroup n }

local attribute [instance] fin.add_comm_semigroup

@[simp]
lemma of_nat_eq_coe (n : ℕ) (a : ℕ) : (of_nat a : fin (n+1)) = a :=
begin
induction a with a ih, { refl },
ext, show (a+1) % (n+1) = fin.val (a+1 : fin (n+1)),
{ rw [fin.val_add, ← ih, of_nat],
exact nat.add_mod _ _ _ }
end

end fin

/-- The integers modulo `n : ℕ`. -/
Expand Down