class
Std.Sat.AIG.RefVec.LawfulZipOperator
(α : Type)
[Hashable α]
[DecidableEq α]
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
:
Instances
theorem
Std.Sat.AIG.RefVec.LawfulZipOperator.denote_prefix_cast_ref
{α : Type}
[Hashable α]
[DecidableEq α]
{assign : α → Bool}
{aig : AIG α}
{input1 input2 : aig.BinaryInput}
{f : (aig : AIG α) → aig.BinaryInput → Entrypoint α}
[LawfulOperator α BinaryInput f]
[LawfulZipOperator α f]
{h : aig.decls.size ≤ (f aig input1).aig.decls.size}
:
⟦assign, f (f aig input1).aig (input2.cast h)⟧ = ⟦assign, f aig input2⟧
instance
Std.Sat.AIG.RefVec.LawfulZipOperator.instMkAndCached
{α : Type}
[Hashable α]
[DecidableEq α]
:
instance
Std.Sat.AIG.RefVec.LawfulZipOperator.instMkOrCached
{α : Type}
[Hashable α]
[DecidableEq α]
:
instance
Std.Sat.AIG.RefVec.LawfulZipOperator.instMkXorCached
{α : Type}
[Hashable α]
[DecidableEq α]
:
instance
Std.Sat.AIG.RefVec.LawfulZipOperator.instMkBEqCached
{α : Type}
[Hashable α]
[DecidableEq α]
:
instance
Std.Sat.AIG.RefVec.LawfulZipOperator.instMkImpCached
{α : Type}
[Hashable α]
[DecidableEq α]
:
structure
Std.Sat.AIG.RefVec.ZipTarget
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : AIG α)
(len : Nat)
:
- input : aig.BinaryRefVec len
- func (aig : AIG α) : aig.BinaryInput → Entrypoint α
- lawful : LawfulOperator α BinaryInput self.func
- chainable : LawfulZipOperator α self.func
Instances For
@[specialize #[]]
def
Std.Sat.AIG.RefVec.zip
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
(aig : AIG α)
(target : ZipTarget aig len)
:
RefVecEntry α len
Equations
- Std.Sat.AIG.RefVec.zip aig target = Std.Sat.AIG.RefVec.zip.go aig 0 Std.Sat.AIG.RefVec.empty ⋯ target.input.lhs target.input.rhs target.func
Instances For
@[irreducible, specialize #[]]
def
Std.Sat.AIG.RefVec.zip.go
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
(aig : AIG α)
(idx : Nat)
(s : aig.RefVec idx)
(hidx : idx ≤ len)
(lhs rhs : aig.RefVec len)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[LawfulZipOperator α f]
:
RefVecEntry α len
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[irreducible]
theorem
Std.Sat.AIG.RefVec.zip.go_le_size
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{aig : AIG α}
(idx : Nat)
(hidx : idx ≤ len)
(s : aig.RefVec idx)
(lhs rhs : aig.RefVec len)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[chainable : LawfulZipOperator α f]
:
theorem
Std.Sat.AIG.RefVec.zip_le_size
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{aig : AIG α}
(target : ZipTarget aig len)
:
@[irreducible]
theorem
Std.Sat.AIG.RefVec.zip.go_decl_eq
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{aig : AIG α}
(i : Nat)
(hi : i ≤ len)
(lhs rhs : aig.RefVec len)
(s : aig.RefVec i)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[chainable : LawfulZipOperator α f]
(idx : Nat)
(h1 : idx < aig.decls.size)
(h2 : idx < (go aig i s hi lhs rhs f).aig.decls.size)
:
instance
Std.Sat.AIG.RefVec.instLawfulVecOperatorZipTargetZip
{α : Type}
[Hashable α]
[DecidableEq α]
:
LawfulVecOperator α ZipTarget fun {len : Nat} => zip
@[irreducible]
theorem
Std.Sat.AIG.RefVec.zip.go_get_aux
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{aig : AIG α}
(curr : Nat)
(hcurr : curr ≤ len)
(s : aig.RefVec curr)
(lhs rhs : aig.RefVec len)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[chainable : LawfulZipOperator α f]
(idx : Nat)
(hidx : idx < curr)
(hfoo : aig.decls.size ≤ (go aig curr s hcurr lhs rhs f).aig.decls.size)
:
theorem
Std.Sat.AIG.RefVec.zip.go_get
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{aig : AIG α}
(curr : Nat)
(hcurr : curr ≤ len)
(s : aig.RefVec curr)
(lhs rhs : aig.RefVec len)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[chainable : LawfulZipOperator α f]
(idx : Nat)
(hidx : idx < curr)
:
theorem
Std.Sat.AIG.RefVec.zip.go_denote_mem_prefix
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{assign : α → Bool}
{aig : AIG α}
(curr : Nat)
(hcurr : curr ≤ len)
(s : aig.RefVec curr)
(lhs rhs : aig.RefVec len)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[chainable : LawfulZipOperator α f]
(start : Nat)
(hstart : start < aig.decls.size)
:
@[irreducible]
theorem
Std.Sat.AIG.RefVec.zip.denote_go
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{assign : α → Bool}
{aig : AIG α}
(curr : Nat)
(hcurr : curr ≤ len)
(s : aig.RefVec curr)
(lhs rhs : aig.RefVec len)
(f : (aig : AIG α) → aig.BinaryInput → Entrypoint α)
[LawfulOperator α BinaryInput f]
[chainable : LawfulZipOperator α f]
(idx : Nat)
(hidx1 : idx < len)
:
@[simp]
theorem
Std.Sat.AIG.RefVec.denote_zip
{α : Type}
[Hashable α]
[DecidableEq α]
{len : Nat}
{assign : α → Bool}
{aig : AIG α}
(target : ZipTarget aig len)
(idx : Nat)
(hidx : idx < len)
: