@@ -91,14 +91,14 @@ instance ordered_comm_group.has_exists_mul_of_le (α : Type u)
91
91
has_exists_mul_of_le α :=
92
92
⟨λ a b hab, ⟨b * a⁻¹, (mul_inv_cancel_comm_assoc a b).symm⟩⟩
93
93
94
- @[to_additive] instance [h : has_inv α] : has_inv (order_dual α) := h
95
- @[to_additive] instance [h : has_div α] : has_div (order_dual α) := h
96
- @[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv (order_dual α) := h
97
- @[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid (order_dual α) := h
98
- @[to_additive] instance [h : group α] : group (order_dual α) := h
99
- @[to_additive] instance [h : comm_group α] : comm_group (order_dual α) := h
100
-
101
- @[to_additive] instance [ordered_comm_group α] : ordered_comm_group (order_dual α) :=
94
+ @[to_additive] instance [h : has_inv α] : has_inv αᵒᵈ := h
95
+ @[to_additive] instance [h : has_div α] : has_div αᵒᵈ := h
96
+ @[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv αᵒᵈ := h
97
+ @[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid αᵒᵈ := h
98
+ @[to_additive] instance [h : group α] : group αᵒᵈ := h
99
+ @[to_additive] instance [h : comm_group α] : comm_group αᵒᵈ := h
100
+
101
+ @[to_additive] instance [ordered_comm_group α] : ordered_comm_group αᵒᵈ :=
102
102
{ .. order_dual.ordered_comm_monoid, .. order_dual.group }
103
103
104
104
section group
@@ -292,7 +292,7 @@ variable (α)
292
292
293
293
/-- `x ↦ x⁻¹` as an order-reversing equivalence. -/
294
294
@[to_additive " `x ↦ -x` as an order-reversing equivalence." , simps]
295
- def order_iso.inv : α ≃o order_dual α :=
295
+ def order_iso.inv : α ≃o αᵒᵈ :=
296
296
{ to_equiv := (equiv.inv α).trans order_dual.to_dual,
297
297
map_rel_iff' := λ a b, @inv_le_inv_iff α _ _ _ _ _ _ }
298
298
@@ -853,7 +853,7 @@ calc a ≤ b * (b⁻¹ * c) : h _ (lt_inv_mul_iff_lt.mpr hc)
853
853
854
854
@[to_additive]
855
855
lemma le_of_forall_lt_one_mul_le (h : ∀ ε < 1 , a * ε ≤ b) : a ≤ b :=
856
- @le_of_forall_one_lt_le_mul (order_dual α) _ _ _ _ _ _ h
856
+ @le_of_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ h
857
857
858
858
@[to_additive]
859
859
lemma le_of_forall_one_lt_div_le (h : ∀ ε : α, 1 < ε → a / ε ≤ b) : a ≤ b :=
@@ -866,7 +866,7 @@ lemma le_iff_forall_one_lt_le_mul : a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε
866
866
867
867
@[to_additive]
868
868
lemma le_iff_forall_lt_one_mul_le : a ≤ b ↔ ∀ ε < 1 , a * ε ≤ b :=
869
- @le_iff_forall_one_lt_le_mul (order_dual α) _ _ _ _ _ _
869
+ @le_iff_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _
870
870
871
871
end densely_ordered
872
872
@@ -897,7 +897,7 @@ multiplication is monotone. -/
897
897
class linear_ordered_comm_group (α : Type u) extends ordered_comm_group α, linear_order α
898
898
899
899
@[to_additive] instance [linear_ordered_comm_group α] :
900
- linear_ordered_comm_group (order_dual α) :=
900
+ linear_ordered_comm_group αᵒᵈ :=
901
901
{ .. order_dual.ordered_comm_group, .. order_dual.linear_order α }
902
902
903
903
section linear_ordered_comm_group
@@ -933,11 +933,11 @@ mul_lt_mul_left' h c
933
933
934
934
@[to_additive min_neg_neg]
935
935
lemma min_inv_inv' (a b : α) : min (a⁻¹) (b⁻¹) = (max a b)⁻¹ :=
936
- eq.symm $ @monotone.map_max α (order_dual α) _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr
936
+ eq.symm $ @monotone.map_max α αᵒᵈ _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr
937
937
938
938
@[to_additive max_neg_neg]
939
939
lemma max_inv_inv' (a b : α) : max (a⁻¹) (b⁻¹) = (min a b)⁻¹ :=
940
- eq.symm $ @monotone.map_min α (order_dual α) _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr
940
+ eq.symm $ @monotone.map_min α αᵒᵈ _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr
941
941
942
942
@[to_additive min_sub_sub_right]
943
943
lemma min_div_div_right' (a b c : α) : min (a / c) (b / c) = min a b / c :=
0 commit comments