Skip to content

Commit

Permalink
Cardinality and Order (#969)
Browse files Browse the repository at this point in the history
* Proof that embeddings are monotone over sums

Also cleaned up several aspects of things.

* Embeddings are monotone over products

* Define cardinal exponentiation

* Minor formatting clean-up

* Removal of some excessive imports

* Renamed annihiliation by 0 to be less vague

* Renamed Cantor's Theorem results

* Cleaned up order property definitions

* Replace renaming

* Require preordering for order properties

* Remove extra level variable that's unused

* Removed redundant implicit level term

* Replace type families with embeddings as subtypes

More consistent with the idea of a "subset" and
also makes reasoning easier by keeping the sigma
types out of the type signatures.

* Cleaned up types for embeddings in order props

* Added function in Embedding to simplify constructs

Sigma-types where the type family is always a prop
embed into the first argument's type

* Minor clean-up

* Better respecting camel case

* Name-cleaning and removing unused imports

* Reintroduce erroneously-removed newline

* Renamed loset to toset, strict loset to loset

Also required losets to be weakly linear

* Fixed paths in top-level file

* Defined induced relations over embeddings

* Better capitalization

* Renamed one more detail

* Moved order properties into their own folders

* Reduced imports for cardinals

* Removed unused apartness import

* Changed capitalization back for conversions

* Proved induced relations preserve order properties

* Loosened embedding to mere function for bounds

* Fixed whitespace

* Fix rebase issues

* Resolve PR comments
  • Loading branch information
LuuBluum authored Sep 16, 2023
1 parent 7d47cf3 commit 85f8a71
Show file tree
Hide file tree
Showing 46 changed files with 2,044 additions and 42 deletions.
2 changes: 1 addition & 1 deletion Cubical/Algebra/DistLattice/BigOps.agda
Original file line number Diff line number Diff line change
Expand Up @@ -29,7 +29,7 @@ open import Cubical.Algebra.CommMonoid
open import Cubical.Algebra.Semilattice
open import Cubical.Algebra.Lattice
open import Cubical.Algebra.DistLattice
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset


private
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/Lattice/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@ open import Cubical.Algebra.CommMonoid
open import Cubical.Algebra.Semilattice
open import Cubical.Algebra.Lattice.Base

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

private
variable
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/OrderedCommMonoid/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,7 +10,7 @@ open import Cubical.Foundations.Structure using (⟨_⟩; str)

open import Cubical.Algebra.CommMonoid.Base

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

private
variable
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/OrderedCommMonoid/PropCompletion.agda
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ open import Cubical.HITs.PropositionalTruncation renaming (rec to propTruncRec;
open import Cubical.Algebra.CommMonoid.Base
open import Cubical.Algebra.OrderedCommMonoid

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

private
variable
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/OrderedCommMonoid/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ open import Cubical.Data.Sigma
open import Cubical.Algebra.CommMonoid
open import Cubical.Algebra.OrderedCommMonoid.Base

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset


private
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/Semilattice/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -34,7 +34,7 @@ open import Cubical.Displayed.Record
open import Cubical.Displayed.Universe

open import Cubical.Relation.Binary
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Reflection.RecordEquiv

Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/ZariskiLattice/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -22,7 +22,7 @@ open import Cubical.Data.FinData
open import Cubical.Data.Unit
open import Cubical.Relation.Nullary
open import Cubical.Relation.Binary
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Algebra.Ring
open import Cubical.Algebra.Ring.Properties
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/ZariskiLattice/StructureSheaf.agda
Original file line number Diff line number Diff line change
Expand Up @@ -34,7 +34,7 @@ open import Cubical.Data.FinData
open import Cubical.Data.Unit
open import Cubical.Relation.Nullary
open import Cubical.Relation.Binary
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Algebra.Ring
open import Cubical.Algebra.Ring.Properties
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/ZariskiLattice/StructureSheafPullback.agda
Original file line number Diff line number Diff line change
Expand Up @@ -34,7 +34,7 @@ open import Cubical.Data.FinData
open import Cubical.Data.Unit
open import Cubical.Relation.Nullary
open import Cubical.Relation.Binary
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Algebra.Ring
open import Cubical.Algebra.Ring.Properties
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Algebra/ZariskiLattice/UniversalProperty.agda
Original file line number Diff line number Diff line change
Expand Up @@ -23,7 +23,7 @@ open import Cubical.Data.FinData
open import Cubical.Data.Unit
open import Cubical.Relation.Nullary
open import Cubical.Relation.Binary
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Algebra.Ring
open import Cubical.Algebra.Ring.Properties
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Categories/DistLatticeSheaf/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,7 +10,7 @@ open import Cubical.Data.Nat using (ℕ)
open import Cubical.Data.Nat.Order
open import Cubical.Data.FinData

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Algebra.Semilattice
open import Cubical.Algebra.Lattice
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Categories/DistLatticeSheaf/ComparisonLemma.agda
Original file line number Diff line number Diff line change
Expand Up @@ -26,7 +26,7 @@ open import Cubical.Data.FinData
open import Cubical.Data.FinData.Order
open import Cubical.Data.Sum

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset
open import Cubical.HITs.PropositionalTruncation as PT

open import Cubical.Algebra.Semilattice
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Categories/DistLatticeSheaf/Diagram.agda
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@ open import Cubical.Data.FinData.Order renaming (_<'Fin_ to _<_)
open import Cubical.Data.Sum

open import Cubical.Relation.Nullary
open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Categories.Category
open import Cubical.Categories.Functor
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Categories/DistLatticeSheaf/Extension.agda
Original file line number Diff line number Diff line change
Expand Up @@ -23,7 +23,7 @@ open import Cubical.Data.FinData
open import Cubical.Data.FinData.Order
open import Cubical.Data.Sum

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset
open import Cubical.HITs.PropositionalTruncation

open import Cubical.Algebra.Semilattice
Expand Down
2 changes: 1 addition & 1 deletion Cubical/Categories/Instances/Poset.agda
Original file line number Diff line number Diff line change
Expand Up @@ -3,7 +3,7 @@ module Cubical.Categories.Instances.Poset where

open import Cubical.Foundations.Prelude

open import Cubical.Relation.Binary.Poset
open import Cubical.Relation.Binary.Order.Poset

open import Cubical.Categories.Category

Expand Down
5 changes: 5 additions & 0 deletions Cubical/Data/Cardinality.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
{-# OPTIONS --safe #-}
module Cubical.Data.Cardinality where

open import Cubical.Data.Cardinality.Base public
open import Cubical.Data.Cardinality.Properties public
54 changes: 54 additions & 0 deletions Cubical/Data/Cardinality/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,54 @@
{-
This file contains:
- Treatment of set truncation as cardinality
as per the HoTT book, section 10.2
-}
{-# OPTIONS --safe #-}
module Cubical.Data.Cardinality.Base where

open import Cubical.HITs.SetTruncation as ∥₂
open import Cubical.Foundations.Prelude
open import Cubical.Foundations.HLevels
open import Cubical.Data.Empty
open import Cubical.Data.Sigma
open import Cubical.Data.Sum
open import Cubical.Data.Unit

private
variable
: Level

-- First, we define a cardinal as the set truncation of Set
Card : Type (ℓ-suc ℓ)
Card {ℓ} = ∥ hSet ℓ ∥₂

-- Verify that it's a set
isSetCard : isSet (Card {ℓ})
isSetCard = isSetSetTrunc

-- Set truncation of a set is its cardinality
card : hSet ℓ Card {ℓ}
card = ∣_∣₂

-- Some special cardinalities
𝟘 : Card {ℓ}
𝟘 = card (⊥* , isProp→isSet isProp⊥*)

𝟙 : Card {ℓ}
𝟙 = card (Unit* , isSetUnit*)

-- Now we define some arithmetic
_+_ : Card {ℓ} Card {ℓ} Card {ℓ}
_+_ = ∥₂.rec2 isSetCard λ (A , isSetA) (B , isSetB)
card ((A ⊎ B) , isSet⊎ isSetA isSetB)

_·_ : Card {ℓ} Card {ℓ} Card {ℓ}
_·_ = ∥₂.rec2 isSetCard λ (A , isSetA) (B , isSetB)
card ((A × B) , isSet× isSetA isSetB)

_^_ : Card {ℓ} Card {ℓ} Card {ℓ}
_^_ = ∥₂.rec2 isSetCard λ (A , isSetA) (B , _)
card ((B A) , isSet→ isSetA)
194 changes: 194 additions & 0 deletions Cubical/Data/Cardinality/Properties.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,194 @@
{-
This file contains:
- Properties of cardinality
- Preordering of cardinalities
-}
{-# OPTIONS --safe #-}
module Cubical.Data.Cardinality.Properties where

open import Cubical.HITs.SetTruncation as ∥₂
open import Cubical.Data.Cardinality.Base

open import Cubical.Algebra.CommSemiring

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Structure
open import Cubical.Functions.Embedding
open import Cubical.Data.Empty as ⊥
open import Cubical.Data.Sigma
open import Cubical.Data.Sum as ⊎
open import Cubical.Data.Unit
open import Cubical.HITs.PropositionalTruncation as ∥₁
open import Cubical.Relation.Binary.Base
open import Cubical.Relation.Binary.Order.Preorder.Base
open import Cubical.Relation.Binary.Order.Properties

private
variable
: Level

-- Cardinality is a commutative semiring
module _ where
private
+Assoc : (A B C : Card {ℓ}) A + (B + C) ≡ (A + B) + C
+Assoc = ∥₂.elim3 (λ _ _ _ isProp→isSet (isSetCard _ _))
λ _ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(sym (isoToPath ⊎-assoc-Iso)))

·Assoc : (A B C : Card {ℓ}) A · (B · C) ≡ (A · B) · C
·Assoc = ∥₂.elim3 (λ _ _ _ isProp→isSet (isSetCard _ _))
λ _ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(sym (isoToPath Σ-assoc-Iso)))

+IdR𝟘 : (A : Card {ℓ}) A + 𝟘 ≡ A
+IdR𝟘 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath ⊎-IdR-⊥*-Iso))

+IdL𝟘 : (A : Card {ℓ}) 𝟘 + A ≡ A
+IdL𝟘 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath ⊎-IdL-⊥*-Iso))

