Skip to content

Commit ad25642

Browse files
Saransh-cppMatthewDaggitt
authored andcommitted
Fixities for Data and Codata (#1985)
1 parent 029556f commit ad25642

File tree

19 files changed

+78
-2
lines changed

19 files changed

+78
-2
lines changed

CHANGELOG.md

+29
Original file line numberDiff line numberDiff line change
@@ -29,6 +29,35 @@ Highlights
2929
Bug-fixes
3030
---------
3131

32+
* The following operators were missing a fixity declaration, which has now
33+
been fixed -
34+
```
35+
infix 4 _ℕ<_ _ℕ≤infinity _ℕ≤_ (Codata.Sized.Conat)
36+
infix 6 _ℕ+_ _+ℕ_ (Codata.Sized.Conat)
37+
infixl 4 _+ _* (Data.List.Kleene.Base)
38+
infixr 4 _++++_ _+++*_ _*+++_ _*++*_ (Data.List.Kleene.Base)
39+
infix 4 _[_]* _[_]+ (Data.List.Kleene.Base)
40+
infix 4 _≢∈_ (Data.List.Membership.Propositional)
41+
infixr 5 _`∷_ (Data.List.Reflection)
42+
infix 4 _≡?_ (Data.List.Relation.Binary.Equality.DecPropositional)
43+
infixr 5 _++ᵖ_ (Data.List.Relation.Binary.Prefix.Heterogeneous)
44+
infixr 5 _++ˢ_ (Data.List.Relation.Binary.Suffix.Heterogeneous)
45+
infixr 5 _++_ _++[] (Data.List.Relation.Ternary.Appending.Propositional)
46+
infixr 5 _∷=_ (Data.List.Relation.Unary.Any)
47+
infixr 5 _++_ (Data.List.Ternary.Appending)
48+
infixr 2 _×-⇔_ _×-↣_ _×-↞_ _×-↠_ _×-↔_ _×-cong_ (Data.Product.Function.NonDependent.Propositional)
49+
infixr 2 _×-⟶_ (Data.Product.Function.NonDependent.Setoid)
50+
infixr 2 _×-equivalence_ _×-injection_ _×-left-inverse_ (Data.Product.Function.NonDependent.Setoid)
51+
infixr 2 _×-surjection_ _×-inverse_ (Data.Product.Function.NonDependent.Setoid)
52+
infixr 1 _⊎-⇔_ _⊎-↣_ _⊎-↞_ _⊎-↠_ _⊎-↔_ _⊎-cong_ (Data.Sum.Function.Propositional)
53+
infixr 1 _⊎-⟶_ (Data.Sum.Function.Setoid)
54+
infixr 1 _⊎-equivalence_ _⊎-injection_ _⊎-left-inverse_ (Data.Sum.Function.Setoid)
55+
infixr 1 _⊎-surjection_ _⊎-inverse_ (Data.Sum.Function.Setoid)
56+
infix 8 _⁻¹ (Data.Parity.Base)
57+
infixr 5 _`∷_ (Data.Vec.Reflection)
58+
infixr 5 _∷=_ (Data.Vec.Membership.Setoid)
59+
```
60+
3261
* In `System.Exit`, the `ExitFailure` constructor is now carrying an integer
3362
rather than a natural. The previous binding was incorrectly assuming that
3463
all exit codes where non-negative.

src/Algebra/Definitions.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -109,7 +109,7 @@ _DistributesOver_ : Op₂ A → Op₂ A → Set _
109109
* DistributesOver + = (* DistributesOverˡ +) × (* DistributesOverʳ +)
110110

111111
_MiddleFourExchange_ : Op₂ A Op₂ A Set _
112-
_*_ MiddleFourExchange _+_ =
112+
_*_ MiddleFourExchange _+_ =
113113
w x y z ((w + x) * (y + z)) ≈ ((w + y) * (x + z))
114114

115115
_IdempotentOn_ : Op₂ A A Set _

src/Codata/Sized/Conat.agda

+3-1
Original file line numberDiff line numberDiff line change
@@ -35,7 +35,7 @@ pred : ∀ {i} {j : Size< i} → Conat i → Conat j
3535
pred zero = zero
3636
pred (suc n) = n .force
3737

38-
infixl 6 _∸_ _+_
38+
infixl 6 _∸_ _+_ _ℕ+_ _+ℕ_
3939
infixl 7 _*_
4040

4141
_∸_ : Conat ∞ Conat ∞
@@ -91,6 +91,8 @@ toℕ (suc n) = suc (toℕ n)
9191
------------------------------------------------------------------------
9292
-- Order wrt to Nat
9393

94+
infix 4 _ℕ<_ _ℕ≤infinity _ℕ≤_
95+
9496
data _ℕ≤_ : Conat ∞ Set where
9597
zℕ≤n : {n} zero ℕ≤ n
9698
sℕ≤s : {k n} k ℕ≤ n .force suc k ℕ≤ suc n

src/Data/List/Kleene/Base.agda

+4
Original file line numberDiff line numberDiff line change
@@ -40,6 +40,7 @@ private
4040
-- simplify certain proofs.
4141

4242
infixr 5 _&_ ∹_
43+
infixl 4 _+ _*
4344

4445
record _+ {a} (A : Set a) : Set a
4546
data _* {a} (A : Set a) : Set a
@@ -101,6 +102,7 @@ module _ (f : B → A → B) where
101102
-- Concatenation
102103

103104
module Concat where
105+
infixr 4 _++++_ _+++*_ _*+++_ _*++*_
104106
_++++_ : A + A + A +
105107
_+++*_ : A + A * A +
106108
_*+++_ : A * A + A +
@@ -270,6 +272,8 @@ module _ (f : A → Maybe B → B) where
270272
------------------------------------------------------------------------
271273
-- Indexing
272274

275+
infix 4 _[_]* _[_]+
276+
273277
_[_]* : A * Maybe A
274278
_[_]+ : A + Maybe A
275279

src/Data/List/Membership/Propositional.agda

+2
Original file line numberDiff line numberDiff line change
@@ -22,6 +22,8 @@ open SetoidMembership (setoid A) public hiding (lose)
2222
------------------------------------------------------------------------
2323
-- Different members
2424

25+
infix 4 _≢∈_
26+
2527
_≢∈_ : {x y : A} {xs} x ∈ xs y ∈ xs Set _
2628
_≢∈_ x∈xs y∈xs = x≡y subst (_∈ _) x≡y x∈xs ≢ y∈xs
2729

src/Data/List/Reflection.agda

+2
Original file line numberDiff line numberDiff line change
@@ -24,6 +24,8 @@ open import Reflection.AST.Argument
2424
`[] : Term
2525
`[] = con (quote List.[]) (2 ⋯⟅∷⟆ [])
2626

27+
infixr 5 _`∷_
28+
2729
_`∷_ : Term Term Term
2830
x `∷ xs = con (quote List._∷_) (2 ⋯⟅∷⟆ x ⟨∷⟩ xs ⟨∷⟩ [])
2931

src/Data/List/Relation/Binary/Equality/DecPropositional.agda

+2
Original file line numberDiff line numberDiff line change
@@ -32,5 +32,7 @@ open DecSetoidEq (decSetoid _≟_) public
3232
------------------------------------------------------------------------
3333
-- Additional proofs
3434

35+
infix 4 _≡?_
36+
3537
_≡?_ : Decidable (_≡_ {A = List A})
3638
_≡?_ = ≡-dec _≟_

src/Data/List/Relation/Binary/Prefix/Heterogeneous.agda

+2
Original file line numberDiff line numberDiff line change
@@ -45,6 +45,8 @@ module _ {a b r s} {A : Set a} {B : Set b} {R : REL A B r} {S : REL A B s} where
4545

4646
module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where
4747

48+
infixr 5 _++ᵖ_
49+
4850
_++ᵖ_ : {as bs} Prefix R as bs suf Prefix R as (bs List.++ suf)
4951
[] ++ᵖ suf = []
5052
(r ∷ rs) ++ᵖ suf = r ∷ (rs ++ᵖ suf)

src/Data/List/Relation/Binary/Suffix/Heterogeneous.agda

+2
Original file line numberDiff line numberDiff line change
@@ -31,6 +31,8 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where
3131
tail (here (_ ∷ rs)) = there (here rs)
3232
tail (there x) = there (tail x)
3333

34+
infixr 5 _++ˢ_
35+
3436
_++ˢ_ : pre {as bs} Suffix R as bs Suffix R as (pre List.++ bs)
3537
[] ++ˢ rs = rs
3638
(x ∷ xs) ++ˢ rs = there (xs ++ˢ rs)

src/Data/List/Relation/Ternary/Appending.agda

+2
Original file line numberDiff line numberDiff line change
@@ -38,6 +38,8 @@ module _ (L : REL A C l) (R : REL B C r) where
3838
------------------------------------------------------------------------
3939
-- Functions manipulating Appending
4040

41+
infixr 5 _++_
42+
4143
_++_ : {cs₁ cs₂ : List C} Pointwise L as cs₁ Pointwise R bs cs₂
4244
Appending L R as bs (cs₁ List.++ cs₂)
4345
[] ++ rs = []++ rs

src/Data/List/Relation/Ternary/Appending/Propositional.agda

+2
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,8 @@ open General public
3232
------------------------------------------------------------------------
3333
-- Definition
3434

35+
infixr 5 _++_ _++[]
36+
3537
_++_ : (as bs : List A) Appending as bs (as List.++ bs)
3638
as ++ bs = Pw.≡⇒Pointwise-≡ refl General.++ Pw.≡⇒Pointwise-≡ refl
3739

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

+2
Original file line numberDiff line numberDiff line change
@@ -62,6 +62,8 @@ index (there pxs) = suc (index pxs)
6262
lookup : {P : Pred A p} Any P xs A
6363
lookup {xs = xs} p = List.lookup xs (index p)
6464

65+
infixr 5 _∷=_
66+
6567
_∷=_ : {P : Pred A p} Any P xs A List A
6668
_∷=_ {xs = xs} x∈xs v = xs List.[ index x∈xs ]∷= v
6769

src/Data/Parity/Base.agda

+2
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,8 @@ data Parity : Set where
2626

2727
-- The opposite parity.
2828

29+
infix 8 _⁻¹
30+
2931
_⁻¹ : Parity Parity
3032
1ℙ ⁻¹ = 0ℙ
3133
0ℙ ⁻¹ = 1ℙ

src/Data/Product/Function/NonDependent/Propositional.agda

+4
Original file line numberDiff line numberDiff line change
@@ -27,6 +27,8 @@ open import Function.Surjection as Surj using (_↠_; module Surjection)
2727

2828
module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
2929

30+
infixr 2 _×-⇔_ _×-↣_ _×-↞_ _×-↠_ _×-↔_
31+
3032
_×-⇔_ : A ⇔ B C ⇔ D (A × C) ⇔ (B × D)
3133
_×-⇔_ A⇔B C⇔D =
3234
Inverse.equivalence Pointwise-≡↔≡ ⟨∘⟩
@@ -64,6 +66,8 @@ module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
6466

6567
module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
6668

69+
infixr 2 _×-cong_
70+
6771
_×-cong_ : {k} A ∼[ k ] B C ∼[ k ] D (A × C) ∼[ k ] (B × D)
6872
_×-cong_ {implication} = λ f g map f g
6973
_×-cong_ {reverse-implication} = λ f g lam (map (app-← f) (app-← g))

src/Data/Product/Function/NonDependent/Setoid.agda

+6
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,8 @@ module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
3333
{C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
3434
where
3535

36+
infixr 2 _×-⟶_
37+
3638
_×-⟶_ : (A ⟶ B) (C ⟶ D) (A ×ₛ C) ⟶ (B ×ₛ D)
3739
_×-⟶_ f g = record
3840
{ _⟨$⟩_ = fg
@@ -76,6 +78,8 @@ module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
7678
{C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
7779
where
7880

81+
infixr 2 _×-equivalence_ _×-injection_ _×-left-inverse_
82+
7983
_×-equivalence_ : Equivalence A B Equivalence C D
8084
Equivalence (A ×ₛ C) (B ×ₛ D)
8185
_×-equivalence_ A⇔B C⇔D = record
@@ -110,6 +114,8 @@ module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
110114
{C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
111115
where
112116

117+
infixr 2 _×-surjection_ _×-inverse_
118+
113119
_×-surjection_ : Surjection A B Surjection C D
114120
Surjection (A ×ₛ C) (B ×ₛ D)
115121
A↠B ×-surjection C↠D = record

src/Data/Sum/Function/Propositional.agda

+4
Original file line numberDiff line numberDiff line change
@@ -24,6 +24,8 @@ open import Function.Surjection as Surj using (_↠_; module Surjection)
2424

2525
module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
2626

27+
infixr 1 _⊎-⇔_ _⊎-↣_ _⊎-↞_ _⊎-↠_ _⊎-↔_
28+
2729
_⊎-⇔_ : A ⇔ B C ⇔ D (A ⊎ C) ⇔ (B ⊎ D)
2830
_⊎-⇔_ A⇔B C⇔D =
2931
Inverse.equivalence (Pointwise-≡↔≡ B D) ⟨∘⟩
@@ -61,6 +63,8 @@ module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
6163

6264
module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
6365

66+
infixr 1 _⊎-cong_
67+
6468
_⊎-cong_ : {k} A ∼[ k ] B C ∼[ k ] D (A ⊎ C) ∼[ k ] (B ⊎ D)
6569
_⊎-cong_ {implication} = map
6670
_⊎-cong_ {reverse-implication} = λ f g lam (map (app-← f) (app-← g))

src/Data/Sum/Function/Setoid.agda

+5
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,8 @@ module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
3232
{C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
3333
where
3434

35+
infixr 1 _⊎-⟶_
36+
3537
_⊎-⟶_ : (A ⟶ B) (C ⟶ D) (A ⊎ₛ C) ⟶ (B ⊎ₛ D)
3638
_⊎-⟶_ f g = record
3739
{ _⟨$⟩_ = fg
@@ -77,6 +79,8 @@ module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
7779
{C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
7880
where
7981

82+
infixr 1 _⊎-equivalence_ _⊎-injection_ _⊎-left-inverse_
83+
8084
_⊎-equivalence_ : Equivalence A B Equivalence C D
8185
Equivalence (A ⊎ₛ C) (B ⊎ₛ D)
8286
A⇔B ⊎-equivalence C⇔D = record
@@ -119,6 +123,7 @@ module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
119123
{C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
120124
where
121125

126+
infixr 1 _⊎-surjection_ _⊎-inverse_
122127
_⊎-surjection_ : Surjection A B Surjection C D
123128
Surjection (A ⊎ₛ C) (B ⊎ₛ D)
124129
A↠B ⊎-surjection C↠D = record

src/Data/Vec/Membership/Setoid.agda

+2
Original file line numberDiff line numberDiff line change
@@ -39,6 +39,8 @@ mapWith∈ : ∀ {b} {B : Set b} {n}
3939
mapWith∈ [] f = []
4040
mapWith∈ (x ∷ xs) f = f (here refl) ∷ mapWith∈ xs (f ∘ there)
4141

42+
infixr 5 _∷=_
43+
4244
_∷=_ : {n} {xs : Vec A n} {x} x ∈ xs A Vec A n
4345
_∷=_ {xs = xs} x∈xs v = xs Vec.[ index x∈xs ]≔ v
4446

src/Data/Vec/Reflection.agda

+2
Original file line numberDiff line numberDiff line change
@@ -25,6 +25,8 @@ open import Reflection.AST.Argument
2525
`[] : Term
2626
`[] = con (quote []) (2 ⋯⟅∷⟆ List.[])
2727

28+
infixr 5 _`∷_
29+
2830
_`∷_ : Term Term Term
2931
_`∷_ x xs = con (quote _∷_) (3 ⋯⟅∷⟆ x ⟨∷⟩ xs ⟨∷⟩ List.[])
3032

0 commit comments

Comments
 (0)