Skip to content

Commit a183081

Browse files
feat: port Topology.Instances.AddCircle (#3984)
Co-authored-by: David Loeffler <d.loeffler.01@cantab.net>
1 parent cfd8426 commit a183081

File tree

4 files changed

+707
-27
lines changed

4 files changed

+707
-27
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2096,6 +2096,7 @@ import Mathlib.Topology.Homotopy.Contractible
20962096
import Mathlib.Topology.Homotopy.Equiv
20972097
import Mathlib.Topology.Homotopy.Path
20982098
import Mathlib.Topology.Inseparable
2099+
import Mathlib.Topology.Instances.AddCircle
20992100
import Mathlib.Topology.Instances.ENNReal
21002101
import Mathlib.Topology.Instances.EReal
21012102
import Mathlib.Topology.Instances.Int

Mathlib/Algebra/CharZero/Quotient.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -53,7 +53,7 @@ theorem nsmul_mem_zmultiples_iff_exists_sub_div {r : R} {n : ℕ} (hn : n ≠ 0)
5353

5454
end AddSubgroup
5555

56-
namespace quotientAddGroup
56+
namespace QuotientAddGroup
5757

5858
theorem zmultiples_zsmul_eq_zsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} {z : ℤ} (hz : z ≠ 0) :
5959
z • ψ = z • θ ↔ ∃ k : Fin z.natAbs, ψ = θ + (k : ℕ) • (p / z : R) := by
@@ -66,13 +66,13 @@ theorem zmultiples_zsmul_eq_zsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} {
6666
simp_rw [← QuotientAddGroup.mk_zsmul, ← QuotientAddGroup.mk_add,
6767
QuotientAddGroup.eq_iff_sub_mem, ← smul_sub, ← sub_sub]
6868
exact AddSubgroup.zsmul_mem_zmultiples_iff_exists_sub_div hz
69-
#align quotient_add_group.zmultiples_zsmul_eq_zsmul_iff quotientAddGroup.zmultiples_zsmul_eq_zsmul_iff
69+
#align quotient_add_group.zmultiples_zsmul_eq_zsmul_iff QuotientAddGroup.zmultiples_zsmul_eq_zsmul_iff
7070

7171
theorem zmultiples_nsmul_eq_nsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} {n : ℕ} (hz : n ≠ 0) :
7272
n • ψ = n • θ ↔ ∃ k : Fin n, ψ = θ + (k : ℕ) • (p / n : R) := by
7373
rw [← coe_nat_zsmul ψ, ← coe_nat_zsmul θ,
7474
zmultiples_zsmul_eq_zsmul_iff (Int.coe_nat_ne_zero.mpr hz), Int.cast_ofNat]
7575
rfl
76-
#align quotient_add_group.zmultiples_nsmul_eq_nsmul_iff quotientAddGroup.zmultiples_nsmul_eq_nsmul_iff
76+
#align quotient_add_group.zmultiples_nsmul_eq_nsmul_iff QuotientAddGroup.zmultiples_nsmul_eq_nsmul_iff
7777

78-
end quotientAddGroup
78+
end QuotientAddGroup

Mathlib/Algebra/Order/ToIntervalMod.lean

Lines changed: 23 additions & 23 deletions
Original file line numberDiff line numberDiff line change
@@ -793,7 +793,7 @@ end Zero
793793

794794
/-- `toIcoMod` as an equiv from the quotient. -/
795795
@[simps symm_apply]
796-
def quotientAddGroup.equivIcoMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃ Set.Ico a (a + p) where
796+
def QuotientAddGroup.equivIcoMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃ Set.Ico a (a + p) where
797797
toFun b :=
798798
⟨(toIcoMod_periodic hp a).lift b, QuotientAddGroup.induction_on' b <| toIcoMod_mem_Ico hp a⟩
799799
invFun := (↑)
@@ -803,23 +803,23 @@ def quotientAddGroup.equivIcoMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃
803803
dsimp
804804
rw [QuotientAddGroup.eq_iff_sub_mem, toIcoMod_sub_self]
805805
apply AddSubgroup.zsmul_mem_zmultiples
806-
#align quotient_add_group.equiv_Ico_mod quotientAddGroup.equivIcoMod
806+
#align quotient_add_group.equiv_Ico_mod QuotientAddGroup.equivIcoMod
807807

