Skip to content
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
2 changes: 1 addition & 1 deletion Cubical/Algebra/OrderedCommRing/Base.agda
Original file line number Diff line number Diff line change
@@ -1,6 +1,6 @@
module Cubical.Algebra.OrderedCommRing.Base where
{-
Definition of an commutative ordered ring.
Definition of an ordered commutative ring.
-}

open import Cubical.Foundations.Prelude
Expand Down
66 changes: 66 additions & 0 deletions Cubical/Algebra/OrderedCommRing/Instances/Int.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,66 @@
module Cubical.Algebra.OrderedCommRing.Instances.Int where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Function
open import Cubical.Foundations.Equiv

open import Cubical.Data.Empty as ⊥

open import Cubical.HITs.PropositionalTruncation

open import Cubical.Data.Int as ℤ
renaming (_+_ to _+ℤ_ ; _-_ to _-ℤ_; -_ to -ℤ_ ; _·_ to _·ℤ_)
open import Cubical.Data.Int.Order
renaming (_<_ to _<ℤ_ ; _≤_ to _≤ℤ_)

open import Cubical.Algebra.CommRing
open import Cubical.Algebra.CommRing.Instances.Int

open import Cubical.Algebra.OrderedCommRing

open import Cubical.Relation.Nullary

open import Cubical.Relation.Binary.Order.StrictOrder
open import Cubical.Relation.Binary.Order.StrictOrder.Instances.Int

open import Cubical.Relation.Binary.Order.Pseudolattice
open import Cubical.Relation.Binary.Order.Pseudolattice.Instances.Int

open CommRingStr
open OrderedCommRingStr
open PseudolatticeStr
open StrictOrderStr

ℤOrderedCommRing : OrderedCommRing ℓ-zero ℓ-zero
fst ℤOrderedCommRing = ℤ
0r (snd ℤOrderedCommRing) = 0
1r (snd ℤOrderedCommRing) = 1
_+_ (snd ℤOrderedCommRing) = _+ℤ_
_·_ (snd ℤOrderedCommRing) = _·ℤ_
-_ (snd ℤOrderedCommRing) = -ℤ_
_<_ (snd ℤOrderedCommRing) = _<ℤ_
_≤_ (snd ℤOrderedCommRing) = _≤ℤ_
isOrderedCommRing (snd ℤOrderedCommRing) = isOrderedCommRingℤ
where
open IsOrderedCommRing

isOrderedCommRingℤ : IsOrderedCommRing 0 1 _+ℤ_ _·ℤ_ -ℤ_ _<ℤ_ _≤ℤ_
isOrderedCommRingℤ .isCommRing = ℤCommRing .snd .isCommRing
isOrderedCommRingℤ .isPseudolattice = ℤ≤Pseudolattice .snd .is-pseudolattice
isOrderedCommRingℤ .isStrictOrder = ℤ<StrictOrder .snd .isStrictOrder
isOrderedCommRingℤ .<-≤-weaken = λ _ _ → <-weaken
isOrderedCommRingℤ .≤≃¬> = λ x y →
propBiimpl→Equiv isProp≤ (isProp¬ (y <ℤ x))
(λ x≤y y<x → isIrrefl< (≤<-trans x≤y y<x))
(λ ¬y<x → case x ≟ y return (λ _ → x ≤ℤ y) of λ {
(lt x<y) → <-weaken x<y ;
(eq x≡y) → subst (x ≤ℤ_) x≡y isRefl≤ ;
(gt y<z) → ⊥.rec (¬y<x y<z) })
isOrderedCommRingℤ .+MonoR≤ = λ _ _ z → ≤-+o {o = z}
isOrderedCommRingℤ .+MonoR< = λ _ _ z → <-+o {o = z}
isOrderedCommRingℤ .posSum→pos∨pos = λ _ _ → ∣_∣₁ ∘ 0<+ _ _
isOrderedCommRingℤ .<-≤-trans = λ _ _ _ → <≤-trans
isOrderedCommRingℤ .≤-<-trans = λ _ _ _ → ≤<-trans
isOrderedCommRingℤ .·MonoR≤ = λ _ _ _ → 0≤o→≤-·o
isOrderedCommRingℤ .·MonoR< = λ _ _ _ → 0<o→<-·o
isOrderedCommRingℤ .0<1 = isRefl≤
13 changes: 13 additions & 0 deletions Cubical/Data/Int/Order.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,6 +10,7 @@ open import Cubical.Data.Int.Properties as ℤ
open import Cubical.Data.Nat as ℕ
open import Cubical.Data.NatPlusOne.Base as ℕ₊₁
open import Cubical.Data.Sigma
open import Cubical.Data.Sum

