Skip to content

Commit

Permalink
chore(CountableDenseLinearOrder): Encodable -> Countable (#10308)
Browse files Browse the repository at this point in the history
  • Loading branch information
urkud committed Feb 8, 2024
1 parent 79c11a0 commit 71b03a2
Showing 1 changed file with 8 additions and 5 deletions.
13 changes: 8 additions & 5 deletions Mathlib/Order/CountableDenseLinearOrder.lean
Expand Up @@ -199,8 +199,9 @@ open PartialIso
-- variable (α β)

/-- Any countable linear order embeds in any nontrivial dense linear order. -/
theorem embedding_from_countable_to_dense [Encodable α] [DenselyOrdered β] [Nontrivial β] :
theorem embedding_from_countable_to_dense [Countable α] [DenselyOrdered β] [Nontrivial β] :
Nonempty (α ↪o β) := by
cases nonempty_encodable α
rcases exists_pair_lt β with ⟨x, y, hxy⟩
cases' exists_between hxy with a ha
haveI : Nonempty (Set.Ioo x y) := ⟨⟨a, ha⟩⟩
Expand All @@ -217,15 +218,17 @@ theorem embedding_from_countable_to_dense [Encodable α] [DenselyOrdered β] [No
#align order.embedding_from_countable_to_dense Order.embedding_from_countable_to_dense

/-- Any two countable dense, nonempty linear orders without endpoints are order isomorphic. -/
theorem iso_of_countable_dense [Encodable α] [DenselyOrdered α] [NoMinOrder α] [NoMaxOrder α]
[Nonempty α] [Encodable β] [DenselyOrdered β] [NoMinOrder β] [NoMaxOrder β] [Nonempty β] :
Nonempty (α ≃o β) :=
theorem iso_of_countable_dense [Countable α] [DenselyOrdered α] [NoMinOrder α] [NoMaxOrder α]
[Nonempty α] [Countable β] [DenselyOrdered β] [NoMinOrder β] [NoMaxOrder β] [Nonempty β] :
Nonempty (α ≃o β) := by
cases nonempty_encodable α
cases nonempty_encodable β
let to_cofinal : Sum α β → Cofinal (PartialIso α β) := fun p ↦
Sum.recOn p (definedAtLeft β) (definedAtRight α)
let our_ideal : Ideal (PartialIso α β) := idealOfCofinals default to_cofinal
let F a := funOfIdeal a our_ideal (cofinal_meets_idealOfCofinals _ to_cofinal (Sum.inl a))
let G b := invOfIdeal b our_ideal (cofinal_meets_idealOfCofinals _ to_cofinal (Sum.inr b))
⟨OrderIso.ofCmpEqCmp (fun a ↦ (F a).val) (fun b ↦ (G b).val) fun a b ↦ by
exact ⟨OrderIso.ofCmpEqCmp (fun a ↦ (F a).val) (fun b ↦ (G b).val) fun a b ↦ by
rcases (F a).prop with ⟨f, hf, ha⟩
rcases (G b).prop with ⟨g, hg, hb⟩
rcases our_ideal.directed _ hf _ hg with ⟨m, _, fm, gm⟩
Expand Down

0 comments on commit 71b03a2

Please sign in to comment.