808808
@[simp]
809-
theorem quotientAddGroup.equivIcoMod_coe (a b : α) :
810-
quotientAddGroup.equivIcoMod hp a ↑b = ⟨toIcoMod hp a b, toIcoMod_mem_Ico hp a _⟩ :=
809+
theorem QuotientAddGroup.equivIcoMod_coe (a b : α) :
810+
QuotientAddGroup.equivIcoMod hp a ↑b = ⟨toIcoMod hp a b, toIcoMod_mem_Ico hp a _⟩ :=
811811
rfl
812-
#align quotient_add_group.equiv_Ico_mod_coe quotientAddGroup.equivIcoMod_coe
812+
#align quotient_add_group.equiv_Ico_mod_coe QuotientAddGroup.equivIcoMod_coe
813813

814814
@[simp]
815-
theorem quotientAddGroup.equivIcoMod_zero (a : α) :
816-
quotientAddGroup.equivIcoMod hp a 0 = ⟨toIcoMod hp a 0, toIcoMod_mem_Ico hp a _⟩ :=
815+
theorem QuotientAddGroup.equivIcoMod_zero (a : α) :
816+
QuotientAddGroup.equivIcoMod hp a 0 = ⟨toIcoMod hp a 0, toIcoMod_mem_Ico hp a _⟩ :=
817817
rfl
818-
#align quotient_add_group.equiv_Ico_mod_zero quotientAddGroup.equivIcoMod_zero
818+
#align quotient_add_group.equiv_Ico_mod_zero QuotientAddGroup.equivIcoMod_zero
819819

820820
/-- `toIocMod` as an equiv from the quotient. -/
821821
@[simps symm_apply]
822-
def quotientAddGroup.equivIocMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃ Set.Ioc a (a + p) where
822+
def QuotientAddGroup.equivIocMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃ Set.Ioc a (a + p) where
823823
toFun b :=
824824
⟨(toIocMod_periodic hp a).lift b, QuotientAddGroup.induction_on' b <| toIocMod_mem_Ioc hp a⟩
825825
invFun := (↑)
@@ -829,19 +829,19 @@ def quotientAddGroup.equivIocMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃
829829
dsimp
830830
rw [QuotientAddGroup.eq_iff_sub_mem, toIocMod_sub_self]
831831
apply AddSubgroup.zsmul_mem_zmultiples
832-
#align quotient_add_group.equiv_Ioc_mod quotientAddGroup.equivIocMod
832+
#align quotient_add_group.equiv_Ioc_mod QuotientAddGroup.equivIocMod
833833

834834
@[simp]
835-
theorem quotientAddGroup.equivIocMod_coe (a b : α) :
836-
quotientAddGroup.equivIocMod hp a ↑b = ⟨toIocMod hp a b, toIocMod_mem_Ioc hp a _⟩ :=
835+
theorem QuotientAddGroup.equivIocMod_coe (a b : α) :
836+
QuotientAddGroup.equivIocMod hp a ↑b = ⟨toIocMod hp a b, toIocMod_mem_Ioc hp a _⟩ :=
837837
rfl
838-
#align quotient_add_group.equiv_Ioc_mod_coe quotientAddGroup.equivIocMod_coe
838+
#align quotient_add_group.equiv_Ioc_mod_coe QuotientAddGroup.equivIocMod_coe
839839

840840
@[simp]
841-
theorem quotientAddGroup.equivIocMod_zero (a : α) :
842-
quotientAddGroup.equivIocMod hp a 0 = ⟨toIocMod hp a 0, toIocMod_mem_Ioc hp a _⟩ :=
841+
theorem QuotientAddGroup.equivIocMod_zero (a : α) :
842+
QuotientAddGroup.equivIocMod hp a 0 = ⟨toIocMod hp a 0, toIocMod_mem_Ioc hp a _⟩ :=
843843
rfl
844-
#align quotient_add_group.equiv_Ioc_mod_zero quotientAddGroup.equivIocMod_zero
844+
#align quotient_add_group.equiv_Ioc_mod_zero QuotientAddGroup.equivIocMod_zero
845845