open import Cubical.Relation.Nullary

Expand Down Expand Up @@ -450,6 +451,18 @@ min≤ {negsuc (suc m)} {negsuc (suc n)} = pred-≤-pred (subst (_≤ negsuc m)
cong₂ ℤ.max (≤→max m≤n) (≤→max o≤s))
(≤max {m = ℤ.max m o} {n = ℤ.max n s})

0<+ : ∀ m n → 0 < m ℤ.+ n → (0 < m) ⊎ (0 < n)
0<+ (pos zero) (pos zero) = ⊥.rec ∘ isIrrefl<
0<+ (pos zero) (pos (suc n)) = inr ∘ subst (0 <_) (sym $ pos0+ _)
0<+ (pos (suc m)) (pos n) = λ _ → inl (suc-≤-suc zero-≤pos)
0<+ (pos zero) (negsuc n) = ⊥.rec ∘ ¬pos≤negsuc ∘ subst (0 <_)
(sym $ pos0+ (negsuc n))
0<+ (pos (suc m)) (negsuc n) = λ _ → inl (suc-≤-suc zero-≤pos)
0<+ (negsuc m) (pos zero) = ⊥.rec ∘ ¬pos≤negsuc
0<+ (negsuc m) (pos (suc n)) = λ _ → inr (suc-≤-suc zero-≤pos)
0<+ (negsuc m) (negsuc n) = ⊥.rec ∘ ¬pos≤negsuc ∘ subst (0 <_)
(sym $ neg+ (suc m) (suc n))

≤Dec : ∀ m n → Dec (m ≤ n)
≤Dec (pos zero) (pos n) = yes zero-≤pos
≤Dec (pos (suc m)) (pos zero) = no ¬-pos<-zero
Expand Down
14 changes: 14 additions & 0 deletions Cubical/Data/Nat/Order.agda
Original file line number Diff line number Diff line change
Expand Up @@ -277,6 +277,20 @@ min-≤-right {zero} {n} = zero-≤
min-≤-right {suc m} {zero} = ≤-refl
min-≤-right {suc m} {suc n} = subst (_≤ _) (sym minSuc) $ suc-≤-suc $ min-≤-right {m} {n}

maxLUB : ∀ {x} → m ≤ x → n ≤ x → max m n ≤ x
maxLUB {zero} {n} _ n≤x = n≤x
maxLUB {suc m} {zero} sm≤x _ = sm≤x
maxLUB {suc m} {suc n} sm≤x sn≤x with m <ᵇ n
... | false = sm≤x
... | true = sn≤x

minGLB : ∀ {x} → x ≤ m → x ≤ n → x ≤ min m n
minGLB {zero} {n} x≤0 _ = x≤0
minGLB {suc m} {zero} _ x≤0 = x≤0
minGLB {suc m} {suc n} x≤sm x≤sn with m <ᵇ n
... | false = x≤sn
... | true = x≤sm

-- Boolean order relations and their conversions to/from ≤ and <

_≤ᵇ_ : ℕ → ℕ → Bool
Expand Down
33 changes: 33 additions & 0 deletions Cubical/Relation/Binary/Order/Pseudolattice/Base.agda
Original file line number Diff line number Diff line change
@@ -1,6 +1,9 @@
module Cubical.Relation.Binary.Order.Pseudolattice.Base where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Function
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.SIP

open import Cubical.Relation.Binary.Base
Expand Down Expand Up @@ -57,3 +60,33 @@ makeIsPseudolattice {_≤_ = _≤_} is-setL is-prop-valued is-refl is-trans is-a
PS : IsPseudolattice _≤_
PS .IsPseudolattice.isPoset = isposet is-setL is-prop-valued is-refl is-trans is-antisym
PS .IsPseudolattice.isPseudolattice = is-meet-semipseudolattice , is-join-semipseudolattice

module _
(P : Poset ℓ ℓ') (_∧_ _∨_ : ⟨ P ⟩ → ⟨ P ⟩ → ⟨ P ⟩) where
open PosetStr (str P) renaming (_≤_ to infix 8 _≤_)
module _
(π₁ : ∀ {a b} → a ∧ b ≤ a)
(π₂ : ∀ {a b} → a ∧ b ≤ b)
(ϕ : ∀ {a b x} → x ≤ a → x ≤ b → x ≤ a ∧ b)
(ι₁ : ∀ {a b} → a ≤ a ∨ b)
(ι₂ : ∀ {a b} → b ≤ a ∨ b)
(ψ : ∀ {a b x} → a ≤ x → b ≤ x → a ∨ b ≤ x) where

makePseudolatticeFromPoset : Pseudolattice ℓ ℓ'
makePseudolatticeFromPoset .fst = ⟨ P ⟩
makePseudolatticeFromPoset .snd .PseudolatticeStr._≤_ = (str P) .PosetStr._≤_
makePseudolatticeFromPoset .snd .PseudolatticeStr.is-pseudolattice = isPL where
isPL : IsPseudolattice _≤_
isPL .IsPseudolattice.isPoset = isPoset
isPL .IsPseudolattice.isPseudolattice .fst a b .fst = a ∧ b
isPL .IsPseudolattice.isPseudolattice .fst a b .snd x = propBiimpl→Equiv
(is-prop-valued _ _)
(isProp× (is-prop-valued _ _) (is-prop-valued _ _))
(λ x≤a∧b → is-trans _ _ _ x≤a∧b π₁ , is-trans _ _ _ x≤a∧b π₂)
(uncurry ϕ)
isPL .IsPseudolattice.isPseudolattice .snd a b .fst = a ∨ b
isPL .IsPseudolattice.isPseudolattice .snd a b .snd x = propBiimpl→Equiv
(is-prop-valued _ _)
(isProp× (is-prop-valued _ _) (is-prop-valued _ _))
(λ a∨b≤x → is-trans _ _ _ ι₁ a∨b≤x , is-trans _ _ _ ι₂ a∨b≤x)
(uncurry ψ)
18 changes: 18 additions & 0 deletions Cubical/Relation/Binary/Order/Pseudolattice/Instances/Int.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,18 @@
module Cubical.Relation.Binary.Order.Pseudolattice.Instances.Int where

open import Cubical.Foundations.Prelude

open import Cubical.Data.Int
open import Cubical.Data.Int.Order renaming (_≤_ to _≤ℤ_)

open import Cubical.Relation.Binary.Order.Poset.Instances.Int
open import Cubical.Relation.Binary.Order.Pseudolattice

ℤ≤Pseudolattice : Pseudolattice ℓ-zero ℓ-zero
ℤ≤Pseudolattice = makePseudolatticeFromPoset ℤ≤Poset min max
min≤
(λ {a b} → subst (_≤ℤ b) (minComm b a) min≤)
(λ {a b} x≤a x≤b → subst (_≤ℤ min a b) (minIdem _) (≤MonotoneMin x≤a x≤b))
≤max
(λ {a b} → subst (b ≤ℤ_) (maxComm b a) ≤max)
(λ {a b} a≤x b≤x → subst (max a b ≤ℤ_) (maxIdem _) (≤MonotoneMax a≤x b≤x))
18 changes: 18 additions & 0 deletions Cubical/Relation/Binary/Order/Pseudolattice/Instances/Nat.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,18 @@
module Cubical.Relation.Binary.Order.Pseudolattice.Instances.Nat where

open import Cubical.Foundations.Prelude

open import Cubical.Data.Nat
open import Cubical.Data.Nat.Order renaming (_≤_ to _≤ℕ_)

open import Cubical.Relation.Binary.Order.Poset.Instances.Nat
open import Cubical.Relation.Binary.Order.Pseudolattice

ℕ≤Pseudolattice : Pseudolattice ℓ-zero ℓ-zero
ℕ≤Pseudolattice = makePseudolatticeFromPoset ℕ≤Poset min max
min-≤-left
(λ {a b} min-≤-right {a} {b})
minGLB
left-≤-max
(λ {a b} right-≤-max {b} {a})
maxLUB