This module provides a basic theory around the naive AIG node creation functions. It is mostly fundamental work for the cached versions later on.
theorem
Std.Sat.AIG.mkGate_decl_eq
{α : Type}
[Hashable α]
[DecidableEq α]
(idx : Nat)
(aig : AIG α)
(input : aig.GateInput)
{h : idx < aig.decls.size}
:
The AIG produced by AIG.mkGate
agrees with the input AIG on all indices that are valid for both.
@[simp]
instance
Std.Sat.AIG.instLawfulOperatorMkAtom
{α : Type}
[Hashable α]
[DecidableEq α]
:
LawfulOperator α (fun (x : AIG α) => α) mkAtom
theorem
Std.Sat.AIG.mkConst_decl_eq
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : AIG α)
(val : Bool)
(idx : Nat)
{h : idx < aig.decls.size}
:
The AIG produced by AIG.mkConst
agrees with the input AIG on all indices that are valid for both.
instance
Std.Sat.AIG.instLawfulOperatorBoolMkConst
{α : Type}
[Hashable α]
[DecidableEq α]
:
LawfulOperator α (fun (x : AIG α) => Bool) mkConst
theorem
Std.Sat.AIG.denote_idx_gate
{α : Type}
[Hashable α]
[DecidableEq α]
{start : Nat}
{assign : α → Bool}
{lhs : Nat}
{linv : Bool}
{rhs : Nat}
{rinv : Bool}
{aig : AIG α}
{hstart : start < aig.decls.size}
(h : aig.decls[start] = Decl.gate lhs rhs linv rinv)
:
If an index contains a Decl.gate
we know how to denote it.
theorem
Std.Sat.AIG.idx_trichotomy
{α : Type}
[Hashable α]
[DecidableEq α]
{start : Nat}
(aig : AIG α)
(hstart : start < aig.decls.size)
{prop : Prop}
(hconst : ∀ (b : Bool), aig.decls[start] = Decl.const b → prop)
(hatom : ∀ (a : α), aig.decls[start] = Decl.atom a → prop)
(hgate : ∀ (lhs rhs : Nat) (linv rinv : Bool), aig.decls[start] = Decl.gate lhs rhs linv rinv → prop)
:
prop
theorem
Std.Sat.AIG.denote_idx_trichotomy
{α : Type}
[Hashable α]
[DecidableEq α]
{start : Nat}
{assign : α → Bool}
{res : Bool}
{aig : AIG α}
{hstart : start < aig.decls.size}
(hconst :
∀ (b : Bool),
aig.decls[start] = Decl.const b → ⟦assign, { aig := aig, ref := { gate := start, hgate := hstart } }⟧ = res)
(hatom :
∀ (a : α), aig.decls[start] = Decl.atom a → ⟦assign, { aig := aig, ref := { gate := start, hgate := hstart } }⟧ = res)
(hgate :
∀ (lhs rhs : Nat) (linv rinv : Bool),
aig.decls[start] = Decl.gate lhs rhs linv rinv →
⟦assign, { aig := aig, ref := { gate := start, hgate := hstart } }⟧ = res)
: