choose
tactic #
Performs Skolemization, that is, given h : ∀ a:α, ∃ b:β, p a b |- G
produces
f : α → β, hf: ∀ a, p a (f a) |- G
.
choose a b h h' using hyp
takes an hypothesis hyp
of the form
∀ (x : X) (y : Y), ∃ (a : A) (b : B), P x y a b ∧ Q x y a b
for some P Q : X → Y → A → B → Prop
and outputs
into context two functions a : X → Y → A
, b : X → Y → B
and two assumptions:
h : ∀ (x : X) (y : Y), P x y (a x y) (b x y)
and
h' : ∀ (x : X) (y : Y), Q x y (a x y) (b x y)
. It also works with dependent versions.
choose! a b h h' using hyp
does the same, except that it will remove dependency of
the functions on propositional arguments if possible. For example if Y
is a proposition
and A
and B
are nonempty in the above example then we will instead get
a : X → A
, b : X → B
, and the assumptions
h : ∀ (x : X) (y : Y), P x y (a x) (b x)
and
h' : ∀ (x : X) (y : Y), Q x y (a x) (b x)
.
Examples: