refactor(library,hott): remove coercions between algebraic structures

They are classes, and mixing coercion with type class resolution is a
recipe for disaster (aka counterintuitive behavior).
This commit is contained in:
Leonardo de Moura 2015-11-11 11:32:05 -08:00
parent fee0cff295
commit 9bedbbb739
22 changed files with 102 additions and 90 deletions

View file

@ -28,6 +28,12 @@ namespace category
structure category [class] (ob : Type) extends parent : precategory ob :=
mk' :: (iso_of_path_equiv : is_univalent parent)
-- Remark: category and precategory are classes. So, the structure command
-- does not create a coercion between them automatically.
-- This coercion is needed for definitions such as category_eq_of_equiv
-- without it, we would have to explicitly use category.to_precategory
attribute category.to_precategory [coercion]
attribute category [multiple_instances]
abbreviation iso_of_path_equiv := @category.iso_of_path_equiv
@ -81,7 +87,7 @@ namespace category
definition Category.to_Precategory [constructor] [coercion] [reducible] (C : Category)
: Precategory :=
Precategory.mk (Category.carrier C) C
Precategory.mk (Category.carrier C) _
definition category.Mk [constructor] [reducible] := Category.mk
definition category.MK [constructor] [reducible] (C : Precategory)

View file

@ -65,7 +65,7 @@ namespace category
attribute Groupoid.struct [instance] [coercion]
definition Groupoid.to_Precategory [coercion] [reducible] (C : Groupoid) : Precategory :=
Precategory.mk (Groupoid.carrier C) C
Precategory.mk (Groupoid.carrier C) _
definition groupoid.Mk [reducible] := Groupoid.mk
definition groupoid.MK [reducible] (C : Precategory) (H : Π (a b : C) (f : a ⟶ b), is_iso f)

View file

@ -26,7 +26,7 @@ namespace category
definition Strict_precategory.to_Precategory [coercion] [reducible]
(C : Strict_precategory) : Precategory :=
Precategory.mk (Strict_precategory.carrier C) C
Precategory.mk (Strict_precategory.carrier C) _
open functor

View file

@ -328,7 +328,7 @@ section discrete_field
(assume H1 : x ≠ 0,
sum.inr (by rewrite [-one_mul, -(inv_mul_cancel H1), mul.assoc, H, mul_zero]))
definition discrete_field.to_integral_domain [instance] [reducible] [coercion] :
definition discrete_field.to_integral_domain [instance] [reducible] :
integral_domain A :=
⦃ integral_domain, s,
eq_zero_or_eq_zero_of_mul_eq_zero := discrete_field.eq_zero_or_eq_zero_of_mul_eq_zero⦄

View file

@ -53,10 +53,10 @@ theorem mul.comm [s : comm_semigroup A] (a b : A) : a * b = b * a :=
!comm_semigroup.mul_comm
theorem mul.left_comm [s : comm_semigroup A] (a b c : A) : a * (b * c) = b * (a * c) :=
binary.left_comm (@mul.comm A s) (@mul.assoc A s) a b c
binary.left_comm (@mul.comm A _) (@mul.assoc A _) a b c
theorem mul.right_comm [s : comm_semigroup A] (a b c : A) : (a * b) * c = (a * c) * b :=
binary.right_comm (@mul.comm A s) (@mul.assoc A s) a b c
binary.right_comm (@mul.comm A _) (@mul.assoc A _) a b c
structure left_cancel_semigroup [class] (A : Type) extends semigroup A :=
(mul_left_cancel : ∀a b c, mul a b = mul a c → b = c)
@ -95,10 +95,10 @@ theorem add.comm [s : add_comm_semigroup A] (a b : A) : a + b = b + a :=
theorem add.left_comm [s : add_comm_semigroup A] (a b c : A) :
a + (b + c) = b + (a + c) :=
binary.left_comm (@add.comm A s) (@add.assoc A s) a b c
binary.left_comm (@add.comm A _) (@add.assoc A _) a b c
theorem add.right_comm [s : add_comm_semigroup A] (a b c : A) : (a + b) + c = (a + c) + b :=
binary.right_comm (@add.comm A s) (@add.assoc A s) a b c
binary.right_comm (@add.comm A _) (@add.assoc A _) a b c
structure add_left_cancel_semigroup [class] (A : Type) extends add_semigroup A :=
(add_left_cancel : ∀a b c, add a b = add a c → b = c)
@ -248,11 +248,11 @@ section group
theorem mul_right_cancel {a b c : A} (H : a * b = c * b) : a = c :=
by rewrite [-mul_inv_cancel_right a b, H, mul_inv_cancel_right]
definition group.to_left_cancel_semigroup [instance] [coercion] [reducible] : left_cancel_semigroup A :=
definition group.to_left_cancel_semigroup [instance] [reducible] : left_cancel_semigroup A :=
⦃ left_cancel_semigroup, s,
mul_left_cancel := @mul_left_cancel A s ⦄
definition group.to_right_cancel_semigroup [instance] [coercion] [reducible] : right_cancel_semigroup A :=
definition group.to_right_cancel_semigroup [instance] [reducible] : right_cancel_semigroup A :=
⦃ right_cancel_semigroup, s,
mul_right_cancel := @mul_right_cancel A s ⦄
@ -363,12 +363,12 @@ section add_group
... = (c + b) + -b : H
... = c : add_neg_cancel_right
definition add_group.to_left_cancel_semigroup [instance] [coercion] [reducible] :
definition add_group.to_left_cancel_semigroup [instance] [reducible] :
add_left_cancel_semigroup A :=
⦃ add_left_cancel_semigroup, s,
add_left_cancel := @add_left_cancel A s ⦄
definition add_group.to_add_right_cancel_semigroup [instance] [coercion] [reducible] :
definition add_group.to_add_right_cancel_semigroup [instance][reducible] :
add_right_cancel_semigroup A :=
⦃ add_right_cancel_semigroup, s,
add_right_cancel := @add_right_cancel A s ⦄

