Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 211b0c0

Browse files
committed
feat(logic/basic): forall2_congr lemmas (#4904)
Some helpful lemmas for working with quantifiers, just other versions of what's already there.
1 parent 0081a5a commit 211b0c0

File tree

1 file changed

+28
-0
lines changed

1 file changed

+28
-0
lines changed

src/logic/basic.lean

Lines changed: 28 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -635,12 +635,40 @@ variables {α : Sort*} {β : Sort*} {p q : α → Prop} {b : Prop}
635635
lemma forall_imp (h : ∀ a, p a → q a) : (∀ a, p a) → ∀ a, q a :=
636636
λ h' a, h a (h' a)
637637

638+
lemma forall₂_congr {p q : α → β → Prop} (h : ∀ a b, p a b ↔ q a b) :
639+
(∀ a b, p a b) ↔ (∀ a b, q a b) :=
640+
forall_congr (λ a, forall_congr (h a))
641+
642+
lemma forall₃_congr {γ : Sort*} {p q : α → β → γ → Prop}
643+
(h : ∀ a b c, p a b c ↔ q a b c) :
644+
(∀ a b c, p a b c) ↔ (∀ a b c, q a b c) :=
645+
forall_congr (λ a, forall₂_congr (h a))
646+
647+
lemma forall₄_congr {γ δ : Sort*} {p q : α → β → γ → δ → Prop}
648+
(h : ∀ a b c d, p a b c d ↔ q a b c d) :
649+
(∀ a b c d, p a b c d) ↔ (∀ a b c d, q a b c d) :=
650+
forall_congr (λ a, forall₃_congr (h a))
651+
638652
lemma Exists.imp (h : ∀ a, (p a → q a)) (p : ∃ a, p a) : ∃ a, q a := exists_imp_exists h p
639653

640654
lemma exists_imp_exists' {p : α → Prop} {q : β → Prop} (f : α → β) (hpq : ∀ a, p a → q (f a))
641655
(hp : ∃ a, p a) : ∃ b, q b :=
642656
exists.elim hp (λ a hp', ⟨_, hpq _ hp'⟩)
643657

658+
lemma exists₂_congr {p q : α → β → Prop} (h : ∀ a b, p a b ↔ q a b) :
659+
(∃ a b, p a b) ↔ (∃ a b, q a b) :=
660+
exists_congr (λ a, exists_congr (h a))
661+
662+
lemma exists₃_congr {γ : Sort*} {p q : α → β → γ → Prop}
663+
(h : ∀ a b c, p a b c ↔ q a b c) :
664+
(∃ a b c, p a b c) ↔ (∃ a b c, q a b c) :=
665+
exists_congr (λ a, exists₂_congr (h a))
666+
667+
lemma exists₄_congr {γ δ : Sort*} {p q : α → β → γ → δ → Prop}
668+
(h : ∀ a b c d, p a b c d ↔ q a b c d) :
669+
(∃ a b c d, p a b c d) ↔ (∃ a b c d, q a b c d) :=
670+
exists_congr (λ a, exists₃_congr (h a))
671+
644672
theorem forall_swap {p : α → β → Prop} : (∀ x y, p x y) ↔ ∀ y x, p x y :=
645673
⟨function.swap, function.swap⟩
646674

0 commit comments

Comments
 (0)