846846
/-!
847847
### The circular order structure on `α ⧸ AddSubgroup.zmultiples p`
@@ -921,7 +921,7 @@ private theorem toIxxMod_trans {x₁ x₂ x₃ x₄ : α}
921921
· rw [not_le] at h₁₂₃ h₂₃₄⊢
922922
exact (h₁₂₃.2.trans_le (toIcoMod_le_toIocMod _ x₃ x₂)).trans h₂₃₄.2
923923

924-
namespace quotientAddGroup
924+
namespace QuotientAddGroup
925925

926926
variable [hp' : Fact (0 < p)]
927927

@@ -932,14 +932,14 @@ theorem btw_coe_iff' {x₁ x₂ x₃ : α} :
932932
Btw.btw (x₁ : α ⧸ AddSubgroup.zmultiples p) x₂ x₃ ↔
933933
toIcoMod hp'.out 0 (x₂ - x₁) ≤ toIocMod hp'.out 0 (x₃ - x₁) :=
934934
Iff.rfl
935-
#align quotient_add_group.btw_coe_iff' quotientAddGroup.btw_coe_iff'
935+
#align quotient_add_group.btw_coe_iff' QuotientAddGroup.btw_coe_iff'
936936

937937
-- maybe harder to use than the primed one?
938938
theorem btw_coe_iff {x₁ x₂ x₃ : α} :
939939
Btw.btw (x₁ : α ⧸ AddSubgroup.zmultiples p) x₂ x₃ ↔
940940
toIcoMod hp'.out x₁ x₂ ≤ toIocMod hp'.out x₁ x₃ :=
941941
by rw [btw_coe_iff', toIocMod_sub_eq_sub, toIcoMod_sub_eq_sub, zero_add, sub_le_sub_iff_right]
942-
#align quotient_add_group.btw_coe_iff quotientAddGroup.btw_coe_iff
942+
#align quotient_add_group.btw_coe_iff QuotientAddGroup.btw_coe_iff
943943

944944
instance circularPreorder : CircularPreorder (α ⧸ AddSubgroup.zmultiples p) where
945945
btw_refl x := show _ ≤ _ by simp [sub_self, hp'.out.le]
@@ -959,10 +959,10 @@ instance circularPreorder : CircularPreorder (α ⧸ AddSubgroup.zmultiples p) w
959959
induction x₄ using QuotientAddGroup.induction_on'
960960
simp_rw [btw_coe_iff] at h₁₂₃ h₂₃₄⊢
961961
apply toIxxMod_trans _ h₁₂₃ h₂₃₄
962-
#align quotient_add_group.circular_preorder quotientAddGroup.circularPreorder
962+
#align quotient_add_group.circular_preorder QuotientAddGroup.circularPreorder
963963

964964
instance circularOrder : CircularOrder (α ⧸ AddSubgroup.zmultiples p) :=
965-
{ quotientAddGroup.circularPreorder with
965+
{ QuotientAddGroup.circularPreorder with
966966
btw_antisymm := fun {x₁ x₂ x₃} h₁₂₃ h₃₂₁ => by
967967
induction x₁ using QuotientAddGroup.induction_on'
968968
induction x₂ using QuotientAddGroup.induction_on'
@@ -977,9 +977,9 @@ instance circularOrder : CircularOrder (α ⧸ AddSubgroup.zmultiples p) :=
977977
induction x₃ using QuotientAddGroup.induction_on'
978978
simp_rw [btw_coe_iff]
979979
apply toIxxMod_total }
980-
#align quotient_add_group.circular_order quotientAddGroup.circularOrder
980+
#align quotient_add_group.circular_order QuotientAddGroup.circularOrder
981981

982-
end quotientAddGroup
982+
end QuotientAddGroup
983983

984984
end Circular
985985

0 commit comments

Comments
 (0)