@@ -10,7 +10,7 @@ module Relation.Nullary.Negation.Core where
1010
1111open import  Data.Empty  using  (⊥; ⊥-elim-irr)
1212open import  Data.Sum.Base  using  (_⊎_; [_,_]; inj₁; inj₂)
13- open import  Function.Base  using  (flip;  _∘_; const)
13+ open import  Function.Base  using  (_∘_; const)
1414open import  Level  using  (Level; _⊔_)
1515
1616private 
@@ -20,37 +20,51 @@ private
2020    Whatever  :  Set  w
2121
2222------------------------------------------------------------------------ 
23- -- Negation . 
23+ -- Definition . 
2424
2525infix  3  ¬_
2626¬_  :  Set  a →  Set  a
2727¬ A =  A →  ⊥
2828
2929------------------------------------------------------------------------ 
30- -- Stability. 
30+ -- Properties, I: as a definition in *minimal logic* 
31+ 
32+ -- Contraposition 
33+ contraposition  :  (A →  B) →  ¬ B →  ¬ A
34+ contraposition f ¬b a =   ¬b (f a)
35+ 
36+ -- Relationship to sum 
37+ infixr  1  _¬-⊎_
38+ _¬-⊎_  :  ¬ A →  ¬ B →  ¬ (A ⊎ B)
39+ _¬-⊎_ =  [_,_]
40+ 
41+ -- Self-contradictory propositions are false by 'diagonalisation' 
42+ contra-diagonal  :  (A →  ¬ A) →  ¬ A
43+ contra-diagonal self a =  self a a
3144
3245-- Double-negation 
3346DoubleNegation  :  Set  a →  Set  a
3447DoubleNegation A =  ¬ ¬ A
3548
3649-- Eta law for double-negation 
37- 
3850¬¬-η  :  A →  ¬ ¬ A
3951¬¬-η a ¬a =  ¬a a
4052
53+ -- Functoriality for double-negation 
54+ ¬¬-map  :  (A →  B) →  ¬ ¬ A →  ¬ ¬ B
55+ ¬¬-map =  contraposition ∘ contraposition
56+ 
57+ ------------------------------------------------------------------------ 
4158-- Stability under double-negation. 
4259Stable  :  Set  a →  Set  a
4360Stable A =  ¬ ¬ A →  A
4461
45- ------------------------------------------------------------------------ 
46- -- Relationship to sum 
47- 
48- infixr  1  _¬-⊎_
49- _¬-⊎_  :  ¬ A →  ¬ B →  ¬ (A ⊎ B)
50- _¬-⊎_ =  [_,_]
62+ -- Negated predicates are stable. 
63+ negated-stable  :  Stable (¬ A)
64+ negated-stable ¬¬¬a a =  ¬¬¬a (¬¬-η a)
5165
5266------------------------------------------------------------------------ 
53- -- Uses of negation  
67+ -- Properties, II: using the *ex falso* rule ⊥-elim  
5468
5569contradiction-irr  :  .A →  .(¬ A) →  Whatever
5670contradiction-irr a ¬a =  ⊥-elim-irr (¬a a)
@@ -65,27 +79,7 @@ contradiction₂ : A ⊎ B → ¬ A → ¬ B → Whatever
6579contradiction₂ (inj₁ a) ¬a ¬b =  contradiction a ¬a
6680contradiction₂ (inj₂ b) ¬a ¬b =  contradiction b ¬b
6781
68- contraposition  :  (A →  B) →  ¬ B →  ¬ A
69- contraposition f ¬b a =  contradiction (f a) ¬b
70- 
71- -- Self-contradictory propositions are false by 'diagonalisation' 
72- 
73- contra-diagonal  :  (A →  ¬ A) →  ¬ A
74- contra-diagonal self a =  self a a
75- 
7682-- Everything is stable in the double-negation monad. 
7783stable  :  ¬ ¬ Stable A
7884stable ¬[¬¬a→a] =  ¬[¬¬a→a] (contradiction (¬[¬¬a→a] ∘ const))
7985
80- -- Negated predicates are stable. 
81- negated-stable  :  Stable (¬ A)
82- negated-stable ¬¬¬a a =  ¬¬¬a (contradiction a)
83- 
84- ¬¬-map  :  (A →  B) →  ¬ ¬ A →  ¬ ¬ B
85- ¬¬-map f =  contraposition (contraposition f)
86- 
87- -- Note also the following use of flip: 
88- private 
89-   note  :  (A →  ¬ B) →  B →  ¬ A
90-   note =  flip
91- 
0 commit comments