Skip to content

Commit 9fe7218

Browse files
committed
1 parent e46fa51 commit 9fe7218

File tree

1 file changed

+34
-44
lines changed

1 file changed

+34
-44
lines changed

Mathlib/Data/Set/Intervals/Basic.lean

Lines changed: 34 additions & 44 deletions
Original file line numberDiff line numberDiff line change
@@ -715,6 +715,40 @@ theorem Iic_inter_Ioc_of_le (h : a ≤ c) : Iic a ∩ Ioc b c = Ioc b a :=
715715
ext fun _ => ⟨fun H => ⟨H.2.1, H.1⟩, fun H => ⟨H.2, H.1, H.2.trans h⟩⟩
716716
#align set.Iic_inter_Ioc_of_le Set.Iic_inter_Ioc_of_le
717717

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+
718752
end Preorder
719753

720754
section PartialOrder
@@ -1014,22 +1048,6 @@ theorem not_mem_Iic : c ∉ Iic b ↔ b < c :=
10141048
not_le
10151049
#align set.not_mem_Iic Set.not_mem_Iic
10161050

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-
10331051
theorem not_mem_Ioi : c ∉ Ioi a ↔ c ≤ a :=
10341052
not_lt
10351053
#align set.not_mem_Ioi Set.not_mem_Ioi
@@ -1038,34 +1056,6 @@ theorem not_mem_Iio : c ∉ Iio b ↔ b ≤ c :=
10381056
not_lt
10391057
#align set.not_mem_Iio Set.not_mem_Iio
10401058

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-
10691059
@[simp]
10701060
theorem compl_Iic : Iic aᶜ = Ioi a :=
10711061
ext fun _ => not_le

0 commit comments

Comments
 (0)