Skip to content

Commit

Permalink
chore: proved Rat.eq_neg_of_add_eq_zero_left
Browse files Browse the repository at this point in the history
  • Loading branch information
AdrienChampion committed Jul 8, 2024
1 parent 013047f commit 02c0907
Showing 1 changed file with 12 additions and 13 deletions.
25 changes: 12 additions & 13 deletions Smt/Reconstruct/Rat/Core.lean
Original file line number Diff line number Diff line change
Expand Up @@ -66,6 +66,8 @@ protected theorem eq_mkRat : ∀ a : Rat, a = mkRat a.num a.den := fun a => by
@[simp]
theorem mk'_zero (d) (h : d ≠ 0) (w) : mk' 0 d h w = 0 := by
congr
apply Nat.coprime_zero_left d |>.mp w


theorem eq_iff_mul_eq_mul {p q : Rat} : p = q ↔ p.num * q.den = q.num * p.den := by
conv =>
Expand Down Expand Up @@ -109,7 +111,7 @@ protected theorem lt_asymm {x y : Rat} : x < y → ¬ y < x := by
rw [h]
| inr h =>
split at h
case inl xnum_0 =>
case isTrue xnum_0 =>
simp [Int.not_lt_of_lt_rev h, xnum_0, h]
case inr xnum_ne_0 =>
let ⟨h, h'⟩ := h
Expand Down Expand Up @@ -357,18 +359,15 @@ protected theorem sub_self : x - x = 0 :=
protected theorem add_neg_self : x + -x = 0 :=
Rat.sub_eq_add_neg x x ▸ Rat.sub_self

-- #TODO use `eq_iff_mul_eq_mul`
protected theorem eq_neg_of_add_eq_zero_left : x + y = 0 → x = - y :=
numDenCasesOn'' x fun nx dx h_dx h_red =>
numDenCasesOn'' y fun ny dy h_dy h_red' => by
simp only [Rat.neg_divInt, Rat.add_def]
numDenCasesOn'' x fun nx dx h_dx h_dx_red =>
numDenCasesOn'' y fun ny dy h_dy h_dy_red => by
simp only [Rat.neg_divInt, Rat.add_def, Neg.neg]
simp only [Rat.neg, normalize_eq_zero]
simp only [eq_iff_mul_eq_mul, ← Int.neg_mul_eq_neg_mul]
intro h
simp [Neg.neg] ; simp [Rat.neg]
simp only [Rat.normalize_eq_zero] at h
let h_dpos : 0 < ↑dx := Int.natCast_pos.mpr (Nat.pos_iff_ne_zero.mpr h_dx)
let h_dpos' : 0 < ↑dy := Int.natCast_pos.mpr (Nat.pos_iff_ne_zero.mpr h_dy)
let h' := Int.neg_eq_of_add_eq_zero h
sorry
apply Int.eq_neg_of_eq_neg
exact Int.neg_eq_of_add_eq_zero h |>.symm

protected theorem le_iff_sub_nonneg (x y : Rat) : x ≤ y ↔ 0 ≤ y - x :=
numDenCasesOn'' x fun nx dx h_dx _ =>
Expand Down Expand Up @@ -399,10 +398,10 @@ protected theorem le_iff_sub_nonneg (x y : Rat) : x ≤ y ↔ 0 ≤ y - x :=
else
simp [h]
split
case inl nb_0 =>
case isTrue nb_0 =>
simp [nb_0, Rat.sub_eq_add_neg, Rat.zero_add, Rat.nonneg_sub_iff_nonpos, ← Rat.num_nonpos]
exact Int.not_lt
case inr nb_nz =>
case isFalse nb_nz =>
simp only [Rat.sub_def, normalize_eq, ← Rat.num_nonneg]
if ny_pos : 0 < ny then
simp [ny_pos]
Expand Down

0 comments on commit 02c0907

Please sign in to comment.