diff --git a/CHANGELOG.md b/CHANGELOG.md index cdbdb88ef1..4cd3b98214 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -769,6 +769,30 @@ Non-backwards compatible changes IO.Instances ``` +### Changes to triple reasoning interface + +* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict + relation is irreflexive. + +* This allows the following new proof combinator: + ```agda + begin-contradiction : (r : x IsRelatedTo x) → {s : True (IsStrict? r)} → A + ``` + that takes a proof that a value is strictly less than itself and then applies the principle of explosion. + +* Specialised versions of this combinator are available in the following derived modules: + ``` + Data.Nat.Properties + Data.Nat.Binary.Properties + Data.Integer.Properties + Data.Rational.Unnormalised.Properties + Data.Rational.Properties + Data.Vec.Relation.Binary.Lex.Strict + Data.Vec.Relation.Binary.Lex.NonStrict + Relation.Binary.Reasoning.StrictPartialOrder + Relation.Binary.Reasoning.PartialOrder + ``` + ### Other * In accordance with changes to the flags in Agda 2.6.3, all modules that previously used @@ -991,6 +1015,8 @@ Major improvements * A new module `Algebra.Lattice.Bundles.Raw` is also introduced. * `RawLattice` has been moved from `Algebra.Lattice.Bundles` to this new module. +* In `Relation.Binary.Reasoning.Base.Triple`, added a new parameter `<-irrefl : Irreflexive _≈_ _<_` + Deprecated modules ------------------ diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda index ef7528495e..5760b050ec 100644 --- a/src/Data/Integer/Properties.agda +++ b/src/Data/Integer/Properties.agda @@ -365,6 +365,7 @@ i≮i = <-irrefl refl module ≤-Reasoning where open import Relation.Binary.Reasoning.Base.Triple ≤-isPreorder + <-irrefl <-trans (resp₂ _<_) <⇒≤ diff --git a/src/Data/List/Relation/Unary/Any/Properties.agda b/src/Data/List/Relation/Unary/Any/Properties.agda index c72bd0707f..aba1a777ee 100644 --- a/src/Data/List/Relation/Unary/Any/Properties.agda +++ b/src/Data/List/Relation/Unary/Any/Properties.agda @@ -254,7 +254,6 @@ Any-×⁻ pq with Prod.map₂ (Prod.map₂ find) (find pq) (p , q) ∎ - to∘from : ∀ pq → Any-×⁺ {xs = xs} (Any-×⁻ pq) ≡ pq to∘from pq with find pq | (λ (f : (proj₁ (find pq) ≡_) ⋐ _) → map∘find pq {f}) diff --git a/src/Data/Nat/Binary/Properties.agda b/src/Data/Nat/Binary/Properties.agda index 0072a2d325..638482131f 100644 --- a/src/Data/Nat/Binary/Properties.agda +++ b/src/Data/Nat/Binary/Properties.agda @@ -38,7 +38,6 @@ open import Relation.Binary.Consequences open import Relation.Binary.Morphism import Relation.Binary.Morphism.OrderMonomorphism as OrderMonomorphism open import Relation.Binary.PropositionalEquality -import Relation.Binary.Reasoning.Base.Triple as InequalityReasoning open import Relation.Nullary using (¬_; yes; no) import Relation.Nullary.Decidable as Dec open import Relation.Nullary.Negation using (contradiction) @@ -607,12 +606,18 @@ x ≤? y with <-cmp x y ------------------------------------------------------------------------ -- Equational reasoning for _≤_ and _<_ -module ≤-Reasoning = InequalityReasoning - ≤-isPreorder - <-trans - (resp₂ _<_) <⇒≤ - <-≤-trans ≤-<-trans - hiding (step-≈; step-≈˘) +module ≤-Reasoning where + + open import Relation.Binary.Reasoning.Base.Triple + ≤-isPreorder + <-irrefl + <-trans + (resp₂ _<_) + <⇒≤ + <-≤-trans + ≤-<-trans + public + hiding (step-≈; step-≈˘) ------------------------------------------------------------------------ -- Properties of _<ℕ_ diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda index be41e093ac..5ed9b64e43 100644 --- a/src/Data/Nat/Properties.agda +++ b/src/Data/Nat/Properties.agda @@ -495,6 +495,7 @@ m<1+n⇒m≤n (s≤s m≤n) = m≤n module ≤-Reasoning where open import Relation.Binary.Reasoning.Base.Triple ≤-isPreorder + <-irrefl <-trans (resp₂ _<_) <⇒≤ diff --git a/src/Data/Rational/Properties.agda b/src/Data/Rational/Properties.agda index 8910cc5998..959573218a 100644 --- a/src/Data/Rational/Properties.agda +++ b/src/Data/Rational/Properties.agda @@ -713,13 +713,16 @@ _>?_ = flip _?_ = flip _ _ _ k>k′ = flip contradiction (irrefl⁺ [ k ]) $ begin-strict + ... | here eq | tri> _ _ k>k′ = begin-contradiction [ k ] ≈⟨ [ eq ]ᴱ ⟩ [ k′ ] <⟨ [ k>k′ ]ᴿ ⟩ [ k ] ∎ - ... | left lp | tri≈ _ k≈k′ _ = flip contradiction (irrefl⁺ [ k ]) $ begin-strict + ... | left lp | tri≈ _ k≈k′ _ = begin-contradiction let k″ = Any.lookup lp .key; k≈k″ = lookup-result lp; (_ , k″ _ _ k>k′ = flip contradiction (irrefl⁺ [ k ]) $ begin-strict + ... | left lp | tri> _ _ k>k′ = begin-contradiction let k″ = Any.lookup lp .key; k≈k″ = lookup-result lp; (_ , k″k′ ]ᴿ ⟩ [ k ] ∎ - ... | right rp | tri< k