lossy-unification.lagda.rst 2.6 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081
  1. ..
  2. ::
  3. module language.lossy-unification where
  4. .. _lossy-unification:
  5. *****************
  6. Lossy Unification
  7. *****************
  8. The option ``--experimental-lossy-unification`` enables an
  9. experimental heuristic in the unification checker intended to improve
  10. its performance for unification problems of the form ``f es₀ = f es₁``,
  11. i.e. unifying two applications of the same defined function, here
  12. ``f``, to possibly different lists of arguments and projections
  13. ``es₀`` and ``es₁``.
  14. The heuristic is sound but not complete.
  15. In particular if Agda accepts code with the flag enabled it should
  16. also accept it without the flag (with enough resources, and possibly
  17. needing extra type annotations).
  18. The option can be used either globally or in an ``OPTIONS`` pragma, in the latter
  19. case it applies only to the current module.
  20. Heuristic
  21. ~~~~~~~~~
  22. When trying to solve the unification problem ``f es₀ = f es₁`` the
  23. heuristic proceeds by trying to solve ``es₀ = es₁``, if that succeeds
  24. the original problem is also solved, otherwise unification proceeds as
  25. without the flag, likely by reducing both ``f es₀`` and ``f es₁``.
  26. Example
  27. ~~~~~~~
  28. Suppose ``f`` adds ``100`` to its input as defined below
  29. .. code-block:: agda
  30. f : ℕ → ℕ
  31. f n = 100 + n
  32. then to unify ``f 2`` and ``f (1 + 1)`` the heuristic would proceed by
  33. unifying ``2`` with ``(1 + 1)``, which quickly succeeds. Without the
  34. flag we might instead first reduce both ``f 2`` and ``f (1 + 1)`` to
  35. ``102`` and then compare those results.
  36. The performance will improve most dramatically when reducing an
  37. application of ``f`` would produce a large term, perhaps an element of
  38. a record type with several fields and/or large embedded proof terms.
  39. Drawbacks
  40. ~~~~~~~~~
  41. The main drawback is that the heursitic is not complete, i.e. it will cause Agda to
  42. ignore some possible solutions to unification variables.
  43. For example if ``f`` is a constant function, then the constraint ``f
  44. ?0 = f 1`` does not uniquely determine ``?0``, but the heuristic will
  45. end up assigning ``1`` to ``?0``.
  46. Such assignments can lead to Agda to report a type error which would
  47. not have been reported without the heuristic. This is because committing to
  48. ``?0 = 1`` might make other constraints unsatifiable.
  49. The other drawback is that in some cases performance of
  50. unification will be worse with the heuristic. Specifically, if
  51. the heuristic will repeatedly attempt to unify lists of arguments ``es₀
  52. = es₁`` while failing.
  53. References
  54. ----------
  55. Slow typechecking of single one-line definition, `issue (#1625) <https://arxiv.org/abs/1611.02108>`_.