feat(logic/quantifiers): tag congruence theorems

This commit is contained in:
Daniel Selsam 2015-11-04 21:56:35 -08:00 committed by Leonardo de Moura
parent 8e5e8e6540
commit 34f4e315ee

View file

@ -126,14 +126,17 @@ end
section
variables {A : Type} {p₁ p₂ : A → Prop} (H : ∀ x, p₁ x ↔ p₂ x)
theorem congr_forall : (∀ x, p₁ x) ↔ (∀ x, p₂ x) :=
theorem congr_forall [congr] : (∀ x, p₁ x) ↔ (∀ x, p₂ x) :=
forall_congr H
theorem congr_exists : (∃ x, p₁ x) ↔ (∃ x, p₂ x) :=
theorem congr_exists [congr] : (∃ x, p₁ x) ↔ (∃ x, p₂ x) :=
exists_congr H
include H
theorem congr_exists_unique : (∃! x, p₁ x) ↔ (∃! x, p₂ x) :=
theorem congr_exists_unique [congr] : (∃! x, p₁ x) ↔ (∃! x, p₂ x) :=
congr_exists (λx, congr_and (H x) (congr_forall
(λy, congr_imp (H y) iff.rfl)))
end
/- We add this attribute again here so that the simplifier finds it before [congr_forall] -/
attribute congr_imp [congr]