10
10
module Data.Rational.Unnormalised.Properties where
11
11
12
12
open import Algebra
13
+ open import Algebra.Apartness
13
14
open import Algebra.Lattice
14
15
import Algebra.Consequences.Setoid as Consequences
15
16
open import Algebra.Consequences.Propositional
16
17
open import Algebra.Construct.NaturalChoice.Base
17
18
import Algebra.Construct.NaturalChoice.MinMaxOp as MinMaxOp
18
19
import Algebra.Lattice.Construct.NaturalChoice.MinMaxOp as LatticeMinMaxOp
20
+ open import Data.Empty using (⊥-elim)
19
21
open import Data.Bool.Base using (T; true; false)
20
22
open import Data.Nat.Base as ℕ using (suc; pred)
21
23
import Data.Nat.Properties as ℕ
@@ -24,7 +26,7 @@ open import Data.Integer.Base as ℤ using (ℤ; +0; +[1+_]; -[1+_]; 0ℤ; 1ℤ;
24
26
open import Data.Integer.Solver renaming (module +-*-Solver to ℤ-solver)
25
27
import Data.Integer.Properties as ℤ
26
28
open import Data.Rational.Unnormalised.Base
27
- open import Data.Product.Base using (_,_)
29
+ open import Data.Product.Base using (_,_; proj₁; proj₂ )
28
30
open import Data.Sum.Base using (_⊎_; [_,_]′; inj₁; inj₂)
29
31
import Data.Sign as Sign
30
32
open import Function.Base using (_on_; _$_; _∘_; flip)
@@ -36,6 +38,8 @@ open import Relation.Binary
36
38
import Relation.Binary.Consequences as BC
37
39
open import Relation.Binary.PropositionalEquality
38
40
import Relation.Binary.Properties.Poset as PosetProperties
41
+ open import Relation.Nullary using (yes; no)
42
+ import Relation.Binary.Reasoning.Setoid as SetoidReasoning
39
43
40
44
open import Algebra.Properties.CommutativeSemigroup ℤ.*-commutativeSemigroup
41
45
@@ -96,6 +100,21 @@ infix 4 _≃?_
96
100
_≃?_ : Decidable _≃_
97
101
p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.* ↧ p)
98
102
103
+ 0≠1 : 0ℚᵘ ≠ 1ℚᵘ
104
+ 0≠1 = Dec.from-no (0ℚᵘ ≃? 1ℚᵘ)
105
+
106
+ ≃-≠-irreflexive : Irreflexive _≃_ _≠_
107
+ ≃-≠-irreflexive x≃y x≠y = x≠y x≃y
108
+
109
+ ≠-symmetric : Symmetric _≠_
110
+ ≠-symmetric x≠y y≃x = x≠y (≃-sym y≃x)
111
+
112
+ ≠-cotransitive : Cotransitive _≠_
113
+ ≠-cotransitive {x} {y} x≠y z with x ≃? z | z ≃? y
114
+ ... | no x≠z | _ = inj₁ x≠z
115
+ ... | yes _ | no z≠y = inj₂ z≠y
116
+ ... | yes x≃z | yes z≃y = ⊥-elim (x≠y (≃-trans x≃z z≃y))
117
+
99
118
≃-isEquivalence : IsEquivalence _≃_
100
119
≃-isEquivalence = record
101
120
{ refl = ≃-refl
@@ -109,6 +128,17 @@ p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.*
109
128
; _≟_ = _≃?_
110
129
}
111
130
131
+ ≠-isApartnessRelation : IsApartnessRelation _≃_ _≠_
132
+ ≠-isApartnessRelation = record
133
+ { irrefl = ≃-≠-irreflexive
134
+ ; sym = ≠-symmetric
135
+ ; cotrans = ≠-cotransitive
136
+ }
137
+
138
+ ≠-tight : Tight _≃_ _≠_
139
+ proj₁ (≠-tight p q) ¬p≠q = Dec.decidable-stable (p ≃? q) ¬p≠q
140
+ proj₂ (≠-tight p q) p≃q p≠q = p≠q p≃q
141
+
112
142
≃-setoid : Setoid 0ℓ 0ℓ
113
143
≃-setoid = record
114
144
{ isEquivalence = ≃-isEquivalence
@@ -119,6 +149,8 @@ p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.*
119
149
{ isDecEquivalence = ≃-isDecEquivalence
120
150
}
121
151
152
+ module ≃-Reasoning = SetoidReasoning ≃-setoid
153
+
122
154
------------------------------------------------------------------------
123
155
-- Properties of -_
124
156
------------------------------------------------------------------------
@@ -1044,6 +1076,12 @@ p≤q⇒0≤q-p {p} {q} p≤q = begin
1044
1076
*-inverseʳ : ∀ p .{{_ : NonZero p}} → p * 1/ p ≃ 1ℚᵘ
1045
1077
*-inverseʳ p = ≃-trans (*-comm p (1/ p)) (*-inverseˡ p)
1046
1078
1079
+ ≠⇒invertible : p ≠ q → Invertible _≃_ 1ℚᵘ _*_ (p - q)
1080
+ ≠⇒invertible {p} {q} p≠q = _ , *-inverseˡ (p - q) , *-inverseʳ (p - q)
1081
+ where instance
1082
+ _ : NonZero (p - q)
1083
+ _ = ≢-nonZero (p≠q ∘ p-q≃0⇒p≃q p q)
1084
+
1047
1085
*-zeroˡ : LeftZero _≃_ 0ℚᵘ _*_
1048
1086
*-zeroˡ p@record{} = *≡* refl
1049
1087
@@ -1053,6 +1091,14 @@ p≤q⇒0≤q-p {p} {q} p≤q = begin
1053
1091
*-zero : Zero _≃_ 0ℚᵘ _*_
1054
1092
*-zero = *-zeroˡ , *-zeroʳ
1055
1093
1094
+ invertible⇒≠ : Invertible _≃_ 1ℚᵘ _*_ (p - q) → p ≠ q
1095
+ invertible⇒≠ {p} {q} (1/p-q , 1/x*x≃1 , x*1/x≃1) p≃q = 0≠1 (begin
1096
+ 0ℚᵘ ≈˘⟨ *-zeroˡ 1/p-q ⟩
1097
+ 0ℚᵘ * 1/p-q ≈˘⟨ *-congʳ (p≃q⇒p-q≃0 p q p≃q) ⟩
1098
+ (p - q) * 1/p-q ≈⟨ x*1/x≃1 ⟩
1099
+ 1ℚᵘ ∎)
1100
+ where open ≃-Reasoning
1101
+
1056
1102
*-distribˡ-+ : _DistributesOverˡ_ _≃_ _*_ _+_
1057
1103
*-distribˡ-+ p@record{} q@record{} r@record{} =
1058
1104
let ↥p = ↥ p; ↧p = ↧ p
@@ -1293,6 +1339,20 @@ nonNeg*nonNeg⇒nonNeg p q = nonNegative
1293
1339
; *-comm = *-comm
1294
1340
}
1295
1341
1342
+ +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ
1343
+ +-*-isHeytingCommutativeRing = record
1344
+ { isCommutativeRing = +-*-isCommutativeRing
1345
+ ; isApartnessRelation = ≠-isApartnessRelation
1346
+ ; #⇒invertible = ≠⇒invertible
1347
+ ; invertible⇒# = invertible⇒≠
1348
+ }
1349
+
1350
+ +-*-isHeytingField : IsHeytingField _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ
1351
+ +-*-isHeytingField = record
1352
+ { isHeytingCommutativeRing = +-*-isHeytingCommutativeRing
1353
+ ; tight = ≠-tight
1354
+ }
1355
+
1296
1356
------------------------------------------------------------------------
1297
1357
-- Algebraic bundles
1298
1358
@@ -1326,6 +1386,16 @@ nonNeg*nonNeg⇒nonNeg p q = nonNegative
1326
1386
{ isCommutativeRing = +-*-isCommutativeRing
1327
1387
}
1328
1388
1389
+ +-*-heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ
1390
+ +-*-heytingCommutativeRing = record
1391
+ { isHeytingCommutativeRing = +-*-isHeytingCommutativeRing
1392
+ }
1393
+
1394
+ +-*-heytingField : HeytingField 0ℓ 0ℓ 0ℓ
1395
+ +-*-heytingField = record
1396
+ { isHeytingField = +-*-isHeytingField
1397
+ }
1398
+
1329
1399
------------------------------------------------------------------------
1330
1400
-- Properties of 1/_
1331
1401
------------------------------------------------------------------------
0 commit comments