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

Commit 084b7e7

Browse files
committed
chore(algebra/order,data/set/intervals): a few more trivial lemmas (#4611)
* a few more lemmas for `has_le.le` and `has_lt.lt` namespaces; * a few lemmas about intersections of intervals; * fix section header in `topology/algebra/module`.
1 parent d11eb84 commit 084b7e7

File tree

3 files changed

+35
-11
lines changed

3 files changed

+35
-11
lines changed

src/algebra/order.lean

Lines changed: 18 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -63,15 +63,29 @@ lemma lt_iff_ne [partial_order α] {x y : α} (h : x ≤ y) : x < y ↔ x ≠ y
6363
lemma le_iff_eq [partial_order α] {x y : α} (h : x ≤ y) : y ≤ x ↔ y = x :=
6464
⟨λ h', h'.antisymm h, eq.le⟩
6565

66+
lemma lt_or_le [linear_order α] {a b : α} (h : a ≤ b) (c : α) : a < c ∨ c ≤ b :=
67+
(lt_or_ge a c).imp id $ λ hc, le_trans hc h
68+
69+
lemma le_or_lt [linear_order α] {a b : α} (h : a ≤ b) (c : α) : a ≤ c ∨ c < b :=
70+
(le_or_gt a c).imp id $ λ hc, lt_of_lt_of_le hc h
71+
72+
lemma le_or_le [linear_order α] {a b : α} (h : a ≤ b) (c : α) : a ≤ c ∨ c ≤ b :=
73+
(h.le_or_lt c).elim or.inl (λ h, or.inr $ le_of_lt h)
74+
6675
end has_le.le
6776

68-
namespace has_lt
69-
namespace lt
77+
namespace has_lt.lt
78+
7079
@[nolint ge_or_gt] -- see Note [nolint_ge]
7180
protected lemma gt [has_lt α] {x y : α} (h : x < y) : y > x := h
7281
protected lemma false [preorder α] {x : α} : x < x → false := lt_irrefl x
73-
end lt
74-
end has_lt
82+
83+
lemma ne' [preorder α] {x y : α} (h : x < y) : y ≠ x := h.ne.symm
84+
85+
lemma lt_or_lt [linear_order α] {x y : α} (h : x < y) (z : α) : x < z ∨ z < y :=
86+
(lt_or_ge z y).elim or.inr (λ hz, or.inl $ h.trans_le hz)
87+
88+
end has_lt.lt
7589

7690
namespace ge
7791
@[nolint ge_or_gt] -- see Note [nolint_ge]
@@ -134,12 +148,6 @@ lemma le_or_lt [linear_order α] : ∀ a b : α, a ≤ b ∨ b < a := le_or_gt
134148
lemma ne.lt_or_lt [linear_order α] {a b : α} (h : a ≠ b) : a < b ∨ b < a :=
135149
lt_or_gt_of_ne h
136150

137-
lemma has_le.le.lt_or_le [linear_order α] {a b : α} (h : a ≤ b) (c : α) : a < c ∨ c ≤ b :=
138-
(lt_or_le a c).imp id $ λ hc, hc.trans h
139-
140-
lemma has_le.le.le_or_lt [linear_order α] {a b : α} (h : a ≤ b) (c : α) : a ≤ c ∨ c < b :=
141-
(le_or_lt a c).imp id $ λ hc, hc.trans_le h
142-
143151
lemma not_lt_iff_eq_or_lt [linear_order α] {a b : α} : ¬ a < b ↔ a = b ∨ b < a :=
144152
not_lt.trans $ le_iff_eq_or_lt.trans $ or_congr eq_comm iff.rfl
145153

src/data/set/intervals/basic.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -804,6 +804,20 @@ end lattice
804804
section decidable_linear_order
805805
variables {α : Type u} [decidable_linear_order α] {a a₁ a₂ b b₁ b₂ c d : α}
806806

807+
lemma Ioc_inter_Ioo_of_left_lt (h : b₁ < b₂) : Ioc a₁ b₁ ∩ Ioo a₂ b₂ = Ioc (max a₁ a₂) b₁ :=
808+
ext $ λ x, by simp [and_assoc, @and.left_comm (x ≤ _),
809+
and_iff_left_iff_imp.2 (λ h', lt_of_le_of_lt h' h)]
810+
811+
lemma Ioc_inter_Ioo_of_right_le (h : b₂ ≤ b₁) : Ioc a₁ b₁ ∩ Ioo a₂ b₂ = Ioo (max a₁ a₂) b₂ :=
812+
ext $ λ x, by simp [and_assoc, @and.left_comm (x ≤ _),
813+
and_iff_right_iff_imp.2 (λ h', (le_trans (le_of_lt h') h))]
814+
815+
lemma Ioo_inter_Ioc_of_left_le (h : b₁ ≤ b₂) : Ioo a₁ b₁ ∩ Ioc a₂ b₂ = Ioo (max a₁ a₂) b₁ :=
816+
by rw [inter_comm, Ioc_inter_Ioo_of_right_le h, max_comm]
817+
818+
lemma Ioo_inter_Ioc_of_right_lt (h : b₂ < b₁) : Ioo a₁ b₁ ∩ Ioc a₂ b₂ = Ioc (max a₁ a₂) b₂ :=
819+
by rw [inter_comm, Ioc_inter_Ioo_of_left_lt h, max_comm]
820+
807821
@[simp] lemma Ico_diff_Iio : Ico a b \ Iio c = Ico (max a c) b :=
808822
ext $ by simp [Ico, Iio, iff_def, max_le_iff] {contextual:=tt}
809823

src/topology/algebra/module.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -191,7 +191,9 @@ notation M ` ≃L[`:50 R `] ` M₂ := continuous_linear_equiv R M M₂
191191
namespace continuous_linear_map
192192

193193
section semiring
194-
/- Properties that hold for non-necessarily commutative semirings. -/
194+
/-!
195+
### Properties that hold for non-necessarily commutative semirings.
196+
-/
195197

196198
variables
197199
{R : Type*} [semiring R]

0 commit comments

Comments
 (0)