View file

@ -13,6 +13,14 @@ open equiv eq equiv.ops is_trunc
namespace algebra
open Group has_mul has_inv
-- we prove under which conditions two groups are equal
-- group and has_mul are classes. So, lean does not automatically generate
-- coercions between them anymore.
-- So, an application such as (@mul A G g h) in the following definition
-- is type incorrect if the coercion is not added (explicitly or implicitly).
local attribute group.to.has_mul [coercion]
local attribute group.to_has_inv [coercion]
universe variable l
variables {A B : Type.{l}}
definition group_eq {G H : group A} (same_mul' : Π(g h : A), @mul A G g h = @mul A H g h)

View file

@ -142,7 +142,7 @@ section
ne_ab eq_ab,
show a < c, from lt_of_le_of_ne le_ac ne_ac
definition order_pair.to_strict_order [instance] [coercion] [reducible] : strict_order A :=
definition order_pair.to_strict_order [instance] [reducible] : strict_order A :=
⦃ strict_order, s, lt_irrefl := lt_irrefl s, lt_trans := lt_trans s ⦄
definition lt_of_lt_of_le [trans] : a < b → b ≤ c → a < c :=
@ -228,7 +228,7 @@ iff.intro
iff.mp !strict_order_with_le.le_iff_lt_or_eq (pr1 H),
show a < b, from sum_resolve_left H1 (pr2 H))
definition strict_order_with_le.to_order_pair [instance] [coercion] [reducible] [s : strict_order_with_le A] :
definition strict_order_with_le.to_order_pair [instance] [reducible] [s : strict_order_with_le A] :
strong_order_pair A :=
⦃ strong_order_pair, s,
le_refl := le_refl s,
@ -263,7 +263,7 @@ section
(assume H, H1 H)
(assume H, sum.rec_on H (assume H', H2 H') (assume H', H3 H'))
definition linear_strong_order_pair.to_linear_order_pair [instance] [coercion] [reducible]
definition linear_strong_order_pair.to_linear_order_pair [instance] [reducible]
: linear_order_pair A :=
⦃ linear_order_pair, s ⦄

View file

