Skip to content

Commit 6e23886

Browse files
alexaricegallais
andauthored
Fix deprecated modules (#2224)
* Fix deprecated modules * [ ci ] Also build deprecated modules * [ ci ] ignore user warnings in test * [ ci ] fix filtering logic Deprecation and safety are not the same thing --------- Co-authored-by: Guillaume Allais <[email protected]>
1 parent 42b2a1a commit 6e23886

File tree

6 files changed

+37
-27
lines changed

6 files changed

+37
-27
lines changed

.github/workflows/ci-ubuntu.yml

+9-1
Original file line numberDiff line numberDiff line change
@@ -151,10 +151,19 @@ jobs:
151151

152152
- name: Test stdlib
153153
run: |
154+
# Including deprecated modules purely for testing
155+
cabal run GenerateEverything -- --include-deprecated
156+
${{ env.AGDA }} -WnoUserWarning --safe EverythingSafe.agda
157+
${{ env.AGDA }} -WnoUserWarning Everything.agda
158+
159+
- name: Prepare HTML index
160+
run: |
161+
# Regenerating the Everything files without the deprecated modules
154162
cabal run GenerateEverything
155163
cp .github/tooling/* .
156164
./index.sh
157165
${{ env.AGDA }} --safe EverythingSafe.agda
166+
${{ env.AGDA }} Everything.agda
158167
${{ env.AGDA }} index.agda
159168
160169
- name: Golden testing
@@ -177,7 +186,6 @@ jobs:
177186
rm -f '${{ env.AGDA_HTML_DIR }}'/*.html
178187
rm -f '${{ env.AGDA_HTML_DIR }}'/*.css
179188
${{ env.AGDA }} --html --html-dir ${{ env.AGDA_HTML_DIR }} index.agda
180-
181189
cp .github/tooling/* .
182190
./landing.sh
183191

GenerateEverything.hs

+18-14
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,7 @@
11
{-# LANGUAGE PatternGuards #-}
22
{-# LANGUAGE PatternSynonyms #-}
33
{-# LANGUAGE RecordWildCards #-}
4+
{-# LANGUAGE MultiWayIf #-}
45

56
import Control.Applicative
67
import Control.Monad
@@ -227,10 +228,10 @@ extractHeader mod = extract
227228

228229
-- | A crude classifier looking for lines containing options
229230

230-
data Status = Deprecated | Unsafe | Safe
231-
deriving (Eq)
231+
data Safety = Unsafe | Safe deriving (Eq)
232+
data Status = Deprecated | Active deriving (Eq)
232233

233-
classify :: FilePath -> [String] -> [String] -> Exc Status
234+
classify :: FilePath -> [String] -> [String] -> Exc (Safety, Status)
234235
classify fp hd ls
235236
-- We start with sanity checks
236237
| isUnsafe && safe = throwError $ fp ++ contradiction "unsafe" "safe"
@@ -239,11 +240,12 @@ classify fp hd ls
239240
| isWithK && not withK = throwError $ fp ++ missingWithK
240241
| not (isWithK || cubicalC) = throwError $ fp ++ uncategorized "as relying on K" "cubical-compatible"
241242
-- And then perform the actual classification
242-
| deprecated = pure $ Deprecated
243-
| isUnsafe = pure $ Unsafe
244-
| safe = pure $ Safe
245-
-- We know that @not (isUnsafe || safe)@, all cases are covered
246-
| otherwise = error "IMPOSSIBLE"
243+
| otherwise = do
244+
let safety = if | safe -> Safe
245+
| isUnsafe -> Unsafe
246+
| otherwise -> error "IMPOSSIBLE"
247+
let status = if deprecated then Deprecated else Active
248+
pure (safety, status)
247249

248250
where
249251

@@ -280,18 +282,20 @@ classify fp hd ls
280282
data LibraryFile = LibraryFile
281283
{ filepath :: FilePath -- ^ FilePath of the source file
282284
, header :: [String] -- ^ All lines in the headers are already prefixed with \"-- \".
283-
, status :: Status -- ^ Safety options used by the module
285+
, safety :: Safety
286+
, status :: Status -- ^ Deprecation status options used by the module
284287
}
285288

286289
analyse :: FilePath -> IO LibraryFile
287290
analyse fp = do
288291
ls <- lines <$> readFileUTF8 fp
289292
hd <- runExc $ extractHeader fp ls
290-
cl <- runExc $ classify fp hd ls
293+
(sf, st) <- runExc $ classify fp hd ls
291294
return $ LibraryFile
292-
{ filepath = fp
293-
, header = hd
294-
, status = cl
295+
{ filepath = fp
296+
, header = hd
297+
, safety = sf
298+
, status = st
295299
}
296300

297301
checkFilePaths :: String -> [FilePath] -> IO ()
@@ -356,7 +360,7 @@ main = do
356360
unlines [ header
357361
, "{-# OPTIONS --safe --guardedness #-}\n"
358362
, mkModule safeOutputFile
359-
, format $ filter ((Unsafe /=) . status) libraryfiles
363+
, format $ filter ((Unsafe /=) . safety) libraryfiles
360364
]
361365

362366
-- | Usage info.

src/Algebra/Operations/Semiring.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,6 @@ open Semiring S
2626
-- Re-exports
2727

2828
open MonoidOperations +-commutativeMonoid public
29-
open import Algebra.Properties.Semiring.Exponentiation S public
30-
open import Algebra.Properties.Semiring.Multiplication S public
31-
using (×1-homo-*; ×′1-homo-*)
29+
open import Algebra.Properties.Semiring.Exp S public
30+
open import Algebra.Properties.Semiring.Mult S public
31+
using (×1-homo-*)

src/Algebra/Properties/BooleanAlgebra.agda

+4-5
Original file line numberDiff line numberDiff line change
@@ -28,8 +28,7 @@ import Algebra.Properties.DistributiveLattice as DistribLatticeProperties
2828

2929
open import Algebra.Structures _≈_
3030
open import Relation.Binary
31-
open import Function.Equality using (_⟨$⟩_)
32-
open import Function.Equivalence using (_⇔_; module Equivalence)
31+
open import Function.Bundles using (module Equivalence; _⇔_)
3332
open import Data.Product.Base using (_,_)
3433

3534
------------------------------------------------------------------------
@@ -47,9 +46,9 @@ replace-equality {_≈′_} ≈⇔≈′ = record
4746
{ isBooleanAlgebra = record
4847
{ isDistributiveLattice = DistributiveLattice.isDistributiveLattice
4948
(DistribLatticeProperties.replace-equality distributiveLattice ≈⇔≈′)
50-
; ∨-complement = ((λ x to ⟨$⟩ ∨-complementˡ x) , λ x to ⟨$⟩ ∨-complementʳ x)
51-
; ∧-complement = ((λ x to ⟨$⟩ ∧-complementˡ x) , λ x to ⟨$⟩ ∧-complementʳ x)
52-
; ¬-cong = λ i≈j to ⟨$⟩ ¬-cong (from ⟨$⟩ i≈j)
49+
; ∨-complement = ((λ x to (∨-complementˡ x)) , λ x to (∨-complementʳ x))
50+
; ∧-complement = ((λ x to (∧-complementˡ x)) , λ x to (∧-complementʳ x))
51+
; ¬-cong = λ i≈j to (¬-cong (from i≈j))
5352
}
5453
} where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
5554
{-# WARNING_ON_USAGE replace-equality

src/Algebra/Properties/DistributiveLattice.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -12,8 +12,7 @@
1212
open import Algebra.Lattice.Bundles
1313
open import Algebra.Lattice.Structures.Biased
1414
open import Relation.Binary
15-
open import Function.Equality
16-
open import Function.Equivalence
15+
open import Function.Bundles using (module Equivalence; _⇔_)
1716
import Algebra.Construct.Subst.Equality as SubstEq
1817

1918
module Algebra.Properties.DistributiveLattice
@@ -44,7 +43,7 @@ replace-equality {_≈′_} ≈⇔≈′ = record
4443
{ isDistributiveLattice = isDistributiveLatticeʳʲᵐ (record
4544
{ isLattice = Lattice.isLattice
4645
(LatticeProperties.replace-equality lattice ≈⇔≈′)
47-
; ∨-distribʳ-∧ = λ x y z to ⟨$⟩ ∨-distribʳ-∧ x y z
46+
; ∨-distribʳ-∧ = λ x y z to (∨-distribʳ-∧ x y z)
4847
})
4948
} where open module E {x y} = Equivalence (≈⇔≈′ {x} {y})
5049
{-# WARNING_ON_USAGE replace-equality

src/Data/Fin/Substitution/Example.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -20,7 +20,7 @@ open import Data.Fin.Substitution.Lemmas
2020
open import Data.Nat.Base hiding (_/_)
2121
open import Data.Fin.Base using (Fin)
2222
open import Data.Vec.Base
23-
open import Relation.Binary.PropositionalEquality.Core as PropEq
23+
open import Relation.Binary.PropositionalEquality as PropEq
2424
using (_≡_; refl; sym; cong; cong₂)
2525
open PropEq.≡-Reasoning
2626
open import Relation.Binary.Construct.Closure.ReflexiveTransitive

0 commit comments

Comments
 (0)