The Fourier-Motzkin elimination procedure
The Fourier-Motzkin procedure is a variable elimination method for linear inequalities. https://en.wikipedia.org/wiki/Fourier%E2%80%93Motzkin_elimination
Given a set of linear inequalities comps = {tᵢ Rᵢ 0}
,
we aim to eliminate a single variable a
from the set.
We partition comps
into comps_pos
, comps_neg
, and comps_zero
,
where comps_pos
contains the comparisons tᵢ Rᵢ 0
in which
the coefficient of a
in tᵢ
is positive, and similar.
For each pair of comparisons tᵢ Rᵢ 0 ∈ comps_pos
, tⱼ Rⱼ 0 ∈ comps_neg
,
we compute coefficients vᵢ, vⱼ ∈ ℕ
such that vᵢ*tᵢ + vⱼ*tⱼ
cancels out a
.
We collect these sums vᵢ*tᵢ + vⱼ*tⱼ R' 0
in a set S
and set comps' = S ∪ comps_zero
,
a new set of comparisons in which a
has been eliminated.
Theorem: comps
and comps'
are equisatisfiable.
We recursively eliminate all variables from the system. If we derive an empty clause 0 < 0
,
we conclude that the original system was unsatisfiable.
Datatypes
The comp_source
and pcomp
datatypes are specific to the FM elimination routine;
they are not shared with other components of linarith
.
- assump : ℕ → linarith.comp_source
- add : linarith.comp_source → linarith.comp_source → linarith.comp_source
- scale : ℕ → linarith.comp_source → linarith.comp_source
comp_source
tracks the source of a comparison.
The atomic source of a comparison is an assumption, indexed by a natural number.
Two comparisons can be added to produce a new comparison,
and one comparison can be scaled by a natural number to produce a new comparison.
Given a comp_source
cs
, cs.flatten
maps an assumption index
to the number of copies of that assumption that appear in the history of cs
.
For example, suppose cs
is produced by scaling assumption 2 by 5,
and adding to that the sum of assumptions 1 and 2.
cs.flatten
maps 1 ↦ 1, 2 ↦ 6
.
Formats a comp_source
for printing.
- c : linarith.comp
- src : linarith.comp_source
- history : native.rb_set ℕ
- effective : native.rb_set ℕ
- implicit : native.rb_set ℕ
- vars : native.rb_set ℕ
A pcomp
stores a linear comparison Σ cᵢ*xᵢ R 0
,
along with information about how this comparison was derived.
The original expressions fed into linarith
are each assigned a unique natural number label.
The historical set pcomp.history
stores the labels of expressions
that were used in deriving the current pcomp
.
Variables are also indexed by natural numbers. The sets pcomp.effective
, pcomp.implicit
,
and pcomp.vars
contain variable indices.
pcomp.vars
contains the variables that appear inpcomp.c
. We store them inpcomp
to avoid recomputing the set, which requires folding over a list. (TODO: is this really needed?)pcomp.effective
contains the variables that have been effectively eliminated frompcomp
. A variablen
is said to be effectively eliminated inpcomp
if the elimination ofn
produced at least one of the ancestors ofpcomp
.pcomp.implicit
contains the variables that have been implicitly eliminated frompcomp
. A variablen
is said to be implicitly eliminated inpcomp
if it satisfies the following properties:- There is some
ancestor
ofpcomp
such thatn
appears inancestor.vars
. n
does not appear inpcomp.vars
.n
was not effectively eliminated.
- There is some
We track these sets in order to compute whether the history of a pcomp
is minimal.
Checking this directly is expensive, but effective approximations can be defined in terms of these
sets. During the variable elimination process, a pcomp
with non-minimal history can be discarded.
Any comparison whose history is not minimal is redundant,
and need not be included in the new set of comparisons.
elimed_ge : ℕ
is a natural number such that all variables with index ≥ elimed_ge
have been
removed from the system.
This test is an overapproximation to minimality. It gives necessary but not sufficient conditions.
If the history of c
is minimal, then c.maybe_minimal
is true,
but c.maybe_minimal
may also be true for some c
with minimal history.
Thus, if c.maybe_minimal
is false, c
is known not to be minimal and must be redundant.
See http://citeseerx.ist.psu.edu/viewdoc/download?doi=10.1.1.51.493&rep=rep1&type=pdf p.13
(Theorem 7).
The condition described there considers only implicitly eliminated variables that have been
officially eliminated from the system. This is not the case for every implicitly eliminated variable.
Consider eliminating z
from {x + y + z < 0, x - y - z < 0}
. The result is the set
{2*x < 0}
; y
is implicitly but not officially eliminated.
This implementation of Fourier-Motzkin elimination processes variables in decreasing order of
indices. Immediately after a step that eliminates variable k
, variable k'
has been eliminated
iff k' ≥ k
. Thus we can compute the intersection of officially and implicitly eliminated variables
by taking the set of implicitly eliminated variables with indices ≥ elimed_ge
.
The comp_source
field is ignored when comparing pcomp
s. Two pcomp
s proving the same
comparison, with different sources, are considered equivalent.
pcomp.scale c n
scales the coefficients of c
by n
and notes this in the comp_source
.
pcomp.add c1 c2 elim_var
creates the result of summing the linear comparisons c1
and c2
,
during the process of eliminating the variable elim_var
.
The computation assumes, but does not enforce, that elim_var
appears in both c1
and c2
and does not appear in the sum.
Computing the sum of the two comparisons is easy; the complicated details lie in tracking the
additional fields of pcomp
.
- The historical set
pcomp.history
ofc1 + c2
is the union of the two historical sets. - We recompute the variables that appear in
c1 + c2
from the newly createdlinexp
, since some may have been implicitly eliminated. - The effectively eliminated variables of
c1 + c2
are the union of the two effective sets, withelim_var
inserted. - The implicitly eliminated variables of
c1 + c2
are those that appear in at least one ofc1.vars
andc2.vars
but not in(c1 + c2).vars
, excludingelim_var
.
pcomp.assump c n
creates a pcomp
whose comparison is c
and whose source is
comp_source.assump n
, that is, c
is derived from the n
th hypothesis.
The history is the singleton set {n}
.
No variables have been eliminated (effectively or implicitly).
Creates an empty set of pcomp
s, sorted using pcomp.cmp
. This should always be used instead
of mk_rb_map
for performance reasons.
Elimination procedure
If c1
and c2
both contain variable a
with opposite coefficients,
produces v1
and v2
such that a
has been cancelled in v1*c1 + v2*c2
.
pelim_var p1 p2
calls elim_var
on the comp
components of p1
and p2
.
If this returns v1
and v2
, it creates a new pcomp
equal to v1*p1 + v2*p2
,
and tracks this in the comp_source
.
A pcomp
represents a contradiction if its comp
field represents a contradiction.
elim_var_with_set a p comps
collects the result of calling pelim_var p p' a
for every p' ∈ comps
.
- max_var : ℕ
- comps : native.rb_set linarith.pcomp
The state for the elimination monad.
max_var
: the largest variable index that has not been eliminated.comps
: a set of comparisons
The elimination procedure proceeds by eliminating variable v
from comps
progressively
in decreasing order.
The linarith monad extends an exceptional monad with a linarith_structure
state.
An exception produces a contradictory pcomp
.
Return the current comparison set.
Throws an exception if a contradictory pcomp
is contained in the current state.
Updates the current state with a new max variable and comparisons,
and calls validate
to check for a contradiction.
split_set_by_var_sign a comps
partitions the set comps
into three parts.
pos
contains the elements ofcomps
in whicha
has a positive coefficient.neg
contains the elements ofcomps
in whicha
has a negative coefficient.not_present
contains the elements ofcomps
in whicha
has coefficient 0.
Returns (pos, neg, not_present)
.
monad.elim_var a
performs one round of Fourier-Motzkin elimination, eliminating the variable a
from the linarith
state.
elim_all_vars
eliminates all variables from the linarith state, leaving it with a set of
ground comparisons. If this succeeds without exception, the original linarith
state is consistent.
mk_linarith_structure hyps vars
takes a list of hypotheses and the largest variable present in
those hypotheses. It produces an initial state for the elimination monad.
produce_certificate hyps vars
tries to derive a contradiction from the comparisons in hyps
by eliminating all variables ≤ max_var
.
If successful, it returns a map coeff : ℕ → ℕ
as a certificate.
This map represents that we can find a contradiction by taking the sum ∑ (coeff i) * hyps[i]
.