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

Commit a15401c

Browse files
committed
refactor(data/polynomial/degree): use with_bot.unbot' instead of option.get_or_else (#17120)
1 parent 3644479 commit a15401c

File tree

3 files changed

+21
-27
lines changed

3 files changed

+21
-27
lines changed

src/data/polynomial/degree/definitions.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -42,7 +42,7 @@ inv_image.wf degree (with_bot.well_founded_lt nat.lt_wf)
4242
instance : has_well_founded R[X] := ⟨_, degree_lt_wf⟩
4343

4444
/-- `nat_degree p` forces `degree p` to ℕ, by defining nat_degree 0 = 0. -/
45-
def nat_degree (p : R[X]) : ℕ := (degree p).get_or_else 0
45+
def nat_degree (p : R[X]) : ℕ := (degree p).unbot' 0
4646

4747
/-- `leading_coeff p` gives the coefficient of the highest power of `X` in `p`-/
4848
def leading_coeff (p : R[X]) : R := coeff p (nat_degree p)
@@ -106,7 +106,7 @@ option.some_inj.1 $ show (nat_degree p : with_bot ℕ) = n,
106106
by rwa [← degree_eq_nat_degree hp0]
107107

108108
@[simp] lemma degree_le_nat_degree : degree p ≤ nat_degree p :=
109-
with_bot.gi_get_or_else_bot.gc.le_u_l _
109+
with_bot.gi_unbot'_bot.gc.le_u_l _
110110

111111
lemma nat_degree_eq_of_degree_eq [semiring S] {q : S[X]} (h : degree p = degree q) :
112112
nat_degree p = nat_degree q :=
@@ -153,20 +153,20 @@ end
153153

154154
lemma degree_ne_of_nat_degree_ne {n : ℕ} :
155155
p.nat_degree ≠ n → degree p ≠ n :=
156-
mt $ λ h, by rw [nat_degree, h, option.get_or_else_coe]
156+
mt $ λ h, by rw [nat_degree, h, with_bot.unbot'_coe]
157157

158158
theorem nat_degree_le_iff_degree_le {n : ℕ} : nat_degree p ≤ n ↔ degree p ≤ n :=
159-
with_bot.get_or_else_bot_le_iff
159+
with_bot.unbot'_bot_le_iff
160160

161161
lemma nat_degree_lt_iff_degree_lt (hp : p ≠ 0) :
162162
p.nat_degree < n ↔ p.degree < ↑n :=
163-
with_bot.get_or_else_bot_lt_iff $ degree_eq_bot.not.mpr hp
163+
with_bot.unbot'_lt_iff $ degree_eq_bot.not.mpr hp
164164

165165
alias nat_degree_le_iff_degree_le ↔ ..
166166

167167
lemma nat_degree_le_nat_degree [semiring S] {q : S[X]} (hpq : p.degree ≤ q.degree) :
168168
p.nat_degree ≤ q.nat_degree :=
169-
with_bot.gi_get_or_else_bot.gc.monotone_l hpq
169+
with_bot.gi_unbot'_bot.gc.monotone_l hpq
170170

171171
@[simp] lemma degree_C (ha : a ≠ 0) : degree (C a) = (0 : with_bot ℕ) :=
172172
by rw [degree, ← monomial_zero_left, support_monomial 0 ha, max_eq_sup_coe, sup_singleton,

src/order/bounded_order.lean

Lines changed: 9 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -693,25 +693,19 @@ lemma map_le_iff [preorder α] [preorder β] (f : α → β) (mono_iff : ∀ {a
693693
| (a : α) ⊥ := by simp only [map_coe, map_bot, coe_ne_bot, not_coe_le_bot _]
694694
| (a : α) (b : α) := by simpa using mono_iff
695695

696-
lemma le_coe_get_or_else [preorder α] : ∀ (a : with_bot α) (b : α), a ≤ a.get_or_else b
697-
| (some a) b := le_refl a
698-
| none b := λ _ h, option.no_confusion h
696+
lemma le_coe_unbot' [preorder α] : ∀ (a : with_bot α) (b : α), a ≤ a.unbot' b
697+
| (a : α) b := le_rfl
698+
| b := bot_le
699699

700-
@[simp] lemma get_or_else_bot (a : α) : option.get_or_else (⊥ : with_bot α) a = a := rfl
701-
702-
lemma get_or_else_bot_le_iff [has_le α] [order_bot α] {a : with_bot α} {b : α} :
703-
a.get_or_else ⊥ ≤ b ↔ a ≤ b :=
700+
lemma unbot'_bot_le_iff [has_le α] [order_bot α] {a : with_bot α} {b : α} :
701+
a.unbot' ⊥ ≤ b ↔ a ≤ b :=
704702
by cases a; simp [none_eq_bot, some_eq_coe]
705703

706-
lemma get_or_else_bot_lt_iff [partial_order α] [order_bot α] {a : with_bot α} {b : α}
707-
(ha : a ≠ ⊥) :
708-
a.get_or_else ⊥ < b ↔ a < b :=
704+
lemma unbot'_lt_iff [has_lt α] {a : with_bot α} {b c : α} (ha : a ≠ ⊥) :
705+
a.unbot' b < c ↔ a < c :=
709706
begin
710-
obtain ⟨a, rfl⟩ := ne_bot_iff_exists.mp ha,
711-
simp only [lt_iff_le_and_ne, get_or_else_bot_le_iff, and.congr_right_iff],
712-
intro h,
713-
apply iff.not,
714-
simp only [with_bot.coe_eq_coe, option.get_or_else_coe, iff_self],
707+
lift a to α using ha,
708+
rw [unbot'_coe, coe_lt_coe]
715709
end
716710

717711
instance [semilattice_sup α] : semilattice_sup (with_bot α) :=

src/order/galois_connection.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -766,11 +766,11 @@ end lift
766766

767767
end galois_coinsertion
768768

769-
/-- If `α` is a partial order with bottom element (e.g., `ℕ`, `ℝ≥0`), then
770-
`λ o : with_bot α, o.get_or_else ⊥` and coercion form a Galois insertion. -/
771-
def with_bot.gi_get_or_else_bot [preorder α] [order_bot α] :
772-
galois_insertion (λ o : with_bot α, o.get_or_else ⊥) coe :=
773-
{ gc := λ a b, with_bot.get_or_else_bot_le_iff,
769+
/-- If `α` is a partial order with bottom element (e.g., `ℕ`, `ℝ≥0`), then `with_bot.unbot' ⊥` and
770+
coercion form a Galois insertion. -/
771+
def with_bot.gi_unbot'_bot [preorder α] [order_bot α] :
772+
galois_insertion (with_bot.unbot' ⊥) (coe : α → with_bot α) :=
773+
{ gc := λ a b, with_bot.unbot'_bot_le_iff,
774774
le_l_u := λ a, le_rfl,
775-
choice := λ o ho, _,
775+
choice := λ o ho, o.unbot' ⊥,
776776
choice_eq := λ _ _, rfl }

0 commit comments

Comments
 (0)