@@ -32,39 +32,76 @@ class IntModule (M : Type u) extends Zero M, Add M, Neg M, Sub M, HMul Int M M w
3232 zero_hmul : ∀ a : M, (0 : Int) * a = 0
3333 one_hmul : ∀ a : M, (1 : Int) * a = a
3434 add_hmul : ∀ n m : Int, ∀ a : M, (n + m) * a = n * a + m * a
35- neg_hmul : ∀ n : Int, ∀ a : M, (-n) * a = - (n * a)
3635 hmul_zero : ∀ n : Int, n * (0 : M) = 0
3736 hmul_add : ∀ n : Int, ∀ a b : M, n * (a + b) = n * a + n * b
3837 mul_hmul : ∀ n m : Int, ∀ a : M, (n * m) * a = n * (m * a)
3938 neg_add_cancel : ∀ a : M, -a + a = 0
4039 sub_eq_add_neg : ∀ a b : M, a - b = a + -b
4140
41+ namespace IntModule
42+
4243attribute [instance 100] IntModule.toZero IntModule.toAdd IntModule.toNeg IntModule.toSub IntModule.toHMul
4344
44- instance IntModule. toNatModule (M : Type u) [i : IntModule M] : NatModule M :=
45+ instance toNatModule (M : Type u) [i : IntModule M] : NatModule M :=
4546 { i with
4647 hMul a x := (a : Int) * x
4748 hmul_zero := by simp [IntModule.hmul_zero]
4849 add_hmul := by simp [IntModule.add_hmul]
4950 hmul_add := by simp [IntModule.hmul_add]
5051 mul_hmul := by simp [IntModule.mul_hmul] }
5152
52- /-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/
53- class Preorder (α : Type u) extends LE α, LT α where
54- le_refl : ∀ a : α, a ≤ a
55- le_trans : ∀ a b c : α, a ≤ b → b ≤ c → a ≤ c
56- lt := fun a b => a ≤ b ∧ ¬b ≤ a
57- lt_iff_le_not_le : ∀ a b : α, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl
58-
59- class IntModule.IsOrdered (M : Type u) [Preorder M] [IntModule M] where
60- neg_le_iff : ∀ a b : M, -a ≤ b ↔ -b ≤ a
61- neg_lt_iff : ∀ a b : M, -a < b ↔ -b < a
62- add_lt_left : ∀ a b c : M, a < b → a + c < b + c
63- add_lt_right : ∀ a b c : M, a < b → c + a < c + b
64- hmul_pos : ∀ (k : Int) (a : M), 0 < a → (0 < k ↔ 0 < k * a)
65- hmul_neg : ∀ (k : Int) (a : M), a < 0 → (0 < k ↔ k * a < 0 )
66- hmul_nonneg : ∀ (k : Int) (a : M), 0 ≤ a → 0 ≤ k → 0 ≤ k * a
67- hmul_nonpos : ∀ (k : Int) (a : M), a ≤ 0 → 0 ≤ k → k * a ≤ 0
53+ variable {M : Type u} [IntModule M]
54+
55+ theorem add_neg_cancel (a : M) : a + -a = 0 := by
56+ rw [add_comm, neg_add_cancel]
57+
58+ theorem add_left_inj {a b : M} (c : M) : a + c = b + c ↔ a = b :=
59+ ⟨fun h => by simpa [add_assoc, add_neg_cancel, add_zero] using (congrArg (· + -c) h),
60+ fun g => congrArg (· + c) g⟩
61+
62+ theorem add_right_inj (a b c : M) : a + b = a + c ↔ b = c := by
63+ rw [add_comm a b, add_comm a c, add_left_inj]
64+
65+ theorem neg_zero : (-0 : M) = 0 := by
66+ rw [← add_left_inj 0 , neg_add_cancel, add_zero]
67+
68+ theorem neg_neg (a : M) : -(-a) = a := by
69+ rw [← add_left_inj (-a), neg_add_cancel, add_neg_cancel]
70+
71+ theorem neg_eq_zero (a : M) : -a = 0 ↔ a = 0 :=
72+ ⟨fun h => by
73+ replace h := congrArg (-·) h
74+ simpa [neg_neg, neg_zero] using h,
75+ fun h => by rw [h, neg_zero]⟩
76+
77+ theorem neg_add (a b : M) : -(a + b) = -a + -b := by
78+ rw [← add_left_inj (a + b), neg_add_cancel, add_assoc (-a), add_comm a b, ← add_assoc (-b),
79+ neg_add_cancel, zero_add, neg_add_cancel]
80+
81+ theorem neg_sub (a b : M) : -(a - b) = b - a := by
82+ rw [sub_eq_add_neg, neg_add, neg_neg, sub_eq_add_neg, add_comm]
83+
84+ theorem sub_self (a : M) : a - a = 0 := by
85+ rw [sub_eq_add_neg, add_neg_cancel]
86+
87+ theorem sub_eq_iff {a b c : M} : a - b = c ↔ a = c + b := by
88+ rw [sub_eq_add_neg]
89+ constructor
90+ next => intro; subst c; rw [add_assoc, neg_add_cancel, add_zero]
91+ next => intro; subst a; rw [add_assoc, add_comm b, neg_add_cancel, add_zero]
92+
93+ theorem sub_eq_zero_iff {a b : M} : a - b = 0 ↔ a = b := by
94+ simp [sub_eq_iff, zero_add]
95+
96+ theorem neg_hmul (n : Int) (a : M) : (-n) * a = - (n * a) := by
97+ apply (add_left_inj (n * a)).mp
98+ rw [← add_hmul, Int.add_left_neg, zero_hmul, neg_add_cancel]
99+
100+ theorem hmul_neg (n : Int) (a : M) : n * (-a) = - (n * a) := by
101+ apply (add_left_inj (n * a)).mp
102+ rw [← hmul_add, neg_add_cancel, neg_add_cancel, hmul_zero]
103+
104+ end IntModule
68105
69106/--
70107Special case of Mathlib's `NoZeroSMulDivisors Nat α`.
0 commit comments