@ -213,12 +213,12 @@ definition ordered_comm_group.le_of_add_le_add_left [s : ordered_comm_group A] {
assert H' : -a + (a + b) ≤ -a + (a + c), from ordered_comm_group.add_le_add_left _ _ H _,
by rewrite *neg_add_cancel_left at H'; exact H'
definition ordered_comm_group.to_ordered_cancel_comm_monoid [instance] [coercion] [reducible]
definition ordered_comm_group.to_ordered_cancel_comm_monoid [instance] [reducible]
[s : ordered_comm_group A] : ordered_cancel_comm_monoid A :=
⦃ ordered_cancel_comm_monoid, s,
add_left_cancel := @add.left_cancel A s,
add_right_cancel := @add.right_cancel A s,
le_of_add_le_add_left := @ordered_comm_group.le_of_add_le_add_left A s
add_left_cancel := @add.left_cancel A _,
add_right_cancel := @add.right_cancel A _,
le_of_add_le_add_left := @ordered_comm_group.le_of_add_le_add_left A _
section
variables [s : ordered_comm_group A] (a b c d e : A)

View file

@ -188,18 +188,18 @@ begin
exact (iff.mp !sub_pos_iff_lt H2)
end
definition ordered_ring.to_ordered_semiring [instance] [coercion] [reducible] [s : ordered_ring A] :
definition ordered_ring.to_ordered_semiring [instance] [reducible] [s : ordered_ring A] :
ordered_semiring A :=
⦃ ordered_semiring, s,
mul_zero := mul_zero,
zero_mul := zero_mul,
add_left_cancel := @add.left_cancel A s,
add_right_cancel := @add.right_cancel A s,
le_of_add_le_add_left := @le_of_add_le_add_left A s,
mul_le_mul_of_nonneg_left := @ordered_ring.mul_le_mul_of_nonneg_left A s,
mul_le_mul_of_nonneg_right := @ordered_ring.mul_le_mul_of_nonneg_right A s,
mul_lt_mul_of_pos_left := @ordered_ring.mul_lt_mul_of_pos_left A s,
mul_lt_mul_of_pos_right := @ordered_ring.mul_lt_mul_of_pos_right A s
add_left_cancel := @add.left_cancel A _,
add_right_cancel := @add.right_cancel A _,
le_of_add_le_add_left := @le_of_add_le_add_left A _,
mul_le_mul_of_nonneg_left := @ordered_ring.mul_le_mul_of_nonneg_left A _,
mul_le_mul_of_nonneg_right := @ordered_ring.mul_le_mul_of_nonneg_right A _,
mul_lt_mul_of_pos_left := @ordered_ring.mul_lt_mul_of_pos_left A _,
mul_lt_mul_of_pos_right := @ordered_ring.mul_lt_mul_of_pos_right A _
section
variable [s : ordered_ring A]
@ -268,19 +268,19 @@ structure linear_ordered_ring [class] (A : Type) extends ordered_ring A, linear_
-- print fields linear_ordered_semiring
definition linear_ordered_ring.to_linear_ordered_semiring [instance] [coercion] [reducible]
definition linear_ordered_ring.to_linear_ordered_semiring [instance] [reducible]
[s : linear_ordered_ring A] :
linear_ordered_semiring A :=
⦃ linear_ordered_semiring, s,
mul_zero := mul_zero,
zero_mul := zero_mul,
add_left_cancel := @add.left_cancel A s,
add_right_cancel := @add.right_cancel A s,
le_of_add_le_add_left := @le_of_add_le_add_left A s,
mul_le_mul_of_nonneg_left := @mul_le_mul_of_nonneg_left A s,
mul_le_mul_of_nonneg_right := @mul_le_mul_of_nonneg_right A s,
mul_lt_mul_of_pos_left := @mul_lt_mul_of_pos_left A s,
mul_lt_mul_of_pos_right := @mul_lt_mul_of_pos_right A s,
add_left_cancel := @add.left_cancel A _,
add_right_cancel := @add.right_cancel A _,
le_of_add_le_add_left := @le_of_add_le_add_left A _,
mul_le_mul_of_nonneg_left := @mul_le_mul_of_nonneg_left A _,
mul_le_mul_of_nonneg_right := @mul_le_mul_of_nonneg_right A _,
mul_lt_mul_of_pos_left := @mul_lt_mul_of_pos_left A _,
mul_lt_mul_of_pos_right := @mul_lt_mul_of_pos_right A _,
le_total := linear_ordered_ring.le_total ⦄
structure linear_ordered_comm_ring [class] (A : Type) extends linear_ordered_ring A, comm_monoid A
@ -321,7 +321,7 @@ lt.by_cases
end))
-- Linearity implies no zero divisors. Doesn't need commutativity.
definition linear_ordered_comm_ring.to_integral_domain [instance] [coercion] [reducible]
definition linear_ordered_comm_ring.to_integral_domain [instance] [reducible]
[s: linear_ordered_comm_ring A] : integral_domain A :=
⦃ integral_domain, s,
eq_zero_or_eq_zero_of_mul_eq_zero :=

