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

Commit 744e79c

Browse files
adomanieric-wieser
andcommitted
feat(algebra/ordered_*, */sub{monoid,group,ring,semiring,field,algebra}): pullback of ordered algebraic structures under an injective map (#6489)
Prove that the following 14 order typeclasses can be pulled back via an injective map (`function.injective.*`), and use them to attach 30 new instances to sub-objects: * `ordered_comm_monoid` (and the implied `ordered_add_comm_monoid`) * `submonoid.to_ordered_comm_monoid` * `submodule.to_ordered_add_comm_monoid` * `ordered_comm_group` (and the implied `ordered_add_comm_group`) * `subgroup.to_ordered_comm_group` * `submodule.to_ordered_add_comm_group` * `ordered_cancel_comm_monoid` (and the implied `ordered_cancel_add_comm_monoid`) * `submonoid.to_ordered_cancel_comm_monoid` * `submodule.to_ordered_cancel_add_comm_monoid` * `linear_ordered_cancel_comm_monoid` (and the implied `linear_ordered_cancel_add_comm_monoid`) * `submonoid.to_linear_ordered_cancel_comm_monoid` * `submodule.to_linear_ordered_cancel_add_comm_monoid` * `linear_ordered_comm_monoid_with_zero` * (no suitable subobject exists for monoid_with_zero) * `linear_ordered_comm_group` (and the implied `linear_ordered_add_comm_group`) * `subgroup.to_linear_ordered_comm_group` * `submodule.to_linear_ordered_add_comm_group` * `ordered_semiring` * `subsemiring.to_ordered_semiring` * `subalgebra.to_ordered_semiring` * `ordered_comm_semiring` * `subsemiring.to_ordered_comm_semiring` * `subalgebra.to_ordered_comm_semiring` * `ordered_ring` * `subring.to_ordered_ring` * `subalgebra.to_ordered_ring` * `ordered_comm_ring` * `subring.to_ordered_comm_ring` * `subalgebra.to_ordered_comm_ring` * `linear_ordered_semiring` * `subring.to_linear_ordered_semiring` * `subalgebra.to_linear_ordered_semiring` * `linear_ordered_ring` * `subring.to_linear_ordered_ring` * `subalgebra.to_linear_ordered_ring` * `linear_ordered_comm_ring` * `subring.to_linear_ordered_comm_ring` * `subalgebra.to_linear_ordered_comm_ring` * `linear_ordered_field` * `subfield.to_linear_ordered_field` Zulip: https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/rings.20from.20subtype Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent 09273ae commit 744e79c

File tree

12 files changed

+322
-0
lines changed

12 files changed

+322
-0
lines changed

src/algebra/algebra/subalgebra.lean

Lines changed: 24 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -166,6 +166,30 @@ instance to_comm_ring {R A}
166166
[comm_ring R] [comm_ring A] [algebra R A] (S : subalgebra R A) :
167167
comm_ring S := S.to_subring.to_comm_ring
168168

169+
instance to_ordered_semiring {R A}
170+
[comm_semiring R] [ordered_semiring A] [algebra R A] (S : subalgebra R A) :
171+
ordered_semiring S := S.to_subsemiring.to_ordered_semiring
172+
instance to_ordered_comm_semiring {R A}
173+
[comm_semiring R] [ordered_comm_semiring A] [algebra R A] (S : subalgebra R A) :
174+
ordered_comm_semiring S := subsemiring.to_ordered_comm_semiring S
175+
instance to_ordered_ring {R A}
176+
[comm_ring R] [ordered_ring A] [algebra R A] (S : subalgebra R A) :
177+
ordered_ring S := S.to_subring.to_ordered_ring
178+
instance to_ordered_comm_ring {R A}
179+
[comm_ring R] [ordered_comm_ring A] [algebra R A] (S : subalgebra R A) :
180+
ordered_comm_ring S := S.to_subring.to_ordered_comm_ring
181+
182+
instance to_linear_ordered_semiring {R A}
183+
[comm_semiring R] [linear_ordered_semiring A] [algebra R A] (S : subalgebra R A) :
184+
linear_ordered_semiring S := S.to_subsemiring.to_linear_ordered_semiring
185+
/-! There is no `linear_ordered_comm_semiring`. -/
186+
instance to_linear_ordered_ring {R A}
187+
[comm_ring R] [linear_ordered_ring A] [algebra R A] (S : subalgebra R A) :
188+
linear_ordered_ring S := S.to_subring.to_linear_ordered_ring
189+
instance to_linear_ordered_comm_ring {R A}
190+
[comm_ring R] [linear_ordered_comm_ring A] [algebra R A] (S : subalgebra R A) :
191+
linear_ordered_comm_ring S := S.to_subring.to_linear_ordered_comm_ring
192+
169193
end
170194

171195
instance algebra : algebra R S :=

src/algebra/linear_ordered_comm_group_with_zero.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -42,6 +42,18 @@ variables [linear_ordered_comm_monoid_with_zero α]
4242
The following facts are true more generally in a (linearly) ordered commutative monoid.
4343
-/
4444

45+
/-- Pullback a `linear_ordered_comm_monoid_with_zero` under an injective map. -/
46+
def function.injective.linear_ordered_comm_monoid_with_zero {β : Type*}
47+
[has_zero β] [has_one β] [has_mul β]
48+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
49+
(mul : ∀ x y, f (x * y) = f x * f y) :
50+
linear_ordered_comm_monoid_with_zero β :=
51+
{ zero_le_one := show f 0 ≤ f 1, by simp only [zero, one,
52+
linear_ordered_comm_monoid_with_zero.zero_le_one],
53+
..linear_order.lift f hf,
54+
..hf.ordered_comm_monoid f one mul,
55+
..hf.comm_monoid_with_zero f zero one mul }
56+
4557
lemma one_le_pow_of_one_le' {n : ℕ} (H : 1 ≤ x) : 1 ≤ x^n :=
4658
begin
4759
induction n with n ih,

src/algebra/module/submodule.lean

Lines changed: 44 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -199,6 +199,50 @@ instance : add_comm_group p :=
199199

200200
end add_comm_group
201201

202+
section ordered_monoid
203+
204+
variables [semiring R]
205+
206+
/-- A submodule of an `ordered_add_comm_monoid` is an `ordered_add_comm_monoid`. -/
207+
instance to_ordered_add_comm_monoid
208+
{M} [ordered_add_comm_monoid M] [semimodule R M] (S : submodule R M) :
209+
ordered_add_comm_monoid S :=
210+
subtype.coe_injective.ordered_add_comm_monoid coe rfl (λ _ _, rfl)
211+
212+
/-- A submodule of an `ordered_cancel_add_comm_monoid` is an `ordered_cancel_add_comm_monoid`. -/
213+
instance to_ordered_cancel_add_comm_monoid
214+
{M} [ordered_cancel_add_comm_monoid M] [semimodule R M] (S : submodule R M) :
215+
ordered_cancel_add_comm_monoid S :=
216+
subtype.coe_injective.ordered_cancel_add_comm_monoid coe rfl (λ _ _, rfl)
217+
218+
/-- A submodule of a `linear_ordered_cancel_add_comm_monoid` is a
219+
`linear_ordered_cancel_add_comm_monoid`. -/
220+
instance to_linear_ordered_cancel_add_comm_monoid
221+
{M} [linear_ordered_cancel_add_comm_monoid M] [semimodule R M] (S : submodule R M) :
222+
linear_ordered_cancel_add_comm_monoid S :=
223+
subtype.coe_injective.linear_ordered_cancel_add_comm_monoid coe rfl (λ _ _, rfl)
224+
225+
end ordered_monoid
226+
227+
section ordered_group
228+
229+
variables [ring R]
230+
231+
/-- A submodule of an `ordered_add_comm_group` is an `ordered_add_comm_group`. -/
232+
instance to_ordered_add_comm_group
233+
{M} [ordered_add_comm_group M] [semimodule R M] (S : submodule R M) :
234+
ordered_add_comm_group S :=
235+
subtype.coe_injective.ordered_add_comm_group coe rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
236+
237+
/-- A submodule of a `linear_ordered_add_comm_group` is a
238+
`linear_ordered_add_comm_group`. -/
239+
instance to_linear_ordered_add_comm_group
240+
{M} [linear_ordered_add_comm_group M] [semimodule R M] (S : submodule R M) :
241+
linear_ordered_add_comm_group S :=
242+
subtype.coe_injective.linear_ordered_add_comm_group coe rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
243+
244+
end ordered_group
245+
202246
end submodule
203247

204248
namespace submodule

src/algebra/ordered_field.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -553,6 +553,18 @@ end
553553
### Miscellaneous lemmas
554554
-/
555555

556+
/-- Pullback a `linear_ordered_field` under an injective map. -/
557+
def function.injective.linear_ordered_field {β : Type*}
558+
[has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_inv β] [has_div β]
559+
[nontrivial β]
560+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
561+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
562+
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y)
563+
(inv : ∀ x, f (x⁻¹) = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) :
564+
linear_ordered_field β :=
565+
{ ..hf.linear_ordered_ring f zero one add mul neg sub,
566+
..hf.field f zero one add mul neg sub inv div}
567+
556568
lemma mul_sub_mul_div_mul_neg_iff (hc : c ≠ 0) (hd : d ≠ 0) :
557569
(a * d - b * c) / (c * d) < 0 ↔ a / c < b / d :=
558570
by rw [mul_comm b c, ← div_sub_div _ _ hc hd, sub_lt_zero]

