@@ -715,6 +715,40 @@ theorem Iic_inter_Ioc_of_le (h : a ≤ c) : Iic a ∩ Ioc b c = Ioc b a :=
715
715
ext fun _ => ⟨fun H => ⟨H.2 .1 , H.1 ⟩, fun H => ⟨H.2 , H.1 , H.2 .trans h⟩⟩
716
716
#align set.Iic_inter_Ioc_of_le Set.Iic_inter_Ioc_of_le
717
717
718
+ theorem not_mem_Icc_of_lt (ha : c < a) : c ∉ Icc a b := fun h => ha.not_le h.1
719
+ #align set.not_mem_Icc_of_lt Set.not_mem_Icc_of_lt
720
+
721
+ theorem not_mem_Icc_of_gt (hb : b < c) : c ∉ Icc a b := fun h => hb.not_le h.2
722
+ #align set.not_mem_Icc_of_gt Set.not_mem_Icc_of_gt
723
+
724
+ theorem not_mem_Ico_of_lt (ha : c < a) : c ∉ Ico a b := fun h => ha.not_le h.1
725
+ #align set.not_mem_Ico_of_lt Set.not_mem_Ico_of_lt
726
+
727
+ theorem not_mem_Ioc_of_gt (hb : b < c) : c ∉ Ioc a b := fun h => hb.not_le h.2
728
+ #align set.not_mem_Ioc_of_gt Set.not_mem_Ioc_of_gt
729
+
730
+ -- Porting note: `simp` can prove this
731
+ -- @[ simp ]
732
+ theorem not_mem_Ioi_self : a ∉ Ioi a := lt_irrefl _
733
+ #align set.not_mem_Ioi_self Set.not_mem_Ioi_self
734
+
735
+ -- Porting note: `simp` can prove this
736
+ -- @[ simp ]
737
+ theorem not_mem_Iio_self : b ∉ Iio b := lt_irrefl _
738
+ #align set.not_mem_Iio_self Set.not_mem_Iio_self
739
+
740
+ theorem not_mem_Ioc_of_le (ha : c ≤ a) : c ∉ Ioc a b := fun h => lt_irrefl _ <| h.1 .trans_le ha
741
+ #align set.not_mem_Ioc_of_le Set.not_mem_Ioc_of_le
742
+
743
+ theorem not_mem_Ico_of_ge (hb : b ≤ c) : c ∉ Ico a b := fun h => lt_irrefl _ <| h.2 .trans_le hb
744
+ #align set.not_mem_Ico_of_ge Set.not_mem_Ico_of_ge
745
+
746
+ theorem not_mem_Ioo_of_le (ha : c ≤ a) : c ∉ Ioo a b := fun h => lt_irrefl _ <| h.1 .trans_le ha
747
+ #align set.not_mem_Ioo_of_le Set.not_mem_Ioo_of_le
748
+
749
+ theorem not_mem_Ioo_of_ge (hb : b ≤ c) : c ∉ Ioo a b := fun h => lt_irrefl _ <| h.2 .trans_le hb
750
+ #align set.not_mem_Ioo_of_ge Set.not_mem_Ioo_of_ge
751
+
718
752
end Preorder
719
753
720
754
section PartialOrder
@@ -1014,22 +1048,6 @@ theorem not_mem_Iic : c ∉ Iic b ↔ b < c :=
1014
1048
not_le
1015
1049
#align set.not_mem_Iic Set.not_mem_Iic
1016
1050
1017
- theorem not_mem_Icc_of_lt (ha : c < a) : c ∉ Icc a b :=
1018
- not_mem_subset Icc_subset_Ici_self <| not_mem_Ici.mpr ha
1019
- #align set.not_mem_Icc_of_lt Set.not_mem_Icc_of_lt
1020
-
1021
- theorem not_mem_Icc_of_gt (hb : b < c) : c ∉ Icc a b :=
1022
- not_mem_subset Icc_subset_Iic_self <| not_mem_Iic.mpr hb
1023
- #align set.not_mem_Icc_of_gt Set.not_mem_Icc_of_gt
1024
-
1025
- theorem not_mem_Ico_of_lt (ha : c < a) : c ∉ Ico a b :=
1026
- not_mem_subset Ico_subset_Ici_self <| not_mem_Ici.mpr ha
1027
- #align set.not_mem_Ico_of_lt Set.not_mem_Ico_of_lt
1028
-
1029
- theorem not_mem_Ioc_of_gt (hb : b < c) : c ∉ Ioc a b :=
1030
- not_mem_subset Ioc_subset_Iic_self <| not_mem_Iic.mpr hb
1031
- #align set.not_mem_Ioc_of_gt Set.not_mem_Ioc_of_gt
1032
-
1033
1051
theorem not_mem_Ioi : c ∉ Ioi a ↔ c ≤ a :=
1034
1052
not_lt
1035
1053
#align set.not_mem_Ioi Set.not_mem_Ioi
@@ -1038,34 +1056,6 @@ theorem not_mem_Iio : c ∉ Iio b ↔ b ≤ c :=
1038
1056
not_lt
1039
1057
#align set.not_mem_Iio Set.not_mem_Iio
1040
1058
1041
- -- Porting note: `simp` can prove this
1042
- -- @[ simp ]
1043
- theorem not_mem_Ioi_self : a ∉ Ioi a :=
1044
- lt_irrefl _
1045
- #align set.not_mem_Ioi_self Set.not_mem_Ioi_self
1046
-
1047
- -- Porting note: `simp` can prove this
1048
- -- @[ simp ]
1049
- theorem not_mem_Iio_self : b ∉ Iio b :=
1050
- lt_irrefl _
1051
- #align set.not_mem_Iio_self Set.not_mem_Iio_self
1052
-
1053
- theorem not_mem_Ioc_of_le (ha : c ≤ a) : c ∉ Ioc a b :=
1054
- not_mem_subset Ioc_subset_Ioi_self <| not_mem_Ioi.mpr ha
1055
- #align set.not_mem_Ioc_of_le Set.not_mem_Ioc_of_le
1056
-
1057
- theorem not_mem_Ico_of_ge (hb : b ≤ c) : c ∉ Ico a b :=
1058
- not_mem_subset Ico_subset_Iio_self <| not_mem_Iio.mpr hb
1059
- #align set.not_mem_Ico_of_ge Set.not_mem_Ico_of_ge
1060
-
1061
- theorem not_mem_Ioo_of_le (ha : c ≤ a) : c ∉ Ioo a b :=
1062
- not_mem_subset Ioo_subset_Ioi_self <| not_mem_Ioi.mpr ha
1063
- #align set.not_mem_Ioo_of_le Set.not_mem_Ioo_of_le
1064
-
1065
- theorem not_mem_Ioo_of_ge (hb : b ≤ c) : c ∉ Ioo a b :=
1066
- not_mem_subset Ioo_subset_Iio_self <| not_mem_Iio.mpr hb
1067
- #align set.not_mem_Ioo_of_ge Set.not_mem_Ioo_of_ge
1068
-
1069
1059
@[simp]
1070
1060
theorem compl_Iic : Iic aᶜ = Ioi a :=
1071
1061
ext fun _ => not_le
0 commit comments