File tree Expand file tree Collapse file tree 5 files changed +17
-7
lines changed
Data/List/Relation/Unary/First Expand file tree Collapse file tree 5 files changed +17
-7
lines changed Original file line number Diff line number Diff line change @@ -77,3 +77,8 @@ Additions to existing modules
7777 ≟-≢ : (m≢n : m ≢ n) → (m ≟ n) ≡ no m≢n
7878 ∸-suc : m ≤ n → suc n ∸ m ≡ suc (n ∸ m)
7979 ```
80+
81+ * In ` Relation.Nullary.Negation.Core `
82+ ``` agda
83+ ¬¬-η : A → ¬ ¬ A
84+ ```
Original file line number Diff line number Diff line change @@ -18,7 +18,7 @@ import Data.Sum.Base as Sum
1818open import Function.Base using (_∘′_; _∘_; id)
1919open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl; _≗_)
2020open import Relation.Nullary.Decidable.Core as Dec
21- open import Relation.Nullary.Negation.Core using (contradiction)
21+ open import Relation.Nullary.Negation.Core using (¬¬-η; contradiction)
2222open import Relation.Nullary.Reflects using (invert)
2323open import Relation.Unary using (Pred; _⊆_; ∁; Irrelevant; Decidable)
2424
@@ -54,7 +54,7 @@ module _ {a p q} {A : Set a} {P : Pred A p} {Q : Pred A q} where
5454module _ {a p q} {A : Set a} {P : Pred A p} {Q : Pred A q} where
5555
5656 All⇒¬First : P ⊆ ∁ Q → All P ⊆ ∁ (First P Q)
57- All⇒¬First p⇒¬q (px ∷ pxs) [ qx ] = contradiction qx ( p⇒¬q px)
57+ All⇒¬First p⇒¬q (px ∷ pxs) [ qx ] = p⇒¬q px qx
5858 All⇒¬First p⇒¬q (_ ∷ pxs) (_ ∷ hf) = All⇒¬First p⇒¬q pxs hf
5959
6060 First⇒¬All : Q ⊆ ∁ P → First P Q ⊆ ∁ (All P)
@@ -97,7 +97,7 @@ module _ {a p} {A : Set a} {P : Pred A p} where
9797
9898 first? : Decidable P → Decidable (First P (∁ P))
9999 first? P? = Dec.fromSum
100- ∘ Sum.map₂ (All⇒¬First contradiction )
100+ ∘ Sum.map₂ (All⇒¬First ¬¬-η )
101101 ∘ first (Dec.toSum ∘ P?)
102102
103103 cofirst? : Decidable P → Decidable (First (∁ P) P)
Original file line number Diff line number Diff line change @@ -28,7 +28,7 @@ open import Relation.Nullary.Recomputable.Core as Recomputable
2828open import Relation.Nullary.Reflects as Reflects
2929 hiding (recompute; recompute-constant)
3030open import Relation.Nullary.Negation.Core
31- using (¬_; Stable; negated-stable; contradiction; DoubleNegation)
31+ using (¬_; ¬¬-η; Stable; negated-stable; contradiction; DoubleNegation)
3232
3333
3434private
@@ -215,7 +215,7 @@ decidable-stable (true because [a]) ¬¬a = invert [a]
215215decidable-stable (false because [¬a]) ¬¬a = contradiction (invert [¬a]) ¬¬a
216216
217217¬-drop-Dec : Dec (¬ ¬ A) → Dec (¬ A)
218- ¬-drop-Dec ¬¬a? = map′ negated-stable contradiction (¬? ¬¬a?)
218+ ¬-drop-Dec ¬¬a? = map′ negated-stable ¬¬-η (¬? ¬¬a?)
219219
220220-- A double-negation-translated variant of excluded middle (or: every
221221-- nullary relation is decidable in the double-negation monad).
Original file line number Diff line number Diff line change @@ -56,7 +56,7 @@ open import Relation.Nullary.Negation.Core public
5656¬¬-Monad : RawMonad {a} DoubleNegation
5757¬¬-Monad = mkRawMonad
5858 DoubleNegation
59- contradiction
59+ ¬¬-η
6060 (λ x f → negated-stable (¬¬-map f x))
6161
6262¬¬-push : DoubleNegation Π[ P ] → Π[ DoubleNegation ∘ P ]
Original file line number Diff line number Diff line change @@ -15,7 +15,7 @@ open import Level using (Level; _⊔_)
1515
1616private
1717 variable
18- a p q w : Level
18+ a w : Level
1919 A B C : Set a
2020 Whatever : Set w
2121
@@ -33,6 +33,11 @@ infix 3 ¬_
3333DoubleNegation : Set a → Set a
3434DoubleNegation A = ¬ ¬ A
3535
36+ -- Eta law for double-negation
37+
38+ ¬¬-η : A → ¬ ¬ A
39+ ¬¬-η a ¬a = ¬a a
40+
3641-- Stability under double-negation.
3742Stable : Set a → Set a
3843Stable A = ¬ ¬ A → A
You can’t perform that action at this time.
0 commit comments