Documentation

Std.Sat.AIG.Relabel

theorem Std.Sat.AIG.Decl.relabel_id_map {α : Type} (decl : Decl α) :
relabel id decl = decl
theorem Std.Sat.AIG.Decl.relabel_comp {α β γ : Type} (decl : Decl α) (g : αβ) (h : βγ) :
relabel (h g) decl = relabel h (relabel g decl)
theorem Std.Sat.AIG.Decl.relabel_const {α β : Type} {idx : Nat} {b : Bool} {decls : Array (Decl α)} {r : αβ} {hidx : idx < decls.size} (h : relabel r decls[idx] = const b) :
decls[idx] = const b
theorem Std.Sat.AIG.Decl.relabel_atom {α β : Type} {idx : Nat} {a : β} {decls : Array (Decl α)} {r : αβ} {hidx : idx < decls.size} (h : relabel r decls[idx] = atom a) :
∃ (x : α), decls[idx] = atom x a = r x
theorem Std.Sat.AIG.Decl.relabel_gate {α β : Type} {idx lhs rhs : Nat} {linv rinv : Bool} {decls : Array (Decl α)} {r : αβ} {hidx : idx < decls.size} (h : relabel r decls[idx] = gate lhs rhs linv rinv) :
decls[idx] = gate lhs rhs linv rinv
def Std.Sat.AIG.relabel {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] (r : αβ) (aig : AIG α) :
AIG β
Equations
Instances For
    @[simp]
    theorem Std.Sat.AIG.relabel_size_eq_size {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {aig : AIG α} {r : αβ} :
    (relabel r aig).decls.size = aig.decls.size
    theorem Std.Sat.AIG.relabel_const {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {idx : Nat} {b : Bool} {aig : AIG α} {r : αβ} {hidx : idx < (relabel r aig).decls.size} (h : (relabel r aig).decls[idx] = Decl.const b) :
    aig.decls[idx] = Decl.const b
    theorem Std.Sat.AIG.relabel_atom {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {idx : Nat} {a : β} {aig : AIG α} {r : αβ} {hidx : idx < (relabel r aig).decls.size} (h : (relabel r aig).decls[idx] = Decl.atom a) :
    ∃ (x : α), aig.decls[idx] = Decl.atom x a = r x
    theorem Std.Sat.AIG.relabel_gate {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {idx lhs rhs : Nat} {linv rinv : Bool} {aig : AIG α} {r : αβ} {hidx : idx < (relabel r aig).decls.size} (h : (relabel r aig).decls[idx] = Decl.gate lhs rhs linv rinv) :
    aig.decls[idx] = Decl.gate lhs rhs linv rinv
    @[simp, irreducible]
    theorem Std.Sat.AIG.denote_relabel {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] (aig : AIG α) (r : αβ) (start : Nat) {hidx : start < (relabel r aig).decls.size} (assign : βBool) :
    assign, { aig := relabel r aig, ref := { gate := start, hgate := hidx } } = assign r, { aig := aig, ref := { gate := start, hgate := } }
    theorem Std.Sat.AIG.unsat_relabel {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {idx : Nat} {aig : AIG α} (r : αβ) {hidx : idx < aig.decls.size} :
    aig.UnsatAt idx hidx(relabel r aig).UnsatAt idx
    theorem Std.Sat.AIG.relabel_unsat_iff {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {idx : Nat} [Nonempty α] {aig : AIG α} {r : αβ} {hidx1 : idx < (relabel r aig).decls.size} {hidx2 : idx < aig.decls.size} (hinj : ∀ (x y : α), x aigy aigr x = r yx = y) :
    (relabel r aig).UnsatAt idx hidx1 aig.UnsatAt idx hidx2
    def Std.Sat.AIG.Entrypoint.relabel {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] (r : αβ) (entry : Entrypoint α) :
    Equations
    Instances For
      @[simp]
      theorem Std.Sat.AIG.Entrypoint.relabel_size_eq {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] {entry : Entrypoint α} {r : αβ} :
      (relabel r entry).aig.decls.size = entry.aig.decls.size
      theorem Std.Sat.AIG.Entrypoint.relabel_unsat_iff {α : Type} [Hashable α] [DecidableEq α] {β : Type} [Hashable β] [DecidableEq β] [Nonempty α] {entry : Entrypoint α} {r : αβ} (hinj : ∀ (x y : α), x entry.aigy entry.aigr x = r yx = y) :
      (relabel r entry).Unsat entry.Unsat