1
1
/-
2
2
Copyright (c) 2016 Jeremy Avigad. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
- Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro
4
+ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Yaël Dillies
5
5
-/
6
6
import order.min_max
7
7
import algebra.order.monoid.cancel.defs
@@ -29,21 +29,23 @@ For short,
29
29
## Typeclasses
30
30
31
31
* `ordered_semiring`: Semiring with a partial order such that `+` and `*` respect `≤`.
32
- * `strict_ordered_semiring`: Semiring with a partial order such that `+` and `*` respects `<`.
32
+ * `strict_ordered_semiring`: Nontrivial semiring with a partial order such that `+` and `*` respects
33
+ `<`.
33
34
* `ordered_comm_semiring`: Commutative semiring with a partial order such that `+` and `*` respect
34
35
`≤`.
35
- * `strict_ordered_comm_semiring`: Commutative semiring with a partial order such that `+` and `* `
36
- respect `<`.
36
+ * `strict_ordered_comm_semiring`: Nontrivial commutative semiring with a partial order such that `+`
37
+ and `*` respect `<`.
37
38
* `ordered_ring`: Ring with a partial order such that `+` respects `≤` and `*` respects `<`.
38
39
* `ordered_comm_ring`: Commutative ring with a partial order such that `+` respects `≤` and
39
40
`*` respects `<`.
40
- * `linear_ordered_semiring`: Semiring with a linear order such that `+` respects `≤` and
41
+ * `linear_ordered_semiring`: Nontrivial semiring with a linear order such that `+` respects `≤` and
41
42
`*` respects `<`.
42
- * `linear_ordered_comm_semiring`: Commutative semiring with a linear order such that `+` respects
43
+ * `linear_ordered_comm_semiring`: Nontrivial commutative semiring with a linear order such that `+`
44
+ respects `≤` and `*` respects `<`.
45
+ * `linear_ordered_ring`: Nontrivial ring with a linear order such that `+` respects `≤` and `*`
46
+ respects `<`.
47
+ * `linear_ordered_comm_ring`: Nontrivial commutative ring with a linear order such that `+` respects
43
48
`≤` and `*` respects `<`.
44
- * `linear_ordered_ring`: Ring with a linear order such that `+` respects `≤` and `*` respects `<`.
45
- * `linear_ordered_comm_ring`: Commutative ring with a linear order such that `+` respects `≤` and
46
- `*` respects `<`.
47
49
* `canonically_ordered_comm_semiring`: Commutative semiring with a partial order such that `+`
48
50
respects `≤`, `*` respects `<`, and `a ≤ b ↔ ∃ c, b = a + c`.
49
51
@@ -57,35 +59,42 @@ immediate predecessors and what conditions are added to each of them.
57
59
- `ordered_add_comm_monoid` & multiplication & `*` respects `≤`
58
60
- `semiring` & partial order structure & `+` respects `≤` & `*` respects `≤`
59
61
* `strict_ordered_semiring`
60
- - `ordered_cancel_add_comm_monoid` & multiplication & `*` respects `<`
61
- - `ordered_semiring` & `+` respects `<` & `*` respects `<`
62
+ - `ordered_cancel_add_comm_monoid` & multiplication & `*` respects `<` & nontriviality
63
+ - `ordered_semiring` & `+` respects `<` & `*` respects `<` & nontriviality
62
64
* `ordered_comm_semiring`
63
65
- `ordered_semiring` & commutativity of multiplication
64
66
- `comm_semiring` & partial order structure & `+` respects `≤` & `*` respects `<`
65
67
* `strict_ordered_comm_semiring`
66
68
- `strict_ordered_semiring` & commutativity of multiplication
67
- - `ordered_comm_semiring` & `+` respects `<` & `*` respects `<`
69
+ - `ordered_comm_semiring` & `+` respects `<` & `*` respects `<` & nontriviality
68
70
* `ordered_ring`
69
- - `strict_ordered_semiring ` & additive inverses
71
+ - `ordered_semiring ` & additive inverses
70
72
- `ordered_add_comm_group` & multiplication & `*` respects `<`
71
73
- `ring` & partial order structure & `+` respects `≤` & `*` respects `<`
74
+ * `strict_ordered_ring`
75
+ - `strict_ordered_semiring` & additive inverses
76
+ - `ordered_semiring` & `+` respects `<` & `*` respects `<` & nontriviality
72
77
* `ordered_comm_ring`
73
78
- `ordered_ring` & commutativity of multiplication
74
79
- `ordered_comm_semiring` & additive inverses
75
80
- `comm_ring` & partial order structure & `+` respects `≤` & `*` respects `<`
81
+ * `strict_ordered_comm_ring`
82
+ - `strict_ordered_comm_semiring` & additive inverses
83
+ - `strict_ordered_ring` & commutativity of multiplication
84
+ - `ordered_comm_ring` & `+` respects `<` & `*` respects `<` & nontriviality
76
85
* `linear_ordered_semiring`
77
- - `strict_ordered_semiring` & totality of the order & nontriviality
86
+ - `strict_ordered_semiring` & totality of the order
78
87
- `linear_ordered_add_comm_monoid` & multiplication & nontriviality & `*` respects `<`
79
88
* `linear_ordered_comm_semiring`
80
- - `strict_ordered_comm_semiring` & totality of the order & nontriviality
89
+ - `strict_ordered_comm_semiring` & totality of the order
81
90
- `linear_ordered_semiring` & commutativity of multiplication
82
91
* `linear_ordered_ring`
83
- - `ordered_ring ` & totality of the order & nontriviality
92
+ - `strict_ordered_ring ` & totality of the order
84
93
- `linear_ordered_semiring` & additive inverses
85
94
- `linear_ordered_add_comm_group` & multiplication & `*` respects `<`
86
95
- `domain` & linear order structure
87
96
* `linear_ordered_comm_ring`
88
- - `ordered_comm_ring ` & totality of the order & nontriviality
97
+ - `strict_ordered_comm_ring ` & totality of the order
89
98
- `linear_ordered_ring` & commutativity of multiplication
90
99
- `linear_ordered_comm_semiring` & additive inverses
91
100
- `is_domain` & linear order structure
@@ -133,10 +142,11 @@ and multiplication by a nonnegative number is monotone. -/
133
142
@[protect_proj, ancestor ordered_ring comm_ring]
134
143
class ordered_comm_ring (α : Type u) extends ordered_ring α, comm_ring α
135
144
136
- /-- A `strict_ordered_semiring` is a semiring with a partial order such that addition is strictly
137
- monotone and multiplication by a positive number is strictly monotone. -/
138
- @[protect_proj, ancestor semiring ordered_cancel_add_comm_monoid]
139
- class strict_ordered_semiring (α : Type u) extends semiring α, ordered_cancel_add_comm_monoid α :=
145
+ /-- A `strict_ordered_semiring` is a nontrivial semiring with a partial order such that addition is
146
+ strictly monotone and multiplication by a positive number is strictly monotone. -/
147
+ @[protect_proj, ancestor semiring ordered_cancel_add_comm_monoid nontrivial]
148
+ class strict_ordered_semiring (α : Type u)
149
+ extends semiring α, ordered_cancel_add_comm_monoid α, nontrivial α :=
140
150
(zero_le_one : (0 : α) ≤ 1 )
141
151
(mul_lt_mul_of_pos_left : ∀ a b c : α, a < b → 0 < c → c * a < c * b)
142
152
(mul_lt_mul_of_pos_right : ∀ a b c : α, a < b → 0 < c → a * c < b * c)
@@ -148,8 +158,8 @@ class strict_ordered_comm_semiring (α : Type u) extends strict_ordered_semiring
148
158
149
159
/-- A `strict_ordered_ring` is a ring with a partial order such that addition is strictly monotone
150
160
and multiplication by a positive number is strictly monotone. -/
151
- @[protect_proj, ancestor ring ordered_add_comm_group]
152
- class strict_ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α :=
161
+ @[protect_proj, ancestor ring ordered_add_comm_group nontrivial ]
162
+ class strict_ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α, nontrivial α :=
153
163
(zero_le_one : 0 ≤ (1 : α))
154
164
(mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b)
155
165
@@ -165,7 +175,7 @@ explore changing this, but be warned that the instances involving `domain` may c
165
175
search loops. -/
166
176
@[protect_proj, ancestor strict_ordered_semiring linear_ordered_add_comm_monoid nontrivial]
167
177
class linear_ordered_semiring (α : Type u)
168
- extends strict_ordered_semiring α, linear_ordered_add_comm_monoid α, nontrivial α
178
+ extends strict_ordered_semiring α, linear_ordered_add_comm_monoid α
169
179
170
180
/-- A `linear_ordered_comm_semiring` is a nontrivial commutative semiring with a linear order such
171
181
that addition is monotone and multiplication by a positive number is strictly monotone. -/
@@ -175,8 +185,8 @@ class linear_ordered_comm_semiring (α : Type*)
175
185
176
186
/-- A `linear_ordered_ring` is a ring with a linear order such that addition is monotone and
177
187
multiplication by a positive number is strictly monotone. -/
178
- @[protect_proj, ancestor strict_ordered_ring linear_order nontrivial ]
179
- class linear_ordered_ring (α : Type u) extends strict_ordered_ring α, linear_order α, nontrivial α
188
+ @[protect_proj, ancestor strict_ordered_ring linear_order]
189
+ class linear_ordered_ring (α : Type u) extends strict_ordered_ring α, linear_order α
180
190
181
191
/-- A `linear_ordered_comm_ring` is a commutative ring with a linear order such that addition is
182
192
monotone and multiplication by a positive number is strictly monotone. -/
@@ -518,29 +528,17 @@ lemma strict_mono.mul (hf : strict_mono f) (hg : strict_mono g) (hf₀ : ∀ x,
518
528
519
529
end monotone
520
530
521
- section nontrivial
522
- variables [nontrivial α]
523
-
524
531
lemma lt_one_add (a : α) : a < 1 + a := lt_add_of_pos_left _ zero_lt_one
525
532
lemma lt_add_one (a : α) : a < a + 1 := lt_add_of_pos_right _ zero_lt_one
526
533
527
534
lemma one_lt_two : (1 : α) < 2 := lt_add_one _
528
535
529
536
lemma lt_two_mul_self (ha : 0 < a) : a < 2 * a := lt_mul_of_one_lt_left ha one_lt_two
530
537
531
- lemma nat.strict_mono_cast : strict_mono (coe : ℕ → α) :=
532
- strict_mono_nat_of_lt_succ $ λ n, by { rw [nat.cast_succ], apply lt_add_one }
533
-
534
- /-- `coe : ℕ → α` as an `order_embedding` -/
535
- @[simps { fully_applied := ff }] def nat.cast_order_embedding : ℕ ↪o α :=
536
- order_embedding.of_strict_mono coe nat.strict_mono_cast
537
-
538
538
@[priority 100 ] -- see Note [lower instance priority]
539
539
instance strict_ordered_semiring.to_no_max_order : no_max_order α :=
540
540
⟨λ a, ⟨a + 1 , lt_add_of_pos_right _ one_pos⟩⟩
541
541
542
- end nontrivial
543
-
544
542
end strict_ordered_semiring
545
543
546
544
section strict_ordered_comm_semiring
0 commit comments