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

Commit 0098286

Browse files
committed
feat(set_theory/ordinal/natural_ops): define natural addition (#14291)
We define the natural addition operation on ordinals. We prove the basic properties, like commutativity, associativity, and cancellativity. We also provide the type synonym `nat_ordinal` for ordinals with natural operations, which allows us to take full advantage of this rich algebraic structure.
1 parent d63246c commit 0098286

File tree

1 file changed

+327
-0
lines changed

1 file changed

+327
-0
lines changed
Lines changed: 327 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,327 @@
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.ordinal.arithmetic
8+
9+
/-!
10+
# Natural operations on ordinals
11+
12+
The goal of this file is to define natural addition and multiplication on ordinals, and provide a
13+
basic API. The natural addition of two ordinals `a ♯ b` is recursively defined as the least ordinal
14+
greater than `a' ♯ b` and `a ♯ b'` for `a' < a` and `b' < b`. The natural multiplication `a ⨳ b` is
15+
likewise recursively defined as the least ordinal such that `a ⨳ b ♯ a' ⨳ b'` is greater than
16+
`a' ⨳ b ♯ a ⨳ b'` for any `a' < a` and `b' < b`.
17+
18+
These operations form a rich algebraic structure: they're commutative, associative, preserve order,
19+
have the usual `0` and `1` from ordinals, and distribute over one another.
20+
21+
Moreover, these operations are the addition and multiplication of ordinals when viewed as
22+
combinatorial `game`s. This makes them particularly useful for game theory.
23+
24+
Finally, both operations admit simple, intuitive descriptions in terms of the Cantor normal form.
25+
The natural addition of two ordinals corresponds to adding their Cantor normal forms as if they were
26+
polynomials in `ω`. Likewise, their natural multiplication corresponds to multiplying the Cantor
27+
normal forms as polynomials.
28+
29+
# Implementation notes
30+
31+
Given the rich algebraic structure of these two operations, we choose to create a type synonym
32+
`nat_ordinal`, where we provide the appropriate instances. However, to avoid casting back and forth
33+
between both types, we attempt to prove and state most results on `ordinal`.
34+
35+
# Todo
36+
37+
- Define natural multiplication and provide a basic API.
38+
- Prove the characterizations of natural addition and multiplication in terms of the Cantor normal
39+
form.
40+
-/
41+
42+
universes u v
43+
44+
open function order
45+
46+
noncomputable theory
47+
48+
/-- A type synonym for ordinals with natural addition and multiplication. -/
49+
def nat_ordinal : Type* := ordinal
50+
51+
instance : linear_order nat_ordinal := ordinal.linear_order
52+
53+
/-- The identity function between `ordinal` and `nat_ordinal`. -/
54+
@[pattern] def ordinal.to_nat_ordinal : ordinal ≃o nat_ordinal := order_iso.refl _
55+
56+
/-- The identity function between `nat_ordinal` and `ordinal`. -/
57+
@[pattern] def nat_ordinal.to_ordinal : nat_ordinal ≃o ordinal := order_iso.refl _
58+
59+
open ordinal
60+
61+
namespace nat_ordinal
62+
63+
variables {a b c : nat_ordinal.{u}}
64+
65+
@[simp] theorem to_ordinal_symm_eq : nat_ordinal.to_ordinal.symm = ordinal.to_nat_ordinal := rfl
66+
@[simp] theorem to_ordinal_to_nat_ordinal (a : nat_ordinal) : a.to_ordinal.to_nat_ordinal = a := rfl
67+
68+
instance : has_zero nat_ordinal := ⟨to_nat_ordinal 0
69+
instance : inhabited nat_ordinal := ⟨0
70+
instance : has_one nat_ordinal := ⟨to_nat_ordinal 1
71+
instance : succ_order nat_ordinal := ordinal.succ_order
72+
73+
theorem lt_wf : @well_founded nat_ordinal (<) := ordinal.lt_wf
74+
instance : has_well_founded nat_ordinal := ordinal.has_well_founded
75+
instance : is_well_order nat_ordinal (<) := ordinal.has_lt.lt.is_well_order
76+
77+
@[simp] theorem to_ordinal_zero : to_ordinal 0 = 0 := rfl
78+
@[simp] theorem to_ordinal_one : to_ordinal 1 = 1 := rfl
79+
80+
@[simp] theorem to_ordinal_eq_zero (a) : to_ordinal a = 0 ↔ a = 0 := iff.rfl
81+
@[simp] theorem to_ordinal_eq_one (a) : to_ordinal a = 1 ↔ a = 1 := iff.rfl
82+
83+
@[simp] theorem to_ordinal_max : (max a b).to_ordinal = max a.to_ordinal b.to_ordinal := rfl
84+
@[simp] theorem to_ordinal_min : (min a b).to_ordinal = min a.to_ordinal b.to_ordinal := rfl
85+
86+
theorem succ_def (a : nat_ordinal) : succ a = (a.to_ordinal + 1).to_nat_ordinal := rfl
87+
88+
/-- A recursor for `nat_ordinal`. Use as `induction x using nat_ordinal.rec`. -/
89+
protected def rec {β : nat_ordinal → Sort*} (h : Π a, β (to_nat_ordinal a)) : Π a, β a :=
90+
λ a, h a.to_ordinal
91+
92+
/-- `ordinal.induction` but for `nat_ordinal`. -/
93+
theorem induction {p : nat_ordinal → Prop} : ∀ i (h : ∀ j, (∀ k, k < j → p k) → p j), p i :=
94+
ordinal.induction
95+
96+
end nat_ordinal
97+
98+
namespace ordinal
99+
100+
variables {a b c : ordinal.{u}}
101+
102+
@[simp] theorem to_nat_ordinal_symm_eq : to_nat_ordinal.symm = nat_ordinal.to_ordinal := rfl
103+
@[simp] theorem to_nat_ordinal_to_ordinal (a : ordinal) : a.to_nat_ordinal.to_ordinal = a := rfl
104+
105+
@[simp] theorem to_nat_ordinal_zero : to_nat_ordinal 0 = 0 := rfl
106+
@[simp] theorem to_nat_ordinal_one : to_nat_ordinal 1 = 1 := rfl
107+
108+
@[simp] theorem to_nat_ordinal_eq_zero (a) : to_nat_ordinal a = 0 ↔ a = 0 := iff.rfl
109+
@[simp] theorem to_nat_ordinal_eq_one (a) : to_nat_ordinal a = 1 ↔ a = 1 := iff.rfl
110+
111+
@[simp] theorem to_nat_ordinal_max :
112+
to_nat_ordinal (max a b) = max a.to_nat_ordinal b.to_nat_ordinal := rfl
113+
@[simp] theorem to_nat_ordinal_min :
114+
(linear_order.min a b).to_nat_ordinal = linear_order.min a.to_nat_ordinal b.to_nat_ordinal := rfl
115+
116+
/-- Natural addition on ordinals `a ♯ b` is recursively defined as the least ordinal greater than
117+
`a' ♯ b` and `a ♯ b'` for all `a' < a` and `b' < b`. In contrast to normal ordinal addition, it is
118+
commutative.
119+
120+
Natural addition can equivalently be characterized as the ordinal resulting from adding up
121+
corresponding coefficients in the Cantor normal forms of `a` and `b`. -/
122+
noncomputable def nadd : ordinal → ordinal → ordinal
123+
| a b := max
124+
(blsub.{u u} a $ λ a' h, nadd a' b)
125+
(blsub.{u u} b $ λ b' h, nadd a b')
126+
using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] }
127+
128+
local infix ` ♯ `:65 := nadd
129+
130+
theorem nadd_def (a b : ordinal) : a ♯ b = max
131+
(blsub.{u u} a $ λ a' h, a' ♯ b)
132+
(blsub.{u u} b $ λ b' h, a ♯ b') :=
133+
by rw nadd
134+
135+
theorem lt_nadd_iff : a < b ♯ c ↔ (∃ b' < b, a ≤ b' ♯ c) ∨ ∃ c' < c, a ≤ b ♯ c' :=
136+
by { rw nadd_def, simp [lt_blsub_iff] }
137+
138+
theorem nadd_le_iff : b ♯ c ≤ a ↔ (∀ b' < b, b' ♯ c < a) ∧ ∀ c' < c, b ♯ c' < a :=
139+
by { rw nadd_def, simp [blsub_le_iff] }
140+
141+
theorem nadd_lt_nadd_left (h : b < c) (a) : a ♯ b < a ♯ c :=
142+
lt_nadd_iff.2 (or.inr ⟨b, h, le_rfl⟩)
143+
144+
theorem nadd_lt_nadd_right (h : b < c) (a) : b ♯ a < c ♯ a :=
145+
lt_nadd_iff.2 (or.inl ⟨b, h, le_rfl⟩)
146+
147+
theorem nadd_le_nadd_left (h : b ≤ c) (a) : a ♯ b ≤ a ♯ c :=
148+
begin
149+
rcases lt_or_eq_of_le h with h | rfl,
150+
{ exact (nadd_lt_nadd_left h a).le },
151+
{ exact le_rfl }
152+
end
153+
154+
theorem nadd_le_nadd_right (h : b ≤ c) (a) : b ♯ a ≤ c ♯ a :=
155+
begin
156+
rcases lt_or_eq_of_le h with h | rfl,
157+
{ exact (nadd_lt_nadd_right h a).le },
158+
{ exact le_rfl }
159+
end
160+
161+
variables (a b)
162+
163+
theorem nadd_comm : ∀ a b, a ♯ b = b ♯ a
164+
| a b := begin
165+
rw [nadd_def, nadd_def, max_comm],
166+
congr; ext c hc;
167+
apply nadd_comm
168+
end
169+
using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] }
170+
171+
theorem blsub_nadd_of_mono {f : Π c < a ♯ b, ordinal.{max u v}}
172+
(hf : ∀ {i j} hi hj, i ≤ j → f i hi ≤ f j hj) : blsub _ f = max
173+
(blsub.{u v} a (λ a' ha', f (a' ♯ b) $ nadd_lt_nadd_right ha' b))
174+
(blsub.{u v} b (λ b' hb', f (a ♯ b') $ nadd_lt_nadd_left hb' a)) :=
175+
begin
176+
apply (blsub_le_iff.2 (λ i h, _)).antisymm (max_le _ _),
177+
{ rcases lt_nadd_iff.1 h with ⟨a', ha', hi⟩ | ⟨b', hb', hi⟩,
178+
{ exact lt_max_of_lt_left ((hf h (nadd_lt_nadd_right ha' b) hi).trans_lt (lt_blsub _ _ _)) },
179+
{ exact lt_max_of_lt_right ((hf h (nadd_lt_nadd_left hb' a) hi).trans_lt (lt_blsub _ _ _)) } },
180+
all_goals
181+
{ apply blsub_le_of_brange_subset.{u u v},
182+
rintro c ⟨d, hd, rfl⟩,
183+
apply mem_brange_self }
184+
end
185+
186+
theorem nadd_assoc : ∀ a b c, a ♯ b ♯ c = a ♯ (b ♯ c)
187+
| a b c := begin
188+
rw [nadd_def a (b ♯ c), nadd_def, blsub_nadd_of_mono, blsub_nadd_of_mono, max_assoc],
189+
{ congr; ext d hd;
190+
apply nadd_assoc },
191+
{ exact λ i j _ _ h, nadd_le_nadd_left h a },
192+
{ exact λ i j _ _ h, nadd_le_nadd_right h c }
193+
end
194+
using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] }
195+
196+
@[simp] theorem nadd_zero : a ♯ 0 = a :=
197+
begin
198+
induction a using ordinal.induction with a IH,
199+
rw [nadd_def, blsub_zero, max_zero_right],
200+
convert blsub_id a,
201+
ext b hb,
202+
exact IH _ hb
203+
end
204+
205+
@[simp] theorem zero_nadd : 0 ♯ a = a :=
206+
by rw [nadd_comm, nadd_zero]
207+
208+
@[simp] theorem nadd_one : a ♯ 1 = succ a :=
209+
begin
210+
induction a using ordinal.induction with a IH,
211+
rw [nadd_def, blsub_one, nadd_zero, max_eq_right_iff, blsub_le_iff],
212+
intros i hi,
213+
rwa [IH i hi, succ_lt_succ_iff]
214+
end
215+
216+
@[simp] theorem one_nadd : 1 ♯ a = succ a :=
217+
by rw [nadd_comm, nadd_one]
218+
219+
theorem nadd_succ : a ♯ succ b = succ (a ♯ b) :=
220+
by rw [←nadd_one (a ♯ b), nadd_assoc, nadd_one]
221+
222+
theorem succ_nadd : succ a ♯ b = succ (a ♯ b) :=
223+
by rw [←one_nadd (a ♯ b), ←nadd_assoc, one_nadd]
224+
225+
@[simp] theorem nadd_nat (n : ℕ) : a ♯ n = a + n :=
226+
begin
227+
induction n with n hn,
228+
{ simp },
229+
{ rw [nat.cast_succ, add_one_eq_succ, nadd_succ, add_succ, hn] }
230+
end
231+
232+
@[simp] theorem nat_nadd (n : ℕ) : ↑n ♯ a = a + n :=
233+
by rw [nadd_comm, nadd_nat]
234+
235+
theorem add_le_nadd : a + b ≤ a ♯ b :=
236+
begin
237+
apply b.limit_rec_on,
238+
{ simp },
239+
{ intros c h,
240+
rwa [add_succ, nadd_succ, succ_le_succ_iff] },
241+
{ intros c hc H,
242+
rw [←is_normal.blsub_eq.{u u} (add_is_normal a) hc, blsub_le_iff],
243+
exact λ i hi, (H i hi).trans_lt (nadd_lt_nadd_left hi a) }
244+
end
245+
246+
end ordinal
247+
248+
open ordinal
249+
250+
namespace nat_ordinal
251+
252+
instance : has_add nat_ordinal := ⟨nadd⟩
253+
254+
instance add_covariant_class_lt :
255+
covariant_class nat_ordinal.{u} nat_ordinal.{u} (+) (<) :=
256+
⟨λ a b c h, nadd_lt_nadd_left h a⟩
257+
258+
instance add_covariant_class_le :
259+
covariant_class nat_ordinal.{u} nat_ordinal.{u} (+) (≤) :=
260+
⟨λ a b c h, nadd_le_nadd_left h a⟩
261+
262+
instance add_contravariant_class_le :
263+
contravariant_class nat_ordinal.{u} nat_ordinal.{u} (+) (≤) :=
264+
⟨λ a b c h, by { by_contra' h', exact h.not_lt (add_lt_add_left h' a) }⟩
265+
266+
instance : ordered_cancel_add_comm_monoid nat_ordinal :=
267+
{ add := (+),
268+
add_assoc := nadd_assoc,
269+
add_left_cancel := λ a b c, add_left_cancel'',
270+
add_le_add_left := λ a b, add_le_add_left,
271+
le_of_add_le_add_left := λ a b c, le_of_add_le_add_left,
272+
zero := 0,
273+
zero_add := zero_nadd,
274+
add_zero := nadd_zero,
275+
add_comm := nadd_comm,
276+
..nat_ordinal.linear_order }
277+
278+
@[simp] theorem add_one_eq_succ : ∀ a : nat_ordinal, a + 1 = succ a := nadd_one
279+
280+
@[simp] theorem to_ordinal_cast_nat (n : ℕ) : to_ordinal n = n :=
281+
begin
282+
induction n with n hn,
283+
{ refl },
284+
{ change nadd (to_ordinal n) 1 = n + 1,
285+
rw hn,
286+
apply nadd_one }
287+
end
288+
289+
end nat_ordinal
290+
291+
open nat_ordinal
292+
293+
namespace ordinal
294+
295+
local infix ` ♯ `:65 := nadd
296+
297+
@[simp] theorem to_nat_ordinal_cast_nat (n : ℕ) : to_nat_ordinal n = n :=
298+
by { rw ←to_ordinal_cast_nat n, refl }
299+
300+
theorem lt_of_nadd_lt_nadd_left : ∀ {a b c}, a ♯ b < a ♯ c → b < c :=
301+
@lt_of_add_lt_add_left nat_ordinal _ _ _
302+
theorem lt_of_nadd_lt_nadd_right : ∀ {a b c}, b ♯ a < c ♯ a → b < c :=
303+
@_root_.lt_of_add_lt_add_right nat_ordinal _ _ _
304+
theorem le_of_nadd_le_nadd_left : ∀ {a b c}, a ♯ b ≤ a ♯ c → b ≤ c :=
305+
@le_of_add_le_add_left nat_ordinal _ _ _
306+
theorem le_of_nadd_le_nadd_right : ∀ {a b c}, b ♯ a ≤ c ♯ a → b ≤ c :=
307+
@le_of_add_le_add_right nat_ordinal _ _ _
308+
309+
theorem nadd_lt_nadd_iff_left : ∀ a {b c}, a ♯ b < a ♯ c ↔ b < c :=
310+
@add_lt_add_iff_left nat_ordinal _ _ _ _
311+
theorem nadd_lt_nadd_iff_right : ∀ a {b c}, b ♯ a < c ♯ a ↔ b < c :=
312+
@add_lt_add_iff_right nat_ordinal _ _ _ _
313+
theorem nadd_le_nadd_iff_left : ∀ a {b c}, a ♯ b ≤ a ♯ c ↔ b ≤ c :=
314+
@add_le_add_iff_left nat_ordinal _ _ _ _
315+
theorem nadd_le_nadd_iff_right : ∀ a {b c}, b ♯ a ≤ c ♯ a ↔ b ≤ c :=
316+
@_root_.add_le_add_iff_right nat_ordinal _ _ _ _
317+
318+
theorem nadd_left_cancel : ∀ {a b c}, a ♯ b = a ♯ c → b = c :=
319+
@_root_.add_left_cancel nat_ordinal _
320+
theorem nadd_right_cancel : ∀ {a b c}, a ♯ b = c ♯ b → a = c :=
321+
@_root_.add_right_cancel nat_ordinal _
322+
theorem nadd_left_cancel_iff : ∀ {a b c}, a ♯ b = a ♯ c ↔ b = c :=
323+
@add_left_cancel_iff nat_ordinal _
324+
theorem nadd_right_cancel_iff : ∀ {a b c}, b ♯ a = c ♯ a ↔ b = c :=
325+
@add_right_cancel_iff nat_ordinal _
326+
327+
end ordinal

0 commit comments

Comments
 (0)