Documentation

Mathlib.Data.Option.Defs

Extra definitions on Option #

This file defines more operations involving Option α. Lemmas about them are located in other files under Mathlib.Data.Option. Other basic operations on Option are defined in the core library.

def Option.traverse {F : Type u → Type v} [Applicative F] {α : Type u_1} {β : Type u} (f : αF β) :
Option αF (Option β)

Traverse an object of Option α with a function f : α → F β for an applicative F.

Equations
Instances For
    def Option.elim' {α : Type u_1} {β : Type u_2} (b : β) (f : αβ) :
    Option αβ

    An elimination principle for Option. It is a nondependent version of Option.rec.

    Equations
    Instances For
      @[simp]
      theorem Option.elim'_none {α : Type u_1} {β : Type u_2} (b : β) (f : αβ) :
      Option.elim' b f none = b
      @[simp]
      theorem Option.elim'_some {α : Type u_1} {β : Type u_2} {a : α} (b : β) (f : αβ) :
      Option.elim' b f (some a) = f a
      theorem Option.elim'_eq_elim {α : Type u_3} {β : Type u_4} (b : β) (f : αβ) (a : Option α) :
      Option.elim' b f a = a.elim b f
      theorem Option.mem_some_iff {α : Type u_3} {a : α} {b : α} :
      a some b b = a
      @[inline]
      def Option.decidableEqNone {α : Type u_1} {o : Option α} :
      Decidable (o = none)

      o = none is decidable even if the wrapped type does not have decidable equality. This is not an instance because it is not definitionally equal to Option.decidableEq. Try to use o.isNone or o.isSome instead.

      Equations
      Instances For
        instance Option.decidableForallMem {α : Type u_1} {p : αProp} [DecidablePred p] (o : Option α) :
        Decidable (∀ (a : α), a op a)
        Equations
        instance Option.decidableExistsMem {α : Type u_1} {p : αProp} [DecidablePred p] (o : Option α) :
        Decidable (∃ (a : α), a o p a)
        Equations
        @[reducible, inline]
        abbrev Option.iget {α : Type u_1} [Inhabited α] :
        Option αα

        Inhabited get function. Returns a if the input is some a, otherwise returns default.

        Equations
        • x.iget = match x with | some x => x | none => default
        Instances For
          theorem Option.iget_some {α : Type u_1} [Inhabited α] {a : α} :
          (some a).iget = a
          @[simp]
          theorem Option.mem_toList {α : Type u_1} {a : α} {o : Option α} :
          a o.toList a o
          instance Option.liftOrGet_isCommutative {α : Type u_1} (f : ααα) [Std.Commutative f] :
          Equations
          • =
          instance Option.liftOrGet_isAssociative {α : Type u_1} (f : ααα) [Std.Associative f] :
          Equations
          • =
          instance Option.liftOrGet_isIdempotent {α : Type u_1} (f : ααα) [Std.IdempotentOp f] :
          Equations
          • =
          instance Option.liftOrGet_isId {α : Type u_1} (f : ααα) :
          Equations
          • =
          def Lean.LOption.toOption {α : Type u_3} :

          Convert undef to none to make an LOption into an Option.

          Equations
          Instances For