Skip to content

Commit

Permalink
chore(order/cover): add covby.eq_of_between (#17146)
Browse files Browse the repository at this point in the history
Also add some trivial lemmas about covby on int, for convenience.
  • Loading branch information
eric-wieser committed Oct 25, 2022
1 parent cc26a6c commit 3b3cd89
Show file tree
Hide file tree
Showing 2 changed files with 11 additions and 0 deletions.
6 changes: 6 additions & 0 deletions src/data/int/succ_pred.lean
Expand Up @@ -55,6 +55,12 @@ instance : is_pred_archimedean ℤ :=

protected lemma covby_iff_succ_eq {m n : ℤ} : m ⋖ n ↔ m + 1 = n := succ_eq_iff_covby.symm

@[simp] lemma sub_one_covby (z : ℤ) : z - 1 ⋖ z :=
by rw [int.covby_iff_succ_eq, sub_add_cancel]

@[simp] lemma covby_add_one (z : ℤ) : z ⋖ z + 1 :=
int.covby_iff_succ_eq.mpr rfl

end int

@[simp, norm_cast] lemma nat.cast_int_covby_iff {a b : ℕ} : (a : ℤ) ⋖ b ↔ a ⋖ b :=
Expand Down
5 changes: 5 additions & 0 deletions src/order/cover.lean
Expand Up @@ -273,6 +273,11 @@ lemma covby.unique_left (ha : a ⋖ c) (hb : b ⋖ c) : a = b :=
lemma covby.unique_right (hb : a ⋖ b) (hc : a ⋖ c) : b = c :=
(hb.ge_of_gt hc.lt).antisymm $ hc.ge_of_gt hb.lt

/-- If `a`, `b`, `c` are consecutive and `a < x < c` then `x = b`. -/
lemma covby.eq_of_between {x : α} (hab : a ⋖ b) (hbc : b ⋖ c) (hax : a < x) (hxc : x < c) :
x = b :=
le_antisymm (le_of_not_lt $ λ h, hbc.2 h hxc) (le_of_not_lt $ hab.2 hax)

end linear_order

namespace set
Expand Down

0 comments on commit 3b3cd89

Please sign in to comment.