src/algebra/ordered_group.lean

Lines changed: 27 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -407,6 +407,20 @@ by simp [div_eq_mul_inv]
407407
@[simp, to_additive] lemma div_lt_self_iff (a : α) {b : α} : a / b < a ↔ 1 < b :=
408408
by simp [div_eq_mul_inv]
409409

410+
/-- Pullback an `ordered_comm_group` under an injective map. -/
411+
@[to_additive function.injective.ordered_add_comm_group
412+
"Pullback an `ordered_add_comm_group` under an injective map."]
413+
def function.injective.ordered_comm_group {β : Type*}
414+
[has_one β] [has_mul β] [has_inv β] [has_div β]
415+
(f : β → α) (hf : function.injective f) (one : f 1 = 1)
416+
(mul : ∀ x y, f (x * y) = f x * f y)
417+
(inv : ∀ x, f (x⁻¹) = (f x)⁻¹)
418+
(div : ∀ x y, f (x / y) = f x / f y) :
419+
ordered_comm_group β :=
420+
{ ..partial_order.lift f hf,
421+
..hf.ordered_comm_monoid f one mul,
422+
..hf.comm_group_div f one mul inv div }
423+
410424
end ordered_comm_group
411425

412426
section ordered_add_comm_group
@@ -570,6 +584,19 @@ instance linear_ordered_comm_group.to_linear_ordered_cancel_comm_monoid :
570584
mul_right_cancel := λ x y z, mul_right_cancel,
571585
..‹linear_ordered_comm_group α› }
572586

