diff --git a/CHANGELOG.md b/CHANGELOG.md index 5a42872be0..60558240eb 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -120,6 +120,34 @@ Bug-fixes infixl 6 _+ℤ_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) infix 4 _≉_ _≈ᵢ_ _≤ᵢ_ (Relation.Binary.Indexed.Homogeneous.Bundles) infixr 5 _∷ᴹ_ _∷⁻¹ᴹ_ (Text.Regex.Search) + infixr 4 _,_ (Data.Refinement) + infixr 4 _,_ (Data.Container.Relation.Binary.Pointwise) + infixr 4 _,_ (Data.Tree.AVL.Value) + infixr 4 _,_ (Foreign.Haskell.Pair) + infixr 4 _,_ (Reflection.AnnotatedAST) + infixr 4 _,_ (Reflection.AST.Traversal) + infixl 6.5 _P′_ _P_ _C′_ _C_ (Data.Nat.Combinatorics.Base) + infixl 1 _>>=-cong_ _≡->>=-cong_ (Effect.Monad.Partiality) + infixl 1 _?>=′_ (Effect.Monad.Predicate) + infixl 1 _>>=-cong_ _>>=-congP_ (Effect.Monad.Partiality.All) + infix 4 _∈FV_ (Reflection.AST.DeBruijn) + infixr 9 _;_ (Relation.Binary.Construct.Composition) + infixl 6 _+²_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) + infixr -1 _atₛ_ (Relation.Binary.Indexed.Heterogeneous.Construct.At) + infixr -1 _atₛ_ (Relation.Binary.Indexed.Homogeneous.Construct.At) + infix 4 _∈_ _∉_ (Relation.Unary.Indexed) + infixr 9 _⍮_ (Relation.Unary.PredicateTransformer) + infix 8 ∼_ (Relation.Unary.PredicateTransformer) + infix 2 _×?_ _⊙?_ (Relation.Unary.Properties) + infix 10 _~? (Relation.Unary.Properties) + infixr 1 _⊎?_ (Relation.Unary.Properties) + infixr 7 _∩?_ (Relation.Unary.Properties) + infixr 6 _∪?_ (Relation.Unary.Properties) + infixl 6 _`⊜_ (Tactic.RingSolver) + infix 8 ⊝_ (Tactic.RingSolver.Core.Expression) + infix 4 _∈ᴿ?_ _∉ᴿ?_ _∈?ε _∈?[_] _∈?[^_] (Text.Regex.Properties) + infix 4 _∈?_ _∉?_ (Text.Regex.Derivative.Brzozowski) + infix 4 _∈_ _∉_ _∈?_ _∉?_ (Text.Regex.String.Unsafe) ``` * In `System.Exit`, the `ExitFailure` constructor is now carrying an integer diff --git a/README/Design/Fixity.agda b/README/Design/Fixity.agda index 87907e49b6..85153c1ec4 100644 --- a/README/Design/Fixity.agda +++ b/README/Design/Fixity.agda @@ -23,6 +23,8 @@ module README.Design.Fixity where -- ternary reasoning infix 1 _⊢_≈_ -- composition infixr 9 _∘_ -- application infixr -1 _$_ _$!_ +-- combinatorics infixl 6.5 _P_ _P′_ _C_ _C′_ +-- pair infixr 4 _,_ -- Reasoning: -- QED infix 3 _∎ diff --git a/src/Data/Container/Relation/Binary/Pointwise.agda b/src/Data/Container/Relation/Binary/Pointwise.agda index 4de4cd0939..2bd20bc1c5 100644 --- a/src/Data/Container/Relation/Binary/Pointwise.agda +++ b/src/Data/Container/Relation/Binary/Pointwise.agda @@ -26,6 +26,7 @@ module _ {s p} (C : Container s p) where constructor _,_ field shape : proj₁ cx ≡ proj₁ cy position : ∀ p → R (proj₂ cx p) (proj₂ cy (subst _ shape p)) + infixr 4 _,_ module _ {s p} {C : Container s p} {x y} {X : Set x} {Y : Set y} {ℓ ℓ′} {R : REL X Y ℓ} {R′ : REL X Y ℓ′} diff --git a/src/Data/Nat/Combinatorics.agda b/src/Data/Nat/Combinatorics.agda index 850cd01857..fc6ef422b4 100644 --- a/src/Data/Nat/Combinatorics.agda +++ b/src/Data/Nat/Combinatorics.agda @@ -75,7 +75,7 @@ nCk≡nC[n∸k] {k} {n} k≤n = begin-equality _ = k !* (n ∸ k) !≢0 _ = (n ∸ k) !* (n ∸ (n ∸ k)) !≢0 -nCk≡nPk/k! : k ≤ n → n C k ≡ (n P k / k !) {{k !≢0}} +nCk≡nPk/k! : k ≤ n → n C k ≡ ((n P k) / k !) {{k !≢0}} nCk≡nPk/k! {k} {n} k≤n = begin-equality n C k ≡⟨ nCk≡n!/k![n-k]! k≤n ⟩ n ! / (k ! * (n ∸ k) !) ≡˘⟨ /-congʳ (*-comm ((n ∸ k)!) (k !)) ⟩ diff --git a/src/Data/Nat/Combinatorics/Base.agda b/src/Data/Nat/Combinatorics/Base.agda index 3ee7066441..149665c801 100644 --- a/src/Data/Nat/Combinatorics/Base.agda +++ b/src/Data/Nat/Combinatorics/Base.agda @@ -25,6 +25,8 @@ open import Data.Nat.Properties using (_!≢0) -- n P k = n ! / (n ∸ k) ! +infixl 6.5 _P′_ _P_ + -- Base definition. Only valid for k ≤ n. _P′_ : ℕ → ℕ → ℕ @@ -44,6 +46,8 @@ n P k = if k ≤ᵇ n then n P′ k else 0 -- n C k = n ! / (k ! * (n ∸ k) !) +infixl 6.5 _C′_ _C_ + -- Base definition. Only valid for k ≤ n. _C′_ : ℕ → ℕ → ℕ diff --git a/src/Data/Nat/Combinatorics/Specification.agda b/src/Data/Nat/Combinatorics/Specification.agda index 0b263eed54..522e3104ed 100644 --- a/src/Data/Nat/Combinatorics/Specification.agda +++ b/src/Data/Nat/Combinatorics/Specification.agda @@ -57,20 +57,20 @@ nP′k≡n[n∸1P′k∸1] : ∀ n k → .{{NonZero n}} → .{{NonZero k}} → nP′k≡n[n∸1P′k∸1] n (suc zero) = refl nP′k≡n[n∸1P′k∸1] n@(suc n-1) k@(suc k-1@(suc k-2)) = begin-equality n P′ k ≡⟨⟩ - (n ∸ k-1) * n P′ k-1 ≡⟨ cong ((n ∸ k-1) *_) (nP′k≡n[n∸1P′k∸1] n k-1) ⟩ - (n ∸ k-1) * (n * n-1 P′ k-2) ≡⟨ x∙yz≈y∙xz (n ∸ k-1) n (n-1 P′ k-2) ⟩ - n * ((n ∸ k-1) * n-1 P′ k-2) ≡⟨⟩ - n * n-1 P′ k-1 ∎ + (n ∸ k-1) * (n P′ k-1) ≡⟨ cong ((n ∸ k-1) *_) (nP′k≡n[n∸1P′k∸1] n k-1) ⟩ + (n ∸ k-1) * (n * (n-1 P′ k-2)) ≡⟨ x∙yz≈y∙xz (n ∸ k-1) n (n-1 P′ k-2) ⟩ + n * ((n ∸ k-1) * (n-1 P′ k-2)) ≡⟨⟩ + n * (n-1 P′ k-1) ∎ where open ≤-Reasoning; open *-CS P′-rec : ∀ {n k} → k ≤ n → .{{NonZero k}} → n P′ k ≡ k * (pred n P′ pred k) + pred n P′ k P′-rec n@{suc n-1} k@{suc k-1} k≤n = begin-equality - n P′ k ≡⟨ nP′k≡n[n∸1P′k∸1] n k ⟩ - n * n-1 P′ k-1 ≡˘⟨ cong (_* n-1 P′ k-1) (m+[n∸m]≡n {k} {n} k≤n) ⟩ - (k + (n ∸ k)) * n-1 P′ k-1 ≡⟨ *-distribʳ-+ (n-1 P′ k-1) k (n ∸ k) ⟩ - k * (n-1 P′ k-1) + (n-1 ∸ k-1) * n-1 P′ k-1 ≡⟨⟩ - k * (n-1 P′ k-1) + n-1 P′ k ∎ + n P′ k ≡⟨ nP′k≡n[n∸1P′k∸1] n k ⟩ + n * (n-1 P′ k-1) ≡˘⟨ cong (_* (n-1 P′ k-1)) (m+[n∸m]≡n {k} {n} k≤n) ⟩ + (k + (n ∸ k)) * (n-1 P′ k-1) ≡⟨ *-distribʳ-+ (n-1 P′ k-1) k (n ∸ k) ⟩ + k * (n-1 P′ k-1) + (n-1 ∸ k-1) * (n-1 P′ k-1) ≡⟨⟩ + k * (n-1 P′ k-1) + (n-1 P′ k) ∎ where open ≤-Reasoning nP′n≡n! : ∀ n → n P′ n ≡ n ! diff --git a/src/Data/Refinement.agda b/src/Data/Refinement.agda index 8671e8593d..dd63e0918a 100644 --- a/src/Data/Refinement.agda +++ b/src/Data/Refinement.agda @@ -23,6 +23,7 @@ record Refinement {a p} (A : Set a) (P : A → Set p) : Set (a ⊔ p) where constructor _,_ field value : A proof : Irrelevant (P value) +infixr 4 _,_ open Refinement public -- The syntax declaration below is meant to mimick set comprehension. diff --git a/src/Data/Tree/AVL/Value.agda b/src/Data/Tree/AVL/Value.agda index ee51150477..6305bd5c2c 100644 --- a/src/Data/Tree/AVL/Value.agda +++ b/src/Data/Tree/AVL/Value.agda @@ -34,6 +34,7 @@ record K&_ {v} (V : Value v) : Set (a ⊔ v) where field key : Key value : Value.family V key +infixr 4 _,_ open K&_ public module _ {v} {V : Value v} where diff --git a/src/Effect/Monad/Partiality.agda b/src/Effect/Monad/Partiality.agda index 5fb6ec3455..ed0e3efbc1 100644 --- a/src/Effect/Monad/Partiality.agda +++ b/src/Effect/Monad/Partiality.agda @@ -517,6 +517,8 @@ module _ {A B : Set s} -- Bind preserves all the relations. + infixl 1 _>>=-cong_ + _>>=-cong_ : ∀ {k} {x₁ x₂ : A ⊥} {f₁ f₂ : A → B ⊥} → let open M in Rel _∼A_ k x₁ x₂ → @@ -578,6 +580,8 @@ module _ {A B : Set ℓ} {_∼_ : B → B → Set ℓ} where -- A variant of _>>=-cong_. + infixl 1 _≡->>=-cong_ + _≡->>=-cong_ : ∀ {k} {x₁ x₂ : A ⊥} {f₁ f₂ : A → B ⊥} → let open M in Rel _≡_ k x₁ x₂ → diff --git a/src/Effect/Monad/Partiality/All.agda b/src/Effect/Monad/Partiality/All.agda index c2db61c031..c5bf2da9ef 100644 --- a/src/Effect/Monad/Partiality/All.agda +++ b/src/Effect/Monad/Partiality/All.agda @@ -44,6 +44,8 @@ data All {A : Set a} (P : A → Set p) : A ⊥ → Set (a ⊔ p) where -- Bind preserves All in the following way: +infixl 1 _>>=-cong_ + _>>=-cong_ : ∀ {p q} {P : A → Set p} {Q : B → Set q} {x : A ⊥} {f : A → B ⊥} → All P x → (∀ {x} → P x → All Q (f x)) → @@ -118,6 +120,8 @@ module Alternative {a p : Level} where _≳⟨_⟩P_ : ∀ x {y} (x≳y : x ≳ y) (p : AllP P y) → AllP P x _⟨_⟩P : ∀ x (p : AllP P x) → AllP P x + infixl 1 _>>=-congP_ + private -- WHNFs. diff --git a/src/Effect/Monad/Predicate.agda b/src/Effect/Monad/Predicate.agda index 1210de7959..7a46fb38a8 100644 --- a/src/Effect/Monad/Predicate.agda +++ b/src/Effect/Monad/Predicate.agda @@ -29,7 +29,7 @@ private record RawPMonad {I : Set i} (M : Pt I (i ⊔ ℓ)) : Set (suc i ⊔ suc ℓ) where - infixl 1 _?>=_ _?>_ _>?>_ + infixl 1 _?>=_ _?>_ _>?>_ _?>=′_ infixr 1 _=