Datatypes for linarith
Some of the data structures here are used in multiple parts of the tactic. We split them into their own file.
This file also contains a few convenient auxiliary functions.
A shorthand for tracing when the trace.linarith
option is set to true.
Linear expressions
A linear expression is a list of pairs of variable indices and coefficients, representing the sum of the products of each coefficient with its corresponding variable.
Some functions on linexp
assume that n : ℕ
occurs at most once as the first element of a pair,
and that the list is sorted in decreasing order of the first argument.
This is not enforced by the type but the operations here preserve it.
Add two linexp
s together componentwise.
Preserves sorting and uniqueness of the first argument.
l.scale c
scales the values in l
by c
without modifying the order or keys.
l.get n
returns the value in l
associated with key n
, if it exists, and none
otherwise.
This function assumes that l
is sorted in decreasing order of the first argument,
that is, it will return none
as soon as it finds a key smaller than n
.
Equations
- linarith.linexp.get n ((a, b) :: t) = ite (a < n) option.none (ite (a = n) (option.some b) (linarith.linexp.get n t))
- linarith.linexp.get n list.nil = option.none
l.contains n
is true iff n
is the first element of a pair in l
.
Equations
l.zfind n
returns the value associated with key n
if there is one, and 0 otherwise.
Equations
- linarith.linexp.zfind n l = linarith.linexp.zfind._match_1 (linarith.linexp.get n l)
- linarith.linexp.zfind._match_1 (option.some v) = v
- linarith.linexp.zfind._match_1 option.none = 0
l.vars
returns the list of variables that occur in l
.
Defines a lex ordering on linexp
. This function is performance critical.
Equations
- linarith.linexp.cmp ((n1, z1) :: t1) ((n2, z2) :: t2) = ite (n1 < n2) ordering.lt (ite (n2 < n1) ordering.gt (ite (z1 < z2) ordering.lt (ite (z2 < z1) ordering.gt (linarith.linexp.cmp t1 t2))))
- linarith.linexp.cmp ((fst, snd) :: tl) list.nil = ordering.gt
- linarith.linexp.cmp list.nil (hd :: tl) = ordering.lt
- linarith.linexp.cmp list.nil list.nil = ordering.eq
Inequalities
- eq : linarith.ineq
- le : linarith.ineq
- lt : linarith.ineq
The three-element type ineq
is used to represent the strength of a comparison between terms.
max R1 R2
computes the strength of the sum of two inequalities. If t1 R1 0
and t2 R2 0
,
then t1 + t2 (max R1 R2) 0
.
Equations
- linarith.ineq.lt.max linarith.ineq.lt = linarith.ineq.lt
- linarith.ineq.lt.max linarith.ineq.le = linarith.ineq.lt
- linarith.ineq.lt.max linarith.ineq.eq = linarith.ineq.lt
- linarith.ineq.le.max linarith.ineq.lt = linarith.ineq.lt
- linarith.ineq.le.max linarith.ineq.le = linarith.ineq.le
- linarith.ineq.le.max linarith.ineq.eq = linarith.ineq.le
- linarith.ineq.eq.max linarith.ineq.lt = linarith.ineq.lt
- linarith.ineq.eq.max linarith.ineq.le = linarith.ineq.le
- linarith.ineq.eq.max linarith.ineq.eq = linarith.ineq.eq
ineq
is ordered eq < le < lt
.
Equations
- linarith.ineq.lt.cmp linarith.ineq.lt = ordering.eq
- linarith.ineq.lt.cmp linarith.ineq.le = ordering.gt
- linarith.ineq.lt.cmp linarith.ineq.eq = ordering.gt
- linarith.ineq.le.cmp linarith.ineq.lt = ordering.lt
- linarith.ineq.le.cmp linarith.ineq.le = ordering.eq
- linarith.ineq.le.cmp linarith.ineq.eq = ordering.gt
- linarith.ineq.eq.cmp linarith.ineq.lt = ordering.lt
- linarith.ineq.eq.cmp linarith.ineq.le = ordering.lt
- linarith.ineq.eq.cmp linarith.ineq.eq = ordering.eq
Prints an ineq
as the corresponding infix symbol.
Equations
- linarith.ineq.lt.to_string = "<"
- linarith.ineq.le.to_string = "≤"
- linarith.ineq.eq.to_string = "="
Finds the name of a multiplicative lemma corresponding to an inequality strength.
Equations
Comparisons with 0
- str : linarith.ineq
- coeffs : linarith.linexp
The main datatype for FM elimination.
Variables are represented by natural numbers, each of which has an integer coefficient.
Index 0 is reserved for constants, i.e. coeffs.find 0
is the coefficient of 1.
The represented term is coeffs.sum (λ ⟨k, v⟩, v * Var[k])
.
str determines the strength of the comparison -- is it < 0, ≤ 0, or = 0?
c.vars
returns the list of variables that appear in the linear expression contained in c
.
comp.coeff_of c a
projects the coefficient of variable a
out of c
.
Equations
- c.coeff_of a = linarith.linexp.zfind a c.coeffs
comp.scale c n
scales the coefficients of c
by n
.
comp.add c1 c2
adds the expressions represented by c1
and c2
.
The coefficient of variable a
in c1.add c2
is the sum of the coefficients of a
in c1
and c2
.
comp
has a lex order. First the ineq
s are compared, then the coeff
s.
A comp
represents a contradiction if its expression has no coefficients and its strength is <,
that is, it represents the fact 0 < 0
.
Parsing into linear form
Control
A preprocessor transforms a proof of a proposition into a proof of a different propositon.
The return type is list expr
, since some preprocessing steps may create multiple new hypotheses,
and some may remove a hypothesis from the list.
A "no-op" preprocessor should return its input as a singleton list.
Some preprocessors need to examine the full list of hypotheses instead of working item by item.
As with preprocessor
, the input to a global_preprocessor
is replaced by, not added to, its output.
Some preprocessors perform branching case splits. A branch
is used to track one of these case
splits. The first component, an expr
, is the goal corresponding to this branch of the split,
given as a metavariable. The list expr
component is the list of hypotheses for linarith
in this branch. Every expr
in this list should be type correct in the context of the associated goal.
Some preprocessors perform branching case splits.
A global_branching_preprocessor
produces a list of branches to run.
Each branch is independent, so hypotheses that appear in multiple branches should be duplicated.
The preprocessor is responsible for making sure that each branch contains the correct goal
metavariable.
A preprocessor
lifts to a global_preprocessor
by folding it over the input list.
A global_preprocessor
lifts to a global_branching_preprocessor
by producing only one branch.
process pp l
runs pp.transform
on l
and returns the result,
tracing the result if trace.linarith
is on.
A certificate_oracle
is a function produce_certificate : list comp → ℕ → tactic (rb_map ℕ ℕ)
.
produce_certificate hyps max_var
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]
.
The default certificate_oracle
used by linarith
is linarith.fourier_motzkin.produce_certificate
- discharger : tactic unit
- restrict_type : option Type
- restrict_type_reflect : reflected c.restrict_type . "apply_instance"
- exfalso : bool
- transparency : tactic.transparency
- split_hypotheses : bool
- split_ne : bool
- preprocessors : option (list linarith.global_branching_preprocessor)
- oracle : option linarith.certificate_oracle
A configuration object for linarith
.
cfg.update_reducibility reduce_semi
will change the transparency setting of cfg
to
semireducible
if reduce_semi
is true. In this case, it also sets the discharger to ring!
,
since this is typically needed when using stronger unification.
Auxiliary functions
These functions are used by multiple modules, so we put them here for accessibility.
parse_into_comp_and_expr e
checks if e
is of the form t < 0
, t ≤ 0
, or t = 0
.
If it is, it returns the comparison along with t
.
mk_single_comp_zero_pf c h
assumes that h
is a proof of t R 0
.
It produces a pair (R', h')
, where h'
is a proof of c*t R' 0
.
Typically R
and R'
will be the same, except when c = 0
, in which case R'
is =
.
If c = 1
, h'
is the same as h
-- specifically, it does not change the type to 1*t R 0
.