feat(library/logic/quantifiers): add 'the'
This commit is contained in:
parent
3035dd7e66
commit
20f6b4c6bd
1 changed files with 56 additions and 21 deletions
|
@ -83,6 +83,41 @@ theorem exists_unique.elim {A : Type} {p : A → Prop} {b : Prop}
|
||||||
obtain w Hw, from H2,
|
obtain w Hw, from H2,
|
||||||
H1 w (and.left Hw) (and.right Hw)
|
H1 w (and.left Hw) (and.right Hw)
|
||||||
|
|
||||||
|
theorem exists_unique_of_exists_of_unique {A : Type} {p : A → Prop}
|
||||||
|
(Hex : ∃ x, p x) (Hunique : ∀ y₁ y₂, p y₁ → p y₂ → y₁ = y₂) :
|
||||||
|
∃! x, p x :=
|
||||||
|
obtain x px, from Hex,
|
||||||
|
exists_unique.intro x px (take y, suppose p y, show y = x, from !Hunique this px)
|
||||||
|
|
||||||
|
theorem exists_of_exists_unique {A : Type} {p : A → Prop} (H : ∃! x, p x) :
|
||||||
|
∃ x, p x :=
|
||||||
|
obtain x Hx, from H,
|
||||||
|
exists.intro x (and.left Hx)
|
||||||
|
|
||||||
|
theorem unique_of_exists_unique {A : Type} {p : A → Prop}
|
||||||
|
(H : ∃! x, p x) {y₁ y₂ : A} (py₁ : p y₁) (py₂ : p y₂) :
|
||||||
|
y₁ = y₂ :=
|
||||||
|
exists_unique.elim H
|
||||||
|
(take x, suppose p x,
|
||||||
|
assume unique : ∀ y, p y → y = x,
|
||||||
|
show y₁ = y₂, from eq.trans (unique _ py₁) (eq.symm (unique _ py₂)))
|
||||||
|
|
||||||
|
/- definite description -/
|
||||||
|
|
||||||
|
section
|
||||||
|
open classical
|
||||||
|
|
||||||
|
noncomputable definition the {A : Type} {p : A → Prop} (H : ∃! x, p x) : A :=
|
||||||
|
some (exists_of_exists_unique H)
|
||||||
|
|
||||||
|
theorem the_spec {A : Type} {p : A → Prop} (H : ∃! x, p x) : p (the H) :=
|
||||||
|
some_spec (exists_of_exists_unique H)
|
||||||
|
|
||||||
|
theorem eq_the {A : Type} {p : A → Prop} (H : ∃! x, p x) {y : A} (Hy : p y) :
|
||||||
|
y = the H :=
|
||||||
|
unique_of_exists_unique H Hy (the_spec H)
|
||||||
|
end
|
||||||
|
|
||||||
/- congruences -/
|
/- congruences -/
|
||||||
|
|
||||||
section
|
section
|
||||||
|
|
Loading…
Reference in a new issue