Skip to content

Commit

Permalink
refactor(Probability/Kernel/CondCdf): mv prod_iInter (#10037)
Browse files Browse the repository at this point in the history
Co-authored-by: Moritz Firsching <firsching@google.com>
  • Loading branch information
mo271 and mo271 committed Jan 27, 2024
1 parent 62acece commit cb8ebaf
Show file tree
Hide file tree
Showing 2 changed files with 8 additions and 8 deletions.
7 changes: 7 additions & 0 deletions Mathlib/Data/Set/Lattice.lean
Expand Up @@ -1989,6 +1989,13 @@ theorem prod_sInter {T : Set (Set β)} (hT : T.Nonempty) (s : Set α) :
simp_rw [singleton_prod, mem_image, iInter_exists, biInter_and', iInter_iInter_eq_right]
#align set.prod_sInter Set.prod_sInter

theorem prod_iInter {s : Set α} {t : ι → Set β} [hι : Nonempty ι] :
(s ×ˢ ⋂ i, t i) = ⋂ i, s ×ˢ t i := by
ext x
simp only [mem_prod, mem_iInter]
exact ⟨fun h i => ⟨h.1, h.2 i⟩, fun h => ⟨(h hι.some).1, fun i => (h i).2⟩⟩
#align prod_Inter Set.prod_iInter

end Prod

section Image2
Expand Down
9 changes: 1 addition & 8 deletions Mathlib/Probability/Kernel/CondCdf.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Rémy Degenne
-/
import Mathlib.Data.Set.Lattice
import Mathlib.MeasureTheory.Measure.Stieltjes
import Mathlib.MeasureTheory.Decomposition.RadonNikodym
import Mathlib.MeasureTheory.Constructions.Prod.Basic
Expand Down Expand Up @@ -66,14 +67,6 @@ theorem sequence_le (a : α) : f (hf.sequence f (Encodable.encode a + 1)) ≤ f

end Directed

-- todo: move to data/set/lattice next to prod_sUnion or prod_sInter
theorem prod_iInter {s : Set α} {t : ι → Set β} [hι : Nonempty ι] :
(s ×ˢ ⋂ i, t i) = ⋂ i, s ×ˢ t i := by
ext x
simp only [mem_prod, mem_iInter]
exact ⟨fun h i => ⟨h.1, h.2 i⟩, fun h => ⟨(h hι.some).1, fun i => (h i).2⟩⟩
#align prod_Inter prod_iInter

theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by
ext1 x
simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff]
Expand Down

0 comments on commit cb8ebaf

Please sign in to comment.