View file

@ -163,7 +163,7 @@ have H : 0 * a + 0 = 0 * a + 0 * a, from calc
... = 0 * a + 0 * a : by rewrite {_*a}ring.right_distrib,
show 0 * a = 0, from (add.left_cancel H)⁻¹
definition ring.to_semiring [instance] [coercion] [reducible] [s : ring A] : semiring A :=
definition ring.to_semiring [instance] [reducible] [s : ring A] : semiring A :=
⦃ semiring, s,
mul_zero := ring.mul_zero,
zero_mul := ring.zero_mul ⦄
@ -242,7 +242,7 @@ end
structure comm_ring [class] (A : Type) extends ring A, comm_semigroup A
definition comm_ring.to_comm_semiring [instance] [coercion] [reducible] [s : comm_ring A] : comm_semiring A :=
definition comm_ring.to_comm_semiring [instance] [reducible] [s : comm_ring A] : comm_semiring A :=
⦃ comm_semiring, s,
mul_zero := mul_zero,
zero_mul := zero_mul ⦄

View file

@ -328,7 +328,7 @@ section discrete_field
(suppose x ≠ 0,
or.inr (by rewrite [-one_mul, -(inv_mul_cancel this), mul.assoc, H, mul_zero]))
definition discrete_field.to_integral_domain [trans_instance] [reducible] [coercion] :
definition discrete_field.to_integral_domain [trans_instance] [reducible] :
integral_domain A :=
⦃ integral_domain, s,
eq_zero_or_eq_zero_of_mul_eq_zero := discrete_field.eq_zero_or_eq_zero_of_mul_eq_zero⦄

View file

@ -31,10 +31,10 @@ theorem mul.comm [s : comm_semigroup A] (a b : A) : a * b = b * a :=
!comm_semigroup.mul_comm
theorem mul.left_comm [s : comm_semigroup A] (a b c : A) : a * (b * c) = b * (a * c) :=
binary.left_comm (@mul.comm A s) (@mul.assoc A s) a b c
binary.left_comm (@mul.comm A _) (@mul.assoc A _) a b c
theorem mul.right_comm [s : comm_semigroup A] (a b c : A) : (a * b) * c = (a * c) * b :=
binary.right_comm (@mul.comm A s) (@mul.assoc A s) a b c
binary.right_comm (@mul.comm A _) (@mul.assoc A _) a b c
structure left_cancel_semigroup [class] (A : Type) extends semigroup A :=
(mul_left_cancel : ∀a b c, mul a b = mul a c → b = c)
@ -70,10 +70,10 @@ theorem add.comm [s : add_comm_semigroup A] (a b : A) : a + b = b + a :=
theorem add.left_comm [s : add_comm_semigroup A] (a b c : A) :
a + (b + c) = b + (a + c) :=
binary.left_comm (@add.comm A s) (@add.assoc A s) a b c
binary.left_comm (@add.comm A _) (@add.assoc A _) a b c
theorem add.right_comm [s : add_comm_semigroup A] (a b c : A) : (a + b) + c = (a + c) + b :=
binary.right_comm (@add.comm A s) (@add.assoc A s) a b c
binary.right_comm (@add.comm A _) (@add.assoc A _) a b c
structure add_left_cancel_semigroup [class] (A : Type) extends add_semigroup A :=
(add_left_cancel : ∀a b c, add a b = add a c → b = c)
@ -309,12 +309,12 @@ section group
... = x ∘c b : Py
... = a : Px)
definition group.to_left_cancel_semigroup [trans_instance] [coercion] [reducible] :
definition group.to_left_cancel_semigroup [trans_instance] [reducible] :
left_cancel_semigroup A :=
⦃ left_cancel_semigroup, s,
mul_left_cancel := @mul_left_cancel A s ⦄
definition group.to_right_cancel_semigroup [trans_instance] [coercion] [reducible] :
definition group.to_right_cancel_semigroup [trans_instance] [reducible] :
right_cancel_semigroup A :=
⦃ right_cancel_semigroup, s,
mul_right_cancel := @mul_right_cancel A s ⦄
@ -437,12 +437,12 @@ section add_group
... = (c + b) + -b : H
... = c : add_neg_cancel_right
definition add_group.to_left_cancel_semigroup [trans_instance] [coercion] [reducible] :
definition add_group.to_left_cancel_semigroup [trans_instance] [reducible] :
add_left_cancel_semigroup A :=
⦃ add_left_cancel_semigroup, s,
add_left_cancel := @add_left_cancel A s ⦄
definition add_group.to_add_right_cancel_semigroup [trans_instance] [coercion] [reducible] :
definition add_group.to_add_right_cancel_semigroup [trans_instance] [reducible] :
add_right_cancel_semigroup A :=
⦃ add_right_cancel_semigroup, s,
add_right_cancel := @add_right_cancel A s ⦄

