Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 3813d4e

Browse files
committed
chore(data/nat/modeq): split out lemmas about lists (#18004)
By splitting out some lemmas relating list-rotation and modular arithmetic, the files `data.nat.modeq` and `data.int.modeq` become portable now.
1 parent cfa57f4 commit 3813d4e

File tree

4 files changed

+70
-60
lines changed

4 files changed

+70
-60
lines changed

archive/miu_language/decision_nec.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Gihan Marasingha
55
-/
66
import .basic
7+
import data.list.count
78
import data.nat.modeq
89
import tactic.ring
910

src/data/list/modeq.lean

Lines changed: 68 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,68 @@
1+
/-
2+
Copyright (c) 2018 Chris Hughes. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Hughes
5+
-/
6+
import data.nat.modeq
7+
import data.list.rotate
8+
9+
/-! # List rotation and modular arithmetic -/
10+
11+
namespace list
12+
variable {α : Type*}
13+
14+
lemma nth_rotate : ∀ {l : list α} {n m : ℕ} (hml : m < l.length),
15+
(l.rotate n).nth m = l.nth ((m + n) % l.length)
16+
| [] n m hml := (nat.not_lt_zero _ hml).elim
17+
| l 0 m hml := by simp [nat.mod_eq_of_lt hml]
18+
| (a::l) (n+1) m hml :=
19+
have h₃ : m < list.length (l ++ [a]), by simpa using hml,
20+
(lt_or_eq_of_le (nat.le_of_lt_succ $ nat.mod_lt (m + n)
21+
(lt_of_le_of_lt (nat.zero_le _) hml))).elim
22+
(λ hml',
23+
have h₁ : (m + (n + 1)) % ((a :: l : list α).length) =
24+
(m + n) % ((a :: l : list α).length) + 1,
25+
from calc (m + (n + 1)) % (l.length + 1) =
26+
((m + n) % (l.length + 1) + 1) % (l.length + 1) :
27+
add_assoc m n 1 ▸ nat.modeq.add_right 1 (nat.mod_mod _ _).symm
28+
... = (m + n) % (l.length + 1) + 1 : nat.mod_eq_of_lt (nat.succ_lt_succ hml'),
29+
have h₂ : (m + n) % (l ++ [a]).length < l.length, by simpa [nat.add_one] using hml',
30+
by rw [list.rotate_cons_succ, nth_rotate h₃, list.nth_append h₂, h₁, list.nth]; simp)
31+
(λ hml',
32+
have h₁ : (m + (n + 1)) % (l.length + 1) = 0,
33+
from calc (m + (n + 1)) % (l.length + 1) = (l.length + 1) % (l.length + 1) :
34+
add_assoc m n 1 ▸ nat.modeq.add_right 1
35+
(hml'.trans (nat.mod_eq_of_lt (nat.lt_succ_self _)).symm)
36+
... = 0 : by simp,
37+
by rw [list.length, list.rotate_cons_succ, nth_rotate h₃, list.length_append,
38+
list.length_cons, list.length, zero_add, hml', h₁, list.nth_concat_length]; refl)
39+
40+
lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : ∀ {l : list α},
41+
(∀ n, l.rotate n = l) ↔ ∃ a, l = list.repeat a l.length
42+
| [] := ⟨λ h, nonempty.elim hα (λ a, ⟨a, by simp⟩), by simp⟩
43+
| (a::l) :=
44+
⟨λ h, ⟨a, list.ext_le (by simp) $ λ n hn h₁,
45+
begin
46+
rw [← option.some_inj, ← list.nth_le_nth],
47+
conv {to_lhs, rw ← h ((list.length (a :: l)) - n)},
48+
rw [nth_rotate hn, add_tsub_cancel_of_le (le_of_lt hn),
49+
nat.mod_self, nth_le_repeat], refl
50+
end⟩,
51+
λ ⟨a, ha⟩ n, ha.symm ▸ list.ext_le (by simp)
52+
(λ m hm h,
53+
have hm' : (m + n) % (list.repeat a (list.length (a :: l))).length < list.length (a :: l),
54+
by rw list.length_repeat; exact nat.mod_lt _ (nat.succ_pos _),
55+
by rw [nth_le_repeat, ← option.some_inj, ← list.nth_le_nth, nth_rotate h, list.nth_le_nth,
56+
nth_le_repeat]; simp * at *)⟩
57+
58+
lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) :
59+
(list.repeat a n).rotate k = list.repeat a n :=
60+
let h : nonempty α := ⟨a⟩ in by exactI rotate_eq_self_iff_eq_repeat.mpr ⟨a, by rw length_repeat⟩ k
61+
62+
lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} :
63+
l.rotate 1 = l ↔ ∃ a : α, l = list.repeat a l.length :=
64+
⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero
65+
(λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n),
66+
λ h, rotate_eq_self_iff_eq_repeat.mpr h 1
67+
68+
end list

src/data/nat/modeq.lean

Lines changed: 0 additions & 60 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Mario Carneiro
55
-/
66
import data.int.gcd
7-
import data.list.rotate
87
import tactic.abel
98

109
/-!
@@ -425,62 +424,3 @@ have help : ∀ (m : ℕ), m < 4 → m % 2 = 1 → m = 1 ∨ m = 3 := dec_trivia
425424
λ h, or.dcases_on h odd_of_mod_four_eq_one odd_of_mod_four_eq_three⟩
426425

427426
end nat
428-
429-
namespace list
430-
variable {α : Type*}
431-
432-
lemma nth_rotate : ∀ {l : list α} {n m : ℕ} (hml : m < l.length),
433-
(l.rotate n).nth m = l.nth ((m + n) % l.length)
434-
| [] n m hml := (nat.not_lt_zero _ hml).elim
435-
| l 0 m hml := by simp [nat.mod_eq_of_lt hml]
436-
| (a::l) (n+1) m hml :=
437-
have h₃ : m < list.length (l ++ [a]), by simpa using hml,
438-
(lt_or_eq_of_le (nat.le_of_lt_succ $ nat.mod_lt (m + n)
439-
(lt_of_le_of_lt (nat.zero_le _) hml))).elim
440-
(λ hml',
441-
have h₁ : (m + (n + 1)) % ((a :: l : list α).length) =
442-
(m + n) % ((a :: l : list α).length) + 1,
443-
from calc (m + (n + 1)) % (l.length + 1) =
444-
((m + n) % (l.length + 1) + 1) % (l.length + 1) :
445-
add_assoc m n 1 ▸ nat.modeq.add_right 1 (nat.mod_mod _ _).symm
446-
... = (m + n) % (l.length + 1) + 1 : nat.mod_eq_of_lt (nat.succ_lt_succ hml'),
447-
have h₂ : (m + n) % (l ++ [a]).length < l.length, by simpa [nat.add_one] using hml',
448-
by rw [list.rotate_cons_succ, nth_rotate h₃, list.nth_append h₂, h₁, list.nth]; simp)
449-
(λ hml',
450-
have h₁ : (m + (n + 1)) % (l.length + 1) = 0,
451-
from calc (m + (n + 1)) % (l.length + 1) = (l.length + 1) % (l.length + 1) :
452-
add_assoc m n 1 ▸ nat.modeq.add_right 1
453-
(hml'.trans (nat.mod_eq_of_lt (nat.lt_succ_self _)).symm)
454-
... = 0 : by simp,
455-
by rw [list.length, list.rotate_cons_succ, nth_rotate h₃, list.length_append,
456-
list.length_cons, list.length, zero_add, hml', h₁, list.nth_concat_length]; refl)
457-
458-
lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : ∀ {l : list α},
459-
(∀ n, l.rotate n = l) ↔ ∃ a, l = list.repeat a l.length
460-
| [] := ⟨λ h, nonempty.elim hα (λ a, ⟨a, by simp⟩), by simp⟩
461-
| (a::l) :=
462-
⟨λ h, ⟨a, list.ext_le (by simp) $ λ n hn h₁,
463-
begin
464-
rw [← option.some_inj, ← list.nth_le_nth],
465-
conv {to_lhs, rw ← h ((list.length (a :: l)) - n)},
466-
rw [nth_rotate hn, add_tsub_cancel_of_le (le_of_lt hn),
467-
nat.mod_self, nth_le_repeat], refl
468-
end⟩,
469-
λ ⟨a, ha⟩ n, ha.symm ▸ list.ext_le (by simp)
470-
(λ m hm h,
471-
have hm' : (m + n) % (list.repeat a (list.length (a :: l))).length < list.length (a :: l),
472-
by rw list.length_repeat; exact nat.mod_lt _ (nat.succ_pos _),
473-
by rw [nth_le_repeat, ← option.some_inj, ← list.nth_le_nth, nth_rotate h, list.nth_le_nth,
474-
nth_le_repeat]; simp * at *)⟩
475-
476-
lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) :
477-
(list.repeat a n).rotate k = list.repeat a n :=
478-
let h : nonempty α := ⟨a⟩ in by exactI rotate_eq_self_iff_eq_repeat.mpr ⟨a, by rw length_repeat⟩ k
479-
480-
lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} :
481-
l.rotate 1 = l ↔ ∃ a : α, l = list.repeat a l.length :=
482-
⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero
483-
(λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n),
484-
λ h, rotate_eq_self_iff_eq_repeat.mpr h 1
485-
486-
end list

src/group_theory/perm/cycle/type.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Thomas Browning
66

77
import algebra.gcd_monoid.multiset
88
import combinatorics.partition
9+
import data.list.modeq
910
import group_theory.perm.cycle.basic
1011
import ring_theory.int.basic
1112
import tactic.linarith

0 commit comments

Comments
 (0)