·IdR𝟙 : (A : Card {ℓ}) A · 𝟙 ≡ A
·IdR𝟙 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath rUnit*×Iso))

·IdL𝟙 : (A : Card {ℓ}) 𝟙 · A ≡ A
·IdL𝟙 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath lUnit*×Iso))

+Comm : (A B : Card {ℓ}) (A + B) ≡ (B + A)
+Comm = ∥₂.elim2 (λ _ _ isProp→isSet (isSetCard _ _))
λ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath ⊎-swap-Iso))

·Comm : (A B : Card {ℓ}) (A · B) ≡ (B · A)
·Comm = ∥₂.elim2 (λ _ _ isProp→isSet (isSetCard _ _))
λ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath Σ-swap-Iso))

·LDist+ : (A B C : Card {ℓ}) A · (B + C) ≡ (A · B) + (A · C)
·LDist+ = ∥₂.elim3 (λ _ _ _ isProp→isSet (isSetCard _ _))
λ _ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath ×DistL⊎Iso))

AnnihilL : (A : Card {ℓ}) 𝟘 · A ≡ 𝟘
AnnihilL = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath (ΣEmpty*Iso λ _ _)))

isCardCommSemiring : IsCommSemiring {ℓ-suc ℓ} 𝟘 𝟙 _+_ _·_
isCardCommSemiring = makeIsCommSemiring isSetCard +Assoc +IdR𝟘 +Comm ·Assoc ·IdR𝟙 ·LDist+ AnnihilL ·Comm

