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

Commit 7967128

Browse files
committed
feat(data/complex/basic): #ℂ = 𝔠 (#12871)
1 parent 584ae9d commit 7967128

File tree

1 file changed

+30
-0
lines changed

1 file changed

+30
-0
lines changed

src/data/complex/cardinality.lean

Lines changed: 30 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,30 @@
1+
/-
2+
Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Violeta Hernández Palacios
5+
-/
6+
7+
import data.complex.basic
8+
import data.real.cardinality
9+
10+
/-!
11+
# The cardinality of the complex numbers
12+
13+
This file shows that the complex numbers have cardinality continuum, i.e. `#ℂ = 𝔠`.
14+
-/
15+
16+
open cardinal set
17+
18+
open_locale cardinal
19+
20+
/-- The cardinality of the complex numbers, as a type. -/
21+
@[simp] theorem mk_complex : #ℂ = 𝔠 :=
22+
by rw [mk_congr complex.equiv_real_prod, mk_prod, lift_id, mk_real, continuum_mul_self]
23+
24+
/-- The cardinality of the complex numbers, as a set. -/
25+
@[simp] lemma mk_univ_complex : #(set.univ : set ℂ) = 𝔠 :=
26+
by rw [mk_univ, mk_complex]
27+
28+
/-- The complex numbers are not countable. -/
29+
lemma not_countable_complex : ¬ countable (set.univ : set ℂ) :=
30+
by { rw [← mk_set_le_omega, not_le, mk_univ_complex], apply cantor }

0 commit comments

Comments
 (0)