Skip to content

Commit

Permalink
Fix
Browse files Browse the repository at this point in the history
  • Loading branch information
TwoFX committed Oct 18, 2024
1 parent 187f4af commit 9dee147
Show file tree
Hide file tree
Showing 17 changed files with 17 additions and 0 deletions.
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Data.HashMap
import Std.Data.HashSet

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/CNF.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.CNF
import Std.Sat.AIG.Basic
import Std.Sat.AIG.Lemmas
Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/Cached.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Basic
import Std.Sat.AIG.Lemmas

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/CachedGates.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Cached
import Std.Sat.AIG.CachedLemmas

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/CachedGatesLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.CachedGates
import Std.Sat.AIG.LawfulOperator

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/CachedLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Cached

/-!
Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/If.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.CachedGatesLemmas
import Std.Sat.AIG.LawfulVecOperator

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/LawfulOperator.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Basic

/-!
Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/LawfulVecOperator.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.LawfulOperator
import Std.Sat.AIG.RefVec

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Basic
import Std.Sat.AIG.LawfulOperator

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/RefVec.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.LawfulOperator
import Std.Sat.AIG.CachedGatesLemmas

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/RefVecOperator.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.RefVecOperator.Map
import Std.Sat.AIG.RefVecOperator.Zip
import Std.Sat.AIG.RefVecOperator.Fold
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/RefVecOperator/Fold.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.RefVec
import Std.Sat.AIG.LawfulVecOperator

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/RefVecOperator/Map.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.RefVec
import Std.Sat.AIG.LawfulVecOperator

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/RefVecOperator/Zip.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.RefVec
import Std.Sat.AIG.LawfulVecOperator

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/Relabel.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Basic
import Std.Sat.AIG.Lemmas

Expand Down
1 change: 1 addition & 0 deletions src/Std/Sat/AIG/RelabelNat.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
prelude
import Std.Sat.AIG.Relabel


Expand Down

0 comments on commit 9dee147

Please sign in to comment.