File tree Expand file tree Collapse file tree 3 files changed +11
-4
lines changed Expand file tree Collapse file tree 3 files changed +11
-4
lines changed Original file line number Diff line number Diff line change @@ -20,6 +20,9 @@ Minor improvements
20
20
* The type of ` Relation.Nullary.Negation.Core.contradiction-irr ` has been further
21
21
weakened so that the negated hypothesis ` ¬ A ` is marked as irrelevant. This is
22
22
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.
23
26
24
27
* Refactored usages of ` +-∸-assoc 1 ` to ` ∸-suc ` in:
25
28
``` agda
@@ -109,5 +112,6 @@ Additions to existing modules
109
112
110
113
* In ` Relation.Nullary.Negation.Core `
111
114
``` agda
112
- ¬¬-η : A → ¬ ¬ A
115
+ ¬¬-η : A → ¬ ¬ A
116
+ contradiction′ : ¬ A → A → Whatever
113
117
```
Original file line number Diff line number Diff line change @@ -72,7 +72,7 @@ open import Relation.Nullary.Negation.Core public
72
72
-- ⊥).
73
73
74
74
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
76
76
77
77
-- The "independence of premise" rule, in the double-negation monad.
78
78
-- 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
82
82
where
83
83
helper : Dec B → Σ[ x ∈ A ] (B → P x)
84
84
helper (yes p) = Product.map₂ const (f p)
85
- helper (no ¬p) = (q , flip contradiction ¬p)
85
+ helper (no ¬p) = (q , contradiction′ ¬p)
86
86
87
87
-- The independence of premise rule for binary sums.
88
88
@@ -91,7 +91,7 @@ independence-of-premise-⊎ {A = A} {B = B} {C = C} f = ¬¬-map helper ¬¬-exc
91
91
where
92
92
helper : Dec A → (A → B) ⊎ (A → C)
93
93
helper (yes p) = Sum.map const const (f p)
94
- helper (no ¬p) = inj₁ (flip contradiction ¬p)
94
+ helper (no ¬p) = inj₁ (contradiction′ ¬p)
95
95
96
96
private
97
97
Original file line number Diff line number Diff line change @@ -58,6 +58,9 @@ contradiction-irr a ¬a = ⊥-elim-irr (¬a a)
58
58
contradiction : A → ¬ A → Whatever
59
59
contradiction a ¬a = contradiction-irr a ¬a
60
60
61
+ contradiction′ : ¬ A → A → Whatever
62
+ contradiction′ ¬a a = contradiction-irr a ¬a
63
+
61
64
contradiction₂ : A ⊎ B → ¬ A → ¬ B → Whatever
62
65
contradiction₂ (inj₁ a) ¬a ¬b = contradiction a ¬a
63
66
contradiction₂ (inj₂ b) ¬a ¬b = contradiction b ¬b
You can’t perform that action at this time.
0 commit comments