Reducing to an interval modulo its length #
This file defines operations that reduce a number (in an Archimedean
LinearOrderedAddCommGroup) to a number in a given interval, modulo the length of that
interval.
Main definitions #
- toIcoDiv hp a b(where- hp : 0 < p): The unique integer such that this multiple of- p, subtracted from- b, is in- Ico a (a + p).
- toIcoMod hp a b(where- hp : 0 < p): Reduce- bto the interval- Ico a (a + p).
- toIocDiv hp a b(where- hp : 0 < p): The unique integer such that this multiple of- p, subtracted from- b, is in- Ioc a (a + p).
- toIocMod hp a b(where- hp : 0 < p): Reduce- bto the interval- Ioc a (a + p).
The unique integer such that this multiple of p, subtracted from b, is in Ico a (a + p).
Equations
- toIcoDiv hp a b = Exists.choose ⋯
Instances For
The unique integer such that this multiple of p, subtracted from b, is in Ioc a (a + p).
Equations
- toIocDiv hp a b = Exists.choose ⋯
Instances For
Note we omit toIcoDiv_zsmul_add' as -m + toIcoDiv hp a b is not very convenient.
Note we omit toIocDiv_zsmul_add' as -m + toIocDiv hp a b is not very convenient.
Links between the Ico and Ioc variants applied to the same element #
Alias of the forward direction of AddCommGroup.modEq_iff_toIcoMod_eq_left.
Alias of the forward direction of AddCommGroup.modEq_iff_toIocMod_eq_right.
If a and b fall within the same cycle WRT c, then they are congruent modulo p.
Alias of the reverse direction of toIcoMod_inj.
If a and b fall within the same cycle WRT c, then they are congruent modulo p.
toIcoMod as an equiv from the quotient.
Equations
- QuotientAddGroup.equivIcoMod hp a = { toFun := fun (b : α ⧸ AddSubgroup.zmultiples p) => ⟨⋯.lift b, ⋯⟩, invFun := fun (x : ↑(Set.Ico a (a + p))) => ↑↑x, left_inv := ⋯, right_inv := ⋯ }
Instances For
toIocMod as an equiv from the quotient.
Equations
- QuotientAddGroup.equivIocMod hp a = { toFun := fun (b : α ⧸ AddSubgroup.zmultiples p) => ⟨⋯.lift b, ⋯⟩, invFun := fun (x : ↑(Set.Ioc a (a + p))) => ↑↑x, left_inv := ⋯, right_inv := ⋯ }
Instances For
The circular order structure on α ⧸ AddSubgroup.zmultiples p #
Equations
- One or more equations did not get rendered due to their size.
Equations
- QuotientAddGroup.circularPreorder = CircularPreorder.mk ⋯ ⋯ ⋯ ⋯
Equations
- QuotientAddGroup.circularOrder = let __src := QuotientAddGroup.circularPreorder; CircularOrder.mk ⋯
Lemmas about unions of translates of intervals #
Alias of iUnion_Ioc_add_intCast.
Alias of iUnion_Ico_add_intCast.
Alias of iUnion_Icc_add_intCast.
Alias of iUnion_Ioc_intCast.
Alias of iUnion_Ico_intCast.
Alias of iUnion_Icc_intCast.