feat(library/standard): add hilbert's choice operator
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
a31457efde
commit
9cb238534d
1 changed files with 22 additions and 0 deletions
22
library/standard/hilbert.lean
Normal file
22
library/standard/hilbert.lean
Normal file
|
@ -0,0 +1,22 @@
|
|||
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
-- Released under Apache 2.0 license as described in the file LICENSE.
|
||||
-- Author: Leonardo de Moura
|
||||
import logic
|
||||
|
||||
variable epsilon {A : Type} {H : inhabited A} (P : A → Bool) : A
|
||||
axiom epsilon_ax {A : Type} {P : A → Bool} (Hex : ∃ a, P a) : P (@epsilon A (inhabited_exists Hex) P)
|
||||
|
||||
theorem epsilon_singleton {A : Type} (a : A) : @epsilon A (inhabited_intro a) (λ x, x = a) = a
|
||||
:= epsilon_ax (exists_intro a (refl a))
|
||||
|
||||
theorem axiom_of_choice {A : Type} {B : A → Type} {R : Π x, B x → Bool} (H : ∀ x, ∃ y, R x y) : ∃ f, ∀ x, R x (f x)
|
||||
:= let f [inline] := λ x, @epsilon _ (inhabited_exists (H x)) (λ y, R x y),
|
||||
H [inline] := take x, epsilon_ax (H x)
|
||||
in exists_intro f H
|
||||
|
||||
theorem skolem {A : Type} {B : A → Type} {P : Π x, B x → Bool} : (∀ x, ∃ y, P x y) ↔ ∃ f, (∀ x, P x (f x))
|
||||
:= iff_intro
|
||||
(assume H : (∀ x, ∃ y, P x y), axiom_of_choice H)
|
||||
(assume H : (∃ f, (∀ x, P x (f x))),
|
||||
take x, obtain (fw : ∀ x, B x) (Hw : ∀ x, P x (fw x)), from H,
|
||||
exists_intro (fw x) (Hw x))
|
Loading…
Reference in a new issue