aboutsummaryrefslogtreecommitdiffhomepage
path: root/doc/sphinx/addendum/omega.rst
blob: 20e40c5507b35a4fc43f7356d32cfdc7cfd0bf76 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
.. _omega:

Omega: a solver for quantifier-free problems in Presburger Arithmetic
=====================================================================

:Author: Pierre Crégut

Description of ``omega``
------------------------

This tactic does not need any parameter:

.. tacn:: omega

``omega`` solves a goal in Presburger arithmetic, i.e. a universally
quantified formula made of equations and inequations. Equations may
be specified either on the type ``nat`` of natural numbers or on
the type ``Z`` of binary-encoded integer numbers. Formulas on
``nat`` are automatically injected into ``Z``.  The procedure
may use any hypothesis of the current proof session to solve the goal.

Multiplication is handled by ``omega`` but only goals where at
least one of the two multiplicands of products is a constant are
solvable. This is the restriction meant by "Presburger arithmetic".

If the tactic cannot solve the goal, it fails with an error message.
In any case, the computation eventually stops.

Arithmetical goals recognized by ``omega``
------------------------------------------

``omega`` applied only to quantifier-free formulas built from the
connectors::

   /\  \/  ~  ->

on atomic formulas. Atomic formulas are built from the predicates::

   =  <  <=  >  >=

on ``nat`` or ``Z``. In expressions of type ``nat``, ``omega`` recognizes::

   +  -  *  S  O  pred

and in expressions of type ``Z``, ``omega`` recognizes numeral constants and::

   +  -  *  Z.succ Z.pred

All expressions of type ``nat`` or ``Z`` not built on these
operators are considered abstractly as if they
were arbitrary variables of type ``nat`` or ``Z``.

Messages from ``omega``
-----------------------

When ``omega`` does not solve the goal, one of the following errors
is generated:

.. exn:: omega can't solve this system

  This may happen if your goal is not quantifier-free (if it is
  universally quantified, try ``intros`` first; if it contains
  existentials quantifiers too, ``omega`` is not strong enough to solve your
  goal). This may happen also if your goal contains arithmetical
  operators unknown from ``omega``. Finally, your goal may be really
  wrong!

.. exn:: omega: Not a quantifier-free goal

  If your goal is universally quantified, you should first apply
  ``intro`` as many time as needed.

.. exn:: omega: Unrecognized predicate or connective: @ident

.. exn:: omega: Unrecognized atomic proposition: ...

.. exn:: omega: Can't solve a goal with proposition variables

.. exn:: omega: Unrecognized proposition

.. exn:: omega: Can't solve a goal with non-linear products

.. exn:: omega: Can't solve a goal with equality on type ...


Using ``omega``
---------------

The ``omega`` tactic does not belong to the core system. It should be
loaded by

.. coqtop:: in

   Require Import Omega.

.. example::

  .. coqtop:: all

     Require Import Omega.

     Open Scope Z_scope.

     Goal forall m n:Z, 1 + 2 * m <> 2 * n.
     intros; omega.
     Abort.

     Goal forall z:Z, z > 0 -> 2 * z + 1 > z.
     intro; omega.
     Abort.


Options
-------

.. opt:: Stable Omega

This deprecated option (on by default) is for compatibility with Coq pre 8.5. It
resets internal name counters to make executions of ``omega`` independent.

.. opt:: Omega UseLocalDefs

This option (on by default) allows ``omega`` to use the bodies of local
variables.

.. opt:: Omega System

This option (off by default) activate the printing of debug information

.. opt:: Omega Action

This option (off by default) activate the printing of debug information

Technical data
--------------

Overview of the tactic
~~~~~~~~~~~~~~~~~~~~~~

 * The goal is negated twice and the first negation is introduced as an hypothesis.
 * Hypothesis are decomposed in simple equations or inequations. Multiple
   goals may result from this phase.
 * Equations and inequations over ``nat`` are translated over
   ``Z``, multiple goals may result from the translation of substraction.
 * Equations and inequations are normalized.
 * Goals are solved by the OMEGA decision procedure.
 * The script of the solution is replayed.

Overview of the OMEGA decision procedure
~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~

The OMEGA decision procedure involved in the ``omega`` tactic uses
a small subset of the decision procedure presented in :cite:`TheOmegaPaper`
Here is an overview, look at the original paper for more information.

 * Equations and inequations are normalized by division by the GCD of their
   coefficients.
 * Equations are eliminated, using the Banerjee test to get a coefficient
   equal to one.
 * Note that each inequation defines a half space in the space of real value
   of the variables.
 * Inequations are solved by projecting on the hyperspace
   defined by cancelling one of the variable.  They are partitioned
   according to the sign of the coefficient of the eliminated
   variable. Pairs of inequations from different classes define a
   new edge in the projection.
 * Redundant inequations are eliminated or merged in new
   equations that can be eliminated by the Banerjee test.
 * The last two steps are iterated until a contradiction is reached
   (success) or there is no more variable to eliminate (failure).

It may happen that there is a real solution and no integer one. The last
steps of the Omega procedure (dark shadow) are not implemented, so the
decision procedure is only partial.

Bugs
----

 * The simplification procedure is very dumb and this results in
   many redundant cases to explore.

 * Much too slow.

 * Certainly other bugs! You can report them to https://coq.inria.fr/bugs/.