Documentation

Lean.Meta.SynthInstance

Instances For
    • mvar : Expr
    • key : Expr
    • instances : Array Instance
    • currInstanceIdx : Nat
    • typeHasMVars : Bool

      typeHasMVars := true if type of mvar contains metavariables. We store this information to implement an optimization that relies on the fact that instances are "morally canonical." That is, we need to find at most one answer for this generator node if the type does not have metavariables.

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

          In tabled resolution, we creating a mapping from goals (e.g., Coe Nat ?x) to answers and waiters. Waiters are consumer nodes that are waiting for answers for a particular node.

          We implement this mapping using a HashMap where the keys are normalized expressions. That is, we replace assignable metavariables with auxiliary free variables of the form _tc.<idx>. We do not declare these free variables in any local context, and we should view them as "normalized names" for metavariables. For example, the term f ?m ?m ?n is normalized as f _tc.0 _tc.0 _tc.1.

          This approach is structural, and we may visit the same goal more than once if the different occurrences are just definitionally equal, but not structurally equal.

          Remark: a metavariable is assignable only if its depth is equal to the metavar context depth.

          @[always_inline]
          Equations
          • One or more equations did not get rendered due to their size.

          Remark: mkTableKey assumes e does not contain assigned metavariables.

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

                  Remark: the SynthInstance.State is not really an extension of Meta.State. The field postponed is not needed, and the field mctx is misleading since synthInstance methods operate over different MetavarContexts simultaneously. That being said, we still use extends because it makes it simpler to move from M to MetaM.

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

                      Return globals and locals instances that may unify with type

                      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

                          Create a new generator node for mvar and add waiter as its waiter. key must be mkTableKey mctx mvarType.

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

                                Create a key for the goal associated with the given metavariable. That is, we create a key for the type of the metavariable.

                                We must instantiate assigned metavariables before we invoke mkTableKey.

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

                                  See getSubgoals and getSubgoalsAux

                                  We use the parameter j to reduce the number of instantiate* invocations. It is the same approach we use at forallTelescope and lambdaTelescope. Given getSubgoalsAux args j subgoals instVal type, we have that type.instantiateRevRange j args.size args does not have loose bound variables.

                                  Instances For

                                    getSubgoals lctx localInsts xs inst creates the subgoals for the instance inst. The subgoals are in the context of the free variables xs, and (lctx, localInsts) is the local context and instances before we added the free variables to it.

                                    This extra complication is required because 1- We want all metavariables created by synthInstance to share the same local context. 2- We want to ensure that applications such as mvar xs are higher order patterns.

                                    The method getGoals create a new metavariable for each parameter of inst. For example, suppose the type of inst is forall (x_1 : A_1) ... (x_n : A_n), B x_1 ... x_n. Then, we create the metavariables ?m_i : forall xs, A_i, and return the subset of these metavariables that are instance implicit arguments, and the expressions:

                                    • inst (?m_1 xs) ... (?m_n xs) (aka instVal)
                                    • B (?m_1 xs) ... (?m_n xs)
                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For

                                      Try to synthesize metavariable mvar using the instance inst. Remark: mctx is set using withMCtx. If it succeeds, the result is a new updated metavariable context and a new list of subgoals. A subgoal is created for each instance implicit parameter of inst.

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

                                        Assign a precomputed answer to mvar. If it succeeds, the result is a new updated metavariable context and a new list of subgoals.

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

                                          Move waiters that are waiting for the given answer to the resume stack.

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

                                              Create a new answer after cNode resolved all subgoals. That is, cNode.subgoals == []. And then, store it in the tabled entries map, and wakeup waiters.

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

                                                Process the next subgoal in the given consumer node.

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

                                                      Try the next instance in the node on the top of the generator stack.

                                                      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

                                                          Given (cNode, answer) on the top of the resume stack, continue execution by using answer to solve the next subgoal.

                                                          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
                                                              Instances For
                                                                Equations
                                                                • One or more equations did not get rendered due to their size.
                                                                Instances For

                                                                  Type class parameters can be annotated with outParam annotations.

                                                                  Given C a_1 ... a_n, we replace a_i with a fresh metavariable ?m_i IF a_i is an outParam. The result is type correct because we reject type class declarations IF it contains a regular parameter X that depends on an out parameter Y.

                                                                  Then, we execute type class resolution as usual. If it succeeds, and metavariables ?m_i have been assigned, we try to unify the original type C a_1 ... a_n with the normalized one.

                                                                  Remark: when maxResultSize? == none, the configuration option synthInstance.maxResultSize is used. Remark: we use a different option for controlling the maximum result size for coercions.

                                                                  def Lean.Meta.synthInstance? (type : Expr) (maxResultSize? : Option Nat := none) :
                                                                  Equations
                                                                  • One or more equations did not get rendered due to their size.
                                                                  Instances For
                                                                    def Lean.Meta.trySynthInstance (type : Expr) (maxResultSize? : Option Nat := none) :

                                                                    Return LOption.some r if succeeded, LOption.none if it failed, and LOption.undef if instance cannot be synthesized right now because type contains metavariables.

                                                                    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
                                                                        def Lean.Meta.synthInstance (type : Expr) (maxResultSize? : Option Nat := none) :
                                                                        Equations
                                                                        • One or more equations did not get rendered due to their size.
                                                                        Instances For