Proof that a cube (in dimension n ≥ 3) cannot be cubed: There does not exist a partition of a cube into finitely many smaller cubes (at least two) of different sizes.
We follow the proof described here: http://www.alaricstephen.com/main-featured/2017/9/28/cubing-a-cube-proof
Given three intervals I, J, K such that J ⊂ I,
neither endpoint of J coincides with an endpoint of I, ¬ (K ⊆ J) and
K does not lie completely to the left nor completely to the right of J.
Then I ∩ K \ J is nonempty.
Equations
- Theorems100.«82».Cube.unitCube = { b := fun (x : Fin n) => 0, w := 1, hw := Theorems100.«82».Cube.unitCube._proof_1 }
Instances For
A finite family of (at least 2) cubes partitioning the unit cube with different sizes
- PairwiseDisjoint : Pairwise (Function.onFun Disjoint (Cube.toSet ∘ cs))
- Injective : Function.Injective (Cube.w ∘ cs)
Instances For
The width of any cube in the partition cannot be 1.
A valley is a square on which cubes in the family of cubes are placed, so that the cubes
completely cover the valley and none of those cubes is partially outside the square.
We also require that no cube on it has the same size as the valley (so that there are at least
two cubes on the valley).
This is the main concept in the formalization.
We prove that the smallest cube on a valley has another valley on the top of it, which
gives an infinite sequence of cubes in the partition, which contradicts the finiteness.
A valley is characterized by a cube c (which is not a cube in the family cs) by considering
the bottom face of c.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The bottom of the unit cube is a valley
If mi lies on the boundary of the valley in dimension j, then this lemma expresses that all
other cubes on the same boundary extend further from the boundary.
More precisely, there is a j-th coordinate x : ℝ in the valley, but not in mi,
such that every cube that shares a (particular) j-th coordinate with mi also contains j-th
coordinate x
mi cannot lie on the boundary of the valley. Otherwise, the cube adjacent to it in the j-th
direction will intersect one of the neighbouring cubes on the same boundary as mi.
We get a sequence of cubes whose size is decreasing
Equations
Instances For
Equations
- Theorems100.«82».decreasingSequence h k = (cs ↑(Theorems100.«82».sequenceOfCubes h k)).w
Instances For
The infinite sequence of cubes contradicts the finiteness of the family.
Dissection of Cubes: A cube cannot be cubed.