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

Commit aa1fa0b

Browse files
committed
feat(data/option/basic): option valued choice operator (#7101)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 6610e8f commit aa1fa0b

File tree

1 file changed

+36
-0
lines changed

1 file changed

+36
-0
lines changed

src/data/option/basic.lean

Lines changed: 36 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -384,4 +384,40 @@ def cases_on' : option α → β → (α → β) → β
384384
cases_on' o (f none) (f ∘ coe) = f o :=
385385
by cases o; refl
386386

387+
section
388+
open_locale classical
389+
390+
/-- An arbitrary `some a` with `a : α` if `α` is nonempty, and otherwise `none`. -/
391+
noncomputable def choice (α : Type*) : option α :=
392+
if h : nonempty α then
393+
some h.some
394+
else
395+
none
396+
397+
lemma choice_eq {α : Type*} [subsingleton α] (a : α) : choice α = some a :=
398+
begin
399+
dsimp [choice],
400+
rw dif_pos (⟨a⟩ : nonempty α),
401+
congr,
402+
end
403+
404+
lemma choice_eq_none {α : Type*} (h : α → false) : choice α = none :=
405+
begin
406+
dsimp [choice],
407+
rw dif_neg (not_nonempty_iff_imp_false.mpr h),
408+
end
409+
410+
lemma choice_is_some_iff_nonempty {α : Type*} : (choice α).is_some ↔ nonempty α :=
411+
begin
412+
fsplit,
413+
{ intro h, exact ⟨option.get h⟩, },
414+
{ rintro ⟨a⟩,
415+
dsimp [choice],
416+
rw dif_pos,
417+
fsplit,
418+
exact ⟨a⟩, },
419+
end
420+
421+
end
422+
387423
end option

0 commit comments

Comments
 (0)