Documentation

Lean.Meta.InferType

Auxiliary function for instantiating the loose bound variables in e with args[start:stop]. This function is similar to instantiateRevRange, but it applies beta-reduction when we instantiate a bound variable with a lambda expression. Example: Given the term #0 a, and start := 0, stop := 1, args := #[fun x => x] the result is a instead of (fun x => x) a. This reduction is useful when we are inferring the type of eliminator-like applications. For example, given (n m : Nat) (f : Nat → Nat) (h : m = n), the type of Eq.subst (motive := fun x => f m = f x) h rfl is motive n which is (fun (x : Nat) => f m = f x) n This function reduces the new application to f m = f n

We use it to implement inferAppType

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    Instances For
      Instances For
        Instances For
          @[inline]
          Instances For
            @[export lean_infer_type]
            Instances For

              isPropQuick e is an "approximate" predicate which returns LBool.true if e is a proposition.

              isProp e returns true if e is a proposition.

              If e contains metavariables, it may not be possible to decide whether is a proposition or not. We return false in this case. We considered using LBool and retuning LBool.undef, but we have no applications for it.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For

                isProofQuick e is an "approximate" predicate which returns LBool.true if e is a proof.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For

                  isTypeQuick e is an "approximate" predicate which returns LBool.true if e is a type.

                  Return true iff the type of e is a Sort _.

                  Instances For

                    Return u iff type is Sort u or As → Sort u.

                    Equations
                    Instances For

                      Return true iff type is Sort _ or As → Sort _.

                      Instances For

                        Return true iff type is Prop or As → Prop.

                        Equations
                        Instances For

                          Return true iff e : Sort _ or e : (forall As, Sort _). Remark: it subsumes isType

                          Instances For

                            Given n and a non-dependent function type α₁ → α₂ → ... → αₙ → Sort u, returns the types α₁, α₂, ..., αₙ. Throws an error if there are not at least n argument types or if a later argument type depends on a prior one (i.e., it's a dependent function type).

                            This can be used to infer the expected type of the alternatives when constructing a MatcherApp.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For

                              Infers the types of the next n parameters that e expects. See arrowDomainsN.

                              Equations
                              Instances For