Skip to content

Commit 35396e4

Browse files
authored
refactor: weaken type of contradiction-irr (#2785)
1 parent 0fa3e60 commit 35396e4

File tree

2 files changed

+6
-2
lines changed

2 files changed

+6
-2
lines changed

CHANGELOG.md

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -17,6 +17,10 @@ Non-backwards compatible changes
1717
Minor improvements
1818
------------------
1919

20+
* The type of `Relation.Nullary.Negation.Core.contradiction-irr` has been further
21+
weakened so that the negated hypothesis `¬ A` is marked as irrelevant. This is
22+
safe to do, in view of `Relation.Nullary.Recomputable.Properties.¬-recompute`.
23+
2024
* Refactored usages of `+-∸-assoc 1` to `∸-suc` in:
2125
```agda
2226
README.Data.Fin.Relation.Unary.Top

src/Relation/Nullary/Negation/Core.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -47,11 +47,11 @@ _¬-⊎_ = [_,_]
4747
------------------------------------------------------------------------
4848
-- Uses of negation
4949

50-
contradiction-irr : .A ¬ A Whatever
50+
contradiction-irr : .A .(¬ A) Whatever
5151
contradiction-irr a ¬a = ⊥-elim-irr (¬a a)
5252

5353
contradiction : A ¬ A Whatever
54-
contradiction a = contradiction-irr a
54+
contradiction a ¬a = contradiction-irr a ¬a
5555

5656
contradiction₂ : A ⊎ B ¬ A ¬ B Whatever
5757
contradiction₂ (inj₁ a) ¬a ¬b = contradiction a ¬a

0 commit comments

Comments
 (0)