587+
/-- Pullback a `linear_ordered_comm_group` under an injective map. -/
588+
@[to_additive function.injective.linear_ordered_add_comm_group
589+
"Pullback a `linear_ordered_add_comm_group` under an injective map."]
590+
def function.injective.linear_ordered_comm_group {β : Type*}
591+
[has_one β] [has_mul β] [has_inv β] [has_div β]
592+
(f : β → α) (hf : function.injective f) (one : f 1 = 1)
593+
(mul : ∀ x y, f (x * y) = f x * f y)
594+
(inv : ∀ x, f (x⁻¹) = (f x)⁻¹)
595+
(div : ∀ x y, f (x / y) = f x / f y) :
596+
linear_ordered_comm_group β :=
597+
{ ..linear_order.lift f hf,
598+
..hf.ordered_comm_group f one mul inv div }
599+
573600
@[to_additive linear_ordered_add_comm_group.add_lt_add_left]
574601
lemma linear_ordered_comm_group.mul_lt_mul_left'
575602
(a b : α) (h : a < b) (c : α) : c * a < c * b :=

src/algebra/ordered_monoid.lean

Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -264,6 +264,21 @@ iff.intro
264264
and.intro ‹a = 1› ‹b = 1›)
265265
(assume ⟨ha', hb'⟩, by rw [ha', hb', mul_one])
266266

