Documentation

Std.Sat.AIG.RefVecOperator.Zip

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.BinaryInputEntrypoint α} [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
    @[specialize #[]]
    def Std.Sat.AIG.RefVec.zip {α : Type} [Hashable α] [DecidableEq α] {len : Nat} (aig : AIG α) (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] :
    Equations
    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.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] :
      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.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] :
        aig.decls.size (go aig idx s hidx lhs rhs f).aig.decls.size
        theorem Std.Sat.AIG.RefVec.zip_le_size {α : Type} [Hashable α] [DecidableEq α] {len : Nat} (aig : AIG α) (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] :
        aig.decls.size (zip aig input func).aig.decls.size
        @[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.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] (idx : Nat) (h1 : idx < aig.decls.size) (h2 : idx < (go aig i s hi lhs rhs f).aig.decls.size) :
        (go aig i s hi lhs rhs f).aig.decls[idx] = aig.decls[idx]
        theorem Std.Sat.AIG.RefVec.zip_decl_eq {α : Type} [Hashable α] [DecidableEq α] {len : Nat} {aig : AIG α} (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] (idx : Nat) (h1 : idx < aig.decls.size) (h2 : idx < (zip aig input func).aig.decls.size) :
        (zip aig input func).aig.decls[idx] = aig.decls[idx]
        theorem Std.Sat.AIG.RefVec.zip_lt_size_of_lt_aig_size {α : Type} [Hashable α] [DecidableEq α] {len x : Nat} (aig : AIG α) (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] (h : x < aig.decls.size) :
        x < (zip aig input func).aig.decls.size
        theorem Std.Sat.AIG.RefVec.zip_le_size_of_le_aig_size {α : Type} [Hashable α] [DecidableEq α] {len x : Nat} (aig : AIG α) (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] (h : x aig.decls.size) :
        x (zip aig input func).aig.decls.size
        theorem Std.Sat.AIG.RefVec.IsPrefix_zip {α : Type} [Hashable α] [DecidableEq α] {len : Nat} {aig : AIG α} (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] :
        IsPrefix aig.decls (zip aig input func).aig.decls
        @[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.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] (idx : Nat) (hidx : idx < curr) (hfoo : aig.decls.size (go aig curr s hcurr lhs rhs f).aig.decls.size) :
        (go aig curr s hcurr lhs rhs f).vec.get idx = (s.get idx hidx).cast hfoo
        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.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] (idx : Nat) (hidx : idx < curr) :
        (go aig curr s hcurr lhs rhs f).vec.get idx = (s.get idx hidx).cast
        theorem Std.Sat.AIG.RefVec.zip.go_denote_mem_prefix {α : Type} [Hashable α] [DecidableEq α] {len : Nat} {assign : αBool} {inv : Bool} {aig : AIG α} (curr : Nat) (hcurr : curr len) (s : aig.RefVec curr) (lhs rhs : aig.RefVec len) (f : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] (start : Nat) (hstart : start < aig.decls.size) :
        assign, { aig := (go aig curr s hcurr lhs rhs f).aig, ref := { gate := start, invert := inv, hgate := } } = assign, { aig := aig, ref := { gate := start, invert := inv, hgate := hstart } }
        @[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.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput f] [LawfulZipOperator α f] (idx : Nat) (hidx1 : idx < len) :
        curr idxassign, { aig := (go aig curr s hcurr lhs rhs f).aig, ref := (go aig curr s hcurr lhs rhs f).vec.get idx hidx1 } = assign, f aig { lhs := lhs.get idx hidx1, rhs := rhs.get idx hidx1 }
        @[simp]
        theorem Std.Sat.AIG.RefVec.denote_zip {α : Type} [Hashable α] [DecidableEq α] {len : Nat} {assign : αBool} {aig : AIG α} (input : aig.BinaryRefVec len) (func : (aig : AIG α) → aig.BinaryInputEntrypoint α) [LawfulOperator α BinaryInput func] [LawfulZipOperator α func] (idx : Nat) (hidx : idx < len) :
        assign, { aig := (zip aig input func).aig, ref := (zip aig input func).vec.get idx hidx } = assign, func aig { lhs := input.lhs.get idx hidx, rhs := input.rhs.get idx hidx }