File tree Expand file tree Collapse file tree 2 files changed +6
-2
lines changed
src/Relation/Nullary/Negation Expand file tree Collapse file tree 2 files changed +6
-2
lines changed Original file line number Diff line number Diff line change @@ -17,6 +17,10 @@ Non-backwards compatible changes
1717Minor 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
Original file line number Diff line number Diff line change @@ -47,11 +47,11 @@ _¬-⊎_ = [_,_]
4747------------------------------------------------------------------------
4848-- Uses of negation
4949
50- contradiction-irr : .A → ¬ A → Whatever
50+ contradiction-irr : .A → .( ¬ A) → Whatever
5151contradiction-irr a ¬a = ⊥-elim-irr (¬a a)
5252
5353contradiction : A → ¬ A → Whatever
54- contradiction a = contradiction-irr a
54+ contradiction a ¬a = contradiction-irr a ¬ a
5555
5656contradiction₂ : A ⊎ B → ¬ A → ¬ B → Whatever
5757contradiction₂ (inj₁ a) ¬a ¬b = contradiction a ¬a
You can’t perform that action at this time.
0 commit comments