Skip to content

Commit

Permalink
chore(*): a few facts about pprod (#10519)
Browse files Browse the repository at this point in the history
Add `equiv.pprod_equiv_prod` and `function.embedding.pprod_map`.
  • Loading branch information
urkud committed Nov 29, 2021
1 parent be48f95 commit 5601833
Show file tree
Hide file tree
Showing 3 changed files with 23 additions and 1 deletion.
8 changes: 8 additions & 0 deletions src/data/equiv/basic.lean
Expand Up @@ -408,6 +408,14 @@ protected def ulift {α : Type v} : ulift.{u} α ≃ α :=
protected def plift : plift α ≃ α :=
⟨plift.down, plift.up, plift.up_down, plift.down_up⟩

/-- `pprod α β` is equivalent to `α × β` -/
@[simps apply symm_apply]
def pprod_equiv_prod {α β : Type*} : pprod α β ≃ α × β :=
{ to_fun := λ x, (x.1, x.2),
inv_fun := λ x, ⟨x.1, x.2⟩,
left_inv := λ ⟨x, y⟩, rfl,
right_inv := λ ⟨x, y⟩, rfl }

/-- equivalence of propositions is the same as iff -/
def of_iff {P Q : Prop} (h : P ↔ Q) : P ≃ Q :=
{ to_fun := h.mp,
Expand Down
11 changes: 10 additions & 1 deletion src/data/pprod.lean
Expand Up @@ -3,12 +3,14 @@ Copyright (c) 2020 Eric Wieser. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Eric Wieser
-/
import logic.basic

/-!
# Extra facts about `pprod`
-/

variables {α : Sort*} {β : Sort*}
open function
variables {α β γ δ : Sort*}

namespace pprod

Expand All @@ -28,3 +30,10 @@ theorem exists' {p : α → β → Prop} : (∃ x : pprod α β, p x.1 x.2) ↔
pprod.exists

end pprod

lemma function.injective.pprod_map {f : α → β} {g : γ → δ} (hf : injective f) (hg : injective g) :
injective (λ x, ⟨f x.1, g x.2⟩ : pprod α γ → pprod β δ) :=
λ ⟨x₁, x₂⟩ ⟨y₁, y₂⟩ h,
have A : _ := congr_arg pprod.fst h,
have B : _ := congr_arg pprod.snd h,
congr_arg2 pprod.mk (hf A) (hg B)
5 changes: 5 additions & 0 deletions src/logic/embedding.lean
Expand Up @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro
import data.equiv.basic
import data.set.basic
import data.sigma.basic
import data.pprod

/-!
# Injective functions
Expand Down Expand Up @@ -199,6 +200,10 @@ def prod_map {α β γ δ : Type*} (e₁ : α ↪ β) (e₂ : γ ↪ δ) : α ×
⇑(e₁.prod_map e₂) = prod.map e₁ e₂ :=
rfl

/-- If `e₁` and `e₂` are embeddings, then so is `λ ⟨a, b⟩, ⟨e₁ a, e₂ b⟩ : pprod α γ → pprod β δ`. -/
def pprod_map {α β γ δ : Sort*} (e₁ : α ↪ β) (e₂ : γ ↪ δ) : pprod α γ ↪ pprod β δ :=
⟨λ x, ⟨e₁ x.1, e₂ x.2⟩, e₁.injective.pprod_map e₂.injective⟩

section sum
open sum

Expand Down

0 comments on commit 5601833

Please sign in to comment.