Skip to content

Commit

Permalink
feat(order/zorn): Variant for Ici (#17053)
Browse files Browse the repository at this point in the history
  • Loading branch information
YaelDillies committed Nov 9, 2022
1 parent bb8b0fd commit fc6a9d1
Showing 1 changed file with 7 additions and 0 deletions.
7 changes: 7 additions & 0 deletions src/order/zorn.lean
Expand Up @@ -129,6 +129,13 @@ begin
exact ⟨z, ⟨hzs, (hcs hy).2.trans $ hz _ hy⟩, hz⟩ } }
end

lemma zorn_nonempty_Ici₀ (a : α)
(ih : ∀ c ⊆ Ici a, is_chain (≤) c → ∀ y ∈ c, ∃ ub, a ≤ ub ∧ ∀ z ∈ c, z ≤ ub) (x : α)
(hax : a ≤ x) :
∃ m, x ≤ m ∧ ∀ z, m ≤ z → z ≤ m :=
let ⟨m, hma, hxm, hm⟩ := zorn_nonempty_preorder₀ (Ici a) (by simpa using ih) x hax in
⟨m, hxm, λ z hmz, hm _ (hax.trans $ hxm.trans hmz) hmz⟩

end preorder

section partial_order
Expand Down

0 comments on commit fc6a9d1

Please sign in to comment.