Skip to content

Commit ea2a7ce

Browse files
authored
[ add ] flipped version of Relation.Nullary.Negation.Core.contradiction (#2797)
* add: flipped version of `contradiction` * add: comment to `CHANGELOG`
1 parent 8add6fd commit ea2a7ce

File tree

3 files changed

+11
-4
lines changed

3 files changed

+11
-4
lines changed

CHANGELOG.md

Lines changed: 5 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -20,6 +20,9 @@ Minor improvements
2020
* The type of `Relation.Nullary.Negation.Core.contradiction-irr` has been further
2121
weakened so that the negated hypothesis `¬ A` is marked as irrelevant. This is
2222
safe to do, in view of `Relation.Nullary.Recomputable.Properties.¬-recompute`.
23+
Furthermore, because the *eager* insertion of implicit arguments during type
24+
inference interacts badly with `contradiction`, we introduce an explicit name
25+
`contradiction′` for its `flip`ped version.
2326

2427
* Refactored usages of `+-∸-assoc 1` to `∸-suc` in:
2528
```agda
@@ -109,5 +112,6 @@ Additions to existing modules
109112

110113
* In `Relation.Nullary.Negation.Core`
111114
```agda
112-
¬¬-η : A → ¬ ¬ A
115+
¬¬-η : A → ¬ ¬ A
116+
contradiction′ : ¬ A → A → Whatever
113117
```

src/Relation/Nullary/Negation.agda

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -72,7 +72,7 @@ open import Relation.Nullary.Negation.Core public
7272
-- ⊥).
7373

7474
call/cc : ((A Whatever) DoubleNegation A) DoubleNegation A
75-
call/cc hyp ¬a = hyp (flip contradiction ¬a) ¬a
75+
call/cc hyp ¬a = hyp (contradiction ¬a) ¬a
7676

7777
-- The "independence of premise" rule, in the double-negation monad.
7878
-- It is assumed that the index set (A) is inhabited.
@@ -82,7 +82,7 @@ independence-of-premise {A = A} {B = B} {P = P} q f = ¬¬-map helper ¬¬-exclu
8282
where
8383
helper : Dec B Σ[ x ∈ A ] (B P x)
8484
helper (yes p) = Product.map₂ const (f p)
85-
helper (no ¬p) = (q , flip contradiction ¬p)
85+
helper (no ¬p) = (q , contradiction ¬p)
8686

8787
-- The independence of premise rule for binary sums.
8888

@@ -91,7 +91,7 @@ independence-of-premise-⊎ {A = A} {B = B} {C = C} f = ¬¬-map helper ¬¬-exc
9191
where
9292
helper : Dec A (A B) ⊎ (A C)
9393
helper (yes p) = Sum.map const const (f p)
94-
helper (no ¬p) = inj₁ (flip contradiction ¬p)
94+
helper (no ¬p) = inj₁ (contradiction ¬p)
9595

9696
private
9797

src/Relation/Nullary/Negation/Core.agda

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -58,6 +58,9 @@ contradiction-irr a ¬a = ⊥-elim-irr (¬a a)
5858
contradiction : A ¬ A Whatever
5959
contradiction a ¬a = contradiction-irr a ¬a
6060

61+
contradiction′ : ¬ A A Whatever
62+
contradiction′ ¬a a = contradiction-irr a ¬a
63+
6164
contradiction₂ : A ⊎ B ¬ A ¬ B Whatever
6265
contradiction₂ (inj₁ a) ¬a ¬b = contradiction a ¬a
6366
contradiction₂ (inj₂ b) ¬a ¬b = contradiction b ¬b

0 commit comments

Comments
 (0)