Skip to content

Commit

Permalink
feat(data/set/basic): prod_subset_iff (#980)
Browse files Browse the repository at this point in the history
* feat(data/set/basic): prod_subset_iff

* syntax
  • Loading branch information
semorrison authored and mergify[bot] committed May 5, 2019
1 parent fbce6e4 commit fc8b08b
Showing 1 changed file with 5 additions and 0 deletions.
5 changes: 5 additions & 0 deletions src/data/set/basic.lean
Expand Up @@ -1254,6 +1254,11 @@ theorem mem_prod_eq {p : α × β} : p ∈ set.prod s t = (p.1 ∈ s ∧ p.2 ∈

lemma mk_mem_prod {a : α} {b : β} (a_in : a ∈ s) (b_in : b ∈ t) : (a, b) ∈ set.prod s t := ⟨a_in, b_in⟩

lemma prod_subset_iff {P : set (α × β)} :
(set.prod s t ⊆ P) ↔ ∀ (x ∈ s) (y ∈ t), (x, y) ∈ P :=
⟨λ h _ xin _ yin, h (mk_mem_prod xin yin),
λ h _ pin, by { cases mem_prod.1 pin with hs ht, simpa using h _ hs _ ht }⟩

@[simp] theorem prod_empty {s : set α} : set.prod s ∅ = (∅ : set (α × β)) :=
ext $ by simp [set.prod]

Expand Down

0 comments on commit fc8b08b

Please sign in to comment.