Skip to content

Commit 41bca5f

Browse files
jamesmckinnaMatthewDaggitt
authored andcommitted
Deprecate inspect (#1930)
1 parent 784aee6 commit 41bca5f

File tree

12 files changed

+99
-65
lines changed

12 files changed

+99
-65
lines changed

CHANGELOG.md

+13-1
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
Version 2.0-dev
22
===============
33

4-
The library has been tested using Agda 2.6.2.
4+
The library has been tested using Agda 2.6.3.
55

66
Highlights
77
----------
@@ -1334,6 +1334,15 @@ Deprecated names
13341334
* This fixes the fact we had picked the wrong name originally. The erased modality
13351335
corresponds to @0 whereas the irrelevance one corresponds to `.`.
13361336

1337+
### Deprecated `Relation.Binary.PropositionalEquality.inspect`
1338+
in favour of `with ... in ...` syntax (issue #1580; PRs #1630, #1930)
1339+
1340+
* In `Relation.Binary.PropositionalEquality`
1341+
both the record type `Reveal_·_is_`
1342+
and its principal mode of use, `inspect`,
1343+
have been deprecated in favour of the new `with ... in ...` syntax.
1344+
See the documentation of [with-abstraction equality](https://agda.readthedocs.io/en/v2.6.3/language/with-abstraction.html#with-abstraction-equality)
1345+
13371346
New modules
13381347
-----------
13391348

@@ -1891,6 +1900,9 @@ Other minor changes
18911900
toℕ-cancel-≤ : toℕ i ℕ.≤ toℕ j → i ≤ j
18921901
toℕ-cancel-< : toℕ i ℕ.< toℕ j → i < j
18931902
1903+
splitAt⁻¹-↑ˡ : splitAt m {n} i ≡ inj₁ j → j ↑ˡ n ≡ i
1904+
splitAt⁻¹-↑ʳ : splitAt m {n} i ≡ inj₂ j → m ↑ʳ j ≡ i
1905+
18941906
toℕ-combine : toℕ (combine i j) ≡ k ℕ.* toℕ i ℕ.+ toℕ j
18951907
combine-injectiveˡ : combine i j ≡ combine k l → i ≡ k
18961908
combine-injectiveʳ : combine i j ≡ combine k l → j ≡ l

README.agda

+6-10
Original file line numberDiff line numberDiff line change
@@ -12,13 +12,14 @@ module README where
1212
-- Liang-Ting Chen, Dominique Devriese, Dan Doel, Érdi Gergő,
1313
-- Zack Grannan, Helmut Grohne, Simon Foster, Liyang Hu, Jason Hu,
1414
-- Patrik Jansson, Alan Jeffrey, Wen Kokke, Evgeny Kotelnikov,
15-
-- Sergei Meshveliani, Eric Mertens, Darin Morrison, Guilhem Moulin,
16-
-- Shin-Cheng Mu, Ulf Norell, Noriyuki Ohkawa, Nicolas Pouillard,
17-
-- Andrés Sicard-Ramírez, Lex van der Stoep, Sandro Stucki, Milo Turner,
18-
-- Noam Zeilberger and other anonymous contributors.
15+
-- James McKinna, Sergei Meshveliani, Eric Mertens, Darin Morrison,
16+
-- Guilhem Moulin, Shin-Cheng Mu, Ulf Norell, Noriyuki Ohkawa,
17+
-- Nicolas Pouillard, Andrés Sicard-Ramírez, Lex van der Stoep,
18+
-- Sandro Stucki, Milo Turner, Noam Zeilberger
19+
-- and other anonymous contributors.
1920
------------------------------------------------------------------------
2021

21-
-- This version of the library has been tested using Agda 2.6.2.
22+
-- This version of the library has been tested using Agda 2.6.3.
2223

2324
-- The library comes with a .agda-lib file, for use with the library
2425
-- management system.
@@ -236,11 +237,6 @@ import README.Debug.Trace
236237

237238
import README.Nary
238239

239-
-- Explaining the inspect idiom: use case, equivalent handwritten
240-
-- auxiliary definitions, and implementation details.
241-
242-
import README.Inspect
243-
244240
-- Explaining how to use the automatic solvers
245241

246242
import README.Tactic.MonoidSolver

README/Inspect.agda

+5-2
Original file line numberDiff line numberDiff line change
@@ -1,14 +1,17 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Explaining how to use the inspect idiom and elaborating on the way
5-
-- it is implemented in the standard library.
4+
-- This module is DEPRECATED.
65
------------------------------------------------------------------------
76

87
{-# OPTIONS --cubical-compatible --safe #-}
98

109
module README.Inspect where
1110

11+
{-# WARNING_ON_IMPORT
12+
"README.Inspect was deprecated in v2.0."
13+
#-}
14+
1215
open import Data.Nat.Base
1316
open import Data.Nat.Properties
1417
open import Data.Product

src/Codata/Sized/Colist/Properties.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -144,10 +144,10 @@ module _ (cons : C → B → C) (alg : A → Maybe (A × B)) where
144144

145145
scanl-unfold : nil a i ⊢ scanl cons nil (unfold alg a)
146146
≈ nil ∷ (λ where .force unfold alg′ (a , nil))
147-
scanl-unfold nil a with alg a | Eq.inspect alg a
148-
... | nothing | [ eq ] = Eq.refl ∷ λ { .force
147+
scanl-unfold nil a with alg a in eq
148+
... | nothing = Eq.refl ∷ λ { .force
149149
sym (fromEq (unfold-nothing (Maybeₚ.map-nothing eq))) }
150-
... | just (a′ , b) | [ eq ] = Eq.refl ∷ λ { .force begin
150+
... | just (a′ , b) = Eq.refl ∷ λ { .force begin
151151
scanl cons (cons nil b) (unfold alg a′)
152152
≈⟨ scanl-unfold (cons nil b) a′ ⟩
153153
(cons nil b ∷ _)

src/Data/Fin/Properties.agda

+15-5
Original file line numberDiff line numberDiff line change
@@ -562,10 +562,20 @@ splitAt-↑ˡ : ∀ m i n → splitAt m (i ↑ˡ n) ≡ inj₁ i
562562
splitAt-↑ˡ (suc m) zero n = refl
563563
splitAt-↑ˡ (suc m) (suc i) n rewrite splitAt-↑ˡ m i n = refl
564564

565+
splitAt⁻¹-↑ˡ : {m} {n} {i} {j} splitAt m {n} i ≡ inj₁ j j ↑ˡ n ≡ i
566+
splitAt⁻¹-↑ˡ {suc m} {n} {0F} {.0F} refl = refl
567+
splitAt⁻¹-↑ˡ {suc m} {n} {suc i} {j} eq with splitAt m i in splitAt[m][i]≡inj₁[j]
568+
... | inj₁ k with refl eq = cong suc (splitAt⁻¹-↑ˡ {i = i} {j = k} splitAt[m][i]≡inj₁[j])
569+
565570
splitAt-↑ʳ : m n i splitAt m (m ↑ʳ i) ≡ inj₂ {B = Fin n} i
566571
splitAt-↑ʳ zero n i = refl
567572
splitAt-↑ʳ (suc m) n i rewrite splitAt-↑ʳ m n i = refl
568573

574+
splitAt⁻¹-↑ʳ : {m} {n} {i} {j} splitAt m {n} i ≡ inj₂ j m ↑ʳ j ≡ i
575+
splitAt⁻¹-↑ʳ {zero} {n} {i} {j} refl = refl
576+
splitAt⁻¹-↑ʳ {suc m} {n} {suc i} {j} eq with splitAt m i in splitAt[m][i]≡inj₂[k]
577+
... | inj₂ k with refl eq = cong suc (splitAt⁻¹-↑ʳ {i = i} {j = k} splitAt[m][i]≡inj₂[k])
578+
569579
splitAt-join : m n i splitAt m (join m n i) ≡ i
570580
splitAt-join m n (inj₁ x) = splitAt-↑ˡ m x n
571581
splitAt-join m n (inj₂ y) = splitAt-↑ʳ m n y
@@ -612,13 +622,13 @@ remQuot-combine {suc n} {k} (suc i) j rewrite splitAt-↑ʳ k (n ℕ.* k) (com
612622
cong (Data.Product.map₁ suc) (remQuot-combine i j)
613623

614624
combine-remQuot : {n} k (i : Fin (n ℕ.* k)) uncurry combine (remQuot {n} k i) ≡ i
615-
combine-remQuot {suc n} k i with splitAt k i | P.inspect (splitAt k) i
616-
... | inj₁ j | P.[ eq ] = begin
625+
combine-remQuot {suc n} k i with splitAt k i in eq
626+
... | inj₁ j = begin
617627
join k (n ℕ.* k) (inj₁ j) ≡˘⟨ cong (join k (n ℕ.* k)) eq ⟩
618628
join k (n ℕ.* k) (splitAt k i) ≡⟨ join-splitAt k (n ℕ.* k) i ⟩
619629
i ∎
620630
where open ≡-Reasoning
621-
... | inj₂ j | P.[ eq ] = begin
631+
... | inj₂ j = begin
622632
k ↑ʳ (uncurry combine (remQuot {n} k j)) ≡⟨ cong (k ↑ʳ_) (combine-remQuot {n} k j) ⟩
623633
join k (n ℕ.* k) (inj₂ j) ≡˘⟨ cong (join k (n ℕ.* k)) eq ⟩
624634
join k (n ℕ.* k) (splitAt k i) ≡⟨ join-splitAt k (n ℕ.* k) i ⟩
@@ -679,8 +689,8 @@ combine-injective i j k l cᵢⱼ≡cₖₗ =
679689
combine-injectiveʳ i j k l cᵢⱼ≡cₖₗ
680690

681691
combine-surjective : (i : Fin (m ℕ.* n)) ∃₂ λ j k combine j k ≡ i
682-
combine-surjective {m} {n} i with remQuot {m} n i | P.inspect (remQuot {m} n) i
683-
... | j , k | P.[ eq ] = j , k , (begin
692+
combine-surjective {m} {n} i with remQuot {m} n i in eq
693+
... | j , k = j , k , (begin
684694
combine j k ≡˘⟨ uncurry (cong₂ combine) (,-injective eq) ⟩
685695
uncurry combine (remQuot {m} n i) ≡⟨ combine-remQuot {m} n i ⟩
686696
i ∎)

src/Data/List/Properties.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -854,9 +854,9 @@ module _ {P : Pred A p} (P? : Decidable P) where
854854

855855
filter-idem : filter P? ∘ filter P? ≗ filter P?
856856
filter-idem [] = refl
857-
filter-idem (x ∷ xs) with does (P? x) | inspect does (P? x)
858-
... | false | _ = filter-idem xs
859-
... | true | P.[ eq ] rewrite eq = cong (x ∷_) (filter-idem xs)
857+
filter-idem (x ∷ xs) with does (P? x) in eq
858+
... | false = filter-idem xs
859+
... | true rewrite eq = cong (x ∷_) (filter-idem xs)
860860

861861
filter-++ : xs ys filter P? (xs ++ ys) ≡ filter P? xs ++ filter P? ys
862862
filter-++ [] ys = refl

src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ open import Level using (Level)
2929
open import Relation.Unary using (Pred)
3030
open import Relation.Binary
3131
open import Relation.Binary.PropositionalEquality as ≡
32-
using (_≡_ ; refl ; cong; cong₂; _≢_; inspect)
32+
using (_≡_ ; refl ; cong; cong₂; _≢_)
3333
open import Relation.Nullary
3434

3535
open PermutationReasoning

src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -37,7 +37,7 @@ open import Level using (Level; _⊔_)
3737
open import Relation.Unary using (Pred; Decidable)
3838
open import Relation.Binary.Properties.Setoid S using (≉-resp₂)
3939
open import Relation.Binary.PropositionalEquality as ≡
40-
using (_≡_ ; refl; sym; cong; cong₂; subst; _≢_; inspect)
40+
using (_≡_ ; refl; sym; cong; cong₂; subst; _≢_)
4141
open import Relation.Nullary.Decidable using (yes; no; does)
4242
open import Relation.Nullary.Negation using (contradiction)
4343

src/Data/List/Relation/Unary/Any/Properties.agda

+4-4
Original file line numberDiff line numberDiff line change
@@ -42,7 +42,7 @@ open import Function.Related as Related using (Kind; Related; SK-sym)
4242
open import Level using (Level)
4343
open import Relation.Binary as B hiding (_⇔_)
4444
open import Relation.Binary.PropositionalEquality as P
45-
using (_≡_; refl; inspect)
45+
using (_≡_; refl)
4646
open import Relation.Unary as U
4747
using (Pred; _⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_)
4848
open import Relation.Nullary using (¬_; _because_; does; ofʸ; ofⁿ; yes; no)
@@ -172,9 +172,9 @@ any⁺ p (there {x = x} pxs) with p x
172172
... | false = any⁺ p pxs
173173

174174
any⁻ : (p : A Bool) xs T (any p xs) Any (T ∘ p) xs
175-
any⁻ p (x ∷ xs) px∷xs with p x | inspect p x
176-
... | true | P.[ eq ] = here (Equivalence.from T-≡ ⟨$⟩ eq)
177-
... | false | _ = there (any⁻ p xs px∷xs)
175+
any⁻ p (x ∷ xs) px∷xs with p x in eq
176+
... | true = here (Equivalence.from T-≡ ⟨$⟩ eq)
177+
... | false = there (any⁻ p xs px∷xs)
178178

179179
any⇔ : {p : A Bool} Any (T ∘ p) xs ⇔ T (any p xs)
180180
any⇔ = equivalence (any⁺ _) (any⁻ _ _)

src/Data/Rational/Properties.agda

+12-12
Original file line numberDiff line numberDiff line change
@@ -1304,24 +1304,24 @@ neg-distribʳ-* = +-*-Monomorphism.neg-distribʳ-* ℚᵘ.+-0-isGroup ℚᵘ.*-i
13041304
------------------------------------------------------------------------
13051305

13061306
p≤q⇒p⊔q≡q : p ≤ q p ⊔ q ≡ q
1307-
p≤q⇒p⊔q≡q {p@record{}} {q@record{}} p≤q with p ≤ᵇ q | inspect (p ≤ᵇ_) q
1308-
... | true | _ = refl
1309-
... | false | [ p≰q ] = contradiction (≤⇒≤ᵇ p≤q) (subst (¬_ ∘ T) (sym p≰q) λ())
1307+
p≤q⇒p⊔q≡q {p@record{}} {q@record{}} p≤q with p ≤ᵇ q in p≰q
1308+
... | true = refl
1309+
... | false = contradiction (≤⇒≤ᵇ p≤q) (subst (¬_ ∘ T) (sym p≰q) λ())
13101310

13111311
p≥q⇒p⊔q≡p : p ≥ q p ⊔ q ≡ p
1312-
p≥q⇒p⊔q≡p {p@record{}} {q@record{}} p≥q with p ≤ᵇ q | inspect (p ≤ᵇ_) q
1313-
... | true | [ p≤q ] = ≤-antisym p≥q (≤ᵇ⇒≤ (subst T (sym p≤q) _))
1314-
... | false | [ p≤q ] = refl
1312+
p≥q⇒p⊔q≡p {p@record{}} {q@record{}} p≥q with p ≤ᵇ q in p≤q
1313+
... | true = ≤-antisym p≥q (≤ᵇ⇒≤ (subst T (sym p≤q) _))
1314+
... | false = refl
13151315

13161316
p≤q⇒p⊓q≡p : p ≤ q p ⊓ q ≡ p
1317-
p≤q⇒p⊓q≡p {p@record{}} {q@record{}} p≤q with p ≤ᵇ q | inspect (p ≤ᵇ_) q
1318-
... | true | _ = refl
1319-
... | false | [ p≰q ] = contradiction (≤⇒≤ᵇ p≤q) (subst (¬_ ∘ T) (sym p≰q) λ())
1317+
p≤q⇒p⊓q≡p {p@record{}} {q@record{}} p≤q with p ≤ᵇ q in p≰q
1318+
... | true = refl
1319+
... | false = contradiction (≤⇒≤ᵇ p≤q) (subst (¬_ ∘ T) (sym p≰q) λ())
13201320

13211321
p≥q⇒p⊓q≡q : p ≥ q p ⊓ q ≡ q
1322-
p≥q⇒p⊓q≡q {p@record{}} {q@record{}} p≥q with p ≤ᵇ q | inspect (p ≤ᵇ_) q
1323-
... | true | [ p≤q ] = ≤-antisym (≤ᵇ⇒≤ (subst T (sym p≤q) _)) p≥q
1324-
... | false | [ p≤q ] = refl
1322+
p≥q⇒p⊓q≡q {p@record{}} {q@record{}} p≥q with p ≤ᵇ q in p≤q
1323+
... | true = ≤-antisym (≤ᵇ⇒≤ (subst T (sym p≤q) _)) p≥q
1324+
... | false = refl
13251325

13261326
⊓-operator : MinOperator ≤-totalPreorder
13271327
⊓-operator = record

src/Data/Vec/Relation/Unary/Any/Properties.agda

+5-5
Original file line numberDiff line numberDiff line change
@@ -333,12 +333,12 @@ module _ {P : Pred A p} where
333333

334334
concat⁺∘concat⁻ : {n m} (xss : Vec (Vec A n) m) (p : Any P (concat xss))
335335
concat⁺ (concat⁻ xss p) ≡ p
336-
concat⁺∘concat⁻ (xs ∷ xss) p with ++⁻ xs p | P.inspect (++⁻ xs) p
337-
... | inj₁ pxs | P.[ p=inj₁ ]
338-
= P.trans (P.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (P.sym p=inj₁))
336+
concat⁺∘concat⁻ (xs ∷ xss) p with ++⁻ xs p in eq
337+
... | inj₁ pxs
338+
= P.trans (P.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (P.sym eq))
339339
$ ++⁺∘++⁻ xs p
340-
... | inj₂ pxss | P.[ p=inj₂ ] rewrite concat⁺∘concat⁻ xss pxss
341-
= P.trans (P.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (P.sym p=inj₂))
340+
... | inj₂ pxss rewrite concat⁺∘concat⁻ xss pxss
341+
= P.trans (P.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (P.sym eq))
342342
$ ++⁺∘++⁻ xs p
343343

344344
concat⁻∘concat⁺ : {n m} {xss : Vec (Vec A n) m} (p : Any (Any P) xss)

src/Relation/Binary/PropositionalEquality.agda

+31-18
Original file line numberDiff line numberDiff line change
@@ -60,24 +60,6 @@ _≗_ {A = A} {B = B} = Setoid._≈_ (A →-setoid B)
6060
(A Setoid.Carrier B) setoid A ⟶ B
6161
→-to-⟶ = :→-to-Π
6262

63-
------------------------------------------------------------------------
64-
-- Inspect
65-
66-
-- Inspect can be used when you want to pattern match on the result r
67-
-- of some expression e, and you also need to "remember" that r ≡ e.
68-
69-
-- See README.Inspect for an explanation of how/why to use this.
70-
71-
record Reveal_·_is_ {A : Set a} {B : A Set b}
72-
(f : (x : A) B x) (x : A) (y : B x) :
73-
Set (a ⊔ b) where
74-
constructor [_]
75-
field eq : f x ≡ y
76-
77-
inspect : {A : Set a} {B : A Set b}
78-
(f : (x : A) B x) (x : A) Reveal f · x is f x
79-
inspect f x = [ refl ]
80-
8163
------------------------------------------------------------------------
8264
-- Propositionality
8365

@@ -120,3 +102,34 @@ module _ (_≟_ : DecidableEquality A) {x y : A} where
120102

121103
≢-≟-identity : (x≢y : x ≢ y) x ≟ y ≡ no x≢y
122104
≢-≟-identity = dec-no (x ≟ y)
105+
106+
107+
108+
109+
------------------------------------------------------------------------
110+
-- DEPRECATED NAMES
111+
------------------------------------------------------------------------
112+
-- Please use the new names as continuing support for the old names is
113+
-- not guaranteed.
114+
115+
-- Version 2.0
116+
117+
record Reveal_·_is_ {A : Set a} {B : A Set b}
118+
(f : (x : A) B x) (x : A) (y : B x) :
119+
Set (a ⊔ b) where
120+
constructor [_]
121+
field eq : f x ≡ y
122+
123+
inspect : {A : Set a} {B : A Set b}
124+
(f : (x : A) B x) (x : A) Reveal f · x is f x
125+
inspect f x = [ refl ]
126+
127+
{-# WARNING_ON_USAGE Reveal_·_is_
128+
"Warning: Reveal_·_is_ was deprecated in v2.0.
129+
Please use new `with ... in` syntax described at https://agda.readthedocs.io/en/v2.6.3/language/with-abstraction.html#with-abstraction-equality instead."
130+
#-}
131+
{-# WARNING_ON_USAGE inspect
132+
"Warning: inspect was deprecated in v2.0.
133+
Please use new `with ... in` syntax described at https://agda.readthedocs.io/en/v2.6.3/language/with-abstraction.html#with-abstraction-equality instead."
134+
#-}
135+

0 commit comments

Comments
 (0)