@@ -551,6 +551,9 @@ instance : inhabited (with_bot α) := ⟨⊥⟩
551
551
lemma coe_injective : injective (coe : α → with_bot α) := option.some_injective _
552
552
@[norm_cast] lemma coe_inj : (a : with_bot α) = b ↔ a = b := option.some_inj
553
553
554
+ protected lemma «forall » {p : with_bot α → Prop } : (∀ x, p x) ↔ p ⊥ ∧ ∀ x : α, p x := option.forall
555
+ protected lemma «exists » {p : with_bot α → Prop } : (∃ x, p x) ↔ p ⊥ ∨ ∃ x : α, p x := option.exists
556
+
554
557
lemma none_eq_bot : (none : with_bot α) = (⊥ : with_bot α) := rfl
555
558
lemma some_eq_coe (a : α) : (some a : with_bot α) = (↑a : with_bot α) := rfl
556
559
@@ -690,12 +693,37 @@ instance [partial_order α] : partial_order (with_bot α) :=
690
693
.. with_bot.preorder }
691
694
692
695
lemma coe_strict_mono [preorder α] : strict_mono (coe : α → with_bot α) := λ a b, some_lt_some.2
696
+ lemma coe_mono [preorder α] : monotone (coe : α → with_bot α) := λ a b, coe_le_coe.2
697
+
698
+ lemma monotone_iff [preorder α] [preorder β] {f : with_bot α → β} :
699
+ monotone f ↔ monotone (f ∘ coe : α → β) ∧ ∀ x : α, f ⊥ ≤ f x :=
700
+ ⟨λ h, ⟨h.comp with_bot.coe_mono, λ x, h bot_le⟩,
701
+ λ h, with_bot.forall.2 ⟨with_bot.forall.2 ⟨λ _, le_rfl, λ x _, h.2 x⟩,
702
+ λ x, with_bot.forall.2 ⟨λ h, (not_coe_le_bot _ h).elim, λ y hle, h.1 (coe_le_coe.1 hle)⟩⟩⟩
703
+
704
+ @[simp] lemma monotone_map_iff [preorder α] [preorder β] {f : α → β} :
705
+ monotone (with_bot.map f) ↔ monotone f :=
706
+ monotone_iff.trans $ by simp [monotone]
707
+
708
+ alias monotone_map_iff ↔ _ _root_.monotone.with_bot_map
709
+
710
+ lemma strict_mono_iff [preorder α] [preorder β] {f : with_bot α → β} :
711
+ strict_mono f ↔ strict_mono (f ∘ coe : α → β) ∧ ∀ x : α, f ⊥ < f x :=
712
+ ⟨λ h, ⟨h.comp with_bot.coe_strict_mono, λ x, h (bot_lt_coe _)⟩,
713
+ λ h, with_bot.forall.2 ⟨with_bot.forall.2 ⟨flip absurd (lt_irrefl _), λ x _, h.2 x⟩,
714
+ λ x, with_bot.forall.2 ⟨λ h, (not_lt_bot h).elim, λ y hle, h.1 (coe_lt_coe.1 hle)⟩⟩⟩
715
+
716
+ @[simp] lemma strict_mono_map_iff [preorder α] [preorder β] {f : α → β} :
717
+ strict_mono (with_bot.map f) ↔ strict_mono f :=
718
+ strict_mono_iff.trans $ by simp [strict_mono, bot_lt_coe]
719
+
720
+ alias strict_mono_map_iff ↔ _ _root_.strict_mono.with_bot_map
693
721
694
722
lemma map_le_iff [preorder α] [preorder β] (f : α → β) (mono_iff : ∀ {a b}, f a ≤ f b ↔ a ≤ b) :
695
723
∀ (a b : with_bot α), a.map f ≤ b.map f ↔ a ≤ b
696
724
| ⊥ _ := by simp only [map_bot, bot_le]
697
725
| (a : α) ⊥ := by simp only [map_coe, map_bot, coe_ne_bot, not_coe_le_bot _]
698
- | (a : α) (b : α) := by simpa using mono_iff
726
+ | (a : α) (b : α) := by simpa only [map_coe, coe_le_coe] using mono_iff
699
727
700
728
lemma le_coe_unbot' [preorder α] : ∀ (a : with_bot α) (b : α), a ≤ a.unbot' b
701
729
| (a : α) b := le_rfl
@@ -872,6 +900,9 @@ meta instance {α : Type} [reflected _ α] [has_reflect α] : has_reflect (with_
872
900
873
901
instance : inhabited (with_top α) := ⟨⊤⟩
874
902
903
+ protected lemma «forall » {p : with_top α → Prop } : (∀ x, p x) ↔ p ⊤ ∧ ∀ x : α, p x := option.forall
904
+ protected lemma «exists » {p : with_top α → Prop } : (∃ x, p x) ↔ p ⊤ ∨ ∃ x : α, p x := option.exists
905
+
875
906
lemma none_eq_top : (none : with_top α) = (⊤ : with_top α) := rfl
876
907
lemma some_eq_coe (a : α) : (some a : with_top α) = (↑a : with_top α) := rfl
877
908
@@ -1153,6 +1184,31 @@ instance [partial_order α] : partial_order (with_top α) :=
1153
1184
.. with_top.preorder }
1154
1185
1155
1186
lemma coe_strict_mono [preorder α] : strict_mono (coe : α → with_top α) := λ a b, some_lt_some.2
1187
+ lemma coe_mono [preorder α] : monotone (coe : α → with_top α) := λ a b, coe_le_coe.2
1188
+
1189
+ lemma monotone_iff [preorder α] [preorder β] {f : with_top α → β} :
1190
+ monotone f ↔ monotone (f ∘ coe : α → β) ∧ ∀ x : α, f x ≤ f ⊤ :=
1191
+ ⟨λ h, ⟨h.comp with_top.coe_mono, λ x, h le_top⟩,
1192
+ λ h, with_top.forall.2 ⟨with_top.forall.2 ⟨λ _, le_rfl, λ x h, (not_top_le_coe _ h).elim⟩,
1193
+ λ x, with_top.forall.2 ⟨λ _, h.2 x, λ y hle, h.1 (coe_le_coe.1 hle)⟩⟩⟩
1194
+
1195
+ @[simp] lemma monotone_map_iff [preorder α] [preorder β] {f : α → β} :
1196
+ monotone (with_top.map f) ↔ monotone f :=
1197
+ monotone_iff.trans $ by simp [monotone]
1198
+
1199
+ alias monotone_map_iff ↔ _ _root_.monotone.with_top_map
1200
+
1201
+ lemma strict_mono_iff [preorder α] [preorder β] {f : with_top α → β} :
1202
+ strict_mono f ↔ strict_mono (f ∘ coe : α → β) ∧ ∀ x : α, f x < f ⊤ :=
1203
+ ⟨λ h, ⟨h.comp with_top.coe_strict_mono, λ x, h (coe_lt_top _)⟩,
1204
+ λ h, with_top.forall.2 ⟨with_top.forall.2 ⟨flip absurd (lt_irrefl _), λ x h, (not_top_lt h).elim⟩,
1205
+ λ x, with_top.forall.2 ⟨λ _, h.2 x, λ y hle, h.1 (coe_lt_coe.1 hle)⟩⟩⟩
1206
+
1207
+ @[simp] lemma strict_mono_map_iff [preorder α] [preorder β] {f : α → β} :
1208
+ strict_mono (with_top.map f) ↔ strict_mono f :=
1209
+ strict_mono_iff.trans $ by simp [strict_mono, coe_lt_top]
1210
+
1211
+ alias strict_mono_map_iff ↔ _ _root_.strict_mono.with_top_map
1156
1212
1157
1213
lemma map_le_iff [preorder α] [preorder β] (f : α → β)
1158
1214
(a b : with_top α) (mono_iff : ∀ {a b}, f a ≤ f b ↔ a ≤ b) :
@@ -1315,27 +1371,6 @@ order_dual.no_min_order (with_bot αᵒᵈ)
1315
1371
1316
1372
end with_top
1317
1373
1318
- section mono
1319
-
1320
- variables [preorder α] [preorder β] {f : α → β}
1321
-
1322
- protected lemma monotone.with_bot_map (hf : monotone f) : monotone (with_bot.map f)
1323
- | ⊥ _ h := bot_le
1324
- | (a : α) ⊥ h := (with_bot.not_coe_le_bot _ h).elim
1325
- | (a : α) (b : α) h := with_bot.coe_le_coe.2 (hf (with_bot.coe_le_coe.1 h))
1326
-
1327
- protected lemma monotone.with_top_map (hf : monotone f) : monotone (with_top.map f) :=
1328
- hf.dual.with_bot_map.dual
1329
-
1330
- protected lemma strict_mono.with_bot_map (hf : strict_mono f) : strict_mono (with_bot.map f)
1331
- | ⊥ (a : α) h := with_bot.bot_lt_coe _
1332
- | (a : α) (b : α) h := with_bot.coe_lt_coe.mpr (hf $ with_bot.coe_lt_coe.mp h)
1333
-
1334
- protected lemma strict_mono.with_top_map (hf : strict_mono f) : strict_mono (with_top.map f) :=
1335
- hf.dual.with_bot_map.dual
1336
-
1337
- end mono
1338
-
1339
1374
/-! ### Subtype, order dual, product lattices -/
1340
1375
1341
1376
namespace subtype
0 commit comments