Documentation

Lean.Elab.App

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

      Erase entry for binderName from namedArgs.

      Equations
      Instances For

        Default application elaborator #

        • ellipsis : Bool

          true if .. was used

        • explicit : Bool

          true if @ modifier was used

        • resultIsOutParamSupport : Bool

          If the result type of an application is the outParam of some local instance, then special support may be needed because type class resolution interacts poorly with coercions in this kind of situation. This flag enables the special support.

          The idea is quite simple, if the result type is the outParam of some local instance, we simply execute synthesizeSyntheticMVarsUsingDefault. We added this feature to make sure examples as follows are correctly elaborated.

          class GetElem (Cont : Type u) (Idx : Type v) (Elem : outParam (Type w)) where
            getElem (xs : Cont) (i : Idx) : Elem
          
          export GetElem (getElem)
          
          instance : GetElem (Array α) Nat α where
            getElem xs i := xs.get ⟨i, sorry⟩
          
          opaque f : Option BoolBool
          opaque g : BoolBool
          
          def bad (xs : Array Bool) : Bool :=
            let x := getElem xs 0
            f x && g x
          

          Without the special support, Lean fails at g x saying x has type Option Bool but is expected to have type Bool. From the user's point of view this is a bug, since let x := getElem xs 0 clearly constrains x to be Bool, but we only obtain this information after we apply the OfNat default instance for 0.

          Before converging to this solution, we have tried to create a "coercion placeholder" when resultIsOutParamSupport = true, but it did not work well in practice. For example, it failed in the example above.

        Instances For

          Auxiliary structure for elaborating the application f args namedArgs.

          • fType : Lean.Expr
          • Remaining regular arguments.

          • remaining named arguments to be processed.

          • expectedType? : Option Lean.Expr
          • etaArgs : Array Lean.Expr

            When named arguments are provided and explicit arguments occurring before them are missing, the elaborator eta-expands the declaration. For example,

            def f (x y : Nat) := x + y
            #check f (y := 5)
            -- fun x => f x 5
            

            etaArgs stores the fresh free variables for implementing the eta-expansion. When .. is used, eta-expansion is disabled, and missing arguments are treated as _.

          • toSetErrorCtx : Array Lean.MVarId

            Metavariables that we need to set the error context using the application being built.

          • instMVars : Array Lean.MVarId

            Metavariables for the instance implicit arguments that have already been processed.

          • propagateExpected : Bool

            The following field is used to implement the propagateExpectedType heuristic. It is set to true true when expectedType still has to be propagated.

          • resultTypeOutParam? : Option Lean.MVarId

            If the result type may be the outParam of some local instance. See comment at Context.resultIsOutParamSupport

          Instances For

            Try to synthesize metavariables are instMVars using type class resolution. The ones that cannot be synthesized yet stay in the instMVars list. Remark: we use this method

            • before trying to apply coercions to function,
            • before unifying the expected type.
            Equations
            • One or more equations did not get rendered due to their size.
            Instances For

              Try to synthesize metavariables are instMVars using type class resolution. The ones that cannot be synthesized yet are registered.

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

                Remove named argument with name binderName from namedArgs.

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

                  Elaborate function application arguments.

                  Eliminator-like function application elaborator #

                  Context of the elab_as_elim elaboration procedure.

                  • expectedType : Lean.Expr
                  • extraArgsPos : Array Nat

                    Position of additional arguments that should be elaborated eagerly because they can contribute to the motive inference procedure. For example, in the following theorem the argument h : a = b should be elaborated eagerly because it contains b which occurs in motive b.

                    theorem Eq.subst' {α} {motive : α → Prop} {a b : α} (h : a = b) : motive a → motive b
                    
                  Instances For

                    State of the elab_as_elim elaboration procedure.

                    • The resultant expression being built.

                    • fType : Lean.Expr

                      `f : fType

                    • User-provided named arguments that still have to be processed.

                    • User-provided arguments that still have to be processed.

                    • discrs : Array Lean.Expr

                      Discriminants processed so far.

                    • instMVars : Array Lean.MVarId

                      Instance implicit arguments collected so far.

                    • idx : Nat

                      Position of the next argument to be processed. We use it to decide whether the argument is the motive or a discriminant.

                    • motive? : Option Lean.Expr

                      Store the metavariable used to represent the motive that will be computed at finalize.

                    Instances For

                      Infer the motive using the expected type by kabstracting the discriminants.

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

                        If the eliminator is over-applied, we "revert" the extra arguments.

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

                          Construct the resulting application after all discriminants have bee elaborated, and we have consumed as many given arguments as possible.

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

                            Return the next argument to be processed. The result is .none if it is an implicit argument which was not provided using a named argument. The result is .undef if args is empty and namedArgs does contain an entry for binderName.

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

                              Set the motive field in the state.

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

                                Push the given expression into the discrs field in the state.

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

                                  Save information for producing error messages.

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

                                    Create an implicit argument using the given BinderInfo.

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

                                      Main loop of the elimAsElab procedure.

                                      Function application elaboration #

                                      def Lean.Elab.Term.elabAppArgs (f : Lean.Expr) (namedArgs : Array Lean.Elab.Term.NamedArg) (args : Array Lean.Elab.Term.Arg) (expectedType? : Option Lean.Expr) (explicit : Bool) (ellipsis : Bool) (resultIsOutParamSupport : optParam Bool true) :

                                      Elaborate a f-application using namedArgs and args as the arguments.

                                      • expectedType? the expected type if available. It is used to propagate typing information only. This method does not ensure the result has this type.
                                      • explicit = true when notation @ is used, and implicit arguments are assumed to be provided at namedArgs and args.
                                      • ellipsis = true when notation .. is used. That is, we add _ for missing arguments.
                                      • resultIsOutParamSupport is used to control whether special support is used when processing applications of functions that return output parameter of some local instance. Example:
                                        GetElem.getElem : {Cont : Type u_1} → {Idx : Type u_2} → {elem : Type u_3} → {dom : cont → idx → Prop} → [self : GetElem cont idx elem dom] → (xs : cont) → (i : idx) → dom xs i → elem
                                        
                                        The result type elem is the output parameter of the local instance self. When this parameter is set to true, we execute synthesizeSyntheticMVarsUsingDefault. For additional details, see comment at ElabAppArgs.resultIsOutParam.
                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For

                                        Return some info if we should elaborate as an eliminator.

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

                                          Collect extra argument positions that must be elaborated eagerly when using elab_as_elim. The idea is that the contribute to motive inference. See comment at ElamElim.Context.extraArgsPos.

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

                                            Auxiliary inductive datatype that represents the resolution of an LVal.

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

                                                Interaction between errToSorry and observing. #

                                                observing x does not check for synthetic sorry's, just an exception. Thus, it may think x worked when it didn't if a synthetic sorry was introduced. We decided that checking for synthetic sorrys at observing is not a good solution because it would not be clear to decide what the "main" error message for the alternative is. When the result contains a synthetic sorry, it is not clear which error message corresponds to the sorry. Moreover, while executing x, many error messages may have been logged. Recall that we need an error per alternative at mergeFailures.

                                                Thus, we decided to set errToSorry to false whenever processing choice nodes and overloaded symbols.

                                                Important: we rely on the property that after errToSorry is set to false, no elaboration function executed by x will reset it to true.

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