Skip to content

Commit 40a2832

Browse files
authored
Added Unnormalised Rational Field Structure (#1959)
1 parent 9a240ee commit 40a2832

File tree

2 files changed

+97
-1
lines changed

2 files changed

+97
-1
lines changed

CHANGELOG.md

+26
Original file line numberDiff line numberDiff line change
@@ -3243,6 +3243,32 @@ This is a full list of proofs that have changed form to use irrelevant instance
32433243
<-weakInduction-startingFrom : P i → (∀ j → P (inject₁ j) → P (suc j)) → ∀ {j} → j ≥ i → P j
32443244
```
32453245

3246+
* Added new module to `Data.Rational.Unnormalised.Properties`
3247+
```agda
3248+
module ≃-Reasoning = SetoidReasoning ≃-setoid
3249+
```
3250+
3251+
* Added new functions to `Data.Rational.Unnormalised.Properties`
3252+
```agda
3253+
0≠1 : 0ℚᵘ ≠ 1ℚᵘ
3254+
≃-≠-irreflexive : Irreflexive _≃_ _≠_
3255+
≠-symmetric : Symmetric _≠_
3256+
≠-cotransitive : Cotransitive _≠_
3257+
≠⇒invertible : p ≠ q → Invertible _≃_ 1ℚᵘ _*_ (p - q)
3258+
```
3259+
3260+
* Added new structures to `Data.Rational.Unnormalised.Properties`
3261+
```agda
3262+
+-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ
3263+
+-*-isHeytingField : IsHeytingField _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ
3264+
```
3265+
3266+
* Added new bundles to `Data.Rational.Unnormalised.Properties`
3267+
```agda
3268+
+-*-heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ
3269+
+-*-heytingField : HeytingField 0ℓ 0ℓ 0ℓ
3270+
```
3271+
32463272
* Added new function to `Data.Vec.Relation.Binary.Pointwise.Inductive`
32473273
```agda
32483274
cong-[_]≔ : Pointwise _∼_ xs ys → Pointwise _∼_ (xs [ i ]≔ p) (ys [ i ]≔ p)

src/Data/Rational/Unnormalised/Properties.agda

+71-1
Original file line numberDiff line numberDiff line change
@@ -10,12 +10,14 @@
1010
module Data.Rational.Unnormalised.Properties where
1111

1212
open import Algebra
13+
open import Algebra.Apartness
1314
open import Algebra.Lattice
1415
import Algebra.Consequences.Setoid as Consequences
1516
open import Algebra.Consequences.Propositional
1617
open import Algebra.Construct.NaturalChoice.Base
1718
import Algebra.Construct.NaturalChoice.MinMaxOp as MinMaxOp
1819
import Algebra.Lattice.Construct.NaturalChoice.MinMaxOp as LatticeMinMaxOp
20+
open import Data.Empty using (⊥-elim)
1921
open import Data.Bool.Base using (T; true; false)
2022
open import Data.Nat.Base as ℕ using (suc; pred)
2123
import Data.Nat.Properties as ℕ
@@ -24,7 +26,7 @@ open import Data.Integer.Base as ℤ using (ℤ; +0; +[1+_]; -[1+_]; 0ℤ; 1ℤ;
2426
open import Data.Integer.Solver renaming (module +-*-Solver to ℤ-solver)
2527
import Data.Integer.Properties as ℤ
2628
open import Data.Rational.Unnormalised.Base
27-
open import Data.Product.Base using (_,_)
29+
open import Data.Product.Base using (_,_; proj₁; proj₂)
2830
open import Data.Sum.Base using (_⊎_; [_,_]′; inj₁; inj₂)
2931
import Data.Sign as Sign
3032
open import Function.Base using (_on_; _$_; _∘_; flip)
@@ -36,6 +38,8 @@ open import Relation.Binary
3638
import Relation.Binary.Consequences as BC
3739
open import Relation.Binary.PropositionalEquality
3840
import Relation.Binary.Properties.Poset as PosetProperties
41+
open import Relation.Nullary using (yes; no)
42+
import Relation.Binary.Reasoning.Setoid as SetoidReasoning
3943

4044
open import Algebra.Properties.CommutativeSemigroup ℤ.*-commutativeSemigroup
4145

@@ -96,6 +100,21 @@ infix 4 _≃?_
96100
_≃?_ : Decidable _≃_
97101
p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.* ↧ p)
98102

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+
99118
≃-isEquivalence : IsEquivalence _≃_
100119
≃-isEquivalence = record
101120
{ refl = ≃-refl
@@ -109,6 +128,17 @@ p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.*
109128
; _≟_ = _≃?_
110129
}
111130

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+
112142
≃-setoid : Setoid 0ℓ 0ℓ
113143
≃-setoid = record
114144
{ isEquivalence = ≃-isEquivalence
@@ -119,6 +149,8 @@ p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.*
119149
{ isDecEquivalence = ≃-isDecEquivalence
120150
}
121151

152+
module ≃-Reasoning = SetoidReasoning ≃-setoid
153+
122154
------------------------------------------------------------------------
123155
-- Properties of -_
124156
------------------------------------------------------------------------
@@ -1044,6 +1076,12 @@ p≤q⇒0≤q-p {p} {q} p≤q = begin
10441076
*-inverseʳ : p .{{_ : NonZero p}} p * 1/ p ≃ 1ℚᵘ
10451077
*-inverseʳ p = ≃-trans (*-comm p (1/ p)) (*-inverseˡ p)
10461078

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+
10471085
*-zeroˡ : LeftZero _≃_ 0ℚᵘ _*_
10481086
*-zeroˡ p@record{} = *≡* refl
10491087

@@ -1053,6 +1091,14 @@ p≤q⇒0≤q-p {p} {q} p≤q = begin
10531091
*-zero : Zero _≃_ 0ℚᵘ _*_
10541092
*-zero = *-zeroˡ , *-zeroʳ
10551093

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+
10561102
*-distribˡ-+ : _DistributesOverˡ_ _≃_ _*_ _+_
10571103
*-distribˡ-+ p@record{} q@record{} r@record{} =
10581104
let ↥p = ↥ p; ↧p = ↧ p
@@ -1293,6 +1339,20 @@ nonNeg*nonNeg⇒nonNeg p q = nonNegative
12931339
; *-comm = *-comm
12941340
}
12951341

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+
12961356
------------------------------------------------------------------------
12971357
-- Algebraic bundles
12981358

@@ -1326,6 +1386,16 @@ nonNeg*nonNeg⇒nonNeg p q = nonNegative
13261386
{ isCommutativeRing = +-*-isCommutativeRing
13271387
}
13281388

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+
13291399
------------------------------------------------------------------------
13301400
-- Properties of 1/_
13311401
------------------------------------------------------------------------

0 commit comments

Comments
 (0)