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

Commit cac4e19

Browse files
committed
feat(set_theory/ordinal_arithmetic): Proved characterization of log (#11192)
As well as a few simple missing lemmas.
1 parent b67857e commit cac4e19

File tree

1 file changed

+48
-1
lines changed

1 file changed

+48
-1
lines changed

src/set_theory/ordinal_arithmetic.lean

Lines changed: 48 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -124,6 +124,9 @@ by simp only [le_antisymm_iff, add_le_add_iff_left]
124124
theorem lt_succ {a b : ordinal} : a < succ b ↔ a ≤ b :=
125125
by rw [← not_le, succ_le, not_lt]
126126

127+
theorem lt_one_iff_zero {a : ordinal} : a < 1 ↔ a = 0 :=
128+
by rw [←succ_zero, lt_succ, ordinal.le_zero]
129+
127130
theorem add_lt_add_iff_left (a) {b c : ordinal} : a + b < a + c ↔ b < c :=
128131
by rw [← not_le, ← not_le, add_le_add_iff_left]
129132

@@ -591,6 +594,12 @@ quotient.induction_on₃ a b c $ λ ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨γ, t, _⟩
591594
{ exact prod.lex.right _ (f.to_rel_embedding.map_rel_iff.2 h') }
592595
end
593596

597+
theorem le_mul_left (a : ordinal) {b : ordinal} (hb : 0 < b) : a ≤ a * b :=
598+
by { convert mul_le_mul_left a (one_le_iff_pos.2 hb), rw mul_one a }
599+
600+
theorem le_mul_right (a : ordinal) {b : ordinal} (hb : 0 < b) : a ≤ b * a :=
601+
by { convert mul_le_mul_right a (one_le_iff_pos.2 hb), rw one_mul a }
602+
594603
theorem mul_le_mul {a b c d : ordinal} (h₁ : a ≤ c) (h₂ : b ≤ d) : a * b ≤ c * d :=
595604
le_trans (mul_le_mul_left _ h₂) (mul_le_mul_right _ h₁)
596605

@@ -1269,7 +1278,7 @@ end
12691278
/-! ### Ordinal logarithm -/
12701279

12711280
/-- The ordinal logarithm is the solution `u` to the equation
1272-
`x = b ^ u * v + w` where `v < b` and `w < b`. -/
1281+
`x = b ^ u * v + w` where `v < b` and `w < b ^ u`. -/
12731282
def log (b : ordinal) (x : ordinal) : ordinal :=
12741283
if h : 1 < b then pred $
12751284
omin {o | x < b^o} ⟨succ x, succ_le.1 (le_power_self _ h)⟩
@@ -1350,6 +1359,44 @@ if b1 : 1 < b then
13501359
le_trans (le_power_self _ b1) (power_log_le b (ordinal.pos_iff_ne_zero.2 x0))
13511360
else by simp only [log_not_one_lt b1, ordinal.zero_le]
13521361

1362+
@[simp] theorem log_one (b : ordinal) : log b 1 = 0 :=
1363+
if hb : 1 < b then by rwa [←lt_one_iff_zero, log_lt hb zero_lt_one, power_one]
1364+
else log_not_one_lt hb 1
1365+
1366+
lemma power_mul_add_pos {b v : ordinal} (hb : 0 < b) (u) (hv : 0 < v) (w) :
1367+
0 < b ^ u * v + w :=
1368+
(power_pos u hb).trans_le ((le_mul_left _ hv).trans (le_add_right _ _))
1369+
1370+
lemma power_mul_add_lt_power_mul_succ {b u w : ordinal} (v : ordinal) (hw : w < b ^ u) :
1371+
b ^ u * v + w < b ^ u * v.succ :=
1372+
by rwa [mul_succ, add_lt_add_iff_left]
1373+
1374+
lemma power_mul_add_lt_power_succ {b u v w : ordinal} (hvb : v < b) (hw : w < b ^ u) :
1375+
b ^ u * v + w < b ^ u.succ :=
1376+
begin
1377+
convert (power_mul_add_lt_power_mul_succ v hw).trans_le (mul_le_mul_left _ (succ_le.2 hvb)),
1378+
exact power_succ b u
1379+
end
1380+
1381+
theorem log_power_mul_add {b u v w : ordinal} (hb : 1 < b) (hv : 0 < v) (hvb : v < b)
1382+
(hw : w < b ^ u) : log b (b ^ u * v + w) = u :=
1383+
begin
1384+
have hpos := power_mul_add_pos (zero_lt_one.trans hb) u hv w,
1385+
by_contra' hne,
1386+
cases lt_or_gt_of_ne hne with h h,
1387+
{ rw log_lt hb hpos at h,
1388+
exact not_le_of_lt h ((le_mul_left _ hv).trans (le_add_right _ _)) },
1389+
{ change _ < _ at h,
1390+
rw [←succ_le, le_log hb hpos] at h,
1391+
exact (not_lt_of_le h) (power_mul_add_lt_power_succ hvb hw) }
1392+
end
1393+
1394+
@[simp] theorem log_power {b : ordinal} (hb : 1 < b) (x : ordinal) : log b (b ^ x) = x :=
1395+
begin
1396+
convert log_power_mul_add hb zero_lt_one hb (power_pos x (zero_lt_one.trans hb)),
1397+
rw [add_zero, mul_one]
1398+
end
1399+
13531400
theorem add_log_le_log_mul {x y : ordinal} (b : ordinal) (x0 : 0 < x) (y0 : 0 < y) :
13541401
log b x + log b y ≤ log b (x * y) :=
13551402
begin

0 commit comments

Comments
 (0)