267+
/-- Pullback an `ordered_comm_monoid` under an injective map. -/
268+
@[to_additive function.injective.ordered_add_comm_monoid
269+
"Pullback an `ordered_add_comm_monoid` under an injective map."]
270+
def function.injective.ordered_comm_monoid {β : Type*}
271+
[has_one β] [has_mul β]
272+
(f : β → α) (hf : function.injective f) (one : f 1 = 1)
273+
(mul : ∀ x y, f (x * y) = f x * f y) :
274+
ordered_comm_monoid β :=
275+
{ mul_le_mul_left := λ a b ab c,
276+
show f (c * a) ≤ f (c * b), by simp [mul, mul_le_mul_left' ab],
277+
lt_of_mul_lt_mul_left :=
278+
λ a b c bc, @lt_of_mul_lt_mul_left' _ _ (f a) _ _ (by rwa [← mul, ← mul]),
279+
..partial_order.lift f hf,
280+
..hf.comm_monoid f one mul }
281+
267282
section mono
268283

269284
variables {β : Type*} [preorder β] {f g : β → α}
@@ -964,6 +979,20 @@ by split; apply le_antisymm; try {assumption};
964979
rw ← hab; simp [ha, hb],
965980
λ ⟨ha', hb'⟩, by rw [ha', hb', mul_one]⟩
966981

982+
/-- Pullback an `ordered_cancel_comm_monoid` under an injective map. -/
983+
@[to_additive function.injective.ordered_cancel_add_comm_monoid
984+
"Pullback an `ordered_cancel_add_comm_monoid` under an injective map."]
985+
def function.injective.ordered_cancel_comm_monoid {β : Type*}
986+
[has_one β] [has_mul β]
987+
(f : β → α) (hf : function.injective f) (one : f 1 = 1)
988+
(mul : ∀ x y, f (x * y) = f x * f y) :
989+
ordered_cancel_comm_monoid β :=
990+
{ le_of_mul_le_mul_left := λ a b c (ab : f (a * b) ≤ f (a * c)),
991+
(by { rw [mul, mul] at ab, exact le_of_mul_le_mul_left' ab }),
992+
..hf.left_cancel_semigroup f mul,
993+
..hf.right_cancel_semigroup f mul,
994+
..hf.ordered_comm_monoid f one mul }
995+
967996
section mono
968997

969998
variables {β : Type*} [preorder β] {f g : β → α}
@@ -1089,6 +1118,17 @@ min_le_iff.2 $ or.inr $ le_mul_of_one_le_left' ha
10891118
lemma max_le_mul_of_one_le {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : max a b ≤ a * b :=
10901119
max_le_iff.2 ⟨le_mul_of_one_le_right' hb, le_mul_of_one_le_left' ha⟩
10911120

1121+
/-- Pullback a `linear_ordered_cancel_comm_monoid` under an injective map. -/
1122+
@[to_additive function.injective.linear_ordered_cancel_add_comm_monoid
1123+
"Pullback a `linear_ordered_cancel_add_comm_monoid` under an injective map."]
1124+
def function.injective.linear_ordered_cancel_comm_monoid {β : Type*}
1125+
[has_one β] [has_mul β]
1126+
(f : β → α) (hf : function.injective f) (one : f 1 = 1)
1127+
(mul : ∀ x y, f (x * y) = f x * f y) :
1128+
linear_ordered_cancel_comm_monoid β :=
1129+
{ ..linear_order.lift f hf,
1130+
..hf.ordered_cancel_comm_monoid f one mul }
1131+
10921132
end linear_ordered_cancel_comm_monoid
10931133

10941134
namespace order_dual

src/algebra/ordered_ring.lean

Lines changed: 85 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -148,6 +148,28 @@ lemma one_le_mul_of_one_le_of_one_le {a b : α} (a1 : 1 ≤ a) (b1 : 1 ≤ b) :
148148
(1 : α) ≤ a * b :=
149149
(mul_one (1 : α)).symm.le.trans (mul_le_mul a1 b1 zero_le_one (zero_le_one.trans a1))
150150

151+
/-- Pullback an `ordered_semiring` under an injective map. -/
152+
def function.injective.ordered_semiring {β : Type*}
153+
[has_zero β] [has_one β] [has_add β] [has_mul β]
154+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
155+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) :
156+
ordered_semiring β :=
157+
{ zero_le_one := show f 0 ≤ f 1, by simp only [zero, one, zero_le_one],
158+
mul_lt_mul_of_pos_left := λ a b c ab c0, show f (c * a) < f (c * b),
159+
begin
160+
rw [mul, mul],
161+
refine mul_lt_mul_of_pos_left ab _,
162+
rwa ← zero,
163+
end,
164+
mul_lt_mul_of_pos_right := λ a b c ab c0, show f (a * c) < f (b * c),
165+
begin
166+
rw [mul, mul],
167+
refine mul_lt_mul_of_pos_right ab _,
168+
rwa ← zero,
169+
end,
170+
..hf.ordered_cancel_add_comm_monoid f zero add,
171+
..hf.semiring f zero one add mul }
172+
151173
section
152174
variable [nontrivial α]
153175

@@ -216,6 +238,15 @@ multiplication with a positive number and addition are monotone. -/
216238
@[protect_proj]
217239
class ordered_comm_semiring (α : Type u) extends ordered_semiring α, comm_semiring α
218240

241+
/-- Pullback an `ordered_comm_semiring` under an injective map. -/
242+
def function.injective.ordered_comm_semiring [ordered_comm_semiring α] {β : Type*}
243+
[has_zero β] [has_one β] [has_add β] [has_mul β]
244+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
245+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) :
246+
ordered_comm_semiring β :=
247+
{ ..hf.comm_semiring f zero one add mul,
248+
..hf.ordered_semiring f zero one add mul }
249+
219250
end ordered_comm_semiring
220251

221252
/--
@@ -444,6 +475,16 @@ instance linear_ordered_semiring.to_no_top_order {α : Type*} [linear_ordered_se
444475
no_top_order α :=
445476
⟨assume a, ⟨a + 1, lt_add_of_pos_right _ zero_lt_one⟩⟩
446477

478+
/-- Pullback a `linear_ordered_semiring` under an injective map. -/
479+
def function.injective.linear_ordered_semiring {β : Type*}
480+
[has_zero β] [has_one β] [has_add β] [has_mul β] [nontrivial β]
481+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
482+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) :
483+
linear_ordered_semiring β :=
484+
{ ..linear_order.lift f hf,
485+
..‹nontrivial β›,
486+
..hf.ordered_semiring f zero one add mul }
487+
447488
end linear_ordered_semiring
448489

449490
section mono
@@ -605,6 +646,17 @@ lemma mul_pos_of_neg_of_neg {a b : α} (ha : a < 0) (hb : b < 0) : 0 < a * b :=
605646
have 0 * b < a * b, from mul_lt_mul_of_neg_right ha hb,
606647
by rwa zero_mul at this
607648

649+
/-- Pullback an `ordered_ring` under an injective map. -/
650+
def function.injective.ordered_ring {β : Type*}
651+
[has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β]
652+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
653+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
654+
(neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) :
655+
ordered_ring β :=
656+
{ mul_pos := λ a b a0 b0, show f 0 < f (a * b), by { rw [zero, mul], apply mul_pos; rwa ← zero },
657+
..hf.ordered_semiring f zero one add mul,
658+
..hf.ring_sub f zero one add mul neg sub }
659+
608660
end ordered_ring
609661

610662
section ordered_comm_ring
@@ -614,6 +666,17 @@ multiplication with a positive number and addition are monotone. -/
614666
@[protect_proj]
615667
class ordered_comm_ring (α : Type u) extends ordered_ring α, ordered_comm_semiring α, comm_ring α
616668

669+
/-- Pullback an `ordered_comm_ring` under an injective map. -/
670+
def function.injective.ordered_comm_ring [ordered_comm_ring α] {β : Type*}
671+
[has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β]
672+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
673+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
674+
(neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) :
675+
ordered_comm_ring β :=
676+
{ ..hf.ordered_comm_semiring f zero one add mul,
677+
..hf.ordered_ring f zero one add mul neg sub,
678+
..hf.comm_ring_sub f zero one add mul neg sub }
679+
617680
end ordered_comm_ring
618681

619682
/-- A `linear_ordered_ring α` is a ring `α` with a linear order such that
@@ -807,6 +870,17 @@ end
807870
lemma abs_le_one_iff_mul_self_le_one : abs a ≤ 1 ↔ a * a ≤ 1 :=
808871
by simpa only [abs_one, one_mul] using @abs_le_iff_mul_self_le α _ a 1
809872

873+
/-- Pullback a `linear_ordered_ring` under an injective map. -/
874+
def function.injective.linear_ordered_ring {β : Type*}
875+
[has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [nontrivial β]
876+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
877+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
878+
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) :
879+
linear_ordered_ring β :=
880+
{ ..linear_order.lift f hf,
881+
..‹nontrivial β›,
882+
..hf.ordered_ring f zero one add mul neg sub }
883+
810884
end linear_ordered_ring
811885

812886
/-- A `linear_ordered_comm_ring α` is a commutative ring `α` with a linear order
@@ -872,6 +946,17 @@ begin
872946
simp [left_distrib, right_distrib, add_assoc, add_comm, add_left_comm, mul_comm, sub_eq_add_neg],
873947
end
874948

949+
/-- Pullback a `linear_ordered_comm_ring` under an injective map. -/
950+
def function.injective.linear_ordered_comm_ring {β : Type*}
951+
[has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [nontrivial β]
952+
(f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1)
953+
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
954+
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) :
955+
linear_ordered_comm_ring β :=
956+
{ ..linear_order.lift f hf,
957+
..‹nontrivial β›,
958+
..hf.ordered_comm_ring f zero one add mul neg sub }
959+
875960
end linear_ordered_comm_ring
876961

877962
/-- Extend `nonneg_add_comm_group` to support ordered rings

src/field_theory/subfield.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -199,6 +199,12 @@ instance to_field : field s :=
199199
subtype.coe_injective.field coe
200200
rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
201201

202+
/-- A subfield of a `linear_ordered_field` is a `linear_ordered_field`. -/
203+
instance to_linear_ordered_field {K} [linear_ordered_field K] (s : subfield K) :
204+
linear_ordered_field s :=
205+
subtype.coe_injective.linear_ordered_field coe
206+
rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
207+
202208
@[simp, norm_cast] lemma coe_add (x y : s) : (↑(x + y) : K) = ↑x + ↑y := rfl
203209
@[simp, norm_cast] lemma coe_sub (x y : s) : (↑(x - y) : K) = ↑x - ↑y := rfl
204210
@[simp, norm_cast] lemma coe_neg (x : s) : (↑(-x) : K) = -↑x := rfl

src/group_theory/subgroup.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -306,6 +306,19 @@ subtype.coe_injective.group_div _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
306306
instance to_comm_group {G : Type*} [comm_group G] (H : subgroup G) : comm_group H :=
307307
subtype.coe_injective.comm_group_div _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
308308

309+
/-- A subgroup of an `ordered_comm_group` is an `ordered_comm_group`. -/
310+
@[to_additive "An `add_subgroup` of an `add_ordered_comm_group` is an `add_ordered_comm_group`."]
311+
instance to_ordered_comm_group {G : Type*} [ordered_comm_group G] (H : subgroup G) :
312+
ordered_comm_group H :=
313+
subtype.coe_injective.ordered_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
314+
315+
/-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/
316+
@[to_additive "An `add_subgroup` of a `linear_ordered_add_comm_group` is a
317+
`linear_ordered_add_comm_group`."]
318+
instance to_linear_ordered_comm_group {G : Type*} [linear_ordered_comm_group G]
319+
(H : subgroup G) : linear_ordered_comm_group H :=
320+
subtype.coe_injective.linear_ordered_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl)
321+
309322
/-- The natural group hom from a subgroup of group `G` to `G`. -/
310323
@[to_additive "The natural group hom from an `add_subgroup` of `add_group` `G` to `G`."]
311324
def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩

0 commit comments

Comments
 (0)