View file

@ -107,7 +107,7 @@ namespace finset
include cmB
theorem mulf_rcomm (f : A → B) : right_commutative (mulf f) :=
right_commutative_compose_right (@has_mul.mul B cmB) f (@mul.right_comm B cmB)
right_commutative_compose_right (@has_mul.mul B _) f (@mul.right_comm B _)
theorem Prodl_eq_Prodl_of_perm (f : A → B) {l₁ l₂ : list A} :
perm l₁ l₂ → Prodl l₁ f = Prodl l₂ f :=

View file

@ -111,7 +111,7 @@ section
private theorem lt_trans (s' : order_pair A) (a b c: A) (lt_ab : a < b) (lt_bc : b < c) : a < c :=
lt_of_lt_of_le lt_ab (le_of_lt lt_bc)
definition order_pair.to_strict_order [trans_instance] [coercion] [reducible] : strict_order A :=
definition order_pair.to_strict_order [trans_instance] [reducible] : strict_order A :=
⦃ strict_order, s, lt_irrefl := lt_irrefl s, lt_trans := lt_trans s ⦄
theorem gt_of_gt_of_ge [trans] (H1 : a > b) (H2 : b ≥ c) : a > c := lt_of_le_of_lt H2 H1
@ -181,7 +181,7 @@ have ne_ac : a ≠ c, from
show false, from ne_of_lt' lt_bc eq_bc,
show a < c, from iff.mpr (lt_iff_le_and_ne) (and.intro le_ac ne_ac)
definition strong_order_pair.to_order_pair [trans_instance] [coercion] [reducible]
definition strong_order_pair.to_order_pair [trans_instance] [reducible]
[s : strong_order_pair A] : order_pair A :=
⦃ order_pair, s,
lt_irrefl := lt_irrefl',
@ -196,7 +196,7 @@ structure linear_order_pair [class] (A : Type) extends order_pair A, linear_weak
structure linear_strong_order_pair [class] (A : Type) extends strong_order_pair A,
linear_weak_order A
definition linear_strong_order_pair.to_linear_order_pair [trans_instance] [coercion] [reducible]
definition linear_strong_order_pair.to_linear_order_pair [trans_instance] [reducible]
[s : linear_strong_order_pair A] : linear_order_pair A :=
⦃ linear_order_pair, s, strong_order_pair.to_order_pair ⦄

View file

@ -384,7 +384,7 @@ section discrete_linear_ordered_field
(assume H' : ¬ y < x,
decidable.inl (le.antisymm (le_of_not_gt H') (le_of_not_gt H))))
definition discrete_linear_ordered_field.to_discrete_field [trans_instance] [reducible] [coercion]
definition discrete_linear_ordered_field.to_discrete_field [trans_instance] [reducible]
: discrete_field A :=
⦃ discrete_field, s, has_decidable_eq := dec_eq_of_dec_lt⦄

View file

@ -210,13 +210,14 @@ theorem ordered_comm_group.lt_of_add_lt_add_left [s : ordered_comm_group A] {a b
assert H' : -a + (a + b) < -a + (a + c), from ordered_comm_group.add_lt_add_left _ _ H _,
by rewrite *neg_add_cancel_left at H'; exact H'
definition ordered_comm_group.to_ordered_cancel_comm_monoid [trans_instance] [coercion] [reducible]
set_option pp.all true
definition ordered_comm_group.to_ordered_cancel_comm_monoid [trans_instance] [reducible]
[s : ordered_comm_group A] : ordered_cancel_comm_monoid A :=
⦃ ordered_cancel_comm_monoid, s,
add_left_cancel := @add.left_cancel A s,
add_right_cancel := @add.right_cancel A s,
le_of_add_le_add_left := @ordered_comm_group.le_of_add_le_add_left A s,
lt_of_add_lt_add_left := @ordered_comm_group.lt_of_add_lt_add_left A s
add_left_cancel := @add.left_cancel A _,
add_right_cancel := @add.right_cancel A _,
le_of_add_le_add_left := @ordered_comm_group.le_of_add_le_add_left A _,
lt_of_add_lt_add_left := @ordered_comm_group.lt_of_add_lt_add_left A _
section
variables [s : ordered_comm_group A] (a b c d e : A)
@ -582,12 +583,12 @@ structure decidable_linear_ordered_comm_group [class] (A : Type)
(add_lt_add_left : ∀ a b, lt a b → ∀ c, lt (add c a) (add c b))
definition decidable_linear_ordered_comm_group.to_ordered_comm_group
[trans_instance] [reducible] [coercion]
[trans_instance] [reducible]
(A : Type) [s : decidable_linear_ordered_comm_group A] : ordered_comm_group A :=
⦃ ordered_comm_group, s,
le_of_lt := @le_of_lt A s,
lt_of_le_of_lt := @lt_of_le_of_lt A s,
lt_of_lt_of_le := @lt_of_lt_of_le A s
le_of_lt := @le_of_lt A _,
lt_of_le_of_lt := @lt_of_le_of_lt A _,
lt_of_lt_of_le := @lt_of_lt_of_le A _
section
variables [s : decidable_linear_ordered_comm_group A]

View file

@ -235,20 +235,20 @@ begin
exact (iff.mp !sub_pos_iff_lt H2)
end
definition ordered_ring.to_ordered_semiring [trans_instance] [coercion] [reducible]
definition ordered_ring.to_ordered_semiring [trans_instance] [reducible]
[s : ordered_ring A] :
ordered_semiring A :=
⦃ ordered_semiring, s,
mul_zero := mul_zero,
zero_mul := zero_mul,
add_left_cancel := @add.left_cancel A s,
add_right_cancel := @add.right_cancel A s,
le_of_add_le_add_left := @le_of_add_le_add_left A s,
mul_le_mul_of_nonneg_left := @ordered_ring.mul_le_mul_of_nonneg_left A s,
mul_le_mul_of_nonneg_right := @ordered_ring.mul_le_mul_of_nonneg_right A s,
mul_lt_mul_of_pos_left := @ordered_ring.mul_lt_mul_of_pos_left A s,
mul_lt_mul_of_pos_right := @ordered_ring.mul_lt_mul_of_pos_right A s,
lt_of_add_lt_add_left := @lt_of_add_lt_add_left A s
add_left_cancel := @add.left_cancel A _,
add_right_cancel := @add.right_cancel A _,
le_of_add_le_add_left := @le_of_add_le_add_left A _,
mul_le_mul_of_nonneg_left := @ordered_ring.mul_le_mul_of_nonneg_left A _,
mul_le_mul_of_nonneg_right := @ordered_ring.mul_le_mul_of_nonneg_right A _,
mul_lt_mul_of_pos_left := @ordered_ring.mul_lt_mul_of_pos_left A _,
mul_lt_mul_of_pos_right := @ordered_ring.mul_lt_mul_of_pos_right A _,
lt_of_add_lt_add_left := @lt_of_add_lt_add_left A _
section
variable [s : ordered_ring A]
@ -317,21 +317,21 @@ structure linear_ordered_ring [class] (A : Type)
extends ordered_ring A, linear_strong_order_pair A :=
(zero_lt_one : lt zero one)
definition linear_ordered_ring.to_linear_ordered_semiring [trans_instance] [coercion] [reducible]
definition linear_ordered_ring.to_linear_ordered_semiring [trans_instance] [reducible]
[s : linear_ordered_ring A] :
linear_ordered_semiring A :=
⦃ linear_ordered_semiring, s,
mul_zero := mul_zero,
zero_mul := zero_mul,
add_left_cancel := @add.left_cancel A s,
add_right_cancel := @add.right_cancel A s,
le_of_add_le_add_left := @le_of_add_le_add_left A s,
mul_le_mul_of_nonneg_left := @mul_le_mul_of_nonneg_left A s,
mul_le_mul_of_nonneg_right := @mul_le_mul_of_nonneg_right A s,
mul_lt_mul_of_pos_left := @mul_lt_mul_of_pos_left A s,
mul_lt_mul_of_pos_right := @mul_lt_mul_of_pos_right A s,
add_left_cancel := @add.left_cancel A _,
add_right_cancel := @add.right_cancel A _,
le_of_add_le_add_left := @le_of_add_le_add_left A _,
mul_le_mul_of_nonneg_left := @mul_le_mul_of_nonneg_left A _,
mul_le_mul_of_nonneg_right := @mul_le_mul_of_nonneg_right A _,
mul_lt_mul_of_pos_left := @mul_lt_mul_of_pos_left A _,
mul_lt_mul_of_pos_right := @mul_lt_mul_of_pos_right A _,
le_total := linear_ordered_ring.le_total,
lt_of_add_lt_add_left := @lt_of_add_lt_add_left A s
lt_of_add_lt_add_left := @lt_of_add_lt_add_left A _
structure linear_ordered_comm_ring [class] (A : Type) extends linear_ordered_ring A, comm_monoid A
@ -371,7 +371,7 @@ lt.by_cases
end))
-- Linearity implies no zero divisors. Doesn't need commutativity.
definition linear_ordered_comm_ring.to_integral_domain [trans_instance] [coercion] [reducible]
definition linear_ordered_comm_ring.to_integral_domain [trans_instance] [reducible]
[s: linear_ordered_comm_ring A] : integral_domain A :=
⦃ integral_domain, s,
eq_zero_or_eq_zero_of_mul_eq_zero :=

View file

@ -174,7 +174,7 @@ have 0 * a + 0 = 0 * a + 0 * a, from calc
... = 0 * a + 0 * a : by rewrite {_*a}ring.right_distrib,
show 0 * a = 0, from (add.left_cancel this)⁻¹
definition ring.to_semiring [trans_instance] [coercion] [reducible] [s : ring A] : semiring A :=
definition ring.to_semiring [trans_instance] [reducible] [s : ring A] : semiring A :=
⦃ semiring, s,
mul_zero := ring.mul_zero,
zero_mul := ring.zero_mul ⦄
@ -259,7 +259,7 @@ end
structure comm_ring [class] (A : Type) extends ring A, comm_semigroup A
definition comm_ring.to_comm_semiring [trans_instance] [coercion] [reducible] [s : comm_ring A] : comm_semiring A :=
definition comm_ring.to_comm_semiring [trans_instance] [reducible] [s : comm_ring A] : comm_semiring A :=
⦃ comm_semiring, s,
mul_zero := mul_zero,
zero_mul := zero_mul ⦄

View file

@ -179,9 +179,9 @@ section
open algebra
omit decR
lemma strongly_sorted_sort [ord : decidable_linear_order A] (l : list A) : strongly_sorted le (sort le l) :=
strongly_sorted_sort_core le.total (@le.trans A ord) le.refl l
strongly_sorted_sort_core le.total (@le.trans A _) le.refl l
lemma sort_eq_of_perm {l₁ l₂ : list A} [ord : decidable_linear_order A] (h : l₁ ~ l₂) : sort le l₁ = sort le l₂ :=
sort_eq_of_perm_core le.total (@le.trans A ord) le.refl (@le.antisymm A ord) h
sort_eq_of_perm_core le.total (@le.trans A _) le.refl (@le.antisymm A _) h
end
end list

View file

@ -52,7 +52,7 @@ variables {A B : Type}
variable [s1 : group A]
definition id_is_iso [instance] : @is_hom_class A A s1 s1 (@id A) :=
is_iso_class.mk (take a b, rfl) injective_id
is_hom_class.mk (take a b, rfl)
variable [s2 : group B]
include s1

View file

@ -818,7 +818,8 @@ struct structure_cmd_fn {
save_def_info(coercion_name);
add_alias(coercion_name);
if (!m_private_parents[i]) {
m_env = add_coercion(m_env, m_p.ios(), coercion_name);
if (!m_modifiers.is_class() || !is_class(m_env, parent_name))
m_env = add_coercion(m_env, m_p.ios(), coercion_name);
if (m_modifiers.is_class() && is_class(m_env, parent_name)) {
// if both are classes, then we also mark coercion_name as an instance
m_env = add_trans_instance(m_env, coercion_name);

View file

@ -139,8 +139,4 @@ calc
a * 1 * b * c = a * b * c : {!mul_right_id}
... = a * (b * c) : !mul_assoc
theorem test6 {M : CommMonoid} (a b c : M) : a * 1 * b * c = a * (b * c) :=
calc
a * 1 * b * c = a * b * c : {!mul_right_id}
... = a * (b * c) : !mul_assoc
end examples