lrat_proof
command #
Defines a macro for producing SAT proofs from CNF / LRAT files. These files are commonly used in the SAT community for writing proofs.
Most SAT solvers support export to DRAT format, but this format can be expensive to reconstruct because it requires recomputing all unit propagation steps. The LRAT format solves this issue by attaching a proof to the deduction of each new clause. (The L in LRAT stands for Linear time verification.) There are several verified checkers for the LRAT format, and the program implemented here makes it possible to use the lean kernel as an LRAT checker as well and expose the results as a standard propositional theorem.
The input to the lrat_proof
command is the name of the theorem to define,
and the statement (written in CNF format) and the proof (in LRAT format).
For example:
lrat_proof foo
"p cnf 2 4 1 2 0 -1 2 0 1 -2 0 -1 -2 0"
"5 -2 0 4 3 0 5 d 3 4 0 6 1 0 5 1 0 6 d 1 0 7 0 5 2 6 0"
produces a theorem:
foo : ∀ (a a_1 : Prop), (¬a ∧ ¬a_1 ∨ a ∧ ¬a_1) ∨ ¬a ∧ a_1 ∨ a ∧ a_1
- You can see the theorem statement by hovering over the word
foo
. - You can use the
example
keyword in place offoo
to avoid generating a theorem. - You can use the
include_str
macro in place of the two strings to load CNF / LRAT files from disk.
A literal is a positive or negative occurrence of an atomic propositional variable. Note that unlike DIMACS, 0 is a valid variable index.
- pos: ℕ → Sat.Literal
- neg: ℕ → Sat.Literal
Instances For
Construct a literal. Positive numbers are translated to positive literals, and negative numbers become negative literals. The input is assumed to be nonzero.
Equations
- Sat.Literal.ofInt i = if i < 0 then Sat.Literal.neg (-i - 1).toNat else Sat.Literal.pos (i - 1).toNat
Instances For
Swap the polarity of a literal.
Equations
- (Sat.Literal.pos i).negate = Sat.Literal.neg i
- (Sat.Literal.neg i).negate = Sat.Literal.pos i
Instances For
Equations
- One or more equations did not get rendered due to their size.
A clause is a list of literals, thought of as a disjunction like a ∨ b ∨ ¬c
.
Equations
Instances For
Equations
- Sat.Clause.cons = List.cons
Instances For
Formula f
subsumes f'
if all the clauses in f'
are in f
.
We use this to prove that all clauses in the formula are subsumed by it.
- prop : ∀ (x : Sat.Clause), x ∈ f' → x ∈ f
Instances For
A valuation is an assignment of values to all the propositional variables.
Equations
- Sat.Valuation = (ℕ → Prop)
Instances For
v.neg lit
asserts that literal lit
is falsified in the valuation.
Equations
- v.neg (Sat.Literal.pos i) = ¬v i
- v.neg (Sat.Literal.neg i) = v i
Instances For
v.satisfies_fmla f
asserts that formula f
is satisfied by the valuation.
A formula is satisfied if all clauses in it are satisfied.
- prop : ∀ (c : Sat.Clause), c ∈ f → v.satisfies c
Instances For
f.proof c
asserts that c
is derivable from f
.
Equations
- f.proof c = ∀ (v : Sat.Valuation), v.satisfies_fmla f → v.satisfies c
Instances For
If f
subsumes c
(i.e. c ∈ f
), then f.proof c
.
The core unit-propagation step.
We have a local context of assumptions ¬l'
(sometimes called an assignment)
and we wish to add ¬l
to the context, that is, we want to prove l
is also falsified.
This is because there is a clause a ∨ b ∨ ¬l
in the global context
such that all literals in the clause are falsified except for ¬l
;
so in the context h₁
where we suppose that ¬l
is falsified,
the clause itself is falsified so we can prove False
.
We continue the proof in h₂
, with the assumption that l
is falsified.
Valuation.mk [a, b, c]
is a valuation which is a
at 0, b
at 1 and c
at 2, and false
everywhere else.
Equations
- Sat.Valuation.mk [] x = False
- Sat.Valuation.mk (a :: tail) 0 = a
- Sat.Valuation.mk (head :: as) n.succ = Sat.Valuation.mk as n
Instances For
If f
is unsatisfiable, and every v
which agrees with ps
implies ¬⟦f⟧_v → p
, then p
.
Equivalently, there exists a valuation v
which agrees with ps
,
and every such valuation yields ¬⟦f⟧_v
because f
is unsatisfiable.
Negation turns AND into OR, so ¬⟦f₁ ∧ f₂⟧_v ≡ ¬⟦f₁⟧_v ∨ ¬⟦f₂⟧_v
.
Reification of a single clause formula.
Asserts that ¬⟦l⟧_v
implies p
.
- prop : v.neg l → p
Instances For
Negation turns OR into AND, so ¬⟦l ∨ c⟧_v ≡ ¬⟦l⟧_v ∧ ¬⟦c⟧_v
.
The reification of a singleton clause ¬⟦l⟧_v ≡ ¬⟦l⟧_v
.
The reification of a positive literal ¬⟦a⟧_v ≡ ¬a
.
The reification of a negative literal ¬⟦¬a⟧_v ≡ a
.
The representation of a global clause.
The list of literals as read from the input file
- expr : Lean.Expr
The clause expression of type
Clause
- proof : Lean.Expr
A proof of
⊢ ctx.proof c
. Note that we do not usehave
statements to cache these proofs: this is literally the proof expression itself. As a result, the proof terms rely heavily on dag-like sharing of the expression, and printing these proof terms directly is likely to crash lean for larger examples.
Instances For
Construct the clause expression from the input list. For example [1, -2]
is translated to
Clause.cons (Literal.pos 1) (Clause.cons (Literal.neg 2) Clause.nil)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Constructs the proofs of ⊢ ctx.proof c
for each clause c
in ctx
.
The proofs are stashed in a HashMap
keyed on the clause ID.
A localized clause reference.
It is the same as Clause
except that the proof is now a local variable.
The list of literals as read from the input file
- expr : Lean.Expr
The clause expression of type
Clause
- depth : ℕ
The bound variable index of the hypothesis asserting
⊢ ctx.proof c
, counting from the outside and 1-based. (We use this numbering because we will need to reference the variable from multiple binder depths.)
Instances For
Construct an individual proof step ⊢ ctx.proof c
.
db
: the current global contextns
,clause
: the new clausepf
: the LRAT proof tracectx
: the main formula
The proof has three steps:
Introduce local assumptions
have h1 : ctx.proof c1 := p1
for each clausec1
referenced in the proof. We actually do all the introductions at once, as in(fun h1 h2 h3 ↦ ...) p1 p2 p3
, because we wantp_i
to not be under any binders to avoid the cost ofinstantiate
during typechecking and get the benefits of dag-like sharing in thepi
(which are themselves previous proof steps which may be large terms). The hypotheses are ingctx
, keyed on the clause ID.Unfold
⊢ ctx.proof [a, b, c]
to∀ v, v.satisfies_fmla ctx → v.neg a → v.neg b → v.neg c → False
andintro v hv ha hb hc
, storing eachha : v.neg a
inlctx
, keyed on the literala
.For each LRAT step
hc : ctx.proof [x, y]
,hc v hv : v.neg x → v.neg y → False
. We look for a literal that is not falsified in the clause. Since it is a unit propagation step, there can be at most one such literal.- If
x
is the non-falsified clause, letx'
denote the negated literal ofx
. Thenx'.negate
reduces tox
, sohnx : v.neg x'.negate |- hc v hv hnx hy : False
, so we construct the termby_cases (fun hnx : v.neg x'.negate ↦ hc v hv hnx hy) (fun hx : v.neg x ↦ ...)
andhx
is added to the local context. - If all clauses are falsified, then we are done:
hc v hv hx hy : False
.
- If
Equations
- One or more equations did not get rendered due to their size.
Instances For
An LRAT step is either an addition or a deletion step.
- add: ℕ → Array ℤ → Array ℤ → Mathlib.Tactic.Sat.LRATStep
An addition step, with the clause ID, the clause literal list, and the proof trace
- del: Array ℕ → Mathlib.Tactic.Sat.LRATStep
A (multiple) deletion step, which deletes all the listed clause IDs from the context
Instances For
Build the main proof of ⊢ ctx.proof []
using the LRAT proof trace.
arr
: The input CNFctx
: The abbreviated formula, a constant likefoo.ctx_1
ctx'
: The definitional expansion of the formula, a tree ofFmla.and
nodessteps
: The input LRAT proof trace
Equations
- One or more equations did not get rendered due to their size.
Instances For
Build the type and value of the reified theorem. This rewrites all the SAT definitions
into standard operators on Prop
, for example if the formula is [[1, 2], [-1, 2], [-2]]
then
this produces a proof of ⊢ ∀ a b : Prop, (a ∧ b) ∨ (¬a ∧ b) ∨ ¬b
. We use the input nvars
to
decide how many quantifiers to use.
Most of the proof is under 2 * nvars + 1
quantifiers
a1 .. an : Prop, v : Valuation, h1 : v 0 ↔ a1, ... hn : v (n-1) ↔ an ⊢ ...
, and we do the index
arithmetic by hand.
- First, we call
reifyFormula ctx'
which returnsa
andpr : reify v ctx' a
- Then we build
fun (v : Valuation) (h1 : v 0 ↔ a1) ... (hn : v (n-1) ↔ an) ↦ pr
- We have to lower expression
a
from step 1 out of the quantifiers by lowering all variable indices bynvars+1
. This is okay becausev
andh1..hn
do not appear ina
. - We construct the expression
ps
, which isa1 .. an : Prop ⊢ [a1, ..., an] : List Prop
refute ctx (hf : ctx.proof []) (fun v h1 .. hn ↦ pr) : a
forces some definitional unfolding sincefun h1 .. hn ↦ pr
should have typeimplies v (reify v ctx a) [a1, ..., an] a
, which involves unfoldingimplies
n times as well asctx ↦ ctx'
.- Finally, we
intro a1 ... an
so that we have a proof of∀ a1 ... an, a
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The v
variable under the a1 ... an, v, h1 ... hn
context
Equations
- Mathlib.Tactic.Sat.buildReify.v nvars = Lean.mkBVar nvars
Instances For
Returns a
and pr : v n ↔ a
given a variable index n
.
These are both lookups into the context
(a0 .. a(n-1) : Prop) (v) (h1 : v 0 ↔ a0) ... (hn : v (n-1) ↔ a(n-1))
.
Equations
- Mathlib.Tactic.Sat.buildReify.reifyVar nvars v = (Lean.mkBVar (2 * nvars - v.rawNatLit?.get!), Lean.mkBVar (nvars - v.rawNatLit?.get! - 1))
Instances For
Parse a natural number
Instances For
Parse an integer
Equations
- One or more equations did not get rendered due to their size.
Instances For
Parse a list of integers terminated by 0
Parse a list of natural numbers terminated by 0
Parse a DIMACS format .cnf
file.
This is not very robust; we assume the file has had comments stripped.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Parse an LRAT file into a list of steps.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Core of fromLRAT
. Constructs the context and main proof definitions,
but not the reification theorem. Returns:
nvars
: the number of variables specified in the CNF filectx
: The abbreviated formula, a constant likefoo.ctx_1
ctx'
: The definitional expansion of the formula, a tree ofFmla.and
nodesproof
: A proof ofctx.proof []
Equations
- One or more equations did not get rendered due to their size.
Instances For
Main entry point. Given strings cnf
and lrat
with unparsed file data, and a name name
,
adds theorem name : type := proof
where type
is a propositional theorem like
∀ (a a_1 : Prop), (¬a ∧ ¬a_1 ∨ a ∧ ¬a_1) ∨ ¬a ∧ a_1 ∨ a ∧ a_1
.
Also creates auxiliaries named name.ctx_1
(for the CNF formula)
and name.proof_1
(for the LRAT proof), with name
itself containing the reification proof.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A macro for producing SAT proofs from CNF / LRAT files. These files are commonly used in the SAT community for writing proofs.
The input to the lrat_proof
command is the name of the theorem to define,
and the statement (written in CNF format) and the proof (in LRAT format).
For example:
lrat_proof foo
"p cnf 2 4 1 2 0 -1 2 0 1 -2 0 -1 -2 0"
"5 -2 0 4 3 0 5 d 3 4 0 6 1 0 5 1 0 6 d 1 0 7 0 5 2 6 0"
produces a theorem:
foo : ∀ (a a_1 : Prop), (¬a ∧ ¬a_1 ∨ a ∧ ¬a_1) ∨ ¬a ∧ a_1 ∨ a ∧ a_1
- You can see the theorem statement by hovering over the word
foo
. - You can use the
example
keyword in place offoo
to avoid generating a theorem. - You can use the
include_str
macro in place of the two strings to load CNF / LRAT files from disk.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A macro for producing SAT proofs from CNF / LRAT files. These files are commonly used in the SAT community for writing proofs.
The input to the from_lrat
term syntax is two string expressions with
the statement (written in CNF format) and the proof (in LRAT format).
For example:
def foo := from_lrat
"p cnf 2 4 1 2 0 -1 2 0 1 -2 0 -1 -2 0"
"5 -2 0 4 3 0 5 d 3 4 0 6 1 0 5 1 0 6 d 1 0 7 0 5 2 6 0"
produces a theorem:
foo : ∀ (a a_1 : Prop), (¬a ∧ ¬a_1 ∨ a ∧ ¬a_1) ∨ ¬a ∧ a_1 ∨ a ∧ a_1
- You can use this term after
have :=
or indef foo :=
to produce the term without constraining the type. - You can use it when a specific type is expected, but it currently does not pay any attention to the shape of the goal and always produces the same theorem, so you can only use this to do alpha renaming.
- You can use the
include_str
macro in place of the two strings to load CNF / LRAT files from disk.
Equations
- One or more equations did not get rendered due to their size.