-- Exponentiation is also well-behaved

^AnnihilR𝟘 : (A : Card {ℓ}) A ^ 𝟘 ≡ 𝟙
^AnnihilR𝟘 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath (iso⊥ _)))
where iso⊥ : A Iso (⊥* A) Unit*
Iso.fun (iso⊥ A) _ = tt*
Iso.inv (iso⊥ A) _ ()
Iso.rightInv (iso⊥ A) _ = refl
Iso.leftInv (iso⊥ A) _ i ()

^IdR𝟙 : (A : Card {ℓ}) A ^ 𝟙 ≡ A
^IdR𝟙 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath (iso⊤ _)))
where iso⊤ : A Iso (Unit* A) A
Iso.fun (iso⊤ _) f = f tt*
Iso.inv (iso⊤ _) a _ = a
Iso.rightInv (iso⊤ _) _ = refl
Iso.leftInv (iso⊤ _) _ = refl

^AnnihilL𝟙 : (A : Card {ℓ}) 𝟙 ^ A ≡ 𝟙
^AnnihilL𝟙 = ∥₂.elim (λ _ isProp→isSet (isSetCard _ _))
λ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath (iso⊤ _)))
where iso⊤ : A Iso (A Unit*) Unit*
Iso.fun (iso⊤ _) _ = tt*
Iso.inv (iso⊤ _) _ _ = tt*
Iso.rightInv (iso⊤ _) _ = refl
Iso.leftInv (iso⊤ _) _ = refl

^LDist+ : (A B C : Card {ℓ}) A ^ (B + C) ≡ (A ^ B) · (A ^ C)
^LDist+ = ∥₂.elim3 (λ _ _ _ isProp→isSet (isSetCard _ _))
λ _ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath Π⊎Iso))

^Assoc· : (A B C : Card {ℓ}) A ^ (B · C) ≡ (A ^ B) ^ C
^Assoc· = ∥₂.elim3 (λ _ _ _ isProp→isSet (isSetCard _ _))
λ _ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath (is _ _ _)))
where is : A B C Iso (B × C A) (C B A)
is A B C = (B × C A) Iso⟨ domIso Σ-swap-Iso ⟩
(C × B A) Iso⟨ curryIso ⟩
(C B A) ∎Iso

^RDist· : (A B C : Card {ℓ}) (A · B) ^ C ≡ (A ^ C) · (B ^ C)
^RDist· = ∥₂.elim3 (λ _ _ _ isProp→isSet (isSetCard _ _))
λ _ _ _ cong ∣_∣₂ (Σ≡Prop (λ _ isPropIsSet)
(isoToPath Σ-Π-Iso))


-- With basic arithmetic done, we can now define an ordering over cardinals
module _ where
private
_≲'_ : Card {ℓ} Card {ℓ} hProp ℓ
_≲'_ = ∥₂.rec2 isSetHProp λ (A , _) (B , _) ∥ A ↪ B ∥₁ , isPropPropTrunc

_≲_ : Rel (Card {ℓ}) (Card {ℓ}) ℓ
A ≲ B = ⟨ A ≲' B ⟩

isPreorder≲ : IsPreorder {ℓ-suc ℓ} _≲_
isPreorder≲
= ispreorder isSetCard prop reflexive transitive
where prop : BinaryRelation.isPropValued _≲_
prop a b = str (a ≲' b)

reflexive : BinaryRelation.isRefl _≲_
reflexive = ∥₂.elim (λ A isProp→isSet (prop A A))
(λ (A , _) ∣ id↪ A ∣₁)

transitive : BinaryRelation.isTrans _≲_
transitive = ∥₂.elim3 (λ x _ z isSetΠ2
λ _ _ isProp→isSet
(prop x z))
(λ (A , _) (B , _) (C , _)
∥₁.map2 λ A↪B B↪C
compEmbedding
B↪C
A↪B)

isLeast𝟘 : {ℓ} isLeast isPreorder≲ (Card {ℓ} , id↪ (Card {ℓ})) (𝟘 {ℓ})
isLeast𝟘 = ∥₂.elim (λ x isProp→isSet (IsPreorder.is-prop-valued
isPreorder≲ 𝟘 x))
(λ _ ∣ ⊥.rec* , (λ ()) ∣₁)

-- Our arithmetic behaves as expected over our preordering
+Monotone≲ : (A B C D : Card {ℓ}) A ≲ C B ≲ D (A + B) ≲ (C + D)
+Monotone≲
= ∥₂.elim4 (λ w x y z isSetΠ2 λ _ _ isProp→isSet (IsPreorder.is-prop-valued
isPreorder≲
(w + x)
(y + z)))
λ (A , _) (B , _) (C , _) (D , _)
∥₁.map2 λ A↪C B↪D ⊎Monotone↪ A↪C B↪D

·Monotone≲ : (A B C D : Card {ℓ}) A ≲ C B ≲ D (A · B) ≲ (C · D)
·Monotone≲
= ∥₂.elim4 (λ w x y z isSetΠ2 λ _ _ isProp→isSet (IsPreorder.is-prop-valued
isPreorder≲
(w · x)
(y · z)))
λ (A , _) (B , _) (C , _) (D , _)
∥₁.map2 λ A↪C B↪D ×Monotone↪ A↪C B↪D
Loading

0 comments on commit 85f8a71

Please sign in to comment.