diff --git a/.github/workflows/lean_action_ci.yml b/.github/workflows/lean_action_ci.yml index a3d0b313..97320b1d 100644 --- a/.github/workflows/lean_action_ci.yml +++ b/.github/workflows/lean_action_ci.yml @@ -17,3 +17,4 @@ jobs: - uses: leanprover/lean-action@v1 with: build-args: "--wfail" + mk_all-check: "true" diff --git a/Cslib.lean b/Cslib.lean index 28350b62..67b5d493 100644 --- a/Cslib.lean +++ b/Cslib.lean @@ -27,6 +27,7 @@ import Cslib.Foundations.Semantics.ReductionSystem.Basic import Cslib.Foundations.Syntax.HasAlphaEquiv import Cslib.Foundations.Syntax.HasSubstitution import Cslib.Foundations.Syntax.HasWellFormed +import Cslib.Init import Cslib.Languages.CCS.Basic import Cslib.Languages.CCS.BehaviouralTheory import Cslib.Languages.CCS.Semantics diff --git a/Cslib/Computability/Languages/Language.lean b/Cslib/Computability/Languages/Language.lean index 8dc96979..1cd97437 100644 --- a/Cslib/Computability/Languages/Language.lean +++ b/Cslib/Computability/Languages/Language.lean @@ -6,6 +6,7 @@ Authors: Ching-Tsun Chou import Mathlib.Computability.Language import Mathlib.Tactic +import Cslib.Init /-! # Language (additional definitions and theorems) diff --git a/Cslib/Computability/Languages/OmegaLanguage.lean b/Cslib/Computability/Languages/OmegaLanguage.lean index 4391664c..cb6d0239 100644 --- a/Cslib/Computability/Languages/OmegaLanguage.lean +++ b/Cslib/Computability/Languages/OmegaLanguage.lean @@ -256,11 +256,11 @@ theorem le_omegaPow_congr [Inhabited α] {l1 l2 : Language α} (h : l1 ≤ l2) : theorem omegaPow_of_sub_one [Inhabited α] : (l - 1)^ω = l^ω := by ext s ; simp -@[simp] +@[simp, nolint simpNF] theorem zero_omegaPow [Inhabited α] : (0 : Language α)^ω = ⊥ := by ext s ; simp -@[simp] +@[simp, nolint simpNF] theorem one_omegaPow [Inhabited α] : (1 : Language α)^ω = ⊥ := by rw [← omegaPow_of_sub_one, tsub_self, zero_omegaPow] diff --git a/Cslib/Foundations/Control/Monad/Free.lean b/Cslib/Foundations/Control/Monad/Free.lean index aae446b9..22da715b 100644 --- a/Cslib/Foundations/Control/Monad/Free.lean +++ b/Cslib/Foundations/Control/Monad/Free.lean @@ -3,6 +3,7 @@ Copyright (c) 2025 Tanner Duve. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Tanner Duve, Eric Wieser -/ +import Cslib.Init import Mathlib.Tactic.Lemma import Mathlib.Tactic.SimpRw diff --git a/Cslib/Foundations/Data/FinFun.lean b/Cslib/Foundations/Data/FinFun.lean index 1b359420..1c1d8062 100644 --- a/Cslib/Foundations/Data/FinFun.lean +++ b/Cslib/Foundations/Data/FinFun.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi, Xueying Qin -/ +-- TODO: Notation conflict with Mathlib.Finsupp (both use →₀) +-- import Cslib.Init import Mathlib.Data.Finset.Basic /-! # Finite functions diff --git a/Cslib/Foundations/Data/HasFresh.lean b/Cslib/Foundations/Data/HasFresh.lean index fdf9d6da..c2c90beb 100644 --- a/Cslib/Foundations/Data/HasFresh.lean +++ b/Cslib/Foundations/Data/HasFresh.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi, Kenny Lau -/ +import Cslib.Init import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Ring.CharZero import Mathlib.Algebra.Ring.Int.Defs diff --git a/Cslib/Foundations/Data/Nat/Segment.lean b/Cslib/Foundations/Data/Nat/Segment.lean index 9e4a3b9c..b4f8098f 100644 --- a/Cslib/Foundations/Data/Nat/Segment.lean +++ b/Cslib/Foundations/Data/Nat/Segment.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Ching-Tsun Chou -/ +import Cslib.Init import Mathlib.Algebra.Order.Sub.Basic import Mathlib.Data.Nat.Nth import Mathlib.Tactic diff --git a/Cslib/Foundations/Data/OmegaSequence/Defs.lean b/Cslib/Foundations/Data/OmegaSequence/Defs.lean index b9deace3..c3e680f8 100644 --- a/Cslib/Foundations/Data/OmegaSequence/Defs.lean +++ b/Cslib/Foundations/Data/OmegaSequence/Defs.lean @@ -3,6 +3,7 @@ Copyright (c) 2025 Ching-Tsun Chou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ching-Tsun Chou, Fabrizio Montesi -/ +import Cslib.Init import Mathlib.Data.Nat.Notation import Mathlib.Data.FunLike.Basic import Mathlib.Logic.Function.Iterate diff --git a/Cslib/Foundations/Data/OmegaSequence/Flatten.lean b/Cslib/Foundations/Data/OmegaSequence/Flatten.lean index 7449b92e..094bed51 100644 --- a/Cslib/Foundations/Data/OmegaSequence/Flatten.lean +++ b/Cslib/Foundations/Data/OmegaSequence/Flatten.lean @@ -121,8 +121,8 @@ theorem append_flatten [Inhabited α] {ls : ωSequence (List α)} (h_ls : ∀ k, induction n generalizing ls <;> grind [tail_eq_drop, take_succ] /-- The length of `(ls.take n).flatten` is `ls.cumLen n`. -/ -@[simp, scoped grind =] -theorem length_flatten_take [Inhabited α] {ls : ωSequence (List α)} (h_ls : ∀ k, (ls k).length > 0) +@[simp, nolint simpNF, scoped grind =] +theorem length_flatten_take {ls : ωSequence (List α)} (h_ls : ∀ k, (ls k).length > 0) (n : ℕ) : (ls.take n).flatten.length = ls.cumLen n := by induction n <;> grind [take_succ'] @@ -169,7 +169,7 @@ theorem toSegs_def (s : ωSequence α) (f : ℕ → ℕ) (n : ℕ) : /-- `(s.toSegs f).cumLen` is `f` itself. -/ @[simp] -theorem segment_toSegs_cumLen [Inhabited α] {f : ℕ → ℕ} +theorem segment_toSegs_cumLen {f : ℕ → ℕ} (hm : StrictMono f) (h0 : f 0 = 0) (s : ωSequence α) : (s.toSegs f).cumLen = f := by ext n diff --git a/Cslib/Foundations/Data/Relation.lean b/Cslib/Foundations/Data/Relation.lean index 6a176151..eb252bb6 100644 --- a/Cslib/Foundations/Data/Relation.lean +++ b/Cslib/Foundations/Data/Relation.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi, Thomas Waring, Chris Henson -/ +import Cslib.Init import Mathlib.Logic.Relation /-! # Relations -/ diff --git a/Cslib/Foundations/Semantics/LTS/Basic.lean b/Cslib/Foundations/Semantics/LTS/Basic.lean index 5831eaca..c8edb4d9 100644 --- a/Cslib/Foundations/Semantics/LTS/Basic.lean +++ b/Cslib/Foundations/Semantics/LTS/Basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi -/ +-- TODO: Causes type elaboration ambiguities in downstream files (Automata) +-- import Cslib.Init import Mathlib.Tactic.Lemma import Mathlib.Data.Finite.Defs import Mathlib.Data.Fintype.Basic diff --git a/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean b/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean index 5350b856..13a06fc5 100644 --- a/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean +++ b/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi, Thomas Waring -/ +import Cslib.Init import Mathlib.Logic.Relation import Mathlib.Util.Notation3 diff --git a/Cslib/Foundations/Syntax/HasAlphaEquiv.lean b/Cslib/Foundations/Syntax/HasAlphaEquiv.lean index d8cf883d..2fe4d491 100644 --- a/Cslib/Foundations/Syntax/HasAlphaEquiv.lean +++ b/Cslib/Foundations/Syntax/HasAlphaEquiv.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi -/ +import Cslib.Init + namespace Cslib /-- Typeclass for the α-equivalence notation `x =α y`. -/ diff --git a/Cslib/Foundations/Syntax/HasSubstitution.lean b/Cslib/Foundations/Syntax/HasSubstitution.lean index 7897a590..5ec1ce84 100644 --- a/Cslib/Foundations/Syntax/HasSubstitution.lean +++ b/Cslib/Foundations/Syntax/HasSubstitution.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi -/ +import Cslib.Init + namespace Cslib /-- Typeclass for substitution relations and access to their notation. -/ diff --git a/Cslib/Foundations/Syntax/HasWellFormed.lean b/Cslib/Foundations/Syntax/HasWellFormed.lean index 70762b93..0db20b2e 100644 --- a/Cslib/Foundations/Syntax/HasWellFormed.lean +++ b/Cslib/Foundations/Syntax/HasWellFormed.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi -/ +import Cslib.Init + namespace Cslib /-- Typeclass for types equipped with a well-formedness predicate. -/ diff --git a/Cslib/Init.lean b/Cslib/Init.lean new file mode 100644 index 00000000..be6a30a3 --- /dev/null +++ b/Cslib/Init.lean @@ -0,0 +1,18 @@ +/- +Copyright (c) 2025 Jesse Alama. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jesse Alama +-/ + +import Mathlib.Init +import Cslib.Foundations.Lint.Basic + +/-! +# CSLib Initialization + +This is the root file in CSLib: it is imported by virtually *all* CSLib files. +For this reason, the imports of this file are carefully curated. + +Similar to Mathlib.Init, this file imports linters that should be active by default +throughout the CSLib library. +-/ diff --git a/Cslib/Languages/CCS/Basic.lean b/Cslib/Languages/CCS/Basic.lean index d8db8941..1c3a099c 100644 --- a/Cslib/Languages/CCS/Basic.lean +++ b/Cslib/Languages/CCS/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi -/ +import Cslib.Init import Mathlib.Init /-! # Calculus of Communicating Systems (CCS) diff --git a/Cslib/Logics/LinearLogic/CLL/Basic.lean b/Cslib/Logics/LinearLogic/CLL/Basic.lean index 55630563..b9fd10f1 100644 --- a/Cslib/Logics/LinearLogic/CLL/Basic.lean +++ b/Cslib/Logics/LinearLogic/CLL/Basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fabrizio Montesi -/ +-- TODO: Syntax elaboration conflicts +-- import Cslib.Init import Aesop import Mathlib.Tactic.ApplyAt import Mathlib.Order.Notation diff --git a/CslibTests.lean b/CslibTests.lean index 56550ac8..257e5c19 100644 --- a/CslibTests.lean +++ b/CslibTests.lean @@ -3,7 +3,7 @@ import CslibTests.CCS import CslibTests.DFA import CslibTests.FreeMonad import CslibTests.HasFresh +import CslibTests.LTS import CslibTests.LambdaCalculus import CslibTests.Lint -import CslibTests.LTS import CslibTests.ReductionSystem diff --git a/lakefile.toml b/lakefile.toml index b6e6928d..b89f5bd9 100644 --- a/lakefile.toml +++ b/lakefile.toml @@ -1,7 +1,7 @@ name = "cslib" version = "0.1.0" defaultTargets = ["Cslib"] -testDriver = "CslibTests" +testDriver = "runTests" [leanOptions] weak.linter.mathlibStandardSet = true @@ -25,3 +25,13 @@ globs = ["Cslib.*"] [[lean_lib]] name = "CslibTests" globs = ["CslibTests.+"] + +[[lean_exe]] +name = "checkInitImports" +srcDir = "scripts" +root = "CheckInitImports" + +[[lean_exe]] +name = "runTests" +srcDir = "scripts" +root = "RunTests" diff --git a/scripts/CheckInitImports.lean b/scripts/CheckInitImports.lean new file mode 100755 index 00000000..00dad0f6 --- /dev/null +++ b/scripts/CheckInitImports.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2025 Jesse Alama. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jesse Alama, Chris Henson +-/ + +import Lean +import Mathlib.Lean.CoreM +import Batteries.Data.List.Basic +import ImportGraph.Imports + +open Lean Core Elab Command + +/-! +# Check Init Imports + +This script checks that all CSLib modules (transitively) import Cslib.Init. +-/ + +/-- Modules with technical constraints preventing Cslib.Init import -/ +def exceptions : List Name := [ + -- Circular dependency (imported by Cslib.Init) + `Cslib.Foundations.Lint.Basic, + `Cslib.Init, + + -- Type elaboration issues in downstream files + `Cslib.Foundations.Semantics.LTS.Basic, + `Cslib.Foundations.Semantics.LTS.TraceEq, + `Cslib.Computability.Automata.DA, + `Cslib.Computability.Automata.DFA, + `Cslib.Computability.Automata.DFAToNFA, + `Cslib.Computability.Automata.EpsilonNFA, + `Cslib.Computability.Automata.EpsilonNFAToNFA, + `Cslib.Computability.Automata.NA, + `Cslib.Computability.Automata.NFA, + `Cslib.Computability.Automata.NFAToDFA, + + -- Notation conflict with Mathlib.Finsupp (→₀) + `Cslib.Foundations.Data.FinFun, + + -- Syntax elaboration conflicts + `Cslib.Logics.LinearLogic.CLL.Basic, + `Cslib.Logics.LinearLogic.CLL.CutElimination, + `Cslib.Logics.LinearLogic.CLL.MProof, + `Cslib.Logics.LinearLogic.CLL.PhaseSemantics.Basic, +] + +def main : IO UInt32 := do + let searchPath ← addSearchPathFromEnv (← getBuiltinSearchPath (← findSysroot)) + CoreM.withImportModules #[`Cslib] (searchPath := searchPath) (trustLevel := 1024) do + let env ← getEnv + let graph := env.importGraph.transitiveClosure + let noInitGraph := + graph.filter (fun name imports => name.getRoot = `Cslib ∧ !imports.contains `Cslib.Init) + let diff := noInitGraph.keys.diff exceptions + if diff.length > 0 then + IO.eprintln s!"error: the following modules do not import `Cslib.Init`: {diff}" + return diff.length.toUInt32 diff --git a/scripts/RunTests.lean b/scripts/RunTests.lean new file mode 100644 index 00000000..1cd4de5a --- /dev/null +++ b/scripts/RunTests.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2025 Jesse Alama. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jesse Alama +-/ + +/-! +# Test Runner + +This script runs all CSLib tests including: +- Building the CslibTests library +- Running init imports validation +-/ + +def main (_ : List String) : IO UInt32 := do + -- build CslibTests + IO.println "building CslibTests..." + let CslibTestsResult ← IO.Process.spawn { + cmd := "lake" + args := #["build", "CslibTests"] + } >>= (·.wait) + + if CslibTestsResult != 0 then + IO.eprintln "\n✗ CslibTests build failed" + return CslibTestsResult + else + println! "" + + -- Run init imports check + IO.println "\nChecking init imports..." + let checkResult ← IO.Process.spawn { + cmd := "lake" + args := #["exe", "checkInitImports"] + } >>= (·.wait) + + if checkResult != 0 then + IO.eprintln "\n✗ Init imports check failed" + return checkResult + + IO.println "\n✓ All tests passed" + return 0