The lawful operator framework provides free theorems around the typeclass LawfulOperator
.
Its definition is based on section 3.3 of the AIGNET paper.
decls1
is a prefix of decls2
- of :: (
- size_le : decls1.size ≤ decls2.size
The prefix may never be longer than the other array.
The prefix and the other array must agree on all elements up until the bound of the prefix
- )
Instances For
If decls1
is a prefix of decls2
and we start evaluating decls2
at an
index in bounds of decls1
we can evaluate at decls1
.
If decls1
is a prefix of decls2
and we start evaluating decls2
at an
index in bounds of decls1
we can evaluate at decls1
.
Equations
- aig.ExtendingEntrypoint = { entry : Std.Sat.AIG.Entrypoint α // aig.decls.size ≤ entry.aig.decls.size }
Instances For
Equations
- aig.ExtendingRefVecEntry len = { ref : Std.Sat.AIG.RefVecEntry α len // aig.decls.size ≤ ref.aig.decls.size }
Instances For
A function f
that takes some aig : AIG α
and an argument of type β aig
is called a lawful
AIG operator if it only extends the AIG
but never modifies already existing nodes.
This guarantees that applying such a function will not change the semantics of any existing parts of the circuit, allowing us to perform local reasoning on the AIG.
- le_size : ∀ (aig : Std.Sat.AIG α) (input : β aig), aig.decls.size ≤ (f aig input).aig.decls.size
- decl_eq : ∀ (aig : Std.Sat.AIG α) (input : β aig) (idx : Nat) (h1 : idx < aig.decls.size) (h2 : idx < (f aig input).aig.decls.size), (f aig input).aig.decls[idx] = aig.decls[idx]