|
| 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 set_theory.zfc.basic |
| 8 | + |
| 9 | +/-! |
| 10 | +# Von Neumann ordinals |
| 11 | +
|
| 12 | +This file works towards the development of von Neumann ordinals, i.e. transitive sets, well-ordered |
| 13 | +under `∈`. We currently only have an initial development of transitive sets. |
| 14 | +
|
| 15 | +Further development can be found on the branch `von_neumann_v2`. |
| 16 | +
|
| 17 | +## Definitions |
| 18 | +
|
| 19 | +- `Set.is_transitive` means that every element of a set is a subset. |
| 20 | +
|
| 21 | +## Todo |
| 22 | +
|
| 23 | +- Define von Neumann ordinals. |
| 24 | +- Define the basic arithmetic operations on ordinals from a purely set-theoretic perspective. |
| 25 | +- Prove the equivalences between these definitions and those provided in |
| 26 | + `set_theory/ordinal/arithmetic.lean`. |
| 27 | +-/ |
| 28 | + |
| 29 | +namespace Set |
| 30 | + |
| 31 | +/-- A transitive set is one where every element is a subset. -/ |
| 32 | +def is_transitive (x : Set) : Prop := ∀ y ∈ x, y ⊆ x |
| 33 | + |
| 34 | +theorem is_transitive.subset_of_mem {x y : Set} (h : x.is_transitive) : y ∈ x → y ⊆ x := h y |
| 35 | + |
| 36 | +theorem is_transitive_iff_mem_trans {z : Set} : |
| 37 | + z.is_transitive ↔ ∀ {x y : Set}, x ∈ y → y ∈ z → x ∈ z := |
| 38 | +⟨λ h x y hx hy, h.subset_of_mem hy hx, λ H x hx y hy, H hy hx⟩ |
| 39 | + |
| 40 | +alias is_transitive_iff_mem_trans ↔ is_transitive.mem_trans _ |
| 41 | + |
| 42 | +theorem is_transitive.sUnion {x : Set} (h : x.is_transitive) : (⋃₀ x).is_transitive := |
| 43 | +λ y hy z hz, begin |
| 44 | + rcases mem_sUnion.1 hy with ⟨w, hw, hw'⟩, |
| 45 | + exact mem_sUnion_of_mem hz (h.mem_trans hw' hw) |
| 46 | +end |
| 47 | + |
| 48 | +theorem is_transitive_iff_sUnion_subset {x : Set} : x.is_transitive ↔ ⋃₀ x ⊆ x := |
| 49 | +⟨λ h y hy, by { rcases mem_sUnion.1 hy with ⟨z, hz, hz'⟩, exact h.mem_trans hz' hz }, |
| 50 | + λ H y hy z hz, H $ mem_sUnion_of_mem hz hy⟩ |
| 51 | + |
| 52 | +alias is_transitive_iff_sUnion_subset ↔ is_transitive.sUnion_subset _ |
| 53 | + |
| 54 | +theorem is_transitive_iff_subset_powerset {x : Set} : x.is_transitive ↔ x ⊆ powerset x := |
| 55 | +⟨λ h y hy, mem_powerset.2 $ h.subset_of_mem hy, λ H y hy z hz, mem_powerset.1 (H hy) hz⟩ |
| 56 | + |
| 57 | +alias is_transitive_iff_subset_powerset ↔ is_transitive.subset_powerset _ |
| 58 | + |
| 59 | +end Set |
0 commit comments