@@ -730,6 +730,42 @@ protected def function.surjective.non_unital_non_assoc_ring
730
730
{ .. hf.add_comm_group f zero add neg sub, .. hf.mul_zero_class f zero mul,
731
731
.. hf.distrib f add mul }
732
732
733
+ @[priority 100 ]
734
+ instance non_unital_non_assoc_ring.to_has_distrib_neg : has_distrib_neg α :=
735
+ { neg := has_neg.neg,
736
+ neg_neg := neg_neg,
737
+ neg_mul := λ a b, (neg_eq_of_add_eq_zero $ by rw [← right_distrib, add_right_neg, zero_mul]).symm,
738
+ mul_neg := λ a b, (neg_eq_of_add_eq_zero $ by rw [← left_distrib, add_right_neg, mul_zero]).symm }
739
+
740
+ lemma mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c :=
741
+ by simpa only [sub_eq_add_neg, neg_mul_eq_mul_neg] using mul_add a b (-c)
742
+
743
+ alias mul_sub_left_distrib ← mul_sub
744
+
745
+ lemma mul_sub_right_distrib (a b c : α) : (a - b) * c = a * c - b * c :=
746
+ by simpa only [sub_eq_add_neg, neg_mul_eq_neg_mul] using add_mul a (-b) c
747
+
748
+ alias mul_sub_right_distrib ← sub_mul
749
+
750
+ variables {a b c d e : α}
751
+
752
+ /-- An iff statement following from right distributivity in rings and the definition
753
+ of subtraction. -/
754
+ theorem mul_add_eq_mul_add_iff_sub_mul_add_eq : a * e + c = b * e + d ↔ (a - b) * e + c = d :=
755
+ calc
756
+ a * e + c = b * e + d ↔ a * e + c = d + b * e : by simp [add_comm]
757
+ ... ↔ a * e + c - b * e = d : iff.intro (λ h, begin rw h, simp end ) (λ h,
758
+ begin rw ← h, simp end )
759
+ ... ↔ (a - b) * e + c = d : begin simp [sub_mul, sub_add_eq_add_sub] end
760
+
761
+ /-- A simplification of one side of an equation exploiting right distributivity in rings
762
+ and the definition of subtraction. -/
763
+ theorem sub_mul_add_eq_of_mul_add_eq_mul_add : a * e + c = b * e + d → (a - b) * e + c = d :=
764
+ assume h,
765
+ calc
766
+ (a - b) * e + c = (a * e + c) - b * e : begin simp [sub_mul, sub_add_eq_add_sub] end
767
+ ... = d : begin rw h, simp [@add_sub_cancel α] end
768
+
733
769
end non_unital_non_assoc_ring
734
770
735
771
/-- An associative but not-necessarily unital ring. -/
@@ -859,40 +895,6 @@ protected def function.surjective.ring
859
895
ring β :=
860
896
{ .. hf.add_comm_group f zero add neg sub, .. hf.monoid f one mul, .. hf.distrib f add mul }
861
897
862
- @[priority 100 ]
863
- instance ring.to_has_distrib_neg : has_distrib_neg α :=
864
- { neg := has_neg.neg,
865
- neg_neg := neg_neg,
866
- neg_mul := λ a b, (neg_eq_of_add_eq_zero $ by rw [← right_distrib, add_right_neg, zero_mul]).symm,
867
- mul_neg := λ a b, (neg_eq_of_add_eq_zero $ by rw [← left_distrib, add_right_neg, mul_zero]).symm }
868
-
869
- lemma mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c :=
870
- by simpa only [sub_eq_add_neg, neg_mul_eq_mul_neg] using mul_add a b (-c)
871
-
872
- alias mul_sub_left_distrib ← mul_sub
873
-
874
- lemma mul_sub_right_distrib (a b c : α) : (a - b) * c = a * c - b * c :=
875
- by simpa only [sub_eq_add_neg, neg_mul_eq_neg_mul] using add_mul a (-b) c
876
-
877
- alias mul_sub_right_distrib ← sub_mul
878
-
879
- /-- An iff statement following from right distributivity in rings and the definition
880
- of subtraction. -/
881
- theorem mul_add_eq_mul_add_iff_sub_mul_add_eq : a * e + c = b * e + d ↔ (a - b) * e + c = d :=
882
- calc
883
- a * e + c = b * e + d ↔ a * e + c = d + b * e : by simp [add_comm]
884
- ... ↔ a * e + c - b * e = d : iff.intro (λ h, begin rw h, simp end ) (λ h,
885
- begin rw ← h, simp end )
886
- ... ↔ (a - b) * e + c = d : begin simp [sub_mul, sub_add_eq_add_sub] end
887
-
888
- /-- A simplification of one side of an equation exploiting right distributivity in rings
889
- and the definition of subtraction. -/
890
- theorem sub_mul_add_eq_of_mul_add_eq_mul_add : a * e + c = b * e + d → (a - b) * e + c = d :=
891
- assume h,
892
- calc
893
- (a - b) * e + c = (a * e + c) - b * e : begin simp [sub_mul, sub_add_eq_add_sub] end
894
- ... = d : begin rw h, simp [@add_sub_cancel α] end
895
-
896
898
end ring
897
899
898
900
namespace units
@@ -928,25 +930,26 @@ lemma is_unit.sub_iff [ring α] {x y : α} :
928
930
namespace ring_hom
929
931
930
932
/-- Ring homomorphisms preserve additive inverse. -/
931
- protected theorem map_neg {α β} [ring α] [ring β] (f : α →+* β) (x : α) : f (-x) = -(f x) :=
933
+ protected theorem map_neg {α β} [non_assoc_ring α] [non_assoc_ring β] (f : α →+* β) (x : α) :
934
+ f (-x) = -(f x) :=
932
935
map_neg f x
933
936
934
937
/-- Ring homomorphisms preserve subtraction. -/
935
- protected theorem map_sub {α β} [ring α] [ring β] (f : α →+* β) (x y : α) :
938
+ protected theorem map_sub {α β} [non_assoc_ring α] [non_assoc_ring β] (f : α →+* β) (x y : α) :
936
939
f (x - y) = (f x) - (f y) := map_sub f x y
937
940
938
941
/-- A ring homomorphism is injective iff its kernel is trivial. -/
939
- theorem injective_iff {α β} [ring α] [non_assoc_semiring β] (f : α →+* β) :
942
+ theorem injective_iff {α β} [non_assoc_ring α] [non_assoc_semiring β] (f : α →+* β) :
940
943
function.injective f ↔ (∀ a, f a = 0 → a = 0 ) :=
941
944
(f : α →+ β).injective_iff
942
945
943
946
/-- A ring homomorphism is injective iff its kernel is trivial. -/
944
- theorem injective_iff' {α β} [ring α] [non_assoc_semiring β] (f : α →+* β) :
947
+ theorem injective_iff' {α β} [non_assoc_ring α] [non_assoc_semiring β] (f : α →+* β) :
945
948
function.injective f ↔ (∀ a, f a = 0 ↔ a = 0 ) :=
946
949
(f : α →+ β).injective_iff'
947
950
948
951
/-- Makes a ring homomorphism from a monoid homomorphism of rings which preserves addition. -/
949
- def mk' {γ} [non_assoc_semiring α] [ring γ] (f : α →* γ)
952
+ def mk' {γ} [non_assoc_semiring α] [non_assoc_ring γ] (f : α →* γ)
950
953
(map_add : ∀ a b : α, f (a + b) = f a + f b) :
951
954
α →+* γ :=
952
955
{ to_fun := f,
0 commit comments