Skip to content

Commit

Permalink
feat(data/set/basic): Laws for n-ary image (#13011)
Browse files Browse the repository at this point in the history
Prove left/right commutativity, distributivity of `set.image2` in the style of `set.image2_assoc`. Also add `forall₃_imp` and `Exists₃.imp`.
  • Loading branch information
YaelDillies committed Mar 28, 2022
1 parent 9480029 commit 7fea719
Show file tree
Hide file tree
Showing 2 changed files with 65 additions and 20 deletions.
51 changes: 44 additions & 7 deletions src/data/set/basic.lean
Expand Up @@ -2328,6 +2328,15 @@ end
@[simp] lemma image2_empty_left : image2 f ∅ t = ∅ := ext $ by simp
@[simp] lemma image2_empty_right : image2 f s ∅ = ∅ := ext $ by simp

lemma nonempty.image2 : s.nonempty → t.nonempty → (image2 f s t).nonempty :=
λ ⟨a, ha⟩ ⟨b, hb⟩, ⟨_, mem_image2_of_mem ha hb⟩

@[simp] lemma image2_nonempty_iff : (image2 f s t).nonempty ↔ s.nonempty ∧ t.nonempty :=
⟨λ ⟨_, a, b, ha, hb, _⟩, ⟨⟨a, ha⟩, b, hb⟩, λ h, h.1.image2 h.2

@[simp] lemma image2_eq_empty_iff : image2 f s t = ∅ ↔ s = ∅ ∨ t = ∅ :=
by simp_rw [←not_nonempty_iff_eq_empty, image2_nonempty_iff, not_and_distrib]

lemma image2_inter_subset_left : image2 f (s ∩ s') t ⊆ image2 f s t ∩ image2 f s' t :=
by { rintro _ ⟨a, b, ⟨h1a, h2a⟩, hb, rfl⟩, split; exact ⟨_, _, ‹_›, ‹_›, rfl⟩ }

Expand Down Expand Up @@ -2360,6 +2369,9 @@ def image3 (g : α → β → γ → δ) (s : set α) (t : set β) (u : set γ)
@[simp] lemma mem_image3 : d ∈ image3 g s t u ↔ ∃ a b c, a ∈ s ∧ b ∈ t ∧ c ∈ u ∧ g a b c = d :=
iff.rfl

lemma image3_mono (hs : s ⊆ s') (ht : t ⊆ t') (hu : u ⊆ u') : image3 g s t u ⊆ image3 g s' t' u' :=
λ x, Exists₃.imp $ λ a b c ⟨ha, hb, hc, hx⟩, ⟨hs ha, ht hb, hu hc, hx⟩

@[congr] lemma image3_congr (h : ∀ (a ∈ s) (b ∈ t) (c ∈ u), g a b c = g' a b c) :
image3 g s t u = image3 g' s t u :=
by { ext x,
Expand All @@ -2385,11 +2397,6 @@ begin
{ rintro ⟨a, b, c, ha, hb, hc, rfl⟩, refine ⟨a, _, ha, ⟨b, c, hb, hc, rfl⟩, rfl⟩ }
end

lemma image2_assoc {ε'} {f : δ → γ → ε} {g : α → β → δ} {f' : α → ε' → ε} {g' : β → γ → ε'}
(h_assoc : ∀ a b c, f (g a b) c = f' a (g' b c)) :
image2 f (image2 g s t) u = image2 f' s (image2 g' t u) :=
by simp only [image2_image2_left, image2_image2_right, h_assoc]

lemma image_image2 (f : α → β → γ) (g : γ → δ) :
g '' image2 f s t = image2 (λ a b, g (f a b)) s t :=
begin
Expand Down Expand Up @@ -2424,8 +2431,38 @@ by simp [nonempty_def.mp h, ext_iff]
@[simp] lemma image2_right (h : s.nonempty) : image2 (λ x y, y) s t = t :=
by simp [nonempty_def.mp h, ext_iff]

lemma nonempty.image2 (hs : s.nonempty) (ht : t.nonempty) : (image2 f s t).nonempty :=
by { cases hs with a ha, cases ht with b hb, exact ⟨f a b, ⟨a, b, ha, hb, rfl⟩⟩ }
lemma image2_assoc {ε'} {f : δ → γ → ε} {g : α → β → δ} {f' : α → ε' → ε} {g' : β → γ → ε'}
(h_assoc : ∀ a b c, f (g a b) c = f' a (g' b c)) :
image2 f (image2 g s t) u = image2 f' s (image2 g' t u) :=
by simp only [image2_image2_left, image2_image2_right, h_assoc]

lemma image2_comm {g : β → α → γ} (h_comm : ∀ a b, f a b = g b a) : image2 f s t = image2 g t s :=
(image2_swap _ _ _).trans $ by simp_rw h_comm

lemma image2_left_comm {δ'} {f : α → δ → ε} {g : β → γ → δ} {f' : α → γ → δ'} {g' : β → δ' → ε}
(h_left_comm : ∀ a b c, f a (g b c) = g' b (f' a c)) :
image2 f s (image2 g t u) = image2 g' t (image2 f' s u) :=
by { rw [image2_swap f', image2_swap f], exact image2_assoc (λ _ _ _, h_left_comm _ _ _) }

lemma image2_right_comm {δ'} {f : δ → γ → ε} {g : α → β → δ} {f' : α → γ → δ'} {g' : δ' → β → ε}
(h_right_comm : ∀ a b c, f (g a b) c = g' (f' a c) b) :
image2 f (image2 g s t) u = image2 g' (image2 f' s u) t :=
by { rw [image2_swap g, image2_swap g'], exact image2_assoc (λ _ _ _, h_right_comm _ _ _) }

lemma image_image2_distrib {α' β'} {g : γ → δ} {f' : α' → β' → δ} {g₁ : α → α'} {g₂ : β → β'}
(h_distrib : ∀ a b, g (f a b) = f' (g₁ a) (g₂ b)) :
(image2 f s t).image g = image2 f' (s.image g₁) (t.image g₂) :=
by simp_rw [image_image2, image2_image_left, image2_image_right, h_distrib]

lemma image_image2_distrib_left {α'} {g : γ → δ} {f' : α' → β → δ} {g' : α → α'}
(h_distrib : ∀ a b, g (f a b) = f' (g' a) b) :
(image2 f s t).image g = image2 f' (s.image g') t :=
(image_image2_distrib h_distrib).trans $ by rw image_id'

lemma image_image2_distrib_right {β'} {g : γ → δ} {f' : α → β' → δ} {g' : β → β'}
(h_distrib : ∀ a b, g (f a b) = f' a (g' b)) :
(image2 f s t).image g = image2 f' s (t.image g') :=
(image_image2_distrib h_distrib).trans $ by rw image_id'

end n_ary_image

Expand Down
34 changes: 21 additions & 13 deletions src/logic/basic.lean
Expand Up @@ -878,7 +878,7 @@ end equality
section quantifiers
variables {α : Sort*}

section congr
section dependent
variables {β : α → Sort*} {γ : Π a, β a → Sort*} {δ : Π a b, γ a b → Sort*}
{ε : Π a b c, δ a b c → Sort*}

Expand All @@ -899,7 +899,7 @@ lemma forall₅_congr {p q : Π a b c d, ε a b c d → Prop}
(∀ a b c d e, p a b c d e) ↔ ∀ a b c d e, q a b c d e :=
forall_congr $ λ a, forall₄_congr $ h a

lemma exists₂_congr {p q : Π a, β a → Prop} (h : ∀ a b, p a b ↔ q a b) :
lemma exists₂_congr {p q : Π a, β a → Prop} (h : ∀ a b, p a b ↔ q a b) :
(∃ a b, p a b) ↔ ∃ a b, q a b :=
exists_congr $ λ a, exists_congr $ h a

Expand All @@ -916,22 +916,30 @@ lemma exists₅_congr {p q : Π a b c d, ε a b c d → Prop}
(∃ a b c d e, p a b c d e) ↔ ∃ a b c d e, q a b c d e :=
exists_congr $ λ a, exists₄_congr $ h a

end congr
lemma forall_imp {p q : α → Prop} (h : ∀ a, p a → q a) : (∀ a, p a) → ∀ a, q a := λ h' a, h a (h' a)

variables {ι β : Sort*} {κ : ι → Sort*} {p q : α → Prop} {b : Prop}
lemma forall₂_imp {p q : Π a, β a → Prop} (h : ∀ a b, p a b → q a b) :
(∀ a b, p a b) → ∀ a b, q a b :=
forall_imp $ λ i, forall_imp $ h i

lemma forall_imp (h : ∀ a, p a → q a) : (∀ a, p a) → ∀ a, q a :=
λ h' a, h a (h' a)
lemma forall₃_imp {p q : Π a b, γ a b → Prop} (h : ∀ a b c, p a b c → q a b c) :
(∀ a b c, p a b c) → ∀ a b c, q a b c :=
forall_imp $ λ a, forall₂_imp $ h a

lemma forall₂_imp {p q : Π i, κ i → Prop} (h : ∀ i j, p i j → q i j) :
(∀ i j, p i j) → ∀ i j, q i j :=
forall_imp $ λ i, forall_imp $ h i
lemma Exists.imp {p q : α → Prop} (h : ∀ a, (p a → q a)) : (∃ a, p a) → ∃ a, q a :=
exists_imp_exists h

lemma Exists.imp (h : ∀ a, (p a → q a)) (p : ∃ a, p a) : ∃ a, q a := exists_imp_exists h p
lemma Exists₂.imp {p q : Π a, β a → Prop} (h : ∀ a b, p a b → q a b) :
(∃ a b, p a b) → ∃ a b, q a b :=
Exists.imp $ λ a, Exists.imp $ h a

lemma Exists₂.imp {p q : Π i, κ i → Prop} (h : ∀ i j, p i j → q i j) :
(∃ i j, p i j) → ∃ i j, q i j :=
Exists.imp $ λ i, Exists.imp $ h i
lemma Exists₃.imp {p q : Π a b, γ a b → Prop} (h : ∀ a b c, p a b c → q a b c) :
(∃ a b c, p a b c) → ∃ a b c, q a b c :=
Exists.imp $ λ a, Exists₂.imp $ h a

end dependent

variables {ι β : Sort*} {κ : ι → Sort*} {p q : α → Prop} {b : Prop}

lemma exists_imp_exists' {p : α → Prop} {q : β → Prop} (f : α → β) (hpq : ∀ a, p a → q (f a))
(hp : ∃ a, p a) : ∃ b, q b :=
Expand Down

0 comments on commit 7fea719

Please sign in to comment.