Documentation

Lean.Meta.Match.MatcherApp.Transform

Given

  • matcherApp match_i As (fun xs => motive[xs]) discrs (fun ys_1 => (alt_1 : motive (C_1[ys_1])) ... (fun ys_n => (alt_n : motive (C_n[ys_n]) remaining, and
  • expression e : B[discrs], Construct the term match_i As (fun xs => B[xs] -> motive[xs]) discrs (fun ys_1 (y : B[C_1[ys_1]]) => alt_1) ... (fun ys_n (y : B[C_n[ys_n]]) => alt_n) e remaining.

We use kabstract to abstract the discriminants from B[discrs].

This method assumes

  • the matcherApp.motive is a lambda abstraction where xs.size == discrs.size
  • each alternative is a lambda abstraction where ys_i.size == matcherApp.altNumParams[i]

This is used in in Lean.Elab.PreDefinition.WF.Fix when replacing recursive calls with calls to the argument provided by fix to refine the termination argument, which may mention major. See there for how to use this function.

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

    Similar to MatcherApp.addArg, but returns none on failure.

    Equations
    Instances For

      Given

      • matcherApp match_i As (fun xs => motive[xs]) discrs (fun ys_1 => (alt_1 : motive (C_1[ys_1])) ... (fun ys_n => (alt_n : motive (C_n[ys_n]) remaining, and
      • a expression B[discrs] (which may not be a type, e.g. n : Nat), returns the expressions fun ys_1 ... ys_i => B[C_1[ys_1]] ... B[C_n[ys_n]],

      This method assumes

      • the matcherApp.motive is a lambda abstraction where xs.size == discrs.size
      • each alternative is a lambda abstraction where ys_i.size == matcherApp.altNumParams[i]

      This is similar to MatcherApp.addArg when you only have an expression to refined, and not a type with a value.

      This is used in in Lean.Elab.PreDefinition.WF.GuessFix when constructing the context of recursive calls to refine the functions' paramter, which may mention major. See there for how to use this function.

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

        A non-failing version of MatcherApp.refineThrough

        Equations
        • One or more equations did not get rendered due to their size.
        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
            def Lean.Meta.MatcherApp.withUserNames {n : TypeType u_1} [MonadControlT Lean.MetaM n] [Monad n] {α : Type} (fvars : Array Lean.Expr) (names : Array Lake.Name) (k : n α) :
            n α

            Sets the user name of the FVars in the local context according to the given array of names.

            If they differ in size the shorter size wins.

            Equations
            Instances For
              def Lean.Meta.MatcherApp.transform {n : TypeType} [MonadLiftT Lean.MetaM n] [MonadControlT Lean.MetaM n] [Monad n] [Lean.MonadError n] [Lean.MonadEnv n] [Lean.MonadLog n] [Lean.AddMessageContext n] [Lean.MonadOptions n] (matcherApp : Lean.Meta.MatcherApp) (useSplitter : optParam Bool false) (addEqualities : optParam (Array Bool) (mkArray matcherApp.discrs.size false)) (onParams : optParam (Lean.Exprn Lean.Expr) pure) (onMotive : optParam (Array Lean.ExprLean.Exprn Lean.Expr) fun (x : Array Lean.Expr) (e : Lean.Expr) => pure e) (onAlt : optParam (Lean.ExprLean.Exprn Lean.Expr) fun (x e : Lean.Expr) => pure e) (onRemaining : optParam (Array Lean.Exprn (Array Lean.Expr)) pure) :

              Performs a possibly type-changing transformation to a MatcherApp.

              • onParams is run on each parameter and discriminant
              • onMotive runs on the body of the motive, and is passed the motive parameters (one for each MatcherApp.discrs)
              • onAlt runs on each alternative, and is passed the expected type of the alternative, as inferred from the motive
              • onRemaining runs on the remaining arguments (and may change their number)

              If useSplitter is true, the matcher is replaced with the splitter. NB: Not all operations on MatcherApp can handle one matcherName is a splitter.

              The array addEqualities, if provided, indicates for which of the discriminants an equality connecting the discriminant to the parameters of the alternative (like in match h : x with …) should be added (if it is isn't already there).

              This function works even if the the type of alternatives do not fit the inferred type. This allows you to post-process the MatcherApp with MatcherApp.inferMatchType, which will infer a type, given all the alternatives.

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

                Given a MatcherApp, replaces the motive with one that is inferred from the actual types of the alternatives.

                For example, given

                (match (motive := NatUnit → ?) n with
                 0 => 1
                 _ => true) ()
                

                (for any ?; the motive’s result type be ignored) will give this type

                (match n with
                 | 0 => Nat
                 | _ => Bool)
                

                The given MatcherApp must not use a splitter in matcherName. The resulting expression will use the splitter corresponding to matcherName (this is necessary for the construction).

                Interally, this needs to reduce the matcher in a given branch; this is done using Split.simpMatchTarget.

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