|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Floris van Doorn. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Floris van Doorn |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module data.prod.tprod |
| 7 | +! leanprover-community/mathlib commit 7b78d1776212a91ecc94cf601f83bdcc46b04213 |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.Data.List.Nodup |
| 12 | +/-! |
| 13 | +# Finite products of types |
| 14 | +
|
| 15 | +This file defines the product of types over a list. For `l : List ι` and `α : ι → Type v` we define |
| 16 | +`List.TProd α l = l.foldr (λ i β, α i × β) PUnit`. |
| 17 | +This type should not be used if `∀ i, α i` or `∀ i ∈ l, α i` can be used instead |
| 18 | +(in the last expression, we could also replace the list `l` by a set or a finset). |
| 19 | +This type is used as an intermediary between binary products and finitary products. |
| 20 | +The application of this type is finitary product measures, but it could be used in any |
| 21 | +construction/theorem that is easier to define/prove on binary products than on finitary products. |
| 22 | +
|
| 23 | +* Once we have the construction on binary products (like binary product measures in |
| 24 | + `MeasureTheory.prod`), we can easily define a finitary version on the type `TProd l α` |
| 25 | + by iterating. Properties can also be easily extended from the binary case to the finitary case |
| 26 | + by iterating. |
| 27 | +* Then we can use the equivalence `List.TProd.piEquivTProd` below (or enhanced versions of it, |
| 28 | + like a `MeasurableEquiv` for product measures) to get the construction on `∀ i : ι, α i`, at |
| 29 | + least when assuming `[Fintype ι] [Encodable ι]` (using `Encodable.sortedUniv`). |
| 30 | + Using `local attribute [instance] Fintype.toEncodable` we can get rid of the argument |
| 31 | + `[Encodable ι]`. |
| 32 | +
|
| 33 | +## Main definitions |
| 34 | +
|
| 35 | +* We have the equivalence `TProd.piEquivTProd : (∀ i, α i) ≃ TProd α l` |
| 36 | + if `l` contains every element of `ι` exactly once. |
| 37 | +* The product of sets is `Set.tprod : (∀ i, Set (α i)) → Set (TProd α l)`. |
| 38 | +-/ |
| 39 | + |
| 40 | + |
| 41 | +open List Function |
| 42 | +universe u v |
| 43 | +variable {ι : Type u} {α : ι → Type v} {i j : ι} {l : List ι} {f : ∀ i, α i} |
| 44 | + |
| 45 | +namespace List |
| 46 | + |
| 47 | +variable (α) |
| 48 | + |
| 49 | +/-- The product of a family of types over a list. -/ |
| 50 | +def TProd (l : List ι) : Type v := |
| 51 | + l.foldr (fun i β => α i × β) PUnit |
| 52 | +#align list.tprod List.TProd |
| 53 | + |
| 54 | +variable {α} |
| 55 | + |
| 56 | +namespace TProd |
| 57 | + |
| 58 | +open List |
| 59 | + |
| 60 | +/-- Turning a function `f : ∀ i, α i` into an element of the iterated product `TProd α l`. -/ |
| 61 | +protected def mk : ∀ (l : List ι) (_f : ∀ i, α i), TProd α l |
| 62 | + | [] => fun _ => PUnit.unit |
| 63 | + | i :: is => fun f => (f i, TProd.mk is f) |
| 64 | +#align list.tprod.mk List.TProd.mk |
| 65 | + |
| 66 | +instance [∀ i, Inhabited (α i)] : Inhabited (TProd α l) := |
| 67 | + ⟨TProd.mk l default⟩ |
| 68 | + |
| 69 | +@[simp] |
| 70 | +theorem fst_mk (i : ι) (l : List ι) (f : ∀ i, α i) : (TProd.mk (i :: l) f).1 = f i := |
| 71 | + rfl |
| 72 | +#align list.tprod.fst_mk List.TProd.fst_mk |
| 73 | + |
| 74 | +@[simp] |
| 75 | +theorem snd_mk (i : ι) (l : List ι) (f : ∀ i, α i) : |
| 76 | + (TProd.mk.{u,v} (i :: l) f).2 = TProd.mk.{u,v} l f := |
| 77 | + rfl |
| 78 | +#align list.tprod.snd_mk List.TProd.snd_mk |
| 79 | + |
| 80 | +variable [DecidableEq ι] |
| 81 | + |
| 82 | +/-- Given an element of the iterated product `l.Prod α`, take a projection into direction `i`. |
| 83 | + If `i` appears multiple times in `l`, this chooses the first component in direction `i`. -/ |
| 84 | +protected def elim : ∀ {l : List ι} (_ : TProd α l) {i : ι} (_ : i ∈ l), α i |
| 85 | + | i :: is, v, j, hj => |
| 86 | + if hji : j = i then by |
| 87 | + subst hji |
| 88 | + exact v.1 |
| 89 | + else TProd.elim v.2 ((List.mem_cons.mp hj).resolve_left hji) |
| 90 | +#align list.tprod.elim List.TProd.elim |
| 91 | + |
| 92 | +@[simp] |
| 93 | +theorem elim_self (v : TProd α (i :: l)) : v.elim (l.mem_cons_self i) = v.1 := by simp [TProd.elim] |
| 94 | +#align list.tprod.elim_self List.TProd.elim_self |
| 95 | + |
| 96 | +@[simp] |
| 97 | +theorem elim_of_ne (hj : j ∈ i :: l) (hji : j ≠ i) (v : TProd α (i :: l)) : |
| 98 | + v.elim hj = TProd.elim v.2 ((List.mem_cons.mp hj).resolve_left hji) := by simp [TProd.elim, hji] |
| 99 | +#align list.tprod.elim_of_ne List.TProd.elim_of_ne |
| 100 | + |
| 101 | +@[simp] |
| 102 | +theorem elim_of_mem (hl : (i :: l).Nodup) (hj : j ∈ l) (v : TProd α (i :: l)) : |
| 103 | + v.elim (mem_cons_of_mem _ hj) = TProd.elim v.2 hj := |
| 104 | + by |
| 105 | + apply elim_of_ne |
| 106 | + rintro rfl |
| 107 | + exact hl.not_mem hj |
| 108 | +#align list.tprod.elim_of_mem List.TProd.elim_of_mem |
| 109 | + |
| 110 | +theorem elim_mk : ∀ (l : List ι) (f : ∀ i, α i) {i : ι} (hi : i ∈ l), (TProd.mk l f).elim hi = f i |
| 111 | + | i :: is, f, j, hj => by |
| 112 | + by_cases hji : j = i |
| 113 | + · subst hji |
| 114 | + simp |
| 115 | + · rw [TProd.elim_of_ne _ hji, snd_mk, elim_mk is] |
| 116 | + termination_by elim_mk l f j hj => l.length |
| 117 | +#align list.tprod.elim_mk List.TProd.elim_mk |
| 118 | + |
| 119 | +@[ext] |
| 120 | +theorem ext : |
| 121 | + ∀ {l : List ι} (_ : l.Nodup) {v w : TProd α l} |
| 122 | + (_ : ∀ (i) (hi : i ∈ l), v.elim hi = w.elim hi), v = w |
| 123 | + | [], _, v, w, _ => PUnit.ext v w |
| 124 | + | i :: is, hl, v, w, hvw => by |
| 125 | + apply Prod.ext; rw [← elim_self v, hvw, elim_self] |
| 126 | + refine' ext (nodup_cons.mp hl).2 fun j hj => _ |
| 127 | + rw [← elim_of_mem hl, hvw, elim_of_mem hl] |
| 128 | +#align list.tprod.ext List.TProd.ext |
| 129 | + |
| 130 | +/-- A version of `TProd.elim` when `l` contains all elements. In this case we get a function into |
| 131 | + `Π i, α i`. -/ |
| 132 | +@[simp] |
| 133 | +protected def elim' (h : ∀ i, i ∈ l) (v : TProd α l) (i : ι) : α i := |
| 134 | + v.elim (h i) |
| 135 | +#align list.tprod.elim' List.TProd.elim' |
| 136 | + |
| 137 | +theorem mk_elim (hnd : l.Nodup) (h : ∀ i, i ∈ l) (v : TProd α l) : TProd.mk l (v.elim' h) = v := |
| 138 | + TProd.ext hnd fun i hi => by simp [elim_mk] |
| 139 | +#align list.tprod.mk_elim List.TProd.mk_elim |
| 140 | + |
| 141 | +/-- Pi-types are equivalent to iterated products. -/ |
| 142 | +def piEquivTProd (hnd : l.Nodup) (h : ∀ i, i ∈ l) : (∀ i, α i) ≃ TProd α l := |
| 143 | + ⟨TProd.mk l, TProd.elim' h, fun f => funext fun i => elim_mk l f (h i), mk_elim hnd h⟩ |
| 144 | +#align list.tprod.pi_equiv_tprod List.TProd.piEquivTProd |
| 145 | + |
| 146 | +end TProd |
| 147 | + |
| 148 | +end List |
| 149 | + |
| 150 | +namespace Set |
| 151 | + |
| 152 | +open List |
| 153 | + |
| 154 | +/-- A product of sets in `TProd α l`. -/ |
| 155 | +@[simp] |
| 156 | +protected def tprod : ∀ (l : List ι) (_t : ∀ i, Set (α i)), Set (TProd α l) |
| 157 | + | [], _ => univ |
| 158 | + | i :: is, t => t i ×ˢ Set.tprod is t |
| 159 | +#align set.tprod Set.tprod |
| 160 | + |
| 161 | +theorem mk_preimage_tprod : |
| 162 | + ∀ (l : List ι) (t : ∀ i, Set (α i)), TProd.mk l ⁻¹' Set.tprod l t = { i | i ∈ l }.pi t |
| 163 | + | [], t => by simp [Set.tprod] |
| 164 | + | i :: l, t => by |
| 165 | + ext f |
| 166 | + have h : TProd.mk l f ∈ Set.tprod l t ↔ ∀ i : ι, i ∈ l → f i ∈ t i := by |
| 167 | + change f ∈ TProd.mk l ⁻¹' Set.tprod l t ↔ f ∈ { x | x ∈ l }.pi t |
| 168 | + rw [mk_preimage_tprod l t] |
| 169 | + |
| 170 | + -- `simp [Set.TProd, TProd.mk, this]` can close this goal but is slow. |
| 171 | + rw [Set.tprod, TProd.mk, mem_preimage, mem_pi, prod_mk_mem_set_prod_eq] |
| 172 | + simp_rw [mem_setOf_eq, mem_cons] |
| 173 | + rw [forall_eq_or_imp, and_congr_right_iff] |
| 174 | + exact fun _ => h |
| 175 | +#align set.mk_preimage_tprod Set.mk_preimage_tprod |
| 176 | + |
| 177 | +theorem elim_preimage_pi [DecidableEq ι] {l : List ι} (hnd : l.Nodup) (h : ∀ i, i ∈ l) |
| 178 | + (t : ∀ i, Set (α i)) : TProd.elim' h ⁻¹' pi univ t = Set.tprod l t := |
| 179 | + by |
| 180 | + have h2 : { i | i ∈ l } = univ := by |
| 181 | + ext i |
| 182 | + simp [h] |
| 183 | + rw [← h2, ← mk_preimage_tprod, preimage_preimage] |
| 184 | + simp only [TProd.mk_elim hnd h] |
| 185 | + dsimp; rfl |
| 186 | +#align set.elim_preimage_pi Set.elim_preimage_pi |
| 187 | + |
| 188 | +end Set |
0 commit comments