Skip to content

Commit 78e11bd

Browse files
Make size parameter on 'replicate' explicit (agda#2120)
1 parent 49d89d2 commit 78e11bd

File tree

18 files changed

+173
-170
lines changed

18 files changed

+173
-170
lines changed

CHANGELOG.md

+52-50
Original file line numberDiff line numberDiff line change
@@ -603,6 +603,7 @@ Non-backwards compatible changes
603603
```agda
604604
WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _
605605
WfRec _<_ P x = ∀ {y} → y < x → P y
606+
```
606607

607608
### Change in the definition of `_≤″_` (issue #1919)
608609

@@ -924,7 +925,7 @@ Non-backwards compatible changes
924925
Data.Vec.Relation.Binary.Lex.NonStrict
925926
Relation.Binary.Reasoning.StrictPartialOrder
926927
Relation.Binary.Reasoning.PartialOrder
927-
```
928+
```
928929

929930
### Other
930931

@@ -1053,60 +1054,61 @@ Non-backwards compatible changes
10531054
`List` module i.e. `lookup` takes its container first and the index (whose type may
10541055
depend on the container value) second.
10551056
This affects modules:
1056-
```
1057-
Codata.Guarded.Stream
1058-
Codata.Guarded.Stream.Relation.Binary.Pointwise
1059-
Codata.Musical.Colist.Base
1060-
Codata.Musical.Colist.Relation.Unary.Any.Properties
1061-
Codata.Musical.Covec
1062-
Codata.Musical.Stream
1063-
Codata.Sized.Colist
1064-
Codata.Sized.Covec
1065-
Codata.Sized.Stream
1066-
Data.Vec.Relation.Unary.All
1067-
Data.Star.Environment
1068-
Data.Star.Pointer
1069-
Data.Star.Vec
1070-
Data.Trie
1071-
Data.Trie.NonEmpty
1072-
Data.Tree.AVL
1073-
Data.Tree.AVL.Indexed
1074-
Data.Tree.AVL.Map
1075-
Data.Tree.AVL.NonEmpty
1076-
Data.Vec.Recursive
1077-
Tactic.RingSolver
1078-
Tactic.RingSolver.Core.NatSet
1079-
```
1057+
```
1058+
Codata.Guarded.Stream
1059+
Codata.Guarded.Stream.Relation.Binary.Pointwise
1060+
Codata.Musical.Colist.Base
1061+
Codata.Musical.Colist.Relation.Unary.Any.Properties
1062+
Codata.Musical.Covec
1063+
Codata.Musical.Stream
1064+
Codata.Sized.Colist
1065+
Codata.Sized.Covec
1066+
Codata.Sized.Stream
1067+
Data.Vec.Relation.Unary.All
1068+
Data.Star.Environment
1069+
Data.Star.Pointer
1070+
Data.Star.Vec
1071+
Data.Trie
1072+
Data.Trie.NonEmpty
1073+
Data.Tree.AVL
1074+
Data.Tree.AVL.Indexed
1075+
Data.Tree.AVL.Map
1076+
Data.Tree.AVL.NonEmpty
1077+
Data.Vec.Recursive
1078+
Tactic.RingSolver
1079+
Tactic.RingSolver.Core.NatSet
1080+
```
10801081

1081-
* Moved & renamed from `Data.Vec.Relation.Unary.All`
1082-
to `Data.Vec.Relation.Unary.All.Properties`:
1083-
```
1084-
lookup ↦ lookup⁺
1085-
tabulate ↦ lookup⁻
1086-
```
1082+
* Moved & renamed from `Data.Vec.Relation.Unary.All`
1083+
to `Data.Vec.Relation.Unary.All.Properties`:
1084+
```
1085+
lookup ↦ lookup⁺
1086+
tabulate ↦ lookup⁻
1087+
```
10871088

1088-
* Renamed in `Data.Vec.Relation.Unary.Linked.Properties`
1089-
and `Codata.Guarded.Stream.Relation.Binary.Pointwise`:
1090-
```
1091-
lookup ↦ lookup⁺
1092-
```
1089+
* Renamed in `Data.Vec.Relation.Unary.Linked.Properties`
1090+
and `Codata.Guarded.Stream.Relation.Binary.Pointwise`:
1091+
```
1092+
lookup ↦ lookup⁺
1093+
```
10931094

1094-
* Added the following new definitions to `Data.Vec.Relation.Unary.All`:
1095-
```
1096-
lookupAny : All P xs → (i : Any Q xs) → (P ∩ Q) (Any.lookup i)
1097-
lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.lookup i)
1098-
lookup : All P xs → (∀ {x} → x ∈ₚ xs → P x)
1099-
lookupₛ : P Respects _≈_ → All P xs → (∀ {x} → x ∈ xs → P x)
1100-
```
1095+
* Added the following new definitions to `Data.Vec.Relation.Unary.All`:
1096+
```
1097+
lookupAny : All P xs → (i : Any Q xs) → (P ∩ Q) (Any.lookup i)
1098+
lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.lookup i)
1099+
lookup : All P xs → (∀ {x} → x ∈ₚ xs → P x)
1100+
lookupₛ : P Respects _≈_ → All P xs → (∀ {x} → x ∈ xs → P x)
1101+
```
11011102

1102-
* `excluded-middle` in `Relation.Nullary.Decidable.Core` has been renamed to
1103-
`¬¬-excluded-middle`.
1103+
* `excluded-middle` in `Relation.Nullary.Decidable.Core` has been renamed to
1104+
`¬¬-excluded-middle`.
11041105

1105-
* `iterate` in `Data.Vec.Base` now takes `n` (the length of `Vec`) as an
1106-
explicit argument.
1107-
```agda
1108-
iterate : (A → A) → A → ∀ n → Vec A n
1109-
```
1106+
* `iterate` and `replicate` in `Data.Vec.Base` and `Data.Vec.Functional`
1107+
now take the length of vector, `n`, as an explicit rather than an implicit argument.
1108+
```agda
1109+
iterate : (A → A) → A → ∀ n → Vec A n
1110+
replicate : ∀ n → A → Vec A n
1111+
```
11101112

11111113
Major improvements
11121114
------------------

src/Algebra/Properties/Monoid/Sum.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -64,15 +64,15 @@ sum-cong-≗ : ∀ {n} → sum {n} Preserves _≗_ ⟶ _≡_
6464
sum-cong-≗ {zero} xs≗ys = P.refl
6565
sum-cong-≗ {suc n} xs≗ys = P.cong₂ _+_ (xs≗ys zero) (sum-cong-≗ (xs≗ys ∘ suc))
6666

67-
sum-replicate : n {x} sum {n} (replicate x) ≈ n × x
67+
sum-replicate : n {x} sum (replicate n x) ≈ n × x
6868
sum-replicate zero = refl
6969
sum-replicate (suc n) = +-congˡ (sum-replicate n)
7070

7171
sum-replicate-idem : {x} _+_ IdempotentOn x
72-
n .{{_ : NonZero n}} sum {n} (replicate x) ≈ x
72+
n .{{_ : NonZero n}} sum (replicate n x) ≈ x
7373
sum-replicate-idem idem n = trans (sum-replicate n) (×-idem idem n)
7474

75-
sum-replicate-zero : n sum {n} (replicate 0#) ≈ 0#
75+
sum-replicate-zero : n sum (replicate n 0#) ≈ 0#
7676
sum-replicate-zero zero = refl
7777
sum-replicate-zero (suc n) = sum-replicate-idem (+-identityˡ 0#) (suc n)
7878

src/Algebra/Solver/CommutativeMonoid.agda

+18-14
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,10 @@ open import Relation.Nullary.Decidable using (Dec)
3333
open CommutativeMonoid M
3434
open EqReasoning setoid
3535

36+
private
37+
variable
38+
n :
39+
3640
------------------------------------------------------------------------
3741
-- Monoid expressions
3842

@@ -55,7 +59,7 @@ Env n = Vec Carrier n
5559
-- The semantics of an expression is a function from an environment to
5660
-- a value.
5761

58-
⟦_⟧ : {n} Expr n Env n Carrier
62+
⟦_⟧ : Expr n Env n Carrier
5963
⟦ var x ⟧ ρ = lookup ρ x
6064
⟦ id ⟧ ρ = ε
6165
⟦ e₁ ⊕ e₂ ⟧ ρ = ⟦ e₁ ⟧ ρ ∙ ⟦ e₂ ⟧ ρ
@@ -70,27 +74,27 @@ Normal n = Vec ℕ n
7074

7175
-- The semantics of a normal form.
7276

73-
⟦_⟧⇓ : {n} Normal n Env n Carrier
77+
⟦_⟧⇓ : Normal n Env n Carrier
7478
⟦ [] ⟧⇓ _ = ε
75-
⟦ n ∷ v ⟧⇓ (a ∷ ρ) = fold (⟦ v ⟧⇓ ρ) (λ b a ∙ b) n
79+
⟦ n ∷ v ⟧⇓ (a ∷ ρ) = fold (⟦ v ⟧⇓ ρ) (a ∙_) n
7680

7781
------------------------------------------------------------------------
7882
-- Constructions on normal forms
7983

8084
-- The empty bag.
8185

82-
empty : {n} Normal n
83-
empty = replicate 0
86+
empty : Normal n
87+
empty = replicate _ 0
8488

8589
-- A singleton bag.
8690

87-
sg : {n} (i : Fin n) Normal n
91+
sg : (i : Fin n) Normal n
8892
sg zero = 1 ∷ empty
8993
sg (suc i) = 0 ∷ sg i
9094

9195
-- The composition of normal forms.
9296

93-
_•_ : {n} (v w : Normal n) Normal n
97+
_•_ : (v w : Normal n) Normal n
9498
[] • [] = []
9599
(l ∷ v) • (m ∷ w) = l + m ∷ v • w
96100

@@ -99,13 +103,13 @@ _•_ : ∀{n} (v w : Normal n) → Normal n
99103

100104
-- The empty bag stands for the unit ε.
101105

102-
empty-correct : {n} : Env n) ⟦ empty ⟧⇓ ρ ≈ ε
106+
empty-correct :: Env n) ⟦ empty ⟧⇓ ρ ≈ ε
103107
empty-correct [] = refl
104108
empty-correct (a ∷ ρ) = empty-correct ρ
105109

106110
-- The singleton bag stands for a single variable.
107111

108-
sg-correct : {n} (x : Fin n) (ρ : Env n) ⟦ sg x ⟧⇓ ρ ≈ lookup ρ x
112+
sg-correct : (x : Fin n) (ρ : Env n) ⟦ sg x ⟧⇓ ρ ≈ lookup ρ x
109113
sg-correct zero (x ∷ ρ) = begin
110114
x ∙ ⟦ empty ⟧⇓ ρ ≈⟨ ∙-congˡ (empty-correct ρ) ⟩
111115
x ∙ ε ≈⟨ identityʳ _ ⟩
@@ -114,7 +118,7 @@ sg-correct (suc x) (m ∷ ρ) = sg-correct x ρ
114118

115119
-- Normal form composition corresponds to the composition of the monoid.
116120

117-
comp-correct : {n} (v w : Normal n) (ρ : Env n)
121+
comp-correct : (v w : Normal n) (ρ : Env n)
118122
⟦ v • w ⟧⇓ ρ ≈ (⟦ v ⟧⇓ ρ ∙ ⟦ w ⟧⇓ ρ)
119123
comp-correct [] [] ρ = sym (identityˡ _)
120124
comp-correct (l ∷ v) (m ∷ w) (a ∷ ρ) = lemma l m (comp-correct v w ρ)
@@ -136,14 +140,14 @@ comp-correct (l ∷ v) (m ∷ w) (a ∷ ρ) = lemma l m (comp-correct v w ρ)
136140

137141
-- A normaliser.
138142

139-
normalise : {n} Expr n Normal n
143+
normalise : Expr n Normal n
140144
normalise (var x) = sg x
141145
normalise id = empty
142146
normalise (e₁ ⊕ e₂) = normalise e₁ • normalise e₂
143147

144148
-- The normaliser preserves the semantics of the expression.
145149

146-
normalise-correct : {n} (e : Expr n) (ρ : Env n)
150+
normalise-correct : (e : Expr n) (ρ : Env n)
147151
⟦ normalise e ⟧⇓ ρ ≈ ⟦ e ⟧ ρ
148152
normalise-correct (var x) ρ = sg-correct x ρ
149153
normalise-correct id ρ = empty-correct ρ
@@ -171,14 +175,14 @@ open module R = Reflection
171175

172176
infix 5 _≟_
173177

174-
_≟_ : {n} (nf₁ nf₂ : Normal n) Dec (nf₁ ≡ nf₂)
178+
_≟_ : (nf₁ nf₂ : Normal n) Dec (nf₁ ≡ nf₂)
175179
nf₁ ≟ nf₂ = Dec.map Pointwise-≡↔≡ (decidable ℕ._≟_ nf₁ nf₂)
176180
where open Pointwise
177181

178182
-- We can also give a sound, but not necessarily complete, procedure
179183
-- for determining if two expressions have the same semantics.
180184

181-
prove′ : {n} (e₁ e₂ : Expr n) Maybe ( ρ ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
185+
prove′ : (e₁ e₂ : Expr n) Maybe ( ρ ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
182186
prove′ e₁ e₂ =
183187
Maybe.map lemma (decToMaybe (normalise e₁ ≟ normalise e₂))
184188
where

src/Algebra/Solver/IdempotentCommutativeMonoid.agda

+16-12
Original file line numberDiff line numberDiff line change
@@ -34,6 +34,10 @@ module Algebra.Solver.IdempotentCommutativeMonoid
3434
open IdempotentCommutativeMonoid M
3535
open EqReasoning setoid
3636

37+
private
38+
variable
39+
n :
40+
3741
------------------------------------------------------------------------
3842
-- Monoid expressions
3943

@@ -71,7 +75,7 @@ Normal n = Vec Bool n
7175

7276
-- The semantics of a normal form.
7377

74-
⟦_⟧⇓ : {n} Normal n Env n Carrier
78+
⟦_⟧⇓ : Normal n Env n Carrier
7579
⟦ [] ⟧⇓ _ = ε
7680
⟦ b ∷ v ⟧⇓ (a ∷ ρ) = if b then a ∙ (⟦ v ⟧⇓ ρ) else (⟦ v ⟧⇓ ρ)
7781

@@ -80,18 +84,18 @@ Normal n = Vec Bool n
8084

8185
-- The empty set.
8286

83-
empty : {n} Normal n
84-
empty = replicate false
87+
empty : Normal n
88+
empty = replicate _ false
8589

8690
-- A singleton set.
8791

88-
sg : {n} (i : Fin n) Normal n
92+
sg : (i : Fin n) Normal n
8993
sg zero = true ∷ empty
9094
sg (suc i) = false ∷ sg i
9195

9296
-- The composition of normal forms.
9397

94-
_•_ : {n} (v w : Normal n) Normal n
98+
_•_ : (v w : Normal n) Normal n
9599
[] • [] = []
96100
(l ∷ v) • (m ∷ w) = (l ∨ m) ∷ v • w
97101

@@ -100,13 +104,13 @@ _•_ : ∀{n} (v w : Normal n) → Normal n
100104

101105
-- The empty set stands for the unit ε.
102106

103-
empty-correct : {n} : Env n) ⟦ empty ⟧⇓ ρ ≈ ε
107+
empty-correct :: Env n) ⟦ empty ⟧⇓ ρ ≈ ε
104108
empty-correct [] = refl
105109
empty-correct (a ∷ ρ) = empty-correct ρ
106110

107111
-- The singleton set stands for a single variable.
108112

109-
sg-correct : {n} (x : Fin n) (ρ : Env n) ⟦ sg x ⟧⇓ ρ ≈ lookup ρ x
113+
sg-correct : (x : Fin n) (ρ : Env n) ⟦ sg x ⟧⇓ ρ ≈ lookup ρ x
110114
sg-correct zero (x ∷ ρ) = begin
111115
x ∙ ⟦ empty ⟧⇓ ρ ≈⟨ ∙-congˡ (empty-correct ρ) ⟩
112116
x ∙ ε ≈⟨ identityʳ _ ⟩
@@ -132,7 +136,7 @@ distr a b c = begin
132136
a ∙ (b ∙ (a ∙ c)) ≈⟨ sym (assoc _ _ _) ⟩
133137
(a ∙ b) ∙ (a ∙ c) ∎
134138

135-
comp-correct : {n} (v w : Normal n) (ρ : Env n)
139+
comp-correct : (v w : Normal n) (ρ : Env n)
136140
⟦ v • w ⟧⇓ ρ ≈ (⟦ v ⟧⇓ ρ ∙ ⟦ w ⟧⇓ ρ)
137141
comp-correct [] [] ρ = sym (identityˡ _)
138142
comp-correct (true ∷ v) (true ∷ w) (a ∷ ρ) =
@@ -149,14 +153,14 @@ comp-correct (false ∷ v) (false ∷ w) (a ∷ ρ) =
149153

150154
-- A normaliser.
151155

152-
normalise : {n} Expr n Normal n
156+
normalise : Expr n Normal n
153157
normalise (var x) = sg x
154158
normalise id = empty
155159
normalise (e₁ ⊕ e₂) = normalise e₁ • normalise e₂
156160

157161
-- The normaliser preserves the semantics of the expression.
158162

159-
normalise-correct : {n} (e : Expr n) (ρ : Env n)
163+
normalise-correct : (e : Expr n) (ρ : Env n)
160164
⟦ normalise e ⟧⇓ ρ ≈ ⟦ e ⟧ ρ
161165
normalise-correct (var x) ρ = sg-correct x ρ
162166
normalise-correct id ρ = empty-correct ρ
@@ -184,14 +188,14 @@ open module R = Reflection
184188

185189
infix 5 _≟_
186190

187-
_≟_ : {n} (nf₁ nf₂ : Normal n) Dec (nf₁ ≡ nf₂)
191+
_≟_ : (nf₁ nf₂ : Normal n) Dec (nf₁ ≡ nf₂)
188192
nf₁ ≟ nf₂ = Dec.map Pointwise-≡↔≡ (decidable Bool._≟_ nf₁ nf₂)
189193
where open Pointwise
190194

191195
-- We can also give a sound, but not necessarily complete, procedure
192196
-- for determining if two expressions have the same semantics.
193197

194-
prove′ : {n} (e₁ e₂ : Expr n) Maybe ( ρ ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
198+
prove′ : (e₁ e₂ : Expr n) Maybe ( ρ ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ)
195199
prove′ e₁ e₂ =
196200
Maybe.map lemma (decToMaybe (normalise e₁ ≟ normalise e₂))
197201
where

src/Codata/Guarded/Stream/Properties.agda

+5-5
Original file line numberDiff line numberDiff line change
@@ -80,11 +80,11 @@ lookup-repeat : ∀ n (a : A) → lookup (repeat a) n ≡ a
8080
lookup-repeat zero a = P.refl
8181
lookup-repeat (suc n) a = lookup-repeat n a
8282

83-
splitAt-repeat : n (a : A) splitAt n (repeat a) ≡ (Vec.replicate a , repeat a)
83+
splitAt-repeat : n (a : A) splitAt n (repeat a) ≡ (Vec.replicate n a , repeat a)
8484
splitAt-repeat zero a = P.refl
8585
splitAt-repeat (suc n) a = cong (Prod.map₁ (a ∷_)) (splitAt-repeat n a)
8686

87-
take-repeat : n (a : A) take n (repeat a) ≡ Vec.replicate a
87+
take-repeat : n (a : A) take n (repeat a) ≡ Vec.replicate n a
8888
take-repeat n a = cong proj₁ (splitAt-repeat n a)
8989

9090
drop-repeat : n (a : A) drop n (repeat a) ≡ repeat a
@@ -115,17 +115,17 @@ zipWith-repeat : ∀ (f : A → B → C) a b →
115115
zipWith-repeat f a b .head = P.refl
116116
zipWith-repeat f a b .tail = zipWith-repeat f a b
117117

118-
chunksOf-repeat : n (a : A) chunksOf n (repeat a) ≈ repeat (Vec.replicate a)
118+
chunksOf-repeat : n (a : A) chunksOf n (repeat a) ≈ repeat (Vec.replicate n a)
119119
chunksOf-repeat n a = begin go where
120120

121121
open ≈-Reasoning
122122

123-
go : chunksOf n (repeat a) ≈∞ repeat (Vec.replicate a)
123+
go : chunksOf n (repeat a) ≈∞ repeat (Vec.replicate n a)
124124
go .head = take-repeat n a
125125
go .tail =
126126
chunksOf n (drop n (repeat a)) ≡⟨ P.cong (chunksOf n) (drop-repeat n a) ⟩
127127
chunksOf n (repeat a) ↺⟨ go ⟩
128-
repeat (Vec.replicate a)
128+
repeat (Vec.replicate n a)
129129

130130
------------------------------------------------------------------------
131131
-- Properties of map

0 commit comments

Comments
 (0)