Skip to content

Commit 58a5282

Browse files
sterrafalreadydone
andcommitted
feat(SetTheory): definition of initial ordinals, ω₁ as an ordinal, ordinal-indexed unions (#6404)
- I setup notation for the first ordinal in each cardinality. - `ω₁` is defined as an ordinal, not as an `out` (cf. `MeasureTheory.CardMeasurableSpace`). - Lemma using the cofinality of `ω₁`. - Lemma on the cardinality of ordinal-indexed `iUnion`s in preparation for material on the Borel hierarchy. Co-authored-by: Pedro Sánchez Terraf <sterraf@users.noreply.github.com> Co-authored-by: Junyan Xu <junyanxu.math@gmail.com>
1 parent a28c357 commit 58a5282

File tree

3 files changed

+76
-1
lines changed

3 files changed

+76
-1
lines changed

Mathlib/SetTheory/Cardinal/Cofinality.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1250,3 +1250,20 @@ theorem lt_cof_power {a b : Cardinal} (ha : ℵ₀ ≤ a) (b1 : 1 < b) : a < cof
12501250
#align cardinal.lt_cof_power Cardinal.lt_cof_power
12511251

12521252
end Cardinal
1253+
1254+
section Omega1
1255+
1256+
namespace Ordinal
1257+
1258+
open Cardinal
1259+
open scoped Ordinal
1260+
1261+
lemma sup_sequence_lt_omega1 {α} [Countable α] (o : α → Ordinal) (ho : ∀ n, o n < ω₁) :
1262+
sup o < ω₁ := by
1263+
apply sup_lt_ord_lift _ ho
1264+
rw [Cardinal.isRegular_aleph_one.cof_eq]
1265+
exact lt_of_le_of_lt mk_le_aleph0 aleph0_lt_aleph_one
1266+
1267+
end Ordinal
1268+
1269+
end Omega1

Mathlib/SetTheory/Cardinal/Ordinal.lean

Lines changed: 46 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -24,7 +24,9 @@ using ordinals.
2424
It is an order isomorphism between ordinals and cardinals.
2525
* The function `Cardinal.aleph` gives the infinite cardinals listed by their
2626
ordinal index. `aleph 0 = ℵ₀`, `aleph 1 = succ ℵ₀` is the first
27-
uncountable cardinal, and so on.
27+
uncountable cardinal, and so on. The notation `ω_` combines the latter with `Cardinal.ord`,
28+
giving an enumeration of (infinite) initial ordinals.
29+
Thus `ω_ 0 = ω` and `ω₁ = ω_ 1` is the first uncountable ordinal.
2830
* The function `Cardinal.beth` enumerates the Beth cardinals. `beth 0 = ℵ₀`,
2931
`beth (succ o) = 2 ^ beth o`, and for a limit ordinal `o`, `beth o` is the supremum of `beth a`
3032
for `a < o`.
@@ -1446,3 +1448,46 @@ theorem extend_function_of_lt {α β : Type*} {s : Set α} (f : s ↪ β) (hs :
14461448
-- end Bit
14471449

14481450
end Cardinal
1451+
1452+
section Initial
1453+
1454+
namespace Ordinal
1455+
1456+
/--
1457+
`ω_ o` is a notation for the *initial ordinal* of cardinality
1458+
`aleph o`. Thus, for example `ω_ 0 = ω`.
1459+
-/
1460+
scoped notation "ω_" o => ord <| aleph o
1461+
1462+
/--
1463+
`ω₁` is the first uncountable ordinal.
1464+
-/
1465+
scoped notation "ω₁" => ord <| aleph 1
1466+
1467+
lemma omega_lt_omega1 : ω < ω₁ := ord_aleph0.symm.trans_lt (ord_lt_ord.mpr (aleph0_lt_aleph_one))
1468+
1469+
section OrdinalIndices
1470+
/-!
1471+
### Cardinal operations with ordinal indices
1472+
1473+
Results on cardinality of ordinal-indexed families of sets.
1474+
-/
1475+
namespace Cardinal
1476+
1477+
open scoped Cardinal
1478+
1479+
/--
1480+
Bounding the cardinal of an ordinal-indexed union of sets.
1481+
-/
1482+
lemma mk_iUnion_Ordinal_le_of_le {β : Type _} {o : Ordinal} {c : Cardinal}
1483+
(ho : o.card ≤ c) (hc : ℵ₀ ≤ c) (A : Ordinal → Set β)
1484+
(hA : ∀ j < o, #(A j) ≤ c) :
1485+
#(⋃ j < o, A j) ≤ c := by
1486+
simp_rw [← mem_Iio, biUnion_eq_iUnion, iUnion, iSup, ← o.enumIsoOut.symm.surjective.range_comp]
1487+
apply ((mk_iUnion_le _).trans _).trans_eq (mul_eq_self hc)
1488+
rw [mk_ordinal_out]
1489+
exact mul_le_mul' ho <| ciSup_le' <| (hA _ <| typein_lt_self ·)
1490+
1491+
end Cardinal
1492+
1493+
end OrdinalIndices

Mathlib/SetTheory/Ordinal/Basic.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1378,6 +1378,19 @@ theorem lt_ord_succ_card (o : Ordinal) : o < (succ o.card).ord :=
13781378
lt_ord.2 <| lt_succ _
13791379
#align cardinal.lt_ord_succ_card Cardinal.lt_ord_succ_card
13801380

1381+
theorem card_le_iff {o : Ordinal} {c : Cardinal} : o.card ≤ c ↔ o < (succ c).ord := by
1382+
rw [lt_ord, lt_succ_iff]
1383+
1384+
/--
1385+
A variation on `Cardinal.lt_ord` using `≤`: If `o` is no greater than the
1386+
initial ordinal of cardinality `c`, then its cardinal is no greater than `c`.
1387+
1388+
The converse, however, is false (for instance, `o = ω+1` and `c = ℵ₀`).
1389+
-/
1390+
lemma card_le_of_le_ord {o : Ordinal} {c : Cardinal} (ho : o ≤ c.ord) :
1391+
o.card ≤ c := by
1392+
rw [← card_ord c]; exact Ordinal.card_le_card ho
1393+
13811394
@[mono]
13821395
theorem ord_strictMono : StrictMono ord :=
13831396
gciOrdCard.strictMono_l

0 commit comments

Comments
 (0)