@@ -243,10 +243,12 @@ instance [partial_order α] : partial_order (with_zero α) := with_bot.partial_o
243
243
244
244
instance [preorder α] : order_bot (with_zero α) := with_bot.order_bot
245
245
246
- lemma zero_le [partial_order α] (a : with_zero α) : 0 ≤ a := order_bot. bot_le a
246
+ lemma zero_le [preorder α] (a : with_zero α) : 0 ≤ a := bot_le
247
247
248
248
lemma zero_lt_coe [preorder α] (a : α) : (0 : with_zero α) < a := with_bot.bot_lt_coe a
249
249
250
+ lemma zero_eq_bot [preorder α] : (0 : with_zero α) = ⊥ := rfl
251
+
250
252
@[simp, norm_cast] lemma coe_lt_coe [preorder α] {a b : α} : (a : with_zero α) < b ↔ a < b :=
251
253
with_bot.coe_lt_coe
252
254
@@ -257,30 +259,28 @@ instance [lattice α] : lattice (with_zero α) := with_bot.lattice
257
259
258
260
instance [linear_order α] : linear_order (with_zero α) := with_bot.linear_order
259
261
260
- lemma mul_le_mul_left {α : Type u} [has_mul α] [preorder α]
262
+ instance covariant_class_mul_le {α : Type u} [has_mul α] [preorder α]
261
263
[covariant_class α α (*) (≤)] :
262
- ∀ (a b : with_zero α),
263
- a ≤ b → ∀ (c : with_zero α), c * a ≤ c * b :=
264
+ covariant_class (with_zero α) (with_zero α) (*) (≤) :=
264
265
begin
265
- rintro (_ | a) (_ | b) h (_ | c);
266
- try { exact λ f hf, option.no_confusion hf },
267
- { exact false.elim (not_lt_of_le h (with_zero.zero_lt_coe a)) },
268
- { simp_rw [some_eq_coe] at h ⊢ ,
269
- norm_cast at h ⊢ ,
270
- exact covariant_class.elim _ h }
266
+ refine ⟨λ a b c hbc, _⟩,
267
+ induction a using with_zero.rec_zero_coe, { exact zero_le _ },
268
+ induction b using with_zero.rec_zero_coe, { exact zero_le _ },
269
+ rcases with_bot.coe_le_iff. 1 hbc with ⟨c, rfl, hbc'⟩ ,
270
+ rw [← coe_mul, ← coe_mul, coe_le_coe] ,
271
+ exact mul_le_mul_left' hbc' a
271
272
end
272
273
273
- lemma lt_of_mul_lt_mul_left {α : Type u} [has_mul α] [partial_order α]
274
+ instance contravariant_class_mul_lt {α : Type u} [has_mul α] [partial_order α]
274
275
[contravariant_class α α (*) (<)] :
275
- ∀ (a b c : with_zero α), a * b < a * c → b < c :=
276
+ contravariant_class (with_zero α) ( with_zero α) (*) (<) :=
276
277
begin
277
- rintro (_ | a) (_ | b) (_ | c) h;
278
- try { exact false.elim (lt_irrefl _ h) },
279
- { exact with_zero.zero_lt_coe c },
280
- { exact false.elim (not_le_of_lt h (with_zero.zero_le _)) },
281
- { simp_rw [some_eq_coe] at h ⊢,
282
- norm_cast at h ⊢,
283
- apply lt_of_mul_lt_mul_left' h }
278
+ refine ⟨λ a b c h, _⟩,
279
+ have := ((zero_le _).trans_lt h).ne',
280
+ lift a to α using left_ne_zero_of_mul this ,
281
+ lift c to α using right_ne_zero_of_mul this ,
282
+ induction b using with_zero.rec_zero_coe,
283
+ exacts [zero_lt_coe _, coe_lt_coe.mpr (lt_of_mul_lt_mul_left' $ coe_lt_coe.mp h)]
284
284
end
285
285
286
286
@[simp] lemma le_max_iff [linear_order α] {a b c : α} :
@@ -292,10 +292,28 @@ by simp only [with_zero.coe_le_coe, le_max_iff]
292
292
by simp only [with_zero.coe_le_coe, min_le_iff]
293
293
294
294
instance [ordered_comm_monoid α] : ordered_comm_monoid (with_zero α) :=
295
- { mul_le_mul_left := with_zero. mul_le_mul_left,
295
+ { mul_le_mul_left := λ _ _, mul_le_mul_left' ,
296
296
..with_zero.comm_monoid_with_zero,
297
297
..with_zero.partial_order }
298
298
299
+ protected lemma covariant_class_add_le [add_zero_class α] [preorder α]
300
+ [covariant_class α α (+) (≤)] (h : ∀ a : α, 0 ≤ a) :
301
+ covariant_class (with_zero α) (with_zero α) (+) (≤) :=
302
+ begin
303
+ refine ⟨λ a b c hbc, _⟩,
304
+ induction a using with_zero.rec_zero_coe,
305
+ { rwa [zero_add, zero_add] },
306
+ induction b using with_zero.rec_zero_coe,
307
+ { rw [add_zero],
308
+ induction c using with_zero.rec_zero_coe,
309
+ { rw [add_zero], exact le_rfl },
310
+ { rw [← coe_add, coe_le_coe],
311
+ exact le_add_of_nonneg_right (h _) } },
312
+ { rcases with_bot.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩,
313
+ rw [← coe_add, ← coe_add, coe_le_coe],
314
+ exact add_le_add_left hbc' a }
315
+ end
316
+
299
317
/-
300
318
Note 1 : the below is not an instance because it requires `zero_le`. It seems
301
319
like a rather pathological definition because α already has a zero.
@@ -310,25 +328,9 @@ See note [reducible non-instances].
310
328
-/
311
329
@[reducible] protected def ordered_add_comm_monoid [ordered_add_comm_monoid α]
312
330
(zero_le : ∀ a : α, 0 ≤ a) : ordered_add_comm_monoid (with_zero α) :=
313
- begin
314
- suffices , refine
315
- { add_le_add_left := this ,
316
- ..with_zero.partial_order,
317
- ..with_zero.add_comm_monoid, .. },
318
- { intros a b h c ca h₂,
319
- cases b with b,
320
- { rw le_antisymm h bot_le at h₂,
321
- exact ⟨_, h₂, le_rfl⟩ },
322
- cases a with a,
323
- { change c + 0 = some ca at h₂,
324
- simp at h₂, simp [h₂],
325
- exact ⟨_, rfl, by simpa using add_le_add_left (zero_le b) _⟩ },
326
- { simp at h,
327
- cases c with c; change some _ = _ at h₂;
328
- simp [-add_comm] at h₂; subst ca; refine ⟨_, rfl, _⟩,
329
- { exact h },
330
- { exact add_le_add_left h _ } } }
331
- end
331
+ { add_le_add_left := @add_le_add_left _ _ _ (with_zero.covariant_class_add_le zero_le),
332
+ ..with_zero.partial_order,
333
+ ..with_zero.add_comm_monoid, .. }
332
334
333
335
end with_zero
334
336
@@ -411,7 +413,7 @@ iff.intro ne_of_gt $ assume hne, lt_of_le_of_ne (one_le _) hne.symm
411
413
@[to_additive] lemma eq_one_or_one_lt : a = 1 ∨ 1 < a :=
412
414
(one_le a).eq_or_lt.imp_left eq.symm
413
415
414
- @[to_additive] lemma exists_pos_mul_of_lt (h : a < b) : ∃ c > 1 , a * c = b :=
416
+ @[to_additive] lemma exists_one_lt_mul_of_lt (h : a < b) : ∃ c (hc : 1 < c) , a * c = b :=
415
417
begin
416
418
obtain ⟨c, hc⟩ := le_iff_exists_mul.1 h.le,
417
419
refine ⟨c, one_lt_iff_ne_one.2 _, hc.symm⟩,
@@ -885,106 +887,105 @@ by { induction x using with_top.rec_top_coe; simp [← coe_add, -with_zero.coe_a
885
887
@[simp] lemma coe_add_eq_top_iff {y : with_top α} : ↑x + y = ⊤ ↔ y = ⊤ :=
886
888
by { induction y using with_top.rec_top_coe; simp [← coe_add, -with_zero.coe_add] }
887
889
888
- variables [preorder α]
889
-
890
- instance covariant_class_add_le [covariant_class α α (+) (≤)] :
890
+ instance covariant_class_add_le [has_le α] [covariant_class α α (+) (≤)] :
891
891
covariant_class (with_top α) (with_top α) (+) (≤) :=
892
892
⟨λ a b c h, begin
893
893
cases a; cases c; try { exact le_top },
894
- cases b,
895
- { exact (not_top_le_coe _ h).elim },
896
- { exact some_le_some.2 (add_le_add_left (some_le_some.1 h) _) }
894
+ rcases le_coe_iff.1 h with ⟨b, rfl, h'⟩,
895
+ exact coe_le_coe.2 (add_le_add_left (coe_le_coe.1 h) _)
897
896
end ⟩
898
897
899
- instance covariant_class_swap_add_le [covariant_class α α (swap (+)) (≤)] :
898
+ instance covariant_class_swap_add_le [has_le α] [ covariant_class α α (swap (+)) (≤)] :
900
899
covariant_class (with_top α) (with_top α) (swap (+)) (≤) :=
901
900
⟨λ a b c h, begin
902
901
cases a; cases c; try { exact le_top },
903
- cases b,
904
- { exact (not_top_le_coe _ h).elim },
905
- { exact some_le_some.2 (add_le_add_right (some_le_some.1 h) _) }
902
+ rcases le_coe_iff.1 h with ⟨b, rfl, h'⟩,
903
+ exact coe_le_coe.2 (add_le_add_right (coe_le_coe.1 h) _)
906
904
end ⟩
907
905
908
- instance contravariant_class_add_lt [contravariant_class α α (+) (<)] :
906
+ instance contravariant_class_add_lt [has_lt α] [ contravariant_class α α (+) (<)] :
909
907
contravariant_class (with_top α) (with_top α) (+) (<) :=
910
908
⟨λ a b c h, begin
911
- cases a; cases b; try { exact (not_top_lt h).elim },
912
- cases c,
909
+ induction a using with_top.rec_top_coe, { exact (not_none_lt _ h).elim },
910
+ induction b using with_top.rec_top_coe, { exact (not_none_lt _ h).elim },
911
+ induction c using with_top.rec_top_coe,
913
912
{ exact coe_lt_top _ },
914
- { exact some_lt_some .2 (lt_of_add_lt_add_left $ some_lt_some .1 h) }
913
+ { exact coe_lt_coe .2 (lt_of_add_lt_add_left $ coe_lt_coe .1 h) }
915
914
end ⟩
916
915
917
- instance contravariant_class_swap_add_lt [contravariant_class α α (swap (+)) (<)] :
916
+ instance contravariant_class_swap_add_lt [has_lt α] [ contravariant_class α α (swap (+)) (<)] :
918
917
contravariant_class (with_top α) (with_top α) (swap (+)) (<) :=
919
918
⟨λ a b c h, begin
920
- cases a; cases b; try { exact (not_top_lt h).elim },
919
+ cases a; cases b; try { exact (not_none_lt _ h).elim },
921
920
cases c,
922
921
{ exact coe_lt_top _ },
923
- { exact some_lt_some .2 (lt_of_add_lt_add_right $ some_lt_some .1 h) }
922
+ { exact coe_lt_coe .2 (lt_of_add_lt_add_right $ coe_lt_coe .1 h) }
924
923
end ⟩
925
924
926
- protected lemma le_of_add_le_add_left [contravariant_class α α (+) (≤)] (ha : a ≠ ⊤)
925
+ protected lemma le_of_add_le_add_left [has_le α] [ contravariant_class α α (+) (≤)] (ha : a ≠ ⊤)
927
926
(h : a + b ≤ a + c) : b ≤ c :=
928
927
begin
929
928
lift a to α using ha,
930
- cases c; try { exact le_top},
931
- cases b, exact (not_top_le_coe _ h).elim,
932
- simp only [some_eq_coe, ← coe_add, coe_le_coe] at h, rw some_le_some ,
929
+ induction c using with_top.rec_top_coe, { exact le_top },
930
+ induction b using with_top.rec_top_coe, { exact (not_top_le_coe _ h).elim } ,
931
+ simp only [← coe_add, coe_le_coe] at h ⊢ ,
933
932
exact le_of_add_le_add_left h
934
933
end
935
934
936
- protected lemma le_of_add_le_add_right [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤)
937
- (h : b + a ≤ c + a) : b ≤ c :=
935
+ protected lemma le_of_add_le_add_right [has_le α] [ contravariant_class α α (swap (+)) (≤)]
936
+ (ha : a ≠ ⊤) ( h : b + a ≤ c + a) : b ≤ c :=
938
937
begin
939
938
lift a to α using ha,
940
939
cases c,
941
940
{ exact le_top },
942
941
cases b,
943
942
{ exact (not_top_le_coe _ h).elim },
944
- { exact some_le_some .2 (le_of_add_le_add_right $ some_le_some .1 h) }
943
+ { exact coe_le_coe .2 (le_of_add_le_add_right $ coe_le_coe .1 h) }
945
944
end
946
945
947
- protected lemma add_lt_add_left [covariant_class α α (+) (<)] (ha : a ≠ ⊤) (h : b < c) :
946
+ protected lemma add_lt_add_left [has_lt α] [ covariant_class α α (+) (<)] (ha : a ≠ ⊤) (h : b < c) :
948
947
a + b < a + c :=
949
948
begin
950
949
lift a to α using ha,
951
- lift b to α using (h.trans_le le_top).ne ,
950
+ rcases lt_iff_exists_coe. 1 h with ⟨b, rfl, h'⟩ ,
952
951
cases c,
953
952
{ exact coe_lt_top _ },
954
- { exact some_lt_some .2 (add_lt_add_left (some_lt_some .1 h) _) }
953
+ { exact coe_lt_coe .2 (add_lt_add_left (coe_lt_coe .1 h) _) }
955
954
end
956
955
957
- protected lemma add_lt_add_right [covariant_class α α (swap (+)) (<)] (ha : a ≠ ⊤) (h : b < c) :
956
+ protected lemma add_lt_add_right [has_lt α] [covariant_class α α (swap (+)) (<)]
957
+ (ha : a ≠ ⊤) (h : b < c) :
958
958
b + a < c + a :=
959
959
begin
960
960
lift a to α using ha,
961
- lift b to α using (h.trans_le le_top).ne ,
961
+ rcases lt_iff_exists_coe. 1 h with ⟨b, rfl, h'⟩ ,
962
962
cases c,
963
963
{ exact coe_lt_top _ },
964
- { exact some_lt_some .2 (add_lt_add_right (some_lt_some .1 h) _) }
964
+ { exact coe_lt_coe .2 (add_lt_add_right (coe_lt_coe .1 h) _) }
965
965
end
966
966
967
- protected lemma add_le_add_iff_left [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)]
967
+ protected lemma add_le_add_iff_left [has_le α] [covariant_class α α (+) (≤)]
968
+ [contravariant_class α α (+) (≤)]
968
969
(ha : a ≠ ⊤) : a + b ≤ a + c ↔ b ≤ c :=
969
970
⟨with_top.le_of_add_le_add_left ha, λ h, add_le_add_left h a⟩
970
971
971
- protected lemma add_le_add_iff_right [covariant_class α α (swap (+)) (≤)]
972
+ protected lemma add_le_add_iff_right [has_le α] [ covariant_class α α (swap (+)) (≤)]
972
973
[contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤) : b + a ≤ c + a ↔ b ≤ c :=
973
974
⟨with_top.le_of_add_le_add_right ha, λ h, add_le_add_right h a⟩
974
975
975
- protected lemma add_lt_add_iff_left [covariant_class α α (+) (<) ] [contravariant_class α α (+) (<)]
976
- (ha : a ≠ ⊤) : a + b < a + c ↔ b < c :=
976
+ protected lemma add_lt_add_iff_left [has_lt α ] [covariant_class α α (+) (<)]
977
+ [contravariant_class α α (+) (<)] (ha : a ≠ ⊤) : a + b < a + c ↔ b < c :=
977
978
⟨lt_of_add_lt_add_left, with_top.add_lt_add_left ha⟩
978
979
979
- protected lemma add_lt_add_iff_right [covariant_class α α (swap (+)) (<)]
980
+ protected lemma add_lt_add_iff_right [has_lt α] [ covariant_class α α (swap (+)) (<)]
980
981
[contravariant_class α α (swap (+)) (<)] (ha : a ≠ ⊤) : b + a < c + a ↔ b < c :=
981
982
⟨lt_of_add_lt_add_right, with_top.add_lt_add_right ha⟩
982
983
983
- protected lemma add_lt_add_of_le_of_lt [covariant_class α α (+) (<)]
984
+ protected lemma add_lt_add_of_le_of_lt [preorder α] [ covariant_class α α (+) (<)]
984
985
[covariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤) (hab : a ≤ b) (hcd : c < d) : a + c < b + d :=
985
986
(with_top.add_lt_add_left ha hcd).trans_le $ add_le_add_right hab _
986
987
987
- protected lemma add_lt_add_of_lt_of_le [covariant_class α α (+) (≤)]
988
+ protected lemma add_lt_add_of_lt_of_le [preorder α] [ covariant_class α α (+) (≤)]
988
989
[covariant_class α α (swap (+)) (<)] (hc : c ≠ ⊤) (hab : a < b) (hcd : c ≤ d) : a + c < b + d :=
989
990
(with_top.add_lt_add_right hc hab).trans_le $ add_le_add_left hcd _
990
991
0 commit comments