Skip to content

Commit 305fba6

Browse files
authored
feat: missing lemmas about Int order/multiplication (#8346)
This PR adds some missing lemmas about consequences of positivity/non-negativity of `a * b : Int`.
1 parent 8300121 commit 305fba6

File tree

1 file changed

+51
-0
lines changed

1 file changed

+51
-0
lines changed

src/Init/Data/Int/Order.lean

Lines changed: 51 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -166,6 +166,9 @@ protected theorem lt_or_le (a b : Int) : a < b ∨ b ≤ a := by rw [← Int.not
166166
protected theorem le_of_not_gt {a b : Int} (h : ¬ a > b) : a ≤ b :=
167167
Int.not_lt.mp h
168168

169+
protected theorem not_lt_of_ge {a b : Int} (h : b ≤ a) : ¬a < b :=
170+
Int.not_lt.mpr h
171+
169172
protected theorem lt_trichotomy (a b : Int) : a < b ∨ a = b ∨ b < a :=
170173
if eq : a = b then .inr <| .inl eq else
171174
if le : a ≤ b then .inl <| Int.lt_iff_le_and_ne.2 ⟨le, eq⟩ else
@@ -1181,6 +1184,54 @@ protected theorem nonpos_of_mul_nonpos_right {a b : Int}
11811184
(h : a * b ≤ 0) (ha : 0 < a) : b ≤ 0 :=
11821185
Int.le_of_not_gt fun hb : b > 0 => Int.not_le_of_gt (Int.mul_pos ha hb) h
11831186

1187+
protected theorem nonneg_of_mul_nonpos_left {a b : Int}
1188+
(h : a * b ≤ 0) (hb : b < 0) : 0 ≤ a :=
1189+
Int.le_of_not_gt fun ha => Int.not_le_of_gt (Int.mul_pos_of_neg_of_neg ha hb) h
1190+
1191+
protected theorem nonneg_of_mul_nonpos_right {a b : Int}
1192+
(h : a * b ≤ 0) (ha : a < 0) : 0 ≤ b :=
1193+
Int.le_of_not_gt fun hb => Int.not_le_of_gt (Int.mul_pos_of_neg_of_neg ha hb) h
1194+
1195+
protected theorem nonpos_of_mul_nonneg_left {a b : Int}
1196+
(h : 0 ≤ a * b) (hb : b < 0) : a ≤ 0 :=
1197+
Int.le_of_not_gt fun ha : a > 0 => Int.not_le_of_gt (Int.mul_neg_of_pos_of_neg ha hb) h
1198+
1199+
protected theorem nonpos_of_mul_nonneg_right {a b : Int}
1200+
(h : 0 ≤ a * b) (ha : a < 0) : b ≤ 0 :=
1201+
Int.le_of_not_gt fun hb : b > 0 => Int.not_le_of_gt (Int.mul_neg_of_neg_of_pos ha hb) h
1202+
1203+
protected theorem pos_of_mul_pos_left {a b : Int}
1204+
(h : 0 < a * b) (hb : 0 < b) : 0 < a :=
1205+
Int.lt_of_not_ge fun ha => Int.not_lt_of_ge (Int.mul_nonpos_of_nonpos_of_nonneg ha (Int.le_of_lt hb)) h
1206+
1207+
protected theorem pos_of_mul_pos_right {a b : Int}
1208+
(h : 0 < a * b) (ha : 0 < a) : 0 < b :=
1209+
Int.lt_of_not_ge fun hb => Int.not_lt_of_ge (Int.mul_nonpos_of_nonneg_of_nonpos (Int.le_of_lt ha) hb) h
1210+
1211+
protected theorem neg_of_mul_neg_left {a b : Int}
1212+
(h : a * b < 0) (hb : 0 < b) : a < 0 :=
1213+
Int.lt_of_not_ge fun ha => Int.not_lt_of_ge (Int.mul_nonneg ha (Int.le_of_lt hb)) h
1214+
1215+
protected theorem neg_of_mul_neg_right {a b : Int}
1216+
(h : a * b < 0) (ha : 0 < a) : b < 0 :=
1217+
Int.lt_of_not_ge fun hb => Int.not_lt_of_ge (Int.mul_nonneg (Int.le_of_lt ha) hb) h
1218+
1219+
protected theorem pos_of_mul_neg_left {a b : Int}
1220+
(h : a * b < 0) (hb : b < 0) : 0 < a :=
1221+
Int.lt_of_not_ge fun ha => Int.not_lt_of_ge (Int.mul_nonneg_of_nonpos_of_nonpos ha (Int.le_of_lt hb)) h
1222+
1223+
protected theorem pos_of_mul_neg_right {a b : Int}
1224+
(h : a * b < 0) (ha : a < 0) : 0 < b :=
1225+
Int.lt_of_not_ge fun hb => Int.not_lt_of_ge (Int.mul_nonneg_of_nonpos_of_nonpos (Int.le_of_lt ha) hb) h
1226+
1227+
protected theorem neg_of_mul_pos_left {a b : Int}
1228+
(h : 0 < a * b) (hb : b < 0) : a < 0 :=
1229+
Int.lt_of_not_ge fun ha => Int.not_lt_of_ge (Int.mul_nonpos_of_nonneg_of_nonpos ha (Int.le_of_lt hb)) h
1230+
1231+
protected theorem neg_of_mul_pos_right {a b : Int}
1232+
(h : 0 < a * b) (ha : a < 0) : b < 0 :=
1233+
Int.lt_of_not_ge fun hb => Int.not_lt_of_ge (Int.mul_nonpos_of_nonpos_of_nonneg (Int.le_of_lt ha) hb) h
1234+
11841235
/- ## sign -/
11851236

11861237
@[simp] theorem sign_zero : sign 0 = 0 := rfl

0 commit comments

Comments
 (0)