From aeb89de300584c8ee9fcf88d0f612add37acfb68 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 19 Jan 2025 01:51:45 +0100 Subject: [PATCH 01/11] marked book references --- Cubical/Data/Fin/Properties.agda | 2 +- Cubical/Data/Int/Order.agda | 43 + Cubical/Data/Int/Properties.agda | 23 +- Cubical/Data/Nat/Mod.agda | 121 +- Cubical/Data/Nat/Order.agda | 44 + Cubical/Data/Nat/Order/Recursive.agda | 28 +- Cubical/Data/Nat/Properties.agda | 4 + Cubical/Data/Rationals/Order.agda | 315 +++- Cubical/Data/Rationals/Order/Properties.agda | 1506 ++++++++++++++++++ Cubical/Data/Rationals/Properties.agda | 266 +++- Cubical/Foundations/HLevels.agda | 22 + Cubical/HITs/CauchyReals/Base.agda | 348 ++++ Cubical/HITs/CauchyReals/Closeness.agda | 880 ++++++++++ Cubical/HITs/CauchyReals/Continuous.agda | 970 +++++++++++ Cubical/HITs/CauchyReals/Derivative.agda | 165 ++ Cubical/HITs/CauchyReals/Inverse.agda | 1135 +++++++++++++ Cubical/HITs/CauchyReals/Lems.agda | 280 ++++ Cubical/HITs/CauchyReals/Lipschitz.agda | 673 ++++++++ Cubical/HITs/CauchyReals/Multiplication.agda | 709 +++++++++ Cubical/HITs/CauchyReals/Order.agda | 466 ++++++ Cubical/HITs/CauchyReals/Sequence.agda | 513 ++++++ Cubical/HITs/SetQuotients/Properties.agda | 120 ++ 22 files changed, 8537 insertions(+), 96 deletions(-) create mode 100644 Cubical/Data/Rationals/Order/Properties.agda create mode 100644 Cubical/HITs/CauchyReals/Base.agda create mode 100644 Cubical/HITs/CauchyReals/Closeness.agda create mode 100644 Cubical/HITs/CauchyReals/Continuous.agda create mode 100644 Cubical/HITs/CauchyReals/Derivative.agda create mode 100644 Cubical/HITs/CauchyReals/Inverse.agda create mode 100644 Cubical/HITs/CauchyReals/Lems.agda create mode 100644 Cubical/HITs/CauchyReals/Lipschitz.agda create mode 100644 Cubical/HITs/CauchyReals/Multiplication.agda create mode 100644 Cubical/HITs/CauchyReals/Order.agda create mode 100644 Cubical/HITs/CauchyReals/Sequence.agda diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index dda477bf43..87ea9a5cd5 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -16,7 +16,7 @@ open import Cubical.HITs.PropositionalTruncation renaming (rec to ∥∥rec) open import Cubical.Data.Fin.Base as Fin open import Cubical.Data.Nat -open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order hiding (<-·sk-cancel) open import Cubical.Data.Empty as Empty open import Cubical.Data.Unit open import Cubical.Data.Sum diff --git a/Cubical/Data/Int/Order.agda b/Cubical/Data/Int/Order.agda index fac79698da..98ddca2e15 100644 --- a/Cubical/Data/Int/Order.agda +++ b/Cubical/Data/Int/Order.agda @@ -9,6 +9,7 @@ open import Cubical.Data.Empty as ⊥ using (⊥) open import Cubical.Data.Int.Base as ℤ open import Cubical.Data.Int.Properties as ℤ open import Cubical.Data.Nat as ℕ +import Cubical.Data.Nat.Order as ℕ open import Cubical.Data.NatPlusOne.Base as ℕ₊₁ open import Cubical.Data.Sigma @@ -498,3 +499,45 @@ negsuc zero ≟ negsuc zero = eq refl negsuc zero ≟ negsuc (suc n) = gt (negsuc-≤-negsuc zero-≤pos) negsuc (suc m) ≟ negsuc zero = lt (negsuc-≤-negsuc zero-≤pos) negsuc (suc m) ≟ negsuc (suc n) = Trichotomy-pred (negsuc m ≟ negsuc n) + +0<_ : ℤ → Type +0< pos zero = ⊥ +0< pos (suc n) = Unit +0< negsuc n = ⊥ + +isProp0< : ∀ n → isProp (0< n) +isProp0< (pos (suc _)) _ _ = refl + +·0< : ∀ m n → 0< m → 0< n → 0< (m ℤ.· n) +·0< (pos (suc m)) (pos (suc n)) _ _ = + subst (0<_) (pos+ (suc n) (m ℕ.· (suc n)) ∙ cong (pos (suc n) ℤ.+_) (pos·pos m (suc n))) _ + +0<·ℕ₊₁ : ∀ m n → 0< (m ℤ.· pos (ℕ₊₁→ℕ n)) → 0< m +0<·ℕ₊₁ (pos (suc m)) n x = _ +0<·ℕ₊₁ (negsuc n₁) (1+ n) x = + ⊥.rec (subst 0<_ (negsuc·pos n₁ (suc n) + ∙ congS -_ (cong (pos (suc n) ℤ.+_) + (sym (pos·pos n₁ (suc n))) ∙ + sym (pos+ (suc n) (n₁ ℕ.· suc n)))) x) + ++0< : ∀ m n → 0< m → 0< n → 0< (m ℤ.+ n) ++0< (pos (suc m)) (pos (suc n)) _ _ = + subst (0<_) (cong sucℤ (pos+ (suc m) n)) _ + +0<→ℕ₊₁ : ∀ n → 0< n → Σ ℕ₊₁ λ m → n ≡ pos (ℕ₊₁→ℕ m) +0<→ℕ₊₁ (pos (suc n)) x = (1+ n) , refl + +min-0< : ∀ m n → 0< m → 0< n → 0< (ℤ.min m n) +min-0< (pos (suc zero)) (pos (suc n)) x x₁ = tt +min-0< (pos (suc (suc n₁))) (pos (suc zero)) x x₁ = tt +min-0< (pos (suc (suc n₁))) (pos (suc (suc n))) x x₁ = + +0< (sucℤ (ℤ.min (pos n₁) (pos n))) 1 (min-0< (pos (suc n₁)) (pos (suc n)) _ _) _ + +ℕ≤→pos-≤-pos : ∀ m n → m ℕ.≤ n → pos m ≤ pos n +ℕ≤→pos-≤-pos m n (k , p) = k , sym (pos+ m k) ∙∙ cong pos (ℕ.+-comm m k) ∙∙ cong pos p + + +0≤x² : ∀ n → 0 ≤ n ℤ.· n +0≤x² (pos n) = subst (0 ≤_) (pos·pos n n) zero-≤pos +0≤x² (negsuc n) = subst (0 ≤_) (pos·pos (suc n) (suc n) + ∙ sym (negsuc·negsuc n n)) zero-≤pos diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index 1fc854cafa..11b8c5ca18 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -11,7 +11,7 @@ open import Cubical.Relation.Nullary open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool -open import Cubical.Data.Nat +open import Cubical.Data.Nat as ℕ hiding (+-assoc ; +-comm ; min ; max ; minComm ; maxComm) renaming (_·_ to _·ℕ_; _+_ to _+ℕ_ ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm ; isEven to isEvenℕ ; isOdd to isOddℕ) @@ -1505,3 +1505,24 @@ sumFinℤ0 n = sumFinGen0 _+_ 0 (λ _ → refl) n (λ _ → 0) λ _ → refl sumFinℤHom : {n : ℕ} (f g : Fin n → ℤ) → sumFinℤ {n = n} (λ x → f x + g x) ≡ sumFinℤ {n = n} f + sumFinℤ {n = n} g sumFinℤHom {n = n} = sumFinGenHom _+_ 0 (λ _ → refl) +Comm +Assoc n + +abs-max : ∀ n → pos (abs n) ≡ max n (- n) +abs-max (pos zero) = refl +abs-max (pos (suc n)) = refl +abs-max (negsuc n) = refl + +min-pos-pos : ∀ m n → min (pos m) (pos n) ≡ pos (ℕ.min m n) +min-pos-pos zero n = refl +min-pos-pos (suc m) zero = refl +min-pos-pos (suc m) (suc n) = cong sucℤ (min-pos-pos m n) + +max-pos-pos : ∀ m n → max (pos m) (pos n) ≡ pos (ℕ.max m n) +max-pos-pos zero n = refl +max-pos-pos (suc m) zero = refl +max-pos-pos (suc m) (suc n) = cong sucℤ (max-pos-pos m n) + + +sign : ℤ → ℤ +sign (pos zero) = 0 +sign (pos (suc n)) = 1 +sign (negsuc n) = -1 diff --git a/Cubical/Data/Nat/Mod.agda b/Cubical/Data/Nat/Mod.agda index 8911504e25..a406893d45 100644 --- a/Cubical/Data/Nat/Mod.agda +++ b/Cubical/Data/Nat/Mod.agda @@ -4,7 +4,8 @@ module Cubical.Data.Nat.Mod where open import Cubical.Foundations.Prelude open import Cubical.Data.Nat open import Cubical.Data.Nat.Order -open import Cubical.Data.Empty +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ -- Defining x mod 0 to be 0. This way all the theorems below are true -- for n : ℕ instead of n : ℕ₊₁. @@ -36,6 +37,9 @@ mod< n = , cong (λ x → fst ind + suc x) (modIndStep n x) ∙ snd ind +<→mod : (n x : ℕ) → x < (suc n) → x mod (suc n) ≡ x +<→mod = modIndBase + mod-rUnit : (n x : ℕ) → x mod n ≡ ((x + n) mod n) mod-rUnit zero x = refl mod-rUnit (suc n) x = @@ -153,6 +157,7 @@ mod-lCancel n x y = ∙∙ mod-rCancel n y x ∙∙ cong (_mod n) (+-comm y (x mod n)) + -- remainder and quotient after division by n -- Again, allowing for 0-division to get nicer syntax remainder_/_ : (x n : ℕ) → ℕ @@ -192,3 +197,117 @@ private test₁ : ((11 + (10 mod 3)) mod 3) ≡ 0 test₁ = refl + + + +·mod : ∀ k n m → (k · n) mod (k · m) + ≡ k · (n mod m) +·mod zero n m = refl +·mod (suc k) n zero = cong ((n + k · n) mod_) (sym (0≡m·0 (suc k))) + ∙ 0≡m·0 (suc k) +·mod (suc k) n (suc m) = ·mod' n n ≤-refl (splitℕ-≤ (suc m) n) + + where + ·mod' : ∀ N n → n ≤ N → ((suc m) ≤ n) ⊎ (n < suc m) → + ((suc k · n) mod (suc k · suc m)) ≡ suc k · (n mod suc m) + ·mod' _ zero x _ = cong (modInd (m + k · suc m)) (sym (0≡m·0 (suc k))) + ∙ 0≡m·0 (suc k) + ·mod' zero (suc n) x _ = ⊥.rec (¬-<-zero x) + ·mod' (suc N) n@(suc n') x (inl (m' , p)) = + let z = ·mod' N m' (≤-trans (m + , +-comm m m' ∙ injSuc (sym (+-suc m' m) ∙ p)) + (pred-≤-pred x)) (splitℕ-≤ _ _) ∙ cong ((suc k) ·_) + (sym (modIndStep m m') ∙ + cong (_mod (suc m)) (+-comm (suc m) m' ∙ p)) + in (cong (λ y → ((suc k · y) mod (suc k · suc m))) (sym p) + ∙ cong {x = (m' + suc m + k · (m' + suc m))} + {suc (m + k · suc m + (m' + k · m'))} + (modInd (m + k · suc m)) + (cong (_+ k · (m' + suc m)) (+-suc m' m ∙ cong suc (+-comm m' m)) + ∙ cong suc + (sym (+-assoc m m' _) + ∙∙ cong (m +_) + (((cong (m' +_) (sym (·-distribˡ k _ _) + ∙ +-comm (k · m') _) ∙ +-assoc m' (k · suc m) (k · m')) + ∙ cong (_+ k · m') (+-comm m' _)) + ∙ sym (+-assoc (k · suc m) m' (k · m')) ) + ∙∙ +-assoc m _ _)) + ∙ modIndStep (m + k · suc m) (m' + k · m')) ∙ z + ·mod' (suc N) n x (inr x₁) = + modIndBase _ _ ( + subst2 _<_ (·-comm n (suc k)) (·-comm _ (suc k)) + (<-·sk {n} {suc m} {k = k} x₁) ) + ∙ cong ((suc k) ·_) (sym (modIndBase _ _ x₁)) + +2≤x→10 zero (suc l) r = ⊥.rec (¬-<-zero r) n∸l>0 (suc n) zero r = suc-≤-suc zero-≤ n∸l>0 (suc n) (suc l) r = n∸l>0 n l (pred-≤-pred r) +[n-m]+m : ∀ m n → m ≤ n → (n ∸ m) + m ≡ n +[n-m]+m zero n _ = +-zero n +[n-m]+m (suc m) zero p = ⊥.rec (¬-<-zero p) +[n-m]+m (suc m) (suc n) p = + +-suc _ _ ∙ cong suc ([n-m]+m m n (pred-≤-pred p)) + -- automation ≤-solver-type : (m n : ℕ) → Trichotomy m n → Type @@ -545,3 +567,25 @@ pattern s : ∀ x → ⟨ openPred (λ y → (y <ᵣ x) , squash₁) ⟩ +openPred> x y = + PT.map (map-snd (λ {q} q 0 x) + + + +·invℝ' : ∀ r 0 Date: Sun, 19 Jan 2025 01:55:27 +0100 Subject: [PATCH 02/11] fixed whitespace violations --- Cubical/Data/Int/Order.agda | 4 +- Cubical/Data/Int/Properties.agda | 2 +- Cubical/Data/Nat/Mod.agda | 14 +- Cubical/Data/Nat/Order.agda | 6 +- Cubical/Data/Rationals/Order.agda | 80 +++--- Cubical/Data/Rationals/Order/Properties.agda | 250 +++++++++---------- Cubical/Data/Rationals/Properties.agda | 10 +- Cubical/Foundations/HLevels.agda | 2 +- Cubical/HITs/CauchyReals/Base.agda | 52 ++-- Cubical/HITs/CauchyReals/Closeness.agda | 148 +++++------ Cubical/HITs/CauchyReals/Continuous.agda | 96 +++---- Cubical/HITs/CauchyReals/Derivative.agda | 15 +- Cubical/HITs/CauchyReals/Inverse.agda | 210 ++++++++-------- Cubical/HITs/CauchyReals/Lems.agda | 68 ++--- Cubical/HITs/CauchyReals/Lipschitz.agda | 100 ++++---- Cubical/HITs/CauchyReals/Multiplication.agda | 78 +++--- Cubical/HITs/CauchyReals/Order.agda | 80 +++--- Cubical/HITs/CauchyReals/Sequence.agda | 76 +++--- Cubical/HITs/SetQuotients/Properties.agda | 22 +- 19 files changed, 656 insertions(+), 657 deletions(-) diff --git a/Cubical/Data/Int/Order.agda b/Cubical/Data/Int/Order.agda index 98ddca2e15..e728308348 100644 --- a/Cubical/Data/Int/Order.agda +++ b/Cubical/Data/Int/Order.agda @@ -512,7 +512,7 @@ isProp0< (pos (suc _)) _ _ = refl ·0< (pos (suc m)) (pos (suc n)) _ _ = subst (0<_) (pos+ (suc n) (m ℕ.· (suc n)) ∙ cong (pos (suc n) ℤ.+_) (pos·pos m (suc n))) _ -0<·ℕ₊₁ : ∀ m n → 0< (m ℤ.· pos (ℕ₊₁→ℕ n)) → 0< m +0<·ℕ₊₁ : ∀ m n → 0< (m ℤ.· pos (ℕ₊₁→ℕ n)) → 0< m 0<·ℕ₊₁ (pos (suc m)) n x = _ 0<·ℕ₊₁ (negsuc n₁) (1+ n) x = ⊥.rec (subst 0<_ (negsuc·pos n₁ (suc n) @@ -533,7 +533,7 @@ min-0< (pos (suc (suc n₁))) (pos (suc zero)) x x₁ = tt min-0< (pos (suc (suc n₁))) (pos (suc (suc n))) x x₁ = +0< (sucℤ (ℤ.min (pos n₁) (pos n))) 1 (min-0< (pos (suc n₁)) (pos (suc n)) _ _) _ -ℕ≤→pos-≤-pos : ∀ m n → m ℕ.≤ n → pos m ≤ pos n +ℕ≤→pos-≤-pos : ∀ m n → m ℕ.≤ n → pos m ≤ pos n ℕ≤→pos-≤-pos m n (k , p) = k , sym (pos+ m k) ∙∙ cong pos (ℕ.+-comm m k) ∙∙ cong pos p diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index 11b8c5ca18..088182ed42 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -1506,7 +1506,7 @@ sumFinℤHom : {n : ℕ} (f g : Fin n → ℤ) → sumFinℤ {n = n} (λ x → f x + g x) ≡ sumFinℤ {n = n} f + sumFinℤ {n = n} g sumFinℤHom {n = n} = sumFinGenHom _+_ 0 (λ _ → refl) +Comm +Assoc n -abs-max : ∀ n → pos (abs n) ≡ max n (- n) +abs-max : ∀ n → pos (abs n) ≡ max n (- n) abs-max (pos zero) = refl abs-max (pos (suc n)) = refl abs-max (negsuc n) = refl diff --git a/Cubical/Data/Nat/Mod.agda b/Cubical/Data/Nat/Mod.agda index a406893d45..1d53e66990 100644 --- a/Cubical/Data/Nat/Mod.agda +++ b/Cubical/Data/Nat/Mod.agda @@ -37,7 +37,7 @@ mod< n = , cong (λ x → fst ind + suc x) (modIndStep n x) ∙ snd ind -<→mod : (n x : ℕ) → x < (suc n) → x mod (suc n) ≡ x +<→mod : (n x : ℕ) → x < (suc n) → x mod (suc n) ≡ x <→mod = modIndBase mod-rUnit : (n x : ℕ) → x mod n ≡ ((x + n) mod n) @@ -208,7 +208,7 @@ private ·mod (suc k) n (suc m) = ·mod' n n ≤-refl (splitℕ-≤ (suc m) n) where - ·mod' : ∀ N n → n ≤ N → ((suc m) ≤ n) ⊎ (n < suc m) → + ·mod' : ∀ N n → n ≤ N → ((suc m) ≤ n) ⊎ (n < suc m) → ((suc k · n) mod (suc k · suc m)) ≡ suc k · (n mod suc m) ·mod' _ zero x _ = cong (modInd (m + k · suc m)) (sym (0≡m·0 (suc k))) ∙ 0≡m·0 (suc k) @@ -236,10 +236,10 @@ private ·mod' (suc N) n x (inr x₁) = modIndBase _ _ ( subst2 _<_ (·-comm n (suc k)) (·-comm _ (suc k)) - (<-·sk {n} {suc m} {k = k} x₁) ) + (<-·sk {n} {suc m} {k = k} x₁) ) ∙ cong ((suc k) ·_) (sym (modIndBase _ _ x₁)) -2≤x→10 (suc n) (suc l) r = n∸l>0 n l (pred-≤-pred r) [n-m]+m : ∀ m n → m ≤ n → (n ∸ m) + m ≡ n [n-m]+m zero n _ = +-zero n [n-m]+m (suc m) zero p = ⊥.rec (¬-<-zero p) -[n-m]+m (suc m) (suc n) p = +[n-m]+m (suc m) (suc n) p = +-suc _ _ ∙ cong suc ([n-m]+m m n (pred-≤-pred p)) -- automation @@ -573,14 +573,14 @@ elimBy≤ : ∀ {ℓ} {A : ℕ → ℕ → Type ℓ} → (∀ x y → x ≤ y → A x y) → ∀ x y → A x y elimBy≤ {A = A} s f n m = ≤CaseInduction {P = A} - (f _ _) (s _ _ ∘ f _ _ ) + (f _ _) (s _ _ ∘ f _ _ ) elimBy≤+ : ∀ {ℓ} {A : ℕ → ℕ → Type ℓ} → (∀ x y → A x y → A y x) → (∀ x y → A x (y + x)) → ∀ x y → A x y elimBy≤+ {A = A} s f = - elimBy≤ s λ x y (y' , p) → subst (A x) p (f x y') + elimBy≤ s λ x y (y' , p) → subst (A x) p (f x y') module Minimal where Least : ∀{ℓ} → (ℕ → Type ℓ) → (ℕ → Type ℓ) diff --git a/Cubical/Data/Rationals/Order.agda b/Cubical/Data/Rationals/Order.agda index 761f31fd87..78828b1a2a 100644 --- a/Cubical/Data/Rationals/Order.agda +++ b/Cubical/Data/Rationals/Order.agda @@ -443,20 +443,20 @@ isTrans≤< = <≤Monotone+ : ∀ m n o s → m < n → o ≤ s → m ℚ.+ o < n ℚ.+ s <≤Monotone+ m n o s x x₁ = - isTrans<≤ (m ℚ.+ o) (n ℚ.+ o) (n ℚ.+ s) (<-+o m n o x) (≤-o+ o s n x₁) + isTrans<≤ (m ℚ.+ o) (n ℚ.+ o) (n ℚ.+ s) (<-+o m n o x) (≤-o+ o s n x₁) ≤ : ∀ x → ⟨ openPred (λ y → (y <ᵣ x) , squash₁) ⟩ +openPred> : ∀ x → ⟨ openPred (λ y → (y <ᵣ x) , squash₁) ⟩ openPred> x y = PT.map (map-snd (λ {q} q x y = ∘S x Date: Sun, 19 Jan 2025 03:53:55 +0100 Subject: [PATCH 03/11] wip --- Cubical/HITs/CauchyReals/Derivative.agda | 100 ++++++++++++++++--- Cubical/HITs/CauchyReals/Inverse.agda | 25 +++++ Cubical/HITs/CauchyReals/Multiplication.agda | 8 ++ Cubical/HITs/CauchyReals/Order.agda | 7 ++ Cubical/HITs/CauchyReals/Sequence.agda | 2 + 5 files changed, 128 insertions(+), 14 deletions(-) diff --git a/Cubical/HITs/CauchyReals/Derivative.agda b/Cubical/HITs/CauchyReals/Derivative.agda index 4082c43f78..44854dc7ef 100644 --- a/Cubical/HITs/CauchyReals/Derivative.agda +++ b/Cubical/HITs/CauchyReals/Derivative.agda @@ -105,26 +105,53 @@ const-lim C x ε = ∣ (1 , decℚ<ᵣ?) , id-lim : ∀ x → at x limitOf (λ r _ → r) is x id-lim x ε = ∣ ε , (λ r x#r p → p ) ∣₁ -_$[_]$_ : {x : ℝ} +_$[_]$_ : (ℝ → ℝ) + → (ℝ → ℝ → ℝ) + → (ℝ → ℝ) + → (ℝ → ℝ) +f $[ _op_ ]$ g = λ r → (f r) op (g r) + +_#[_]$_ : {x : ℝ} → (∀ r → x # r → ℝ) → (ℝ → ℝ → ℝ) → (∀ r → x # r → ℝ) → (∀ r → x # r → ℝ) -f $[ _op_ ]$ g = λ r x → (f r x) op (g r x) +f #[ _op_ ]$ g = λ r x → (f r x) op (g r x) +-lim : ∀ x f g F G → at x limitOf f is F → at x limitOf g is G - → at x limitOf (f $[ _+ᵣ_ ]$ g) is (F +ᵣ G) + → at x limitOf (f #[ _+ᵣ_ ]$ g) is (F +ᵣ G) +-lim x f g F G fL gL ε = PT.map2 (λ (δ , p) (δ' , p') → (_ , ℝ₊min δ δ') , λ r x#r x₁ → - let u = p r x#r (isTrans<≤ᵣ _ _ _ x₁ {!min≤ᵣ!}) - u' = p' r x#r (isTrans<≤ᵣ _ _ _ x₁ {!min≤ᵣ!}) - in {!u!}) - (fL (ε ₊/ᵣ₊ 2)) - (gL (ε ₊/ᵣ₊ 2)) + let u = p r x#r (isTrans<≤ᵣ _ _ _ x₁ (min≤ᵣ _ _)) + u' = p' r x#r (isTrans<≤ᵣ _ _ _ x₁ (min≤ᵣ' _ _)) + in subst2 _<ᵣ_ + (cong absᵣ (sym L𝐑.lem--041)) + (x·rat[α]+x·rat[β]=x (fst ε)) + (isTrans≤<ᵣ _ _ _ + (absᵣ-triangle _ _) + (<ᵣMonotone+ᵣ _ _ _ _ u u'))) + (fL (ε ₊·ᵣ (rat [ 1 / 2 ] , decℚ<ᵣ?))) + (gL (ε ₊·ᵣ (rat [ 1 / 2 ] , decℚ<ᵣ?))) + + +·-lim : ∀ x f g F G + → at x limitOf f is F + → at x limitOf g is G + → at x limitOf (f #[ _·ᵣ_ ]$ g) is (F ·ᵣ G) +·-lim x f g F G fL gL ε = + PT.map2 (λ (δ , p) (δ' , p') → + (_ , ℝ₊min δ δ') , + λ r x#r x₁ → + let u = p r x#r (isTrans<≤ᵣ _ _ _ x₁ (min≤ᵣ _ _)) + u' = p' r x#r (isTrans<≤ᵣ _ _ _ x₁ (min≤ᵣ' _ _)) + in {!!}) + (fL ε) + (gL ε) + At_limitOf_ : (x : ℝ) → (∀ r → x # r → ℝ) → Type At x limitOf f = Σ _ (at x limitOf f is_) @@ -155,10 +182,55 @@ idDerivative x = subst (at 0 limitOf_is 1) cong (_·ᵣ (invℝ r 0#r)) (sym (L𝐑.lem--063))) (const-lim 1 0) - - --- derivative-^ⁿ : ∀ n x → --- derivativeOf (_^ⁿ (suc n)) at x is (fromNat n ·ᵣ (x ^ⁿ n)) --- derivative-^ⁿ zero x ε = {!!} --- derivative-^ⁿ (suc n) x ε = {!!} ++Derivative : ∀ x f f'x g g'x + → derivativeOf f at x is f'x + → derivativeOf g at x is g'x + → derivativeOf (f $[ _+ᵣ_ ]$ g) at x is (f'x +ᵣ g'x) ++Derivative x f f'x g g'x F G = + subst {x = (differenceAt f x) #[ _+ᵣ_ ]$ (differenceAt g x)} + {y = (differenceAt (f $[ _+ᵣ_ ]$ g) x)} + (at 0 limitOf_is (f'x +ᵣ g'x)) + (funExt₂ λ h 0#h → + sym (·DistR+ _ _ _) ∙ + cong (_·ᵣ (invℝ h 0#h)) + (sym L𝐑.lem--041)) (+-lim _ _ _ _ _ F G) + +C·Derivative : ∀ C x f f'x + → derivativeOf f at x is f'x + → derivativeOf ((C ·ᵣ_) ∘S f) at x is (C ·ᵣ f'x) +C·Derivative C x f f'x F = + subst {x = λ h 0#h → C ·ᵣ differenceAt f x h 0#h} + {y = (differenceAt ((C ·ᵣ_) ∘S f) x)} + (at 0 limitOf_is (C ·ᵣ f'x)) + (funExt₂ λ h 0#h → + ·ᵣAssoc _ _ _ ∙ + cong (_·ᵣ (invℝ h 0#h)) (·DistL- _ _ _)) + (·-lim _ _ _ _ _ (const-lim C 0) F) + +substDer : ∀ {x f' f g} → (∀ r → f r ≡ g r) + → derivativeOf f at x is f' + → derivativeOf g at x is f' +substDer = {!!} + +·Derivative : ∀ x f f'x g g'x + → derivativeOf f at x is f'x + → derivativeOf g at x is g'x + → derivativeOf (f $[ _·ᵣ_ ]$ g) at x is + (f'x ·ᵣ (g x) +ᵣ (f x) ·ᵣ g'x) +·Derivative x f f'x g g'x F G = + let z = +Derivative x {!!} (f'x ·ᵣ (g x)) + {!!} ((f x) ·ᵣ g'x) + {!substDer ? + (C·Derivative (g x) x f f'x F)!} + (C·Derivative (f x) x g g'x G) + in {!!} + -- let z = {!subst {x = (differenceAt f x) #[ _+ᵣ_ ]$ (differenceAt g x)} + -- {y = (differenceAt (f $[ _+ᵣ_ ]$ g) x)} + -- (at 0 limitOf_is (f'x +ᵣ g'x))!} + -- in {!!} + +-- -- derivative-^ⁿ : ∀ n x → +-- -- derivativeOf (_^ⁿ (suc n)) at x is (fromNat n ·ᵣ (x ^ⁿ n)) +-- -- derivative-^ⁿ zero x ε = {!!} +-- -- derivative-^ⁿ (suc n) x ε = {!!} diff --git a/Cubical/HITs/CauchyReals/Inverse.agda b/Cubical/HITs/CauchyReals/Inverse.agda index eadb4162a8..770f9b103b 100644 --- a/Cubical/HITs/CauchyReals/Inverse.agda +++ b/Cubical/HITs/CauchyReals/Inverse.agda @@ -642,6 +642,23 @@ minDistMaxᵣ x y y' = ≡Continuous _ _ ℝ₊min : (m : ℝ₊) (n : ℝ₊) → 0 <ᵣ minᵣ (fst m) (fst n) ℝ₊min (m , Date: Mon, 20 Jan 2025 17:36:29 +0100 Subject: [PATCH 04/11] derivative --- Cubical/Data/Rationals/Order/Properties.agda | 7 +- Cubical/HITs/CauchyReals/Derivative.agda | 212 ++++++++++++++---- Cubical/HITs/CauchyReals/Inverse.agda | 155 +++++++------ Cubical/HITs/CauchyReals/Lems.agda | 9 +- Cubical/HITs/CauchyReals/Multiplication.agda | 5 +- Cubical/HITs/CauchyReals/Order.agda | 13 ++ Cubical/HITs/CauchyReals/Sequence.agda | 127 ++++++++--- .../PropositionalTruncation/Properties.agda | 5 +- 8 files changed, 386 insertions(+), 147 deletions(-) diff --git a/Cubical/Data/Rationals/Order/Properties.agda b/Cubical/Data/Rationals/Order/Properties.agda index ae3a1301f5..a899394064 100644 --- a/Cubical/Data/Rationals/Order/Properties.agda +++ b/Cubical/Data/Rationals/Order/Properties.agda @@ -1235,7 +1235,12 @@ getPosRatio L₁ L₂ = (cong ℕ₊₁→ℤ (·₊₁-assoc a' b' c')) - +·MaxDistrℚ' : ∀ x y z → 0 ≤ z → (max x y) · z ≡ max (x · z) (y · z) +·MaxDistrℚ' x y z = + ⊎.rec (λ p → cong ((max x y) ·_) (sym p) ∙ + ·AnnihilR (max x y) ∙ cong₂ max (sym (·AnnihilR x) ∙ cong (x ·_) p) + (sym (·AnnihilR y) ∙ cong (y ·_) p)) + (·MaxDistrℚ x y z ∘ <→0< z) ∘ (≤→≡⊎< 0 z) ≤Monotone·-onNonNeg : ∀ x x' y y' → x ≤ x' → diff --git a/Cubical/HITs/CauchyReals/Derivative.agda b/Cubical/HITs/CauchyReals/Derivative.agda index 44854dc7ef..d7d6b920d9 100644 --- a/Cubical/HITs/CauchyReals/Derivative.agda +++ b/Cubical/HITs/CauchyReals/Derivative.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --lossy-unification #-} module Cubical.HITs.CauchyReals.Derivative where @@ -84,19 +84,36 @@ import Cubical.Algebra.Ring as RP --- Rℝ = (CR.CommRing→Ring --- (_ , CR.commringstr 0 1 _+ᵣ_ _·ᵣ_ -ᵣ_ IsCommRingℝ)) --- -- module CRℝ = ? - --- module 𝐑 = CR.CommRingTheory (_ , CR.commringstr 0 1 _+ᵣ_ _·ᵣ_ -ᵣ_ IsCommRingℝ) --- module 𝐑' = RP.RingTheory Rℝ - - at_limitOf_is_ : (x : ℝ) → (∀ r → x # r → ℝ) → ℝ → Type at x limitOf f is L = ∀ (ε : ℝ₊) → ∃[ δ ∈ ℝ₊ ] (∀ r x#r → absᵣ (x -ᵣ r) <ᵣ fst δ → absᵣ (L -ᵣ f r x#r) <ᵣ fst ε) +substLim : ∀ {x f f' L} + → (∀ r x#r → f r x#r ≡ f' r x#r) + → at x limitOf f is L → at x limitOf f' is L +substLim {x} {L = L} p = subst (at x limitOf_is L) (funExt₂ p) + +IsContinuousLim : ∀ f x → IsContinuous f → + at x limitOf (λ r _ → f r) is (f x) +IsContinuousLim f x cx = uncurry + λ ε → (PT.rec squash₁ + λ (q , 0 Date: Mon, 20 Jan 2025 17:36:51 +0100 Subject: [PATCH 05/11] derivative --- Cubical/HITs/CauchyReals/Derivative.agda | 38 ++++++++++---------- Cubical/HITs/CauchyReals/Inverse.agda | 6 ++-- Cubical/HITs/CauchyReals/Lems.agda | 6 ++-- Cubical/HITs/CauchyReals/Multiplication.agda | 2 +- Cubical/HITs/CauchyReals/Sequence.agda | 14 ++++---- 5 files changed, 33 insertions(+), 33 deletions(-) diff --git a/Cubical/HITs/CauchyReals/Derivative.agda b/Cubical/HITs/CauchyReals/Derivative.agda index d7d6b920d9..7bec390bcb 100644 --- a/Cubical/HITs/CauchyReals/Derivative.agda +++ b/Cubical/HITs/CauchyReals/Derivative.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --lossy-unification #-} module Cubical.HITs.CauchyReals.Derivative where @@ -91,7 +91,7 @@ at x limitOf f is L = substLim : ∀ {x f f' L} → (∀ r x#r → f r x#r ≡ f' r x#r) - → at x limitOf f is L → at x limitOf f' is L + → at x limitOf f is L → at x limitOf f' is L substLim {x} {L = L} p = subst (at x limitOf_is L) (funExt₂ p) IsContinuousLim : ∀ f x → IsContinuous f → @@ -104,13 +104,13 @@ IsContinuousLim f x cx = uncurry λ r x#r x₁ → isTrans<ᵣ _ _ _ (fst (∼≃abs<ε _ _ _) (X r (invEq (∼≃abs<ε _ _ _) x₁))) q<ε ) - (cx x (q , ℚ.<→0< q (<ᵣ→<ℚ 0 q 0 Date: Wed, 22 Jan 2025 12:03:12 +0100 Subject: [PATCH 06/11] sequences --- Cubical/Data/Rationals/Order/Properties.agda | 8 + Cubical/HITs/CauchyReals/Base.agda | 36 --- Cubical/HITs/CauchyReals/Closeness.agda | 31 +-- Cubical/HITs/CauchyReals/Continuous.agda | 79 +------ Cubical/HITs/CauchyReals/Derivative.agda | 163 +++++++------ Cubical/HITs/CauchyReals/Inverse.agda | 65 +---- Cubical/HITs/CauchyReals/Lipschitz.agda | 71 ++---- Cubical/HITs/CauchyReals/Multiplication.agda | 48 +--- Cubical/HITs/CauchyReals/Order.agda | 33 +-- Cubical/HITs/CauchyReals/Sequence.agda | 235 +++++++++++++++---- 10 files changed, 344 insertions(+), 425 deletions(-) diff --git a/Cubical/Data/Rationals/Order/Properties.agda b/Cubical/Data/Rationals/Order/Properties.agda index a899394064..eae8aa3ff8 100644 --- a/Cubical/Data/Rationals/Order/Properties.agda +++ b/Cubical/Data/Rationals/Order/Properties.agda @@ -204,6 +204,14 @@ floor-frac x with 0 ≟ x p , e' +ceilℚ₊ : (q : ℚ₊) → Σ[ k ∈ ℕ ] (fst q) < fromNat k +ceilℚ₊ q = suc (fst (fst (floor-fracℚ₊ q))) , + subst2 (_<_) -- (fromNat (suc (fst (fst (floor-fracℚ₊ q))))) + (ℚ.+Comm _ _ ∙ fst (snd (floor-fracℚ₊ q))) + (ℚ.ℕ+→ℚ+ _ _) + (<-+o _ _ (fromNat (fst (fst (floor-fracℚ₊ q)))) + (snd (snd (snd (floor-fracℚ₊ q))))) + sign : ℚ → ℚ diff --git a/Cubical/HITs/CauchyReals/Base.agda b/Cubical/HITs/CauchyReals/Base.agda index 93b5ecbf77..387d29eaf2 100644 --- a/Cubical/HITs/CauchyReals/Base.agda +++ b/Cubical/HITs/CauchyReals/Base.agda @@ -2,50 +2,14 @@ module Cubical.HITs.CauchyReals.Base where open import Cubical.Foundations.Prelude hiding (Path) -open import Cubical.Foundations.Structure open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv hiding (_■) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Functions.FunExtEquiv -import Cubical.Functions.Logic as L - - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Bool.Base using () renaming (Bool to 𝟚 ; true to 1̂ ; false to 0̂) -open import Cubical.Data.Nat as ℕ hiding (_·_;_+_) -open import Cubical.Data.Nat.Order.Recursive as OR -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.Unit open import Cubical.Data.Int as ℤ -import Cubical.Data.Int.Order as ℤ -open import Cubical.Data.Maybe as Mb -open import Cubical.Data.Sigma hiding (Path) -open import Cubical.Data.List as L -open import Cubical.Data.List using () renaming (List to ⟦_⟧) -open import Cubical.Foundations.Interpolate -open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary - -open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _//_) - open import Cubical.Data.Rationals as ℚ open import Cubical.Data.Rationals.Order as ℚ open import Cubical.Data.NatPlusOne -open import Cubical.Foundations.Path -open import Cubical.Foundations.CartesianKanOps - -open import Cubical.HITs.CauchyReals.Lems - -import Agda.Builtin.Reflection as R -open import Cubical.Reflection.Base - private variable diff --git a/Cubical/HITs/CauchyReals/Closeness.agda b/Cubical/HITs/CauchyReals/Closeness.agda index c8f2c8898e..979c90f6e1 100644 --- a/Cubical/HITs/CauchyReals/Closeness.agda +++ b/Cubical/HITs/CauchyReals/Closeness.agda @@ -5,46 +5,17 @@ open import Cubical.Foundations.Prelude hiding (Path) open import Cubical.Foundations.Structure open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv hiding (_■) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence open import Cubical.Functions.FunExtEquiv import Cubical.Functions.Logic as L - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Bool.Base using () renaming (Bool to 𝟚 ; true to 1̂ ; false to 0̂) -open import Cubical.Data.Nat as ℕ hiding (_·_;_+_) -open import Cubical.Data.Nat.Order.Recursive as OR -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.Unit -open import Cubical.Data.Int as ℤ -import Cubical.Data.Int.Order as ℤ -open import Cubical.Data.Maybe as Mb open import Cubical.Data.Sigma hiding (Path) -open import Cubical.Data.List as L -open import Cubical.Data.List using () renaming (List to ⟦_⟧) -open import Cubical.Foundations.Interpolate -open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary - -open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _//_) - open import Cubical.Data.Rationals as ℚ open import Cubical.Data.Rationals.Order as ℚ open import Cubical.Data.Rationals.Order.Properties as ℚ -open import Cubical.Data.NatPlusOne -open import Cubical.Foundations.Path -open import Cubical.Foundations.CartesianKanOps - - -import Agda.Builtin.Reflection as R -open import Cubical.Reflection.Base +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.CauchyReals.Base open import Cubical.HITs.CauchyReals.Lems diff --git a/Cubical/HITs/CauchyReals/Continuous.agda b/Cubical/HITs/CauchyReals/Continuous.agda index d9dba5df19..ed2a299cb6 100644 --- a/Cubical/HITs/CauchyReals/Continuous.agda +++ b/Cubical/HITs/CauchyReals/Continuous.agda @@ -5,56 +5,26 @@ module Cubical.HITs.CauchyReals.Continuous where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv hiding (_■) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Functions.FunExtEquiv - -import Cubical.Functions.Logic as L - -open import Cubical.Algebra.CommRing.Instances.Int open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Bool.Base using () renaming (Bool to 𝟚 ; true to 1̂ ; false to 0̂) open import Cubical.Data.Nat as ℕ hiding (_·_;_+_) -open import Cubical.Data.Nat.Order.Recursive as OR -import Cubical.Data.Nat.Order as ℕ -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.Unit open import Cubical.Data.Int as ℤ using (pos) import Cubical.Data.Int.Order as ℤ -open import Cubical.Data.Maybe as Mb -open import Cubical.Data.Sigma hiding (Path) -open import Cubical.Data.List as L -open import Cubical.Data.List using () renaming (List to ⟦_⟧) -open import Cubical.Foundations.Interpolate -open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary +open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _//_) -open import Cubical.Data.Rationals using (ℚ ; [_/_]) -open import Cubical.Data.Rationals.Order using - ( _ℚ₊+_ ; 0<_ ; ℚ₊ ; _ℚ₊·_ ; ℚ₊≡) -import Cubical.Data.Rationals as ℚ -import Cubical.Data.Rationals.Order as ℚ -open import Cubical.Data.Rationals.Order.Properties as ℚ using (invℚ₊) +open import Cubical.Data.Rationals as ℚ using (ℚ ; [_/_]) +open import Cubical.Data.Rationals.Order as ℚ using + ( _ℚ₊+_ ; 0<_ ; ℚ₊ ; _ℚ₊·_ ; ℚ₊≡) +open import Cubical.Data.Rationals.Order.Properties as ℚ + using (invℚ₊;/2₊;/3₊;/4₊) open import Cubical.Data.NatPlusOne -open import Cubical.Foundations.Path -open import Cubical.Foundations.CartesianKanOps - - -import Agda.Builtin.Reflection as R -open import Cubical.Reflection.Base - - -import Cubical.Algebra.CommRing as CR open import Cubical.HITs.CauchyReals.Base open import Cubical.HITs.CauchyReals.Lems @@ -76,7 +46,7 @@ open import Cubical.HITs.CauchyReals.Order zz = sym (congLim' _ _ _ λ q → - sym $ x₁ (ℚ.invℚ₊ (_//_.[ pos 1 , (1+ 0) ] , tt) ℚ₊· q)) + sym $ x₁ (ℚ.invℚ₊ ([ pos 1 / (1+ 0) ] , tt) ℚ₊· q)) in _∙_ {x = maxᵣ 0 (lim (λ q → (absᵣ (x (ℚ.invℚ₊ 1 ℚ₊· q)))) y0)} zz0 zz @@ -321,31 +291,6 @@ denseℚinℝ u v = --- x+[y-x]n/k≡y-[y-x]m/k : ∀ x y n m k → (n ℕ.+ m ≡ ℕ₊₁→ℕ k) → --- x ℚ.+ (y ℚ.- x) ℚ.· [ pos n / k ] --- ≡ y ℚ.- ((y ℚ.- x) ℚ.· [ pos m / k ]) --- x+[y-x]n/k≡y-[y-x]m/k u v n m k p = --- ((cong (u ℚ.+_) (ℚ.·DistR+ _ _ _ ∙ ℚ.+Comm _ _) --- ∙ ℚ.+Assoc _ _ _) ∙∙ --- cong₂ ℚ._+_ --- {!!} --- {!!} --- -- distℚ! u ·[ --- -- (ge1 +ge ((neg-ge ge1) ·ge --- -- ge[ [ pos 1 / 1+ 1 ] ])) --- -- ≡ (neg-ge ((neg-ge ge1) ·ge --- -- ge[ [ pos 1 / 1+ 1 ] ])) ] --- -- distℚ! v ·[ (( ge[ [ pos 1 / 1+ 1 ] ])) --- -- ≡ (ge1 +ge neg-ge ( --- -- ge[ [ pos 1 / 1+ 1 ] ])) ] --- ∙∙ (ℚ.+Comm _ _ ∙ sym (ℚ.+Assoc v --- (ℚ.- (v ℚ.· [ pos m / k ])) --- (ℚ.- ((ℚ.- u) ℚ.· [ pos m / k ]))) --- ∙ cong (ℚ._+_ v) --- (sym (ℚ.·DistL+ -1 _ _)) ∙ cong (ℚ._-_ v) --- (sym (ℚ.·DistR+ v (ℚ.- u) [ pos m / k ])) )) - - x+[y-x]/2≡y-[y-x]/2 : ∀ x y → x ℚ.+ (y ℚ.- x) ℚ.· [ 1 / 2 ] ≡ y ℚ.- ((y ℚ.- x) ℚ.· [ 1 / 2 ]) @@ -617,14 +562,14 @@ restrSq n q r x x₁ ε x₂ = zz x₂ (ℚ.0≤abs (q ℚ.+ r)) ((ℚ.0≤abs (q ℚ.- r))) -0<ℕ₊₁ : ∀ n m → 0 ℚ.< (_//_.[ ℚ.ℕ₊₁→ℤ n , m ]) -0<ℕ₊₁ n m = ℚ.0<→< (_//_.[ ℚ.ℕ₊₁→ℤ n , m ]) tt +0<ℕ₊₁ : ∀ n m → 0 ℚ.< ([ ℚ.ℕ₊₁→ℤ n / m ]) +0<ℕ₊₁ n m = ℚ.0<→< ([ ℚ.ℕ₊₁→ℤ n / m ]) tt -1/n Date: Wed, 22 Jan 2025 17:28:32 +0100 Subject: [PATCH 07/11] exp --- Cubical/Data/Rationals/Order/Properties.agda | 19 ++ Cubical/HITs/CauchyReals/Sequence.agda | 331 ++++++++++++++++--- 2 files changed, 309 insertions(+), 41 deletions(-) diff --git a/Cubical/Data/Rationals/Order/Properties.agda b/Cubical/Data/Rationals/Order/Properties.agda index eae8aa3ff8..71c6fee18c 100644 --- a/Cubical/Data/Rationals/Order/Properties.agda +++ b/Cubical/Data/Rationals/Order/Properties.agda @@ -1289,6 +1289,9 @@ invℚ₊≡invℚ q p = cong₂ _·_ (fst (<→sign (fst q)) (0<ℚ₊ q) ) (cong (fst ∘ invℚ₊) (ℚ₊≡ (sym (abs'≡abs (fst q)) ∙ absPos (fst q) ((0<ℚ₊ q))))) ∙ ·IdL (fst (invℚ₊ q)) +fromNat-invℚ : ∀ n p → invℚ [ pos (suc n) / 1 ] p ≡ [ 1 / 1+ n ] +fromNat-invℚ n p = invℚ₊≡invℚ _ _ + invℚ-pos : ∀ x y → 0 < x → 0 < invℚ x y invℚ-pos x y z = @@ -1358,12 +1361,28 @@ q /ℚ[ r , 0#r ] = q · (invℚ r 0#r) ℚ-[x·y]/y x r 0#r = sym (·Assoc _ _ _) ∙∙ cong (x ·_) (ℚ-y/y r 0#r) ∙∙ ·IdR x +ℚ-[x/y]·y : ∀ x r → (0#r : 0 # r) → ((x /ℚ[ r , 0#r ]) · r) ≡ x +ℚ-[x/y]·y x r 0#r = sym (·Assoc _ _ _) ∙∙ + cong (x ·_) (·Comm _ _ ∙ ℚ-y/y r 0#r) ∙∙ ·IdR x + + ℚ-x·y≡z→x≡z/y : ∀ x q r → (0#r : 0 # r) → (x · r) ≡ q → x ≡ q /ℚ[ r , 0#r ] ℚ-x·y≡z→x≡z/y x q r 0#r p = sym (ℚ-[x·y]/y x r 0#r ) ∙ cong (_/ℚ[ r , 0#r ]) p +ℚ-x/y Date: Wed, 22 Jan 2025 19:59:27 +0100 Subject: [PATCH 08/11] whitespace fix --- Cubical/Data/Rationals/Order/Properties.agda | 4 +- Cubical/HITs/CauchyReals/Derivative.agda | 2 +- Cubical/HITs/CauchyReals/Sequence.agda | 248 ++++++++++--------- 3 files changed, 139 insertions(+), 115 deletions(-) diff --git a/Cubical/Data/Rationals/Order/Properties.agda b/Cubical/Data/Rationals/Order/Properties.agda index 71c6fee18c..a2a610e84d 100644 --- a/Cubical/Data/Rationals/Order/Properties.agda +++ b/Cubical/Data/Rationals/Order/Properties.agda @@ -1374,7 +1374,7 @@ q /ℚ[ r , 0#r ] = q · (invℚ r 0#r) ℚ-x/y Date: Thu, 23 Jan 2025 00:22:45 +0100 Subject: [PATCH 09/11] whitespace fix --- Cubical/HITs/CauchyReals.agda | 12 +++++++++++ Cubical/HITs/CauchyReals/Continuous.agda | 2 +- Cubical/HITs/CauchyReals/Derivative.agda | 2 +- Cubical/HITs/CauchyReals/Inverse.agda | 5 +++-- Cubical/HITs/CauchyReals/Lipschitz.agda | 9 +++++++- Cubical/HITs/CauchyReals/Multiplication.agda | 19 ++++++++++++++++- Cubical/HITs/CauchyReals/Order.agda | 2 +- Cubical/HITs/CauchyReals/Sequence.agda | 22 +++++--------------- 8 files changed, 49 insertions(+), 24 deletions(-) create mode 100644 Cubical/HITs/CauchyReals.agda diff --git a/Cubical/HITs/CauchyReals.agda b/Cubical/HITs/CauchyReals.agda new file mode 100644 index 0000000000..200d8e2dd7 --- /dev/null +++ b/Cubical/HITs/CauchyReals.agda @@ -0,0 +1,12 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.CauchyReals where + +open import Cubical.HITs.CauchyReals.Base public +open import Cubical.HITs.CauchyReals.Closeness public +open import Cubical.HITs.CauchyReals.Lipschitz public +open import Cubical.HITs.CauchyReals.Order public +open import Cubical.HITs.CauchyReals.Continuous public +open import Cubical.HITs.CauchyReals.Multiplication public +open import Cubical.HITs.CauchyReals.Inverse public +open import Cubical.HITs.CauchyReals.Sequence public +open import Cubical.HITs.CauchyReals.Derivative public diff --git a/Cubical/HITs/CauchyReals/Continuous.agda b/Cubical/HITs/CauchyReals/Continuous.agda index ed2a299cb6..2f5ab88109 100644 --- a/Cubical/HITs/CauchyReals/Continuous.agda +++ b/Cubical/HITs/CauchyReals/Continuous.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.CauchyReals.Continuous where diff --git a/Cubical/HITs/CauchyReals/Derivative.agda b/Cubical/HITs/CauchyReals/Derivative.agda index e1df521a55..b628d33591 100644 --- a/Cubical/HITs/CauchyReals/Derivative.agda +++ b/Cubical/HITs/CauchyReals/Derivative.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.CauchyReals.Derivative where diff --git a/Cubical/HITs/CauchyReals/Inverse.agda b/Cubical/HITs/CauchyReals/Inverse.agda index fa8d23a00a..bcdd9b9648 100644 --- a/Cubical/HITs/CauchyReals/Inverse.agda +++ b/Cubical/HITs/CauchyReals/Inverse.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --lossy-unification --safe #-} module Cubical.HITs.CauchyReals.Inverse where @@ -43,7 +43,7 @@ open import Cubical.HITs.CauchyReals.Multiplication Rℝ = (CR.CommRing→Ring (_ , CR.commringstr 0 1 _+ᵣ_ _·ᵣ_ -ᵣ_ IsCommRingℝ)) - +-- module CRℝ = ? module 𝐑 = CR.CommRingTheory (_ , CR.commringstr 0 1 _+ᵣ_ _·ᵣ_ -ᵣ_ IsCommRingℝ) module 𝐑' = RP.RingTheory Rℝ @@ -1020,6 +1020,7 @@ sign·absᵣ r = ∘diag $ ∙ cong rat (cong (ℚ._· ℚ.sign r) (sym (ℚ.abs'≡abs r)) ∙ ℚ.sign·abs r) ) r +-- HoTT Theorem (11.3.47) abstract invℝ : ∀ r → 0 # r → ℝ diff --git a/Cubical/HITs/CauchyReals/Lipschitz.agda b/Cubical/HITs/CauchyReals/Lipschitz.agda index da6a06b6d6..45592123f6 100644 --- a/Cubical/HITs/CauchyReals/Lipschitz.agda +++ b/Cubical/HITs/CauchyReals/Lipschitz.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.CauchyReals.Lipschitz where @@ -98,6 +98,13 @@ lim-surj = PT.map (map-snd (eqℝ _ _)) ∘ (Elimℝ-Prop.go w) -- TODO : (Lemma 11.3.11) +-- HoTT-11-3-11 : ∀ {ℓ} (A : Type ℓ) (isSetA : isSet A) → +-- (f : (Σ[ x ∈ (ℚ₊ → ℝ) ] (∀ ( δ ε : ℚ₊) → x δ ∼[ δ ℚ₊+ ε ] x ε)) +-- → A) → +-- (∀ u v → uncurry lim u ≡ uncurry lim v → f u ≡ f v) +-- → ∃![ g ∈ (ℝ → A) ] f ≡ g ∘ uncurry lim +-- HoTT-11-3-11 A isSetA f p = +-- Lipschitz-ℚ→ℚ : ℚ₊ → (ℚ → ℚ) → Type Lipschitz-ℚ→ℚ L f = diff --git a/Cubical/HITs/CauchyReals/Multiplication.agda b/Cubical/HITs/CauchyReals/Multiplication.agda index 7bc59279c0..d11378d2f6 100644 --- a/Cubical/HITs/CauchyReals/Multiplication.agda +++ b/Cubical/HITs/CauchyReals/Multiplication.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.CauchyReals.Multiplication where @@ -686,3 +686,20 @@ cont₂·ᵣWP P f g fC gC = IsContinuousWP∘' P _ _^ⁿ_ : ℝ → ℕ → ℝ x ^ⁿ zero = 1 x ^ⁿ suc n = (x ^ⁿ n) ·ᵣ x + + +·absᵣ : ∀ x y → absᵣ (x ·ᵣ y) ≡ absᵣ x ·ᵣ absᵣ y +·absᵣ x = ≡Continuous _ _ + ((IsContinuous∘ _ _ IsContinuousAbsᵣ (IsContinuous·ᵣL x) + )) + (IsContinuous∘ _ _ (IsContinuous·ᵣL (absᵣ x)) + IsContinuousAbsᵣ) + λ y' → + ≡Continuous _ _ + ((IsContinuous∘ _ _ IsContinuousAbsᵣ (IsContinuous·ᵣR (rat y')) + )) + (IsContinuous∘ _ _ (IsContinuous·ᵣR (absᵣ (rat y'))) + IsContinuousAbsᵣ) + (λ x' → + cong absᵣ (sym (rat·ᵣrat _ _)) ∙∙ + cong rat (sym (ℚ.abs'·abs' _ _)) ∙∙ rat·ᵣrat _ _) x diff --git a/Cubical/HITs/CauchyReals/Order.agda b/Cubical/HITs/CauchyReals/Order.agda index d41ae8f99d..bd88e12a10 100644 --- a/Cubical/HITs/CauchyReals/Order.agda +++ b/Cubical/HITs/CauchyReals/Order.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.CauchyReals.Order where diff --git a/Cubical/HITs/CauchyReals/Sequence.agda b/Cubical/HITs/CauchyReals/Sequence.agda index fae5d31605..fd7054d972 100644 --- a/Cubical/HITs/CauchyReals/Sequence.agda +++ b/Cubical/HITs/CauchyReals/Sequence.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.CauchyReals.Sequence where @@ -42,22 +42,6 @@ open import Cubical.HITs.CauchyReals.Inverse -·absᵣ : ∀ x y → absᵣ (x ·ᵣ y) ≡ absᵣ x ·ᵣ absᵣ y -·absᵣ x = ≡Continuous _ _ - ((IsContinuous∘ _ _ IsContinuousAbsᵣ (IsContinuous·ᵣL x) - )) - (IsContinuous∘ _ _ (IsContinuous·ᵣL (absᵣ x)) - IsContinuousAbsᵣ) - λ y' → - ≡Continuous _ _ - ((IsContinuous∘ _ _ IsContinuousAbsᵣ (IsContinuous·ᵣR (rat y')) - )) - (IsContinuous∘ _ _ (IsContinuous·ᵣR (absᵣ (rat y'))) - IsContinuousAbsᵣ) - (λ x' → - cong absᵣ (sym (rat·ᵣrat _ _)) ∙∙ - cong rat (sym (ℚ.abs'·abs' _ _)) ∙∙ rat·ᵣrat _ _) x - Seq : Type Seq = ℕ → ℝ @@ -554,6 +538,7 @@ isCauchyAprox f = (δ ε : ℚ₊) → lim' : ∀ x → isCauchyAprox x → ℝ lim' x y = lim x λ δ ε → (invEq (∼≃abs<ε _ _ _)) (y δ ε) +-- HoTT 11.3.49 fromCauchySequence : ∀ s → IsCauchySequence s → ℝ fromCauchySequence s ics = @@ -586,6 +571,9 @@ fromCauchySequence s ics = (subst ((fst ε) ℚ.<_) (ℚ.+Comm _ _) (ℚ.<+ℚ₊' (fst ε) (fst ε) δ (ℚ.isRefl≤ (fst ε)))))) + +-- TODO HoTT 11.3.50. + fromCauchySequence' : ∀ s → IsCauchySequence' s → ℝ fromCauchySequence' s ics = lim x y From 1143bc9cf9e83727cc7cca78c27574280cf75762 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Thu, 23 Jan 2025 00:40:22 +0100 Subject: [PATCH 10/11] uncomment some working code --- Cubical/HITs/CauchyReals/Derivative.agda | 36 ++++++++++++------------ 1 file changed, 18 insertions(+), 18 deletions(-) diff --git a/Cubical/HITs/CauchyReals/Derivative.agda b/Cubical/HITs/CauchyReals/Derivative.agda index b628d33591..59778e28ba 100644 --- a/Cubical/HITs/CauchyReals/Derivative.agda +++ b/Cubical/HITs/CauchyReals/Derivative.agda @@ -14,7 +14,7 @@ open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Data.NatPlusOne - +open import Cubical.Data.Nat as ℕ hiding (_·_;_+_) open import Cubical.Data.Rationals as ℚ using (ℚ ; [_/_]) open import Cubical.Data.Rationals.Order as ℚ using @@ -354,23 +354,23 @@ C·Derivative' C x f f'x F = -- → IsContinuous f -- hasDer→isCont f f' df ε = {!df!} --- -- derivative-^ⁿ : ∀ n x → --- -- derivativeOf (_^ⁿ (suc n)) at x --- -- is (fromNat (suc n) ·ᵣ (x ^ⁿ n)) --- -- derivative-^ⁿ zero x = --- -- substDer₂ --- -- (λ _ → sym (·IdL _)) --- -- (sym (·IdL _)) --- -- (idDerivative x) --- -- derivative-^ⁿ (suc n) x = --- -- substDer₂ (λ _ → refl) --- -- (+ᵣComm _ _ ∙ cong₂ _+ᵣ_ --- -- (·ᵣComm _ _) (sym (·ᵣAssoc _ _ _)) ∙ --- -- sym (·DistR+ _ _ _) ∙ --- -- cong (_·ᵣ ((x ^ⁿ n) ·ᵣ idfun ℝ x)) --- -- (cong rat (ℚ.ℕ+→ℚ+ _ _))) --- -- (·Derivative _ _ _ _ _ IsContinuousId --- -- (derivative-^ⁿ n x) (idDerivative x)) +derivative-^ⁿ : ∀ n x → + derivativeOf (_^ⁿ (suc n)) at x + is (fromNat (suc n) ·ᵣ (x ^ⁿ n)) +derivative-^ⁿ zero x = + substDer₂ + (λ _ → sym (·IdL _)) + (sym (·IdL _)) + (idDerivative x) +derivative-^ⁿ (suc n) x = + substDer₂ (λ _ → refl) + (+ᵣComm _ _ ∙ cong₂ _+ᵣ_ + (·ᵣComm _ _) (sym (·ᵣAssoc _ _ _)) ∙ + sym (·DistR+ _ _ _) ∙ + cong (_·ᵣ ((x ^ⁿ n) ·ᵣ idfun ℝ x)) + (cong rat (ℚ.ℕ+→ℚ+ _ _))) + (·Derivative _ _ _ _ _ IsContinuousId + (derivative-^ⁿ n x) (idDerivative x)) -- -- -- chainRule : ∀ x f f'gx g g'x -- -- -- → derivativeOf g at x is g'x From 18f833478487b581732fc2884f102453e2e3fd62 Mon Sep 17 00:00:00 2001 From: Marcin Jan Turek-Grzybowski Date: Fri, 27 Jun 2025 00:24:12 +0200 Subject: [PATCH 11/11] Squashed commit of the following: MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit commit d08dabfacb486fb77a2ded077a3f0593968412e8 Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 22:32:26 2025 +0200 works commit 74b6ee0df2aaba26be3f09f81b2020470b27cc44 Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 22:29:59 2025 +0200 gi commit f89d4d7f172ae7da24a385efdd783209a5b2bdea Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 22:29:15 2025 +0200 gi commit 3528c1128f48573748c240b3b5d3ea27fee5e0e5 Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 20:04:00 2025 +0200 cleanup commit 4028c0010c1f0f48f54577ce418185648b2b9902 Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 20:02:46 2025 +0200 cleanup commit 02f1aa892a739e61be52f2a609d0bd5891133639 Merge: 07914eea ac62dd4f Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 19:09:38 2025 +0200 pre-merge commit 07914eeae43b77109983039e6d2e6c67fbae3f2f Author: Marcin Jan Turek-Grzybowski Date: Thu Jun 26 18:59:53 2025 +0200 integration done commit ac62dd4f433ed1a45f9bf944e9acb7beb8a4a7c9 Author: anshwad10 <109362320+anshwad10@users.noreply.github.com> Date: Wed Jun 25 17:16:13 2025 +0530 Add Rezk Completion by HIT (#1188) * Add Rezk Completion by HIT * Update Construction.agda * Removed trailing whitespaces * Remove trailing whitespace * Replace HIT with GroupoidQuotient * shorter proof of `inc-surj` * fix whitespace * Forgot to import GroupoidQuotients * add implicit argument commit b15ec4ebde4bc5bf98c3bc671683e85fc02044b3 Author: Marcin Jan Turek-Grzybowski Date: Tue Jun 24 06:45:56 2025 +0200 wip commit 55479efdb10a2999dfec0bec50e4411a788d44ab Author: Marcin Jan Turek-Grzybowski Date: Wed May 28 16:00:42 2025 +0200 wip - dirty commit ecf1bbbfc2231f7700684255a85646a8a3e2e8c4 Author: Laine Taffin Altman Date: Fri May 23 03:17:16 2025 -0700 Fix makefile race condition (#1210) If you use `make -j gen-everythings listings`, currently there is a near-certainty that the `listings` task will fail because the Everythings modules haven't been generated yet. This PR enforces the dependency. commit 43183c8037a67f1f29b45dff86050dbc39ca290e Author: michael <1700346+mzhang28@users.noreply.github.com> Date: Thu May 22 17:01:37 2025 -0500 Five lemma (#1166) * exact sequences * complete five proof * make five safe * fix newline * exactness over a successor structure * working again * fix whitespace... * clean up, just five * clean up * make implicit + rename isExact so the names are more descriptive * fix whitespace commit 63d187dc6f1d6dca388f60dc7fc94f81072ad2cf Author: anshwad10 <109362320+anshwad10@users.noreply.github.com> Date: Tue May 20 22:52:24 2025 +0530 Formalize Theorem 7.2.2 in the HoTT Book (#1180) * Formalize Theorem 7.2.2 in the HoTT Book * move Theorem 7.2.2 to Cubical.Relation.Binary.Properties * Remove a few extra spaces This is not really necessary but it looks better commit 06e62f4cfabb0e693e3c399d24a10348fa578eab Author: michael <1700346+mzhang28@users.noreply.github.com> Date: Tue May 20 07:24:54 2025 -0500 Composition of left module homomorphisms (#1176) * implement composition of left module homomorphisms * pull out to top level * fix whitespace commit 78203ee8e494275bf1b8b90c78b3851a28d8e150 Author: anshwad10 <109362320+anshwad10@users.noreply.github.com> Date: Mon May 19 21:41:29 2025 +0530 Remove duplicated lemma (#1189) * Remove duplicated lemma #1017 * Insert brackets commit 2738401afdec79dbaf580131e27b0cee8d16b7b0 Author: LorenzoMolena <164308953+LorenzoMolena@users.noreply.github.com> Date: Mon May 19 17:52:30 2025 +0200 Strict order equational reasoning (#1203) * Add QuosetReasoning.agda Co-authored-by: Ettore Forigo * add missing `no-eta-equality` in SubRelation, make SubRelation instance private --------- Co-authored-by: Ettore Forigo commit 69b70a2fbad125a7737221540c4f0bb9a5b54f9b Author: anshwad10 <109362320+anshwad10@users.noreply.github.com> Date: Mon May 19 21:19:57 2025 +0530 Add inducedFun for EM1 (#1208) * Add files via upload * Delete Base.agda No idea what this file is doing here * Delete Properties.agda Must have accidentally put this here when I was editing dagger-cat * Add inducedFun for EM1 commit 22458b709ee5609daab0a58ef5c80fc232ce42d7 Author: LeeeeT Date: Mon May 19 18:46:33 2025 +0300 Add missing minus symbol (#1207) commit 3291fc1d8cae1409cc54d4f9cca7366158b573d7 Author: Felix Cherubini Date: Mon May 19 17:24:35 2025 +0200 Release for agda 2.7 (#1206) * test new ci script * fix * use fix-whitespace action * blindly try stuff * blindly try stuff * blindly try stuff * untabify... * Revert "blindly try stuff" This reverts commit 6fe71f3e792cd31cf807509a6ba66ca2f5560eec. # Conflicts: # .github/workflows/ci-ubuntu.yml * use fix-whitespace action * whitespace * whitespace * fix * fix * fix * fix * Fix path of fix-whitespace * Remove unnecessary step to generate user-manual.pdf * Bump base in cubical-utils.cabal * bump version number everywhere, update README and RELEASE * trailing newline needed? * update flake --------- Co-authored-by: Andreas Abel Co-authored-by: Naïm Favier commit eef2ed612ad858f790b802d83ddf7a95946d80f1 Author: Marcin Jan Turek-Grzybowski Date: Thu May 15 07:17:58 2025 +0200 wip commit 71f32f05da9aaa85c2f5c4ef760d0304522aec55 Author: Andreas Abel Date: Fri May 9 21:51:25 2025 +0200 Prep for Agda 2.8.0: remove some spurious `private` commit d5b7bcec99497e2b409a84c5d55fa512b10fdb47 Author: Marcin Grzybowski Date: Mon Apr 21 20:46:48 2025 +0200 sync commit 35d29193477fe168edfe4bff66af984068bf85e4 Merge: 385066b8 cb0510c9 Author: Evan Cavallo Date: Fri Mar 21 06:08:36 2025 +0100 Merge pull request #1157 from Trebor-Huang/devalapurkar-haine Devalapurkar & Haine commit cb0510c90297244bfb527788d292d7ea565a1927 Author: ecavallo Date: Thu Mar 20 18:40:56 2025 +0100 cosmetic changes commit 52cf2a1725f7d8a35528368ac59d204aa768c6ed Author: ecavallo Date: Thu Mar 20 18:40:42 2025 +0100 more informative type signatures commit 01c36f4846e396da0868c8eb3f3c8f9f36bfa960 Author: ecavallo Date: Thu Mar 20 18:40:23 2025 +0100 prefer copattern matching to record syntax commit fa0eba6064f0d7eb38f0332e41803339c491254b Author: ecavallo Date: Thu Mar 20 13:59:26 2025 +0100 make pushout to wedge equivalence opaque commit 4155afda46ffc7ed0eacc49e5b22dad8eefb3f70 Merge: 01fa4208 385066b8 Author: ecavallo Date: Tue Mar 18 14:01:34 2025 +0100 Merge branch 'master' into devalapurkar-haine commit 385066b8128a8cedcc53133bc1ccaa27f00dfa03 Author: Loïc Pujet <38562414+loic-p@users.noreply.github.com> Date: Tue Mar 4 16:53:03 2025 +0100 Reduced homology of CW complexes (#1175) * Reduced homology of CW complexes * explicit description of the augmentation map commit 90c3244f4ea0da5dccee314d60f34e1aacb1a9dd Author: Marcin Grzybowski Date: Fri Feb 14 16:53:06 2025 +0100 wip commit e8ff0c1e18fd200d0c6a2d70f07811134df7d3cf Author: anshwad10 <109362320+anshwad10@users.noreply.github.com> Date: Fri Feb 14 21:06:39 2025 +0530 Move `factorial` from `Cubical.Data.Fin.LehmerCode` to `Cubical.Data.Nat.Properties` (#1184) * Update LehmerCode.agda * Update Properties.agda add `factorial` here * Update LehmerCode.agda made it export `factorial` so nothing else in the library that depends on `factorial` being exported breaks * Removed dependency on LehmerCode exporting `factorial` * Removed dependency on LehmerCode exporting `factorial` * Removed dependency on LehmerCode exporting `factorial` * Remove the public export of `factorial` from LehmerCode.agda commit 3d5bb7d6bf32d4b6835a8b6fd99fa0eb0985f914 Author: Brad Dragun <35275808+LuuBluum@users.noreply.github.com> Date: Tue Feb 11 23:02:11 2025 -0700 Fix missed rename of StrictPoset to Quoset (#1185) commit c16edadee9ca56448546574ab14e34c1f2f875b8 Author: anshwad10 <109362320+anshwad10@users.noreply.github.com> Date: Sat Feb 8 21:32:52 2025 +0530 Remove duplicated factorial function (#1183) `_!` is defined in `Cubical.Data.Nat.Properties`, `factorial` is defined in `Cubical.Data.Fin.LehmerCode`, and they both have identical definitions. commit f717d461c5ac05db4d4325ff1ef45ea6f728efd3 Author: Brad Dragun <35275808+LuuBluum@users.noreply.github.com> Date: Fri Feb 7 09:27:29 2025 -0700 Basic Order Theory (#1154) * Add inverses to monoids/comm monoids * Add meets and joins to order properties * Rewrite isConnected and rename isStronglyConnected * Decidable total and linear orders imply discrete * Remove discrete requirement * Move decidable->discrete from toset to poset * Define binary meets/joins and prove properties * The negation of a linear order is a poset * Define bounds on a poset * Mild refactoring Rename preorder to proset; rename strict poset to quoset (quasiorder) and add strict orders as quasiorders with weak linearity * Remove unnecessary constructors * Define tight relations * Reintroduce constructors * Lattice basics * Distributive lattices * Replace compEquiv usages * Total orders are distributive pseudolattices * Add pseudolattice assumption to make more useful * Mappings on posets * Downsets and upsets are preserved * Defined complete lattices * Duals and closures * Embeddings form a complete lattice * Dual closures * Bicomplete subsets * Poset equivalences * Galois connections * Fix typo * Fix whitespace * Fix build failures * Move PosetDownset to where principalDownsets are * Fix build failures * Split mappings and subsets * Fix whitespace * Proofs regarding residuals and involutions * Cleaner proof commit f3444859771d09cde7b5483415ed141ae78eefd4 Author: Marcin Grzybowski Date: Thu Jan 30 15:55:08 2025 +0100 wip commit 01fa4208b2110765b2cdf93219756124550eff4d Author: Trebor-Huang Date: Thu Sep 19 22:16:54 2024 +0800 safe flags commit 54a2b1fa49ef4a26a62ff52f0f3ba9fa412ebe36 Author: Trebor-Huang Date: Thu Sep 19 21:42:05 2024 +0800 Fix whitespace commit a181af86d2afc3f756f164ccd95b24e02550ac77 Merge: f0d49bc4 53e400ec Author: Trebor-Huang Date: Thu Sep 19 21:36:50 2024 +0800 Merge remote-tracking branch 'origin/master' into devalapurkar-haine commit f0d49bc456b0798603ef0ed6e9bdc88b9a12ff84 Author: Trebor-Huang Date: Thu Sep 19 21:27:40 2024 +0800 James commit d99570a63326ea5471a694e0f55a00aefef93d45 Author: Trebor-Huang Date: Thu Sep 19 21:10:57 2024 +0800 Hilton–Milnor commit c983de796ea23119cfdc910eb41c3d7a77525962 Author: Trebor-Huang Date: Thu Sep 19 20:36:36 2024 +0800 Splitting of loopspace commit 08a56c498329612d3d6e9d4ae27b6655a91b2873 Author: Trebor-Huang Date: Thu Sep 19 20:21:45 2024 +0800 Oddly specific sigma lemma commit 51fa7d02ec85a68ec34a0fb4c405e2298b1cecf9 Author: Trebor-Huang Date: Thu Sep 19 20:12:35 2024 +0800 Remove my version in favor of the other commit dee0c5608ea276a7c97744ecbacc9eee05260df1 Author: Trebor-Huang Date: Sun Aug 18 00:08:50 2024 +0800 extend commutative squares commit 98bbb7a28cdb7d3465455615ad5fadf98970dfaf Author: Trebor-Huang Date: Sun Aug 18 00:08:43 2024 +0800 tweak commit a4b23dd2b9dd08b3d614d447aee9ba435a47977f Author: Trebor-Huang Date: Thu Aug 15 17:45:42 2024 +0800 Rename commit dac0c0ef7abc0c21de26287154dba5c5f5e0f028 Author: Trebor-Huang Date: Tue Aug 13 02:07:32 2024 +0800 Pushout using Unit* commit f9fc7e32263dfe5f5f23b26d2bac87fb9a9c0e86 Author: Trebor-Huang Date: Mon Aug 12 13:50:15 2024 +0800 Use PushoutSquare instead commit 72a75579ac9152673b56bda44f67c9352f7483d1 Author: Trebor-Huang Date: Mon Aug 12 03:58:36 2024 +0800 Universes commit 9d50340ea6072c87a4207d38077740545c08585b Author: Trebor-Huang Date: Mon Aug 12 03:18:06 2024 +0800 Susp (X∙ ⋀ Y∙) ≡ join X Y commit 4bc6517beba1c4ad1cd5a79b6364acf1dd748e50 Author: Trebor-Huang Date: Mon Aug 12 02:51:38 2024 +0800 Pushout⋁≃Unit commit 90f1ca9f37e5b2587f34f55e879cfbcc3d1fe78c Author: Trebor-Huang Date: Mon Aug 12 02:50:52 2024 +0800 Remove redundant definition commit fa2da1a05ecd55bd99203ece9c4cbc1781794cb2 Author: Trebor-Huang Date: Mon Aug 12 02:50:36 2024 +0800 projection functions commit 54038ae9fb28f9a30810c7518246c9819e5bf18f Author: Trebor-Huang Date: Mon Aug 12 02:31:37 2024 +0800 join inclusions are null commit 313b74e0dead200f337708449f2cb1dca68d1cd3 Author: Trebor-Huang Date: Mon Aug 12 00:22:40 2024 +0800 Use pushout squares for cofibInl-⋁ commit d0a06cdb284a25dd095c93c78ca78121ed9e05d9 Author: Trebor-Huang Date: Sun Aug 11 21:40:16 2024 +0800 pushout along identity commit b4e74c14767031ec00ebbb43696186a53b3b030c Author: Trebor-Huang Date: Sun Aug 11 21:08:23 2024 +0800 Susp is Pushout 1 <- X -> 1 commit 5dae63277ab3ff9ae977d2747d33597a08973a08 Author: aljungstrom Date: Tue Sep 17 01:01:18 2024 +0200 fixes commit e13f9abc2333c0281f21b4d2237e4cb7f8030ad2 Merge: 10aea90c e2370fbb Author: aljungstrom Date: Mon Sep 16 22:22:28 2024 +0200 Merge remote-tracking branch 'master/master' into cellular_pointed commit 10aea90c18c24d09626a0ec1c7d317c7e0fb5f2f Author: aljungstrom Date: Tue Jun 4 14:42:34 2024 +0200 minor commit 1dff2d524093f32c9d05500463c09ee781f2c30e Merge: 46482afb a0b4ba3e Author: aljungstrom Date: Tue Jun 4 14:39:52 2024 +0200 merge commit 46482afb26dd76ebaa06e2b57c1134e98b541990 Author: aljungstrom Date: Tue Jun 4 14:22:49 2024 +0200 connected clean commit 787f34f42311454c48db5def54b9154b6d564842 Author: aljungstrom Date: Tue Jun 4 14:12:57 2024 +0200 connected done? commit 73afa8ec3c639433f8d6b701c7d72b9083e6f9dc Author: aljungstrom Date: Tue Jun 4 02:59:48 2024 +0200 cleaning commit 0cb567811fcf46b77da666ac4eac2522039d9f89 Author: aljungstrom Date: Tue Jun 4 02:59:42 2024 +0200 cleaning commit 6bc36641e78cc27c04811c24236e6eb792ee9656 Author: aljungstrom Date: Fri May 31 18:51:51 2024 +0200 pretty much done commit e4f400342907ec03159b741b3a43900777e1b819 Author: aljungstrom Date: Fri May 31 04:09:48 2024 +0200 almost commit c7acbeee22598a9dc201e68b79e659bd8b394ad0 Author: aljungstrom Date: Wed May 29 18:11:46 2024 +0200 stuff commit 77623bb0a40f5be302202c60e9f3ae5791082c90 Author: aljungstrom Date: Mon May 27 03:16:45 2024 +0200 wip... commit dac2fd33e3aaabcb936104afc4e0685f130c0dc6 Author: aljungstrom Date: Thu May 23 18:08:06 2024 +0200 stuff commit afe51a465bd2d086d2461e1d80b44719d539bb5b Author: aljungstrom Date: Wed May 8 18:20:52 2024 +0200 p2 commit 87fa4c021063e4ec37fa39d4b7bfe78e2c44a881 Author: aljungstrom Date: Tue May 7 18:15:55 2024 +0200 done? commit eb6b12e6b4d9c7b00ffe4dec96d86f17320dd173 Author: aljungstrom Date: Tue May 7 10:12:34 2024 +0200 Pointed commit 3784dcd118ef25e2123535a01df66d95097b2de6 Author: aljungstrom Date: Thu May 2 11:34:11 2024 +0200 comments commit beccacd97a903f3140d19c9ac50dfcff40f89b2e Author: aljungstrom Date: Mon Mar 4 12:03:53 2024 +0100 ojdå commit 3e53282ae92499ae69f28596cd195e55055a76b3 Author: aljungstrom Date: Mon Mar 4 11:31:38 2024 +0100 broken code commit ac5249e76295d97d20f3f827158e0f7b994e5150 Author: aljungstrom Date: Mon Mar 4 09:39:57 2024 +0100 wups commit 1b005a64708be0b27eceefd6c5f1628385f00ae7 Author: aljungstrom Date: Mon Mar 4 02:22:00 2024 +0100 ugh commit 2d3fe1c02ac12dcbebc159a47f37a7fb11d08580 Author: aljungstrom Date: Mon Mar 4 01:48:45 2024 +0100 wups commit b1f9e1dbc9c857a821f8be5b24ac1ec46ef51724 Author: aljungstrom Date: Mon Mar 4 01:30:09 2024 +0100 readme commit 6e259aff9109e5c799b9eb15ef45578ffb3965cb Author: aljungstrom Date: Mon Mar 4 00:38:02 2024 +0100 cleanup commit a3f6a9e42ef58043ebf485f00085137f64d1ffc4 Merge: 6ce70a49 f4745a23 4ba3c8eb Author: aljungstrom Date: Sun Mar 3 21:13:21 2024 +0100 merge commit 4ba3c8eb525864245fa8d2e532827e3bc0631b55 Author: aljungstrom Date: Sun Mar 3 20:55:52 2024 +0100 stuff commit a3bb47e1a71da858c66ee029b4a5dd53a9bfd620 Author: aljungstrom Date: Sun Mar 3 20:48:17 2024 +0100 done? commit 53bb0f29ffe6bfe80f6e4256081b2531b21de3d0 Author: aljungstrom Date: Sun Mar 3 03:52:29 2024 +0100 stuff commit 54c7919a18c48e72ccf111c9179fe0f5c3da0a4a Author: aljungstrom Date: Fri Mar 1 19:08:20 2024 +0100 stuff commit 42d6b770bb45031ce19e5f486000dc946cb9ecfe Author: aljungstrom Date: Thu Feb 29 15:34:05 2024 +0100 stuff commit df55d5508e8e8029cb3d147636f8a454a008aed9 Author: aljungstrom Date: Thu Feb 29 02:49:34 2024 +0100 More commit ae6e813d48cb6620094243de4d03e2afaad907be Author: aljungstrom Date: Tue Feb 27 00:10:56 2024 +0100 stuf commit b6941207463c57b23399f97d1709f0d635845e16 Author: aljungstrom Date: Mon Feb 26 16:42:26 2024 +0100 stuff commit ea5fefdda2f7ceb38d0cd1bc799ec4cdff448dac Author: aljungstrom Date: Sun Feb 25 23:02:26 2024 +0100 stuff commit f4745a23aa8462ca80e936c25ccc41c648fe980d Merge: 1c0396ca 227b10b9 Author: aljungstrom Date: Wed Jan 31 10:49:17 2024 +0100 Merge remote-tracking branch 'origin/master' commit 90714817e9e9115cdcc0d50fb9f839d17bfef30d Author: aljungstrom Date: Wed Jan 31 10:46:56 2024 +0100 stuff commit e47bad54dc36571f9fc7f41f551e347131991430 Author: aljungstrom Date: Wed Jan 24 19:53:25 2024 +0100 wups commit f7524ebc3fed41d408c158e93748ec2f3bbae462 Author: aljungstrom Date: Tue Jan 23 19:44:57 2024 +0100 ugh... commit b7716a46e9017d052b7d0a6bd1b2db4db16d40a4 Author: aljungstrom Date: Tue Jan 23 19:19:56 2024 +0100 fixes commit 7313a311f150ab352c767dc716c9f115bce8d46a Author: aljungstrom Date: Tue Jan 23 15:04:18 2024 +0100 wups commit 7875e3ee00816291bf03351f0ea3bfce3fc356fb Author: aljungstrom Date: Tue Jan 23 14:51:05 2024 +0100 wups commit 74a85212fdbfedc27cc2f3c9e2c846246065ec3d Author: aljungstrom Date: Tue Jan 23 14:47:14 2024 +0100 wups commit 5e8b305aec47a6cac3167a2262d0874de6884643 Author: aljungstrom Date: Tue Jan 23 14:17:35 2024 +0100 done? commit 74596b53b1c3b6d96b13ad5f468386ea76769155 Author: aljungstrom Date: Tue Jan 23 03:52:06 2024 +0100 .. commit 2b8b5fda7e43eb304284c54383710752de3b2f10 Author: aljungstrom Date: Mon Jan 22 18:50:46 2024 +0100 more commit 2fef4ef8ccaf61ec6105bb9c6201281bd305eb47 Author: aljungstrom Date: Mon Jan 22 18:36:10 2024 +0100 stuff commit c0345787321d5730595763e4d0f1dc706eebd6af Author: aljungstrom Date: Thu Jan 18 17:17:24 2024 +0100 stuff commit 9d1915dd861102c72f1a497bf2d7e9b9e0295edf Author: aljungstrom Date: Wed Jan 17 02:53:26 2024 +0100 fix commit 227b10b9d16e0bdb866941a1eb9195e481c13999 Author: aljungstrom Date: Thu Jun 8 09:45:56 2023 +0200 w commit 3acc6b862989c2d0e5d0f58137d1b9c2a144a48a Author: aljungstrom Date: Thu Jun 8 09:45:28 2023 +0200 w commit ea75092e20dc74fd00e788088ad7a35cbe5dfc97 Author: aljungstrom Date: Thu Jun 8 09:44:37 2023 +0200 w commit 52429ff1823be41f0ae64830314d4c8e07906d94 Merge: 93417102 814d54b0 Author: aljungstrom Date: Thu Jun 8 09:38:02 2023 +0200 Merge branch 'newM' commit 934171020cab0a9fc68f298d4c34a9958c621ca1 Author: aljungstrom Date: Fri Jun 2 15:04:54 2023 +0200 ganea stuff commit 17f3fc1f5dbed5cf98bb5f8f58d9d4ed56aa81b5 Author: aljungstrom Date: Mon Apr 24 17:13:52 2023 +0200 rme commit aadde33a93f44d140376682ec774a7d460ac183f Author: aljungstrom Date: Mon Apr 24 17:12:41 2023 +0200 wups commit e8244ac1f3364abe7adb794f27d805d466437524 Author: aljungstrom Date: Mon Apr 24 17:10:45 2023 +0200 duplicate commit cc6ced25804997022e855a458511bf6ce1b38247 Merge: b4b6efb7 f0ad815d Author: aljungstrom Date: Mon Apr 24 17:10:10 2023 +0200 done commit b4b6efb72c8bce27d4ae6c9b5275a846868349c1 Merge: 93d7248f 9acdecfa Author: aljungstrom Date: Tue Feb 1 04:02:50 2022 +0100 m commit 93d7248fb129e34702bd7d6a472a4eb9ba55072a Author: aljungstrom Date: Tue Feb 1 04:00:15 2022 +0100 t --- .github/workflows/ci-ubuntu.yml | 37 +- CITATION.cff | 4 +- .../AbGroup/Instances/FreeAbGroup.agda | 47 + Cubical/Algebra/Algebra/Properties.agda | 2 +- Cubical/Algebra/ChainComplex/Homology.agda | 42 +- Cubical/Algebra/CommMonoid/Properties.agda | 16 + .../Localisation/UniversalProperty.agda | 10 +- Cubical/Algebra/DistLattice/Downset.agda | 1 + Cubical/Algebra/Group/Exact.agda | 30 +- Cubical/Algebra/Group/Five.agda | 265 ++ Cubical/Algebra/Group/MorphismProperties.agda | 15 + Cubical/Algebra/Module/Properties.agda | 49 + Cubical/Algebra/Monoid/Base.agda | 10 + Cubical/Algebra/Ring/Properties.agda | 43 + .../ZariskiLattice/Properties.agda | 1 + .../ZariskiLattice/StructureSheaf.agda | 3 +- .../StructureSheafPullback.agda | 3 +- Cubical/CW/ChainComplex.agda | 102 +- Cubical/CW/Homology.agda | 228 +- Cubical/CW/Homotopy.agda | 267 +- Cubical/CW/Map.agda | 82 +- Cubical/CW/Subcomplex.agda | 51 +- .../Categories/Constructions/SubObject.agda | 14 +- .../DistLatticeSheaf/Extension.agda | 9 +- .../RezkCompletion/Construction.agda | 387 ++- Cubical/Data/Cardinal/Properties.agda | 7 +- Cubical/Data/Fin/Inductive/Properties.agda | 6 + Cubical/Data/Fin/LehmerCode.agda | 6 +- Cubical/Data/FinSet/Cardinality.agda | 2 +- Cubical/Data/FinSet/Constructors.agda | 2 +- Cubical/Data/Int/Base.agda | 19 +- Cubical/Data/Int/Order.agda | 16 +- Cubical/Data/Int/Properties.agda | 21 +- Cubical/Data/Nat/Order.agda | 111 + Cubical/Data/Nat/Properties.agda | 8 +- Cubical/Data/Ordinal/Properties.agda | 10 +- Cubical/Data/Rationals/Order.agda | 116 +- Cubical/Data/Rationals/Order/Properties.agda | 278 +- Cubical/Data/Rationals/Properties.agda | 117 + Cubical/Data/Sum/Properties.agda | 17 +- Cubical/Data/SumFin/Properties.agda | 2 +- .../ZCohomologyOld/Properties.agda | 2 +- Cubical/Foundations/HLevels.agda | 9 + Cubical/Foundations/Pointed/Properties.agda | 8 + Cubical/Foundations/Prelude.agda | 4 + Cubical/Foundations/Transport.agda | 4 + Cubical/Functions/Embedding.agda | 76 +- Cubical/Functions/Image.agda | 7 +- Cubical/Functions/Preimage.agda | 58 + Cubical/HITs/CauchyReals/Base.agda | 2 +- Cubical/HITs/CauchyReals/Bisect.agda | 1338 +++++++++ Cubical/HITs/CauchyReals/BisectApprox.agda | 1775 +++++++++++ Cubical/HITs/CauchyReals/Closeness.agda | 22 +- Cubical/HITs/CauchyReals/Continuous.agda | 1457 ++++++--- Cubical/HITs/CauchyReals/Derivative.agda | 721 ++++- Cubical/HITs/CauchyReals/ExpLog.agda | 209 ++ Cubical/HITs/CauchyReals/Exponentiation.agda | 2331 +++++++++++++++ .../HITs/CauchyReals/ExponentiationDer.agda | 1584 ++++++++++ .../HITs/CauchyReals/ExponentiationMore.agda | 2252 ++++++++++++++ Cubical/HITs/CauchyReals/Integration.agda | 2608 +++++++++++++++++ Cubical/HITs/CauchyReals/IntegrationMore.agda | 1266 ++++++++ Cubical/HITs/CauchyReals/Inverse.agda | 1814 ++++++++---- Cubical/HITs/CauchyReals/Lems.agda | 89 + Cubical/HITs/CauchyReals/Lipschitz.agda | 20 + Cubical/HITs/CauchyReals/MeanValue.agda | 610 ++++ Cubical/HITs/CauchyReals/Multiplication.agda | 1267 ++++---- Cubical/HITs/CauchyReals/NthRoot.agda | 581 ++++ Cubical/HITs/CauchyReals/Order.agda | 798 +++-- Cubical/HITs/CauchyReals/Sequence.agda | 1259 +++++--- Cubical/HITs/James/Stable.agda | 87 + Cubical/HITs/Join/Properties.agda | 9 + Cubical/HITs/Pushout/Properties.agda | 64 + Cubical/HITs/SetTruncation/Properties.agda | 35 +- Cubical/HITs/SmashProduct/Base.agda | 7 +- Cubical/HITs/Susp/Properties.agda | 44 + Cubical/HITs/Susp/SuspProduct.agda | 236 ++ Cubical/HITs/Wedge/Base.agda | 8 + Cubical/HITs/Wedge/Properties.agda | 102 +- .../Homotopy/EilenbergMacLane/Properties.agda | 28 + Cubical/Homotopy/HiltonMilnor.agda | 96 + Cubical/Homotopy/Loopspace.agda | 78 + Cubical/README.agda | 2 +- Cubical/Relation/Binary/Base.agda | 30 +- Cubical/Relation/Binary/Order.agda | 6 +- .../Relation/Binary/Order/Apartness/Base.agda | 4 +- .../Binary/Order/Apartness/Properties.agda | 18 + Cubical/Relation/Binary/Order/Loset/Base.agda | 6 +- .../Binary/Order/Loset/Properties.agda | 93 +- Cubical/Relation/Binary/Order/Poset/Base.agda | 20 +- .../Order/Poset/Instances/Embedding.agda | 145 + .../Relation/Binary/Order/Poset/Mappings.agda | 1228 ++++++++ .../Binary/Order/Poset/Properties.agda | 583 +++- .../Relation/Binary/Order/Poset/Subset.agda | 249 ++ Cubical/Relation/Binary/Order/Preorder.agda | 5 - .../Relation/Binary/Order/Preorder/Base.agda | 135 - .../Binary/Order/Preorder/Properties.agda | 57 - Cubical/Relation/Binary/Order/Properties.agda | 302 -- Cubical/Relation/Binary/Order/Proset.agda | 5 + .../Relation/Binary/Order/Proset/Base.agda | 138 + .../Binary/Order/Proset/Properties.agda | 434 +++ Cubical/Relation/Binary/Order/Quoset.agda | 5 + .../Relation/Binary/Order/Quoset/Base.agda | 126 + .../Binary/Order/Quoset/Properties.agda | 80 + .../Binary/Order/QuosetReasoning.agda | 138 + .../Relation/Binary/Order/StrictOrder.agda | 5 + .../Binary/Order/StrictOrder/Base.agda | 130 + .../Binary/Order/StrictOrder/Properties.agda | 120 + .../Relation/Binary/Order/StrictPoset.agda | 5 - .../Binary/Order/StrictPoset/Base.agda | 126 - .../Binary/Order/StrictPoset/Properties.agda | 94 - Cubical/Relation/Binary/Order/Toset/Base.agda | 6 +- .../Binary/Order/Toset/Properties.agda | 337 ++- .../Binary/Order/Woset/Properties.agda | 32 +- .../Binary/Order/Woset/Simulation.agda | 2 +- Cubical/Relation/Binary/Properties.agda | 28 +- GNUmakefile | 2 +- README.md | 6 +- RELEASE.md | 2 + cubical-utils.cabal | 4 +- cubical.agda-lib | 2 +- flake.lock | 66 +- flake.nix | 7 +- 122 files changed, 26705 insertions(+), 3499 deletions(-) create mode 100644 Cubical/Algebra/Group/Five.agda create mode 100644 Cubical/Functions/Preimage.agda create mode 100644 Cubical/HITs/CauchyReals/Bisect.agda create mode 100644 Cubical/HITs/CauchyReals/BisectApprox.agda create mode 100644 Cubical/HITs/CauchyReals/ExpLog.agda create mode 100644 Cubical/HITs/CauchyReals/Exponentiation.agda create mode 100644 Cubical/HITs/CauchyReals/ExponentiationDer.agda create mode 100644 Cubical/HITs/CauchyReals/ExponentiationMore.agda create mode 100644 Cubical/HITs/CauchyReals/Integration.agda create mode 100644 Cubical/HITs/CauchyReals/IntegrationMore.agda create mode 100644 Cubical/HITs/CauchyReals/MeanValue.agda create mode 100644 Cubical/HITs/CauchyReals/NthRoot.agda create mode 100644 Cubical/HITs/James/Stable.agda create mode 100644 Cubical/HITs/Susp/SuspProduct.agda create mode 100644 Cubical/Homotopy/HiltonMilnor.agda create mode 100644 Cubical/Relation/Binary/Order/Poset/Instances/Embedding.agda create mode 100644 Cubical/Relation/Binary/Order/Poset/Mappings.agda create mode 100644 Cubical/Relation/Binary/Order/Poset/Subset.agda delete mode 100644 Cubical/Relation/Binary/Order/Preorder.agda delete mode 100644 Cubical/Relation/Binary/Order/Preorder/Base.agda delete mode 100644 Cubical/Relation/Binary/Order/Preorder/Properties.agda delete mode 100644 Cubical/Relation/Binary/Order/Properties.agda create mode 100644 Cubical/Relation/Binary/Order/Proset.agda create mode 100644 Cubical/Relation/Binary/Order/Proset/Base.agda create mode 100644 Cubical/Relation/Binary/Order/Proset/Properties.agda create mode 100644 Cubical/Relation/Binary/Order/Quoset.agda create mode 100644 Cubical/Relation/Binary/Order/Quoset/Base.agda create mode 100644 Cubical/Relation/Binary/Order/Quoset/Properties.agda create mode 100644 Cubical/Relation/Binary/Order/QuosetReasoning.agda create mode 100644 Cubical/Relation/Binary/Order/StrictOrder.agda create mode 100644 Cubical/Relation/Binary/Order/StrictOrder/Base.agda create mode 100644 Cubical/Relation/Binary/Order/StrictOrder/Properties.agda delete mode 100644 Cubical/Relation/Binary/Order/StrictPoset.agda delete mode 100644 Cubical/Relation/Binary/Order/StrictPoset/Base.agda delete mode 100644 Cubical/Relation/Binary/Order/StrictPoset/Properties.agda diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml index d8b3b347c4..011f755977 100644 --- a/.github/workflows/ci-ubuntu.yml +++ b/.github/workflows/ci-ubuntu.yml @@ -42,9 +42,9 @@ on: ######################################################################## env: - AGDA_BRANCH: v2.6.4.1 - GHC_VERSION: 9.4.7 - CABAL_VERSION: 3.6.2.0 + AGDA_BRANCH: v2.7.0.1 + GHC_VERSION: 9.10.2 + CABAL_VERSION: 3.12.1.0 CABAL_INSTALL: cabal install --overwrite-policy=always --ghc-options='-O1 +RTS -M6G -RTS' CACHE_PATHS: | ~/.cabal/packages @@ -67,7 +67,7 @@ jobs: # absolutely necessary i.e. if we change either the version of Agda, # ghc, or cabal that we want to use for the build. - name: Restore external dependencies cache - uses: actions/cache/restore@v3 + uses: actions/cache/restore@v4 id: cache-external-restore with: path: ${{ env.CACHE_PATHS }} @@ -82,23 +82,13 @@ jobs: run: | git clone https://github.com/agda/agda --branch ${{ env.AGDA_BRANCH }} --depth=1 cd agda - mkdir -p doc - touch doc/user-manual.pdf ${{ env.CABAL_INSTALL }} cd .. rm -rf agda - - name: Download and install fix-whitespace - if: steps.cache-external-restore.outputs.cache-hit != 'true' - run: | - git clone https://github.com/agda/fix-whitespace --depth=1 - cd fix-whitespace - ${{ env.CABAL_INSTALL }} fix-whitespace.cabal - cd .. - - name: Save external dependencies cache if: steps.cache-external-restore.outputs.cache-hit != 'true' - uses: actions/cache/save@v3 + uses: actions/cache/save@v4 id: cache-external-save with: path: ${{ env.CACHE_PATHS }} @@ -110,14 +100,17 @@ jobs: # By default github actions do not pull the repo - name: Checkout cubical - uses: actions/checkout@v3 + uses: actions/checkout@v4 + + - name: Get and run fix-whitespace + uses: andreasabel/fix-whitespace-action@v1 ######################################################################## ## TESTING ######################################################################## - name: Restore library cache - uses: actions/cache/restore@v3 + uses: actions/cache/restore@v4 id: cache-library-restore with: path: ./_build @@ -125,19 +118,19 @@ jobs: restore-keys: | library-${{ runner.os }}-${{ env.GHC_VERSION }}-${{ env.CABAL_VERSION }}-${{ env.AGDA_BRANCH }}- - - name: Put cabal programs in PATH + - name: Put installed programs in PATH run: echo "~/.cabal/bin" >> "${GITHUB_PATH}" - name: Test cubical run: | make test \ AGDA_EXEC='~/.cabal/bin/agda -WnoUnsupportedIndexedMatch -W error' \ - FIX_WHITESPACE='~/.cabal/bin/fix-whitespace' \ + FIX_WHITESPACE='~/.local/fix-whitespace/bin/fix-whitespace' \ EVERYTHINGS='cabal run Everythings' - name: Save library cache if: steps.cache-library-restore.outputs.cache-hit != 'true' - uses: actions/cache/save@v3 + uses: actions/cache/save@v4 id: cache-library-save with: path: ./_build @@ -152,12 +145,12 @@ jobs: run: | make listings \ AGDA_EXEC='~/.cabal/bin/agda -WnoUnsupportedIndexedMatch -W error' \ - FIX_WHITESPACE='~/.cabal/bin/fix-whitespace' \ + FIX_WHITESPACE='~/.local/fix-whitespace/bin/fix-whitespace' \ EVERYTHINGS='cabal run Everythings' - name: Deploy to GitHub Pages if: github.event_name == 'push' && github.ref_name == 'master' - uses: JamesIves/github-pages-deploy-action@4.1.8 + uses: JamesIves/github-pages-deploy-action@v4 with: branch: gh-pages folder: html diff --git a/CITATION.cff b/CITATION.cff index 048a89cdaf..da5f7ddd2c 100644 --- a/CITATION.cff +++ b/CITATION.cff @@ -3,6 +3,6 @@ message: "If you use this software, please cite it as below." authors: - name: "The Agda Community" title: "Cubical Agda Library" -version: 0.7 -date-released: 2024-02-12 +version: 0.8 +date-released: 2025-05-09 url: "https://github.com/agda/cubical" diff --git a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda index 53f6b48a79..f16052a98e 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -19,6 +19,7 @@ open import Cubical.HITs.FreeAbGroup open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.Pi open import Cubical.Algebra.AbGroup.Instances.Int +open import Cubical.Algebra.AbGroup.Instances.DirectProduct open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties @@ -413,3 +414,49 @@ agreeOnℤFinGenerator→≡ {n} {m} {ϕ} {ψ} idr = λ f p → IsGroupHom.presinv (snd ϕ) f ∙∙ (λ i x → -ℤ (p i x)) ∙∙ sym (IsGroupHom.presinv (snd ψ) f))) + +-- +sumCoefficients : (n : ℕ) → AbGroupHom (ℤ[Fin n ]) (ℤ[Fin 1 ]) +fst (sumCoefficients n) v = λ _ → sumFinℤ v +snd (sumCoefficients n) = makeIsGroupHom (λ x y → funExt λ _ → sumFinℤHom x y) + +ℤFinProductIso : (n m : ℕ) → Iso (ℤ[Fin (n +ℕ m) ] .fst) ((AbDirProd ℤ[Fin n ] ℤ[Fin m ]) .fst) +ℤFinProductIso n m = iso sum→product product→sum product→sum→product sum→product→sum + where + sum→product : (ℤ[Fin (n +ℕ m) ] .fst) → ((AbDirProd ℤ[Fin n ] ℤ[Fin m ]) .fst) + sum→product x = ((λ (a , Ha) → x (a , <→<ᵗ (≤-trans (<ᵗ→< Ha) (≤SumLeft {n}{m})))) + , λ (a , Ha) → x (n +ℕ a , <→<ᵗ (<-k+ {a}{m}{n} (<ᵗ→< Ha)))) + + product→sum : ((AbDirProd ℤ[Fin n ] ℤ[Fin m ]) .fst) → (ℤ[Fin (n +ℕ m) ] .fst) + product→sum (x , y) (a , Ha) with (a ≟ᵗ n) + ... | lt H = x (a , H) + ... | eq H = y (a ∸ n , <→<ᵗ (subst (a ∸ n <_) (∸+ m n) (<-∸-< a (n +ℕ m) n (<ᵗ→< Ha) (subst (λ a → a < n +ℕ m) H (<ᵗ→< Ha))))) + ... | gt H = y (a ∸ n , <→<ᵗ (subst (a ∸ n <_) (∸+ m n) (<-∸-< a (n +ℕ m) n (<ᵗ→< Ha) (<ᵗ→< (<ᵗ-trans {n}{a}{n +ℕ m} H Ha))))) + + product→sum→product : ∀ x → sum→product (product→sum x) ≡ x + product→sum→product (x , y) = ≡-× (funExt (λ (a , Ha) → lemx a Ha)) (funExt (λ (a , Ha) → lemy a Ha)) + where + lemx : (a : ℕ) (Ha : a <ᵗ n) → fst (sum→product (product→sum (x , y))) (a , Ha) ≡ x (a , Ha) + lemx a Ha with (a ≟ᵗ n) + ... | lt H = cong x (Fin≡ (a , H) (a , Ha) refl) + ... | eq H = rec (¬m<ᵗm (subst (λ a → a <ᵗ n) H Ha)) + ... | gt H = rec (¬m<ᵗm (<ᵗ-trans Ha H)) + + lemy : (a : ℕ) (Ha : a <ᵗ m) → snd (sum→product (product→sum (x , y))) (a , Ha) ≡ y (a , Ha) + lemy a Ha with ((n +ℕ a) ≟ᵗ n) + ... | lt H = rec (¬m cong (RezkIsoToPath x x) pathToIso-refl ∙ RezkIsoToPathId x + + RezkIsoToPathβ : ∀ x y f → pathToIso (RezkIsoToPath x y f) ≡ f + RezkIsoToPathβ = elimProp₂ (λ x y → isPropΠ λ _ → isSet-CatIso x y _ _) λ x y f → + CatIso≡ _ _ $ transportRefl _ ∙ C.⋆IdL _ + + isUnivalentRezk : isUnivalent Rezk + isUnivalentRezk .univ x y = isoToIsEquiv $ + iso pathToIso (RezkIsoToPath x y) (RezkIsoToPathβ x y) (RezkIsoToPathη x y) + + →Rezk : Functor C Rezk + →Rezk .F-ob = inc + →Rezk .F-hom = idfun _ + →Rezk .F-id = refl + →Rezk .F-seq f g = refl + + →RezkFF : isFullyFaithful →Rezk + →RezkFF x y = idIsEquiv _ + + →RezkSurj : isEssentiallySurj →Rezk + →RezkSurj = elimProp (λ x → isPropPropTrunc) λ x → ∣ x , idCatIso ∣₁ + + isWkEquiv→Rezk : isWeakEquivalence →Rezk + isWkEquiv→Rezk .fullfaith = →RezkFF + isWkEquiv→Rezk .esssurj = →RezkSurj + + isRezkCompletion→Rezk : isRezkCompletion →Rezk + isRezkCompletion→Rezk = makeIsRezkCompletion isUnivalentRezk isWkEquiv→Rezk diff --git a/Cubical/Data/Cardinal/Properties.agda b/Cubical/Data/Cardinal/Properties.agda index 1c7371cf0c..d4041dafed 100644 --- a/Cubical/Data/Cardinal/Properties.agda +++ b/Cubical/Data/Cardinal/Properties.agda @@ -25,8 +25,7 @@ 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 +open import Cubical.Relation.Binary.Order.Proset private variable @@ -171,9 +170,9 @@ module _ where λ (A , _) (B , _) (C , _) → ∥₁.map2 λ A↪B B↪C → compEmbedding B↪C A↪B - isPreorder≲ : IsPreorder {ℓ-suc ℓ} _≲_ + isPreorder≲ : IsProset {ℓ-suc ℓ} _≲_ isPreorder≲ - = ispreorder isSetCard isPropValued≲ isRefl≲ isTrans≲ + = isproset isSetCard isPropValued≲ isRefl≲ isTrans≲ isLeast𝟘 : ∀{ℓ} → isLeast isPreorder≲ (Card {ℓ} , id↪ (Card {ℓ})) (𝟘 {ℓ}) isLeast𝟘 = ∥₂.elim (λ x → isProp→isSet (isPropValued≲ 𝟘 x)) diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index 9b553fd878..7ad1ab2d90 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -25,6 +25,12 @@ open import Cubical.Relation.Nullary open import Cubical.Algebra.AbGroup.Base using (move4) +Fin≡ : {m : ℕ} (a b : Fin m) → fst a ≡ fst b → a ≡ b +Fin≡ {m} (a , Ha) (b , Hb) H i = + (H i , hcomp (λ j → λ { (i = i0) → Ha + ; (i = i1) → isProp<ᵗ {b}{m} (transp (λ j → H j <ᵗ m) i0 Ha) Hb j }) + (transp (λ j → H (i ∧ j) <ᵗ m) (~ i) Ha)) + fsuc-injectSuc : {m : ℕ} (n : Fin m) → injectSuc {n = suc m} (fsuc {n = m} n) ≡ fsuc (injectSuc n) fsuc-injectSuc {m = suc m} (x , p) = refl diff --git a/Cubical/Data/Fin/LehmerCode.agda b/Cubical/Data/Fin/LehmerCode.agda index 1f2da9db8b..eb0c99a4d3 100644 --- a/Cubical/Data/Fin/LehmerCode.agda +++ b/Cubical/Data/Fin/LehmerCode.agda @@ -191,11 +191,7 @@ encode = equivFun lehmerEquiv decode : LehmerCode n → Fin n ≃ Fin n decode = invEq lehmerEquiv -factorial : ℕ → ℕ -factorial zero = 1 -factorial (suc n) = suc n · factorial n - -lehmerFinEquiv : LehmerCode n ≃ Fin (factorial n) +lehmerFinEquiv : LehmerCode n ≃ Fin (n !) lehmerFinEquiv {zero} = isContr→Equiv isContrLehmerZero isContrFin1 lehmerFinEquiv {suc n} = _ ≃⟨ invEquiv lehmerSucEquiv ⟩ _ ≃⟨ ≃-× (idEquiv _) lehmerFinEquiv ⟩ diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index 0c67330496..4aa0849f89 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -379,7 +379,7 @@ module _ module _ (X : FinSet ℓ ) where - cardAut : card (_ , isFinSetAut X) ≡ LehmerCode.factorial (card X) + cardAut : card (_ , isFinSetAut X) ≡ (card X !) cardAut = refl module _ diff --git a/Cubical/Data/FinSet/Constructors.agda b/Cubical/Data/FinSet/Constructors.agda index e4b95450dc..6ae21544ee 100644 --- a/Cubical/Data/FinSet/Constructors.agda +++ b/Cubical/Data/FinSet/Constructors.agda @@ -168,7 +168,7 @@ module _ (X : FinSet ℓ) where isFinSetAut : isFinSet (X .fst ≃ X .fst) - isFinSetAut = LehmerCode.factorial (card X) , + isFinSetAut = card X ! , Prop.map (λ p → isFinOrd≃ (X .fst) (card X , p) .snd) (X .snd .snd) isFinSetTypeAut : isFinSet (X .fst ≡ X .fst) diff --git a/Cubical/Data/Int/Base.agda b/Cubical/Data/Int/Base.agda index 0576b0f241..22fdd45573 100644 --- a/Cubical/Data/Int/Base.agda +++ b/Cubical/Data/Int/Base.agda @@ -6,7 +6,7 @@ module Cubical.Data.Int.Base where open import Cubical.Foundations.Prelude open import Cubical.Data.Bool -open import Cubical.Data.Nat hiding (_+_ ; _·_) renaming (isEven to isEvenℕ ; isOdd to isOddℕ) +open import Cubical.Data.Nat as ℕ hiding (_+_ ; _·_) renaming (isEven to isEvenℕ ; isOdd to isOddℕ) open import Cubical.Data.Fin.Inductive.Base infix 8 -_ @@ -60,6 +60,13 @@ _+_ : ℤ → ℤ → ℤ m + pos n = m +pos n m + negsuc n = m +negsuc n +_+f_ : ℤ → ℤ → ℤ +pos n +f pos n₁ = pos (n ℕ.+ n₁) +negsuc n +f negsuc n₁ = negsuc (suc (n ℕ.+ n₁)) +n +f m = n + m +-- pos n +f negsuc n₁ = {!!} +-- negsuc n +f pos n₁ = {!!} + -_ : ℤ → ℤ - pos zero = pos zero - pos (suc n) = negsuc n @@ -74,6 +81,16 @@ pos (suc n) · m = m + pos n · m negsuc zero · m = - m negsuc (suc n) · m = - m + negsuc n · m +_·f_ : ℤ → ℤ → ℤ +pos n ·f pos n₁ = pos (n ℕ.· n₁) +pos zero ·f negsuc n₁ = pos zero +pos (suc n) ·f negsuc n₁ = negsuc (suc n ℕ.· suc n₁) +negsuc n ·f pos zero = pos zero +negsuc n ·f pos (suc n₁) = negsuc (suc n ℕ.· suc n₁) +negsuc n ·f negsuc n₁ = pos (suc n ℕ.· suc n₁) + + + -- Natural number and negative integer literals for ℤ open import Cubical.Data.Nat.Literals public diff --git a/Cubical/Data/Int/Order.agda b/Cubical/Data/Int/Order.agda index e728308348..0422ce48bb 100644 --- a/Cubical/Data/Int/Order.agda +++ b/Cubical/Data/Int/Order.agda @@ -4,7 +4,7 @@ module Cubical.Data.Int.Order where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function - +open import Cubical.Foundations.Equiv open import Cubical.Data.Empty as ⊥ using (⊥) open import Cubical.Data.Int.Base as ℤ open import Cubical.Data.Int.Properties as ℤ @@ -536,6 +536,20 @@ min-0< (pos (suc (suc n₁))) (pos (suc (suc n))) x x₁ = ℕ≤→pos-≤-pos : ∀ m n → m ℕ.≤ n → pos m ≤ pos n ℕ≤→pos-≤-pos m n (k , p) = k , sym (pos+ m k) ∙∙ cong pos (ℕ.+-comm m k) ∙∙ cong pos p +ℕ≥→negsuc-≤-negsuc : ∀ m n → m ℕ.≤ n → negsuc n ≤ negsuc m +ℕ≥→negsuc-≤-negsuc m n = negsuc-≤-negsuc ∘ ℕ≤→pos-≤-pos m n + +pos-≤-pos→ℕ≤ : ∀ m n → pos m ≤ pos n → m ℕ.≤ n +pos-≤-pos→ℕ≤ m n (k , p) = k , injPos ((pos+ k m ∙ ℤ.+Comm (pos k) (pos m)) ∙ p) + +pos-≤-pos≃ℕ≤ : ∀ m n → (pos m ≤ pos n) ≃ (m ℕ.≤ n) +pos-≤-pos≃ℕ≤ m n = propBiimpl→Equiv isProp≤ ℕ.isProp≤ + (pos-≤-pos→ℕ≤ _ _) (ℕ≤→pos-≤-pos _ _) + +pos-<-pos≃ℕ< : ∀ m n → (pos m < pos n) ≃ (m ℕ.< n) +pos-<-pos≃ℕ< m n = propBiimpl→Equiv (isProp< {pos m} {pos n}) ℕ.isProp≤ + (pos-≤-pos→ℕ≤ _ _) (ℕ≤→pos-≤-pos (suc m) n) + 0≤x² : ∀ n → 0 ≤ n ℤ.· n 0≤x² (pos n) = subst (0 ≤_) (pos·pos n n) zero-≤pos diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index 088182ed42..81582a24ab 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -1,6 +1,7 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --v=tc.conv:20 #-} module Cubical.Data.Int.Properties where + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism @@ -1526,3 +1527,21 @@ sign : ℤ → ℤ sign (pos zero) = 0 sign (pos (suc n)) = 1 sign (negsuc n) = -1 + +sign·abs : ∀ m → sign m · pos (abs m) ≡ m +sign·abs (pos zero) = refl +sign·abs (pos (suc n)) = refl +sign·abs (negsuc n) = refl + +-- abs[sign[m]·pos[n]]≡n : ∀ m n → abs (sign m · pos n) ≡ n +-- abs[sign[m]·pos[n]]≡n m n = {!!} + + +-- ·'≡· : ∀ n m → n · m ≡ n ·' m +-- ·'≡· (pos n) (pos n₁) = sym (pos·pos n n₁) +-- ·'≡· (pos zero) (negsuc n₁) = refl +-- ·'≡· (pos (suc n)) (negsuc n₁) = {!!} +-- ·'≡· (negsuc n) (pos zero) = ·AnnihilR (negsuc n) +-- ·'≡· (negsuc n) (pos (suc n₁)) = negsuc·pos n (suc n₁) ∙ +-- {!!} +-- ·'≡· (negsuc n) (negsuc n₁) = negsuc·negsuc n n₁ ∙ sym (pos·pos (suc n) (suc n₁)) diff --git a/Cubical/Data/Nat/Order.agda b/Cubical/Data/Nat/Order.agda index 0b6808ac7c..ec3f6c1f9e 100644 --- a/Cubical/Data/Nat/Order.agda +++ b/Cubical/Data/Nat/Order.agda @@ -58,6 +58,9 @@ isProp≤ {m} {n} (k , p) (l , q) zero-≤ : 0 ≤ n zero-≤ {n} = n , +-zero n +zero-<-suc : 0 < suc n +zero-<-suc {n} = n , +-comm n 1 + suc-≤-suc : m ≤ n → suc m ≤ suc n suc-≤-suc (k , p) = k , (+-suc k _) ∙ (cong suc p) @@ -147,6 +150,10 @@ suc-< p = pred-≤-pred (≤-suc p) subst2 _≤_ (·-comm m k) (·-comm n k) (≤-·k p) +≤monotone· : ∀ {m n k l} → m ≤ n → k ≤ l → m · k ≤ n · l +≤monotone· {m} {n} {k} {l} m≤n k≤l = + ≤-trans (≤-k· {k = m} k≤l) (≤-·k m≤n) + <-k+-cancel : k + m < k + n → m < n <-k+-cancel {k} {m} {n} = ≤-k+-cancel ∘ subst (_≤ k + n) (sym (+-suc k m)) @@ -229,6 +236,9 @@ predℕ-≤-predℕ {suc m} {suc n} ineq = pred-≤-pred ineq (d + suc m) · suc k ≡⟨ cong (_· suc k) r ⟩ n · suc k ∎ +<-·sk' : 0 < k → m < n → m · k < n · k +<-·sk' {zero} {m} {n} 0<0 = ⊥.rec (¬-<-zero 0<0) +<-·sk' {suc k} {m} {n} _ = <-·sk {m} {n} {k} <-·sk-cancel : ∀ {m n k} → m · suc k < n · suc k → m < n <-·sk-cancel {n = zero} x = ⊥.rec (¬-<-zero x) @@ -589,3 +599,104 @@ module Minimal where isPropLeast : {P : ℕ → Type ℓ} → (∀ m → isProp (P m)) → ∀ m → isProp (Least P m) isPropLeast pP m = isPropΣ (pP m) (λ _ → isPropΠ3 λ _ _ _ → isProp⊥) + +_ + + private ·CommR : (a b c : ℤ) → a ℤ.· b ℤ.· c ≡ a ℤ.· c ℤ.· b ·CommR a b c = sym (ℤ.·Assoc a b c) ∙ cong (a ℤ.·_) (ℤ.·Comm b c) ∙ ℤ.·Assoc a c b @@ -166,6 +168,19 @@ data Trichotomy (m n : ℚ) : Type₀ where eq : m ≡ n → Trichotomy m n gt : m > n → Trichotomy m n +record TrichotomyRec {ℓ : Level} (n : ℚ) (P : ℚ → Type ℓ) : Type ℓ where + no-eta-equality + field + lt-case : ∀ m → (p : m < n) → P m + eq-case : P n + gt-case : ∀ m → (p : m > n) → P m + + go : ∀ m → (t : Trichotomy m n) → P m + go m (lt p) = lt-case m p + go m (eq p) = subst P (sym p) eq-case + go m (gt p) = gt-case m p + + module _ where open BinaryRelation @@ -218,8 +233,8 @@ module _ where isAsym< : isAsym _<_ isAsym< = isIrrefl×isTrans→isAsym _<_ (isIrrefl< , isTrans<) - isStronglyConnected≤ : isStronglyConnected _≤_ - isStronglyConnected≤ = + isTotal≤ : isTotal _≤_ + isTotal≤ = elimProp2 {P = λ a b → (a ≤ b) ⊔′ (b ≤ a)} (λ _ _ → isPropPropTrunc) λ a b → ∣ lem a b ∣₁ @@ -232,12 +247,34 @@ module _ where isConnected< : isConnected _<_ isConnected< = - elimProp2 {P = λ a b → ¬ a ≡ b → (a < b) ⊔′ (b < a)} - (λ _ _ → isProp→ isPropPropTrunc) - λ a b ¬a≡b → ∣ lem a b ¬a≡b ∣₁ + elimProp2 {P = λ a b → (¬ a < b) × (¬ b < a) → a ≡ b} + (λ a b → isProp→ (isSetℚ a b)) + lem + where + lem : (a b : ℤ.ℤ × ℕ₊₁) → (¬ [ a ] < [ b ]) × (¬ [ b ] < [ a ]) → [ a ] ≡ [ b ] + lem (a , b) (c , d) (¬ad b) + + +isIncrasingℙ : (P : ℙ ℚ) → (∀ x → x ∈ P → ℚ) → Type₀ +isIncrasingℙ P f = ∀ x x∈ y y∈ → x ℚ.< y → f x x∈ ℚ.< f y y∈ + +isNondecrasingℙ : (P : ℙ ℚ) → (∀ x → x ∈ P → ℚ) → Type₀ +isNondecrasingℙ P f = ∀ x x∈ y y∈ → x ℚ.≤ y → f x x∈ ℚ.≤ f y y∈ + +isIncrasingℙ→injective : (P : ℙ ℚ) → (f : ∀ x → x ∈ P → ℚ) → + isIncrasingℙ P f → + ∀ x x∈ x' x'∈ + → f x x∈ ≡ f x' x'∈ → x ≡ x' +isIncrasingℙ→injective P f incrF x x∈ x' x'∈ = + ⊎.rec const (⊎.rec (w x x∈ x' x'∈) ((sym ∘_) ∘ (_∘ sym) ∘ w x' x'∈ x x∈)) + (ℚ.≡⊎# x x') + + where + w : ∀ x x∈ x' x'∈ → x ℚ.< x' → f x x∈ ≡ f x' x'∈ → x ≡ x' + w x x∈ x' x'∈ x : ∀ x → ⟨ openPred (λ y → (y <ᵣ x) , squash₁) ⟩ +openPred> : ∀ x → ⟨ openPred (λ y → (y <ᵣ x) , isProp<ᵣ _ _) ⟩ openPred> x y = PT.map (map-snd (λ {q} q x y = ∘S x : ℝ → ℙ ℝ +pred> x y = (x <ᵣ y) , isProp<ᵣ _ _ + +pred≤< : ℝ → ℝ → ℙ ℝ +pred≤< a b x = ((a ≤ᵣ x) × (x <ᵣ b)) , isProp× (isProp≤ᵣ _ _) (isProp<ᵣ _ _) + +pred<≤ : ℝ → ℝ → ℙ ℝ +pred<≤ a b x = ((a <ᵣ x) × (x ≤ᵣ b)) , isProp× (isProp<ᵣ _ _) (isProp≤ᵣ _ _) diff --git a/Cubical/HITs/CauchyReals/Sequence.agda b/Cubical/HITs/CauchyReals/Sequence.agda index fd7054d972..73840e0540 100644 --- a/Cubical/HITs/CauchyReals/Sequence.agda +++ b/Cubical/HITs/CauchyReals/Sequence.agda @@ -7,6 +7,12 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset + +import Cubical.Functions.Logic as L + open import Cubical.Data.Bool as 𝟚 hiding (_≤_) open import Cubical.Data.Nat as ℕ hiding (_·_;_+_) @@ -39,137 +45,86 @@ open import Cubical.HITs.CauchyReals.Multiplication open import Cubical.HITs.CauchyReals.Inverse +Lipschitz-ℚ→ℝℙ : ℚ₊ → (P : ℙ ℝ) → (∀ x → rat x ∈ P → ℝ) → Type +Lipschitz-ℚ→ℝℙ L P f = + (∀ q q∈ r r∈ → (ε : ℚ₊) → + ℚ.abs (q ℚ.- r) ℚ.< (fst ε) → absᵣ (f q q∈ -ᵣ f r r∈) + <ᵣ rat (fst (L ℚ₊· ε ))) +extend-Lipshitzℚ→ℝ : ∀ L → ∀ a b → (a ℚ.≤ b) → ∀ f → + Lipschitz-ℚ→ℝℙ L (intervalℙ (rat a) (rat b)) f → + Σ[ f' ∈ (ℚ → ℝ) ] + (Lipschitz-ℚ→ℝ L f' × (∀ x x∈ → f' x ≡ f x x∈ )) -Seq : Type -Seq = ℕ → ℝ - +extend-Lipshitzℚ→ℝ L a b a≤b f li = + (λ x → f (ℚ.clamp a b x) (∈ℚintervalℙ→∈intervalℙ _ _ _ + (clam∈ℚintervalℙ a b a≤b x))) , + w , (λ x x∈ → cong (uncurry f) + (Σ≡Prop (∈-isProp (intervalℙ (rat a) (rat b) ∘ rat)) + (ℚ.inClamps a b x (≤ᵣ→≤ℚ _ _ (fst x∈)) (≤ᵣ→≤ℚ _ _ (snd x∈)) ))) -≤ᵣ-·ᵣo : ∀ m n o → 0 ≤ᵣ o → m ≤ᵣ n → (m ·ᵣ o ) ≤ᵣ (n ·ᵣ o) -≤ᵣ-·ᵣo m n o 0≤o m≤n = subst (λ o → - (m ·ᵣ o ) ≤ᵣ (n ·ᵣ o)) 0≤o (w ∙ - cong (_·ᵣ maxᵣ (rat [ pos 0 / 1+ 0 ]) o) m≤n) where - w : maxᵣ (m ·ᵣ maxᵣ 0 o ) (n ·ᵣ maxᵣ 0 o) ≡ (maxᵣ m n ·ᵣ maxᵣ 0 o) - w = ≡Continuous _ _ - (cont₂maxᵣ _ _ - ((IsContinuous∘ _ _ - (IsContinuous·ᵣL m) (IsContinuousMaxL 0) )) - ((IsContinuous∘ _ _ - (IsContinuous·ᵣL n) (IsContinuousMaxL 0) ))) - (IsContinuous∘ _ _ - (IsContinuous·ᵣL _) (IsContinuousMaxL 0) ) - (λ o' → - ≡Continuous _ _ - (IsContinuous∘ _ _ (IsContinuousMaxR ((n ·ᵣ maxᵣ 0 (rat o')))) - (IsContinuous·ᵣR (maxᵣ 0 (rat o'))) ) - (IsContinuous∘ _ _ (IsContinuous·ᵣR (maxᵣ 0 (rat o'))) - (IsContinuousMaxR n)) - (λ m' → - ≡Continuous - (λ n → maxᵣ (rat m' ·ᵣ maxᵣ 0 (rat o') ) (n ·ᵣ maxᵣ 0 (rat o'))) - (λ n → maxᵣ (rat m') n ·ᵣ maxᵣ 0 (rat o')) - ((IsContinuous∘ _ _ - (IsContinuousMaxL (((rat m') ·ᵣ maxᵣ 0 (rat o')))) - (IsContinuous·ᵣR (maxᵣ 0 (rat o'))) )) - (IsContinuous∘ _ _ ((IsContinuous·ᵣR (maxᵣ 0 (rat o')))) - (IsContinuousMaxL (rat m'))) - (λ n' → - (cong₂ maxᵣ (sym (rat·ᵣrat _ _)) (sym (rat·ᵣrat _ _))) - ∙∙ cong rat (sym (ℚ.·MaxDistrℚ' m' n' (ℚ.max 0 o') - (ℚ.≤max 0 o'))) ∙∙ - rat·ᵣrat _ _ - ) n) m) o - -≤ᵣ-o·ᵣ : ∀ m n o → 0 ≤ᵣ o → m ≤ᵣ n → (o ·ᵣ m ) ≤ᵣ (o ·ᵣ n) -≤ᵣ-o·ᵣ m n o 0≤o p = subst2 _≤ᵣ_ - (·ᵣComm _ _) - (·ᵣComm _ _) - (≤ᵣ-·ᵣo m n o 0≤o p) - -≤ᵣ₊Monotone·ᵣ : ∀ m n o s → 0 ≤ᵣ n → 0 ≤ᵣ o - → m ≤ᵣ n → o ≤ᵣ s - → (m ·ᵣ o) ≤ᵣ (n ·ᵣ s) -≤ᵣ₊Monotone·ᵣ m n o s 0≤n 0≤o m≤n o≤s = - isTrans≤ᵣ _ _ _ (≤ᵣ-·ᵣo m n o 0≤o m≤n) - (≤ᵣ-o·ᵣ o s n 0≤n o≤s) - - - - -ℝ₊· : (m : ℝ₊) (n : ℝ₊) → 0 <ᵣ (fst m) ·ᵣ (fst n) -ℝ₊· (m , 0 1 --> Y + ↓ ↓ ↓ + 1 -> Σ X -> ? = B + -} + + module _ where + open PushoutPasteLeft + (rotatePushoutSquare (_ , SuspPushoutSquare ℓ-zero ℓ-zero X)) + {B = (Susp X , north) ⋁ (Y , y₀)} + (const y₀) inr inl (funExt (λ _ → push _)) + + squareL : commSquare + squareL = totSquare + + pushoutSquareL : PushoutSquare + pushoutSquareL = squareL , isPushoutRightSquare→isPushoutTotSquare + (snd (⋁-PushoutSquare _ _ ℓ-zero)) + + module _ where + open PushoutPasteDown + (_ , SuspPushoutSquare ℓ-zero ℓ-zero X) + {B = (Susp X , north) ⋁ (Y , y₀)} + (const y₀) inr inl (funExt (λ _ → sym (push _))) + + squareR : commSquare + squareR = totSquare + + pushoutSquareR : PushoutSquare + pushoutSquareR = squareR , isPushoutBottomSquare→isPushoutTotSquare + (snd (rotatePushoutSquare (⋁-PushoutSquare _ _ ℓ-zero))) + +module WedgePushout {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') where + + open 3x3-span + span : 3x3-span + span .A40 = typ A + span .A42 = typ A + span .A44 = typ A + span .A20 = Unit + span .A22 = Unit + span .A24 = Unit + span .A00 = typ B + span .A02 = Unit + span .A04 = typ C + span .f30 _ = pt A + span .f32 _ = pt A + span .f34 _ = pt A + span .f10 _ = pt B + span .f12 _ = _ + span .f14 _ = pt C + span .f41 = idfun (typ A) + span .f21 = _ + span .f01 _ = pt B + span .f43 = idfun (typ A) + span .f23 = _ + span .f03 _ = pt C + span .H11 _ = refl + span .H13 _ = refl + span .H31 _ = refl + span .H33 _ = refl + + A□2≃ : A□2 span ≃ typ A + A□2≃ = invEquiv (pushoutIdfunEquiv' _) + + f : typ A → B ⋁ A + f = inr + + g : typ A → C ⋁ A + g = inr + + A□○≃ : A□○ span ≃ Pushout f g + A□○≃ = pushoutEquiv _ _ _ _ + A□2≃ (idEquiv _) (idEquiv _) + (funExt λ + { (inl x) → push _ + ; (inr x) → refl + ; (push a i) j → sq (~ i) (~ j) }) + (funExt λ + { (inl x) → push _ + ; (inr x) → refl + ; (push a i) j → sq (~ i) (~ j) }) + where + sq : ∀ {ℓ} {C : Type ℓ} {c : C} + → Square refl (sym (Pushout.push _)) + refl (sym (push {f = λ x → c} {g = λ _ → pt A} + tt ∙ refl)) + sq = slideSquare (sym (rUnit _)) + + A4□≃ : A4□ span ≃ typ A + A4□≃ = invEquiv (pushoutIdfunEquiv _) + + A2□≃ : A2□ span ≃ Unit + A2□≃ = invEquiv (pushoutIdfunEquiv _) + + A2□isContr : isContr (A2□ span) + A2□isContr = isOfHLevelRespectEquiv 0 (pushoutIdfunEquiv _) isContrUnit + + A○□≃ : A○□ span ≃ (B ⋁∙ₗ C) ⋁ A + A○□≃ = pushoutEquiv _ _ _ _ + A2□≃ (idEquiv _) A4□≃ + (lemma A2□isContr refl) (lemma A2□isContr refl) + where + -- functions on contractible domains are determined by the basepoint + lemma : ∀ {ℓ ℓ'} + → {A : Type ℓ} {B : Type ℓ'} + → {f g : A → B} + → (c : isContr A) + → f (c .fst) ≡ g (c .fst) + → f ≡ g + lemma {f = f} {g} c p = funExt λ x + → cong f (sym (c .snd x)) + ∙ p + ∙ cong g (c .snd x) + + opaque + extendWedgeEq : Pushout f g ≃ ((B ⋁∙ₗ C) ⋁ A) + extendWedgeEq = invEquiv A□○≃ ∙ₑ pathToEquiv (3x3-lemma span) ∙ₑ A○□≃ + + wedgePushout : PushoutSquare + wedgePushout = extendPushoutSquare + (pushoutToSquare record { f1 = f ; f3 = g }) + extendWedgeEq + + +module _ {ℓ ℓ'} (X : Pointed ℓ) (Y : Pointed ℓ') where + open commSquare + open 3-span + + {- + X × Y ---> Y --------> 1 + ↓ A ↓ C ↓ + X ---> Σ(X∧Y) --> ΣY ∨ Σ(X∧Y) + ↓ B ↓ D ↓ + 1 -> ΣX ∨ Σ(X∧Y) --> ? = ΣX ∨ ΣY ∨ Σ(X∧Y) + -} + + pushoutA₀ : PushoutSquare -- inrP maps to south, but we want north + pushoutA₀ = extendPushoutSquare + (pushoutToSquare (3span fst snd)) + (pathToEquiv (joinPushout≡join _ _) ∙ₑ + invEquiv (isoToEquiv (SmashJoinIso {A = X} {B = Y}))) + + squareA : I → commSquare + squareA i = record + { commSquare (pushoutA₀ .fst) + ; inrP = λ _ → merid (inl _) i + ; comm = transp (λ j → + Path (typ X × typ Y → Susp (X ⋀ Y)) + (λ _ → north) (λ _ → merid (inl _) (i ∨ ~ j))) + i (pushoutA₀ .fst .commSquare.comm) + } + + pushoutA : PushoutSquare + pushoutA = squareA i0 , + subst isPushoutSquare (λ i → squareA (~ i)) (pushoutA₀ .snd) + + module _ where + open PushoutNull (typ X) (Susp (X ⋀ Y)) north + open PushoutPasteDown pushoutA + (squareL .sp .f1) + (squareL .inlP) + (squareL .inrP) + (squareL .commSquare.comm) + + pushoutB : PushoutSquare + pushoutB = pushoutSquareL + + pushoutAB : PushoutSquare + pushoutAB = _ , isPushoutBottomSquare→isPushoutTotSquare (pushoutB .snd) + + cofib-snd : cofib snd ≃ Susp∙ (typ X) ⋁ Susp∙ (X ⋀ Y) + cofib-snd = pushoutEquiv _ _ _ _ + (idEquiv _) Unit≃Unit* (idEquiv _) + refl refl + ∙ₑ (_ , pushoutAB .snd) + + pushoutC : PushoutSquare + pushoutC = PushoutNull.pushoutSquareR (typ Y) (Susp (X ⋀ Y)) north + + pushoutD : PushoutSquare + pushoutD = WedgePushout.wedgePushout + (Susp∙ (X ⋀ Y)) (Susp∙ (typ X)) (Susp∙ (typ Y)) + + pushoutCD : PushoutSquare + pushoutCD = _ , isPushoutBottomSquare→isPushoutTotSquare (pushoutD .snd) + where + open PushoutPasteDown pushoutC + (pushoutD .fst .sp .f1) + (pushoutD .fst .inlP) + (pushoutD .fst .inrP) + (pushoutD .fst .commSquare.comm) + + pushoutTot : PushoutSquare + pushoutTot = _ , isPushoutRightSquare→isPushoutTotSquare (pushoutCD .snd) + where + open PushoutPasteLeft pushoutAB + (pushoutCD .fst .sp .f3) + (pushoutCD .fst .inrP) + (pushoutCD .fst .inlP) + (pushoutCD .fst .commSquare.comm) + + SuspProduct : (Susp∙ (typ X) ⋁∙ₗ Susp∙ (typ Y)) ⋁ Susp∙ (X ⋀ Y) + ≃ Susp (typ X × typ Y) + SuspProduct = invEquiv (_ , pushoutTot .snd) + ∙ₑ (_ , SuspPushoutSquare ℓ-zero ℓ-zero (typ X × typ Y)) diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index 192c7aedcc..80c1c98a03 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -34,6 +34,14 @@ _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointe (f ∨→ g) (inr x) = fst g x (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ +⋁proj₁ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → A ⋁ B → typ A +⋁proj₁ A B = idfun∙ A ∨→ const∙ B A + +⋁proj₂ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → A ⋁ B → typ B +⋁proj₂ A B = const∙ A B ∨→ idfun∙ B + -- Pointed version ∨→∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → (f : A →∙ C) (g : B →∙ C) → ((A ⋁∙ₗ B) →∙ C) diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index c357e94d00..d2099fe4c4 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -12,7 +12,6 @@ open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Data.Sum as ⊎ -open import Cubical.HITs.Pushout.Base open import Cubical.HITs.Wedge.Base open import Cubical.HITs.Susp open import Cubical.HITs.Pushout @@ -44,6 +43,40 @@ Iso.inv ⋁-commIso = ⋁-commFun Iso.rightInv ⋁-commIso = ⋁-commFun² Iso.leftInv ⋁-commIso = ⋁-commFun² +-- Pushout square using Unit* for convenience +⋁-PushoutSquare : ∀ (A : Pointed ℓ) (B : Pointed ℓ') ℓ'' → PushoutSquare +⋁-PushoutSquare A B ℓ'' .fst = cSq + where + open commSquare + open 3-span + cSq : commSquare + cSq .sp .A0 = typ A + cSq .sp .A2 = Unit* {ℓ''} + cSq .sp .A4 = typ B + cSq .sp .f1 _ = pt A + cSq .sp .f3 _ = pt B + cSq .P = A ⋁ B + cSq .inlP = inl + cSq .inrP = inr + cSq .comm = funExt λ _ → push _ +⋁-PushoutSquare A B ℓ'' .snd = + isoToIsEquiv (iso _ inv lInv rInv) + where + inv : _ + inv (inl a) = inl a + inv (inr b) = inr b + inv (push _ i) = push _ i + + rInv : _ + rInv (inl a) = refl + rInv (inr b) = refl + rInv (push _ i) = refl + + lInv : _ + lInv (inl a) = refl + lInv (inr b) = refl + lInv (push _ i) = refl + -- cofibre of A --inl→ A ⋁ B is B cofibInl-⋁ : {A : Pointed ℓ} {B : Pointed ℓ'} → Iso (cofib {B = A ⋁ B} inl) (fst B) @@ -548,3 +581,70 @@ module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : ⋁-cofib-Iso = compIso (compIso (invIso A○□Iso) (invIso (3x3-Iso inst))) A□○Iso + +{- + We prove the square + X ⋁ Y --> X + ↓ ↓ + Y ----> * + is a pushout. +-} + +module _ (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') where + + open 3-span + open commSquare + + private + weirdSquare : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {Y : Type ℓ'} {Z : Type ℓ''} + → (f : X → Y) (g : Y → Z) → commSquare + weirdSquare f g .sp = 3span (idfun _) f + weirdSquare f g .P = _ + weirdSquare f g .inlP = f + weirdSquare f g .inrP = idfun _ + weirdSquare f g .comm = refl ∙ refl ∙ refl + + weirdPushoutSquare : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {Y : Type ℓ'} {Z : Type ℓ''} + → (f : X → Y) (g : Y → Z) → isPushoutSquare (weirdSquare f g) + weirdPushoutSquare f g = isoToIsEquiv (iso _ inr (λ _ → refl) + λ { (inl x) → sym (push _) + ; (inr x) → refl + ; (push a i) j → subst + (λ t → Square (sym (push _)) refl t (push _)) + (cong (cong Pushout.inr) (lUnit _ ∙ lUnit _)) + (λ i j → push a (i ∨ ~ j)) + i j + }) + + {- + The proof proceeds by applying the pasting lemma twice: + 1 ----> Y + ↓ ↓ + X --> X ⋁ Y --> X + ↓ mid ↓ bot ↓ + 1 ----> Y ----> 1 + -} + + midPushout : isPushoutSquare _ + midPushout = isPushoutTotSquare→isPushoutBottomSquare + (weirdPushoutSquare _ (terminal Y)) + where + open PushoutPasteDown + (pushoutToSquare (3span (const x₀) (const y₀))) + (terminal X) (const y₀) (⋁proj₂ X∙ Y∙) refl + + -- slight help to the unifier here + botPushout : isPushoutSquare record { comm = refl } + botPushout = isPushoutTotSquare→isPushoutBottomSquare $ + rotatePushoutSquare (record { comm = refl } , isoToIsEquiv + (iso _ inl (λ _ → refl) λ { + (inl _) i → inl _ + ; (inr a) i → push a i + ; (push a j) i → push a (i ∧ j) + })) .snd + where + open PushoutPasteDown (rotatePushoutSquare (_ , midPushout)) + (⋁proj₁ X∙ Y∙) (terminal X) (terminal Y) refl + + Pushout⋁≃Unit : Pushout (⋁proj₁ X∙ Y∙) (⋁proj₂ X∙ Y∙) ≃ Unit + Pushout⋁≃Unit = _ , botPushout diff --git a/Cubical/Homotopy/EilenbergMacLane/Properties.agda b/Cubical/Homotopy/EilenbergMacLane/Properties.agda index 953ce48309..1d613b4eaa 100644 --- a/Cubical/Homotopy/EilenbergMacLane/Properties.agda +++ b/Cubical/Homotopy/EilenbergMacLane/Properties.agda @@ -812,6 +812,25 @@ module _ where →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) (isOfHLevel↑∙∙'' n m l) +-- A homomorphism φ : G → H of (not necessarily abelian) groups induces a homomorphism +-- φ' : K(G,1) → K(H,1) + +inducedFun-EM₁ : {G : Group ℓ} {H : Group ℓ'} (f : GroupHom G H) → EM₁ G → EM₁ H +inducedFun-EM₁ f = EMrec _ emsquash embase (λ g → emloop (fst f g)) λ g h i j → hcomp (λ k → λ where + (i = i0) → emloop (f .fst g) j + (i = i1) → emloop (f .snd .pres· g h (~ k)) j + (j = i0) → embase + (j = i1) → emloop (f .fst h) i + ) (emcomp (f .fst g) (f. fst h) i j) + where open IsGroupHom + +inducedFun-EM₁-id : {G : Group ℓ} (x : EM₁ G) → inducedFun-EM₁ idGroupHom x ≡ x +inducedFun-EM₁-id = elimSet _ (λ _ → emsquash _ _) refl λ x i → refl + +inducedFun-EM₁-comp : {G : Group ℓ} {H : Group ℓ'} {L : Group ℓ''} (f : GroupHom G H) (g : GroupHom H L) (x : EM₁ G) + → inducedFun-EM₁ (compGroupHom f g) x ≡ inducedFun-EM₁ g (inducedFun-EM₁ f x) +inducedFun-EM₁-comp f g = elimSet _ (λ _ → emsquash _ _) refl λ x i → refl + -- A homomorphism φ : G → H of AbGroups induces a homomorphism -- φ' : K(G,n) → K(H,n) inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} @@ -1015,6 +1034,15 @@ Iso.leftInv (inducedFun-EM-rawIso e n) = h n ; south → refl ; (merid a i) k → merid (p a k) i} +Iso→EM₁Iso : {G : Group ℓ} {H : Group ℓ'} → GroupIso G H → Iso (EM₁ G) (EM₁ H) +Iso→EM₁Iso e .Iso.fun = inducedFun-EM₁ (GroupIso→GroupHom e) +Iso→EM₁Iso e .Iso.inv = inducedFun-EM₁ (GroupIso→GroupHom (invGroupIso e)) +Iso→EM₁Iso e .Iso.rightInv = elimSet _ (λ _ → emsquash _ _) refl λ x i j → emloop (e .fst .Iso.rightInv x j) i +Iso→EM₁Iso e .Iso.leftInv = elimSet _ (λ _ → emsquash _ _) refl λ x i j → emloop (e .fst .Iso.leftInv x j) i + +Equiv→EM₁Equiv : {G : Group ℓ} {H : Group ℓ'} → GroupEquiv G H → EM₁ G ≃ EM₁ H +Equiv→EM₁Equiv e = isoToEquiv (Iso→EM₁Iso (GroupEquiv→GroupIso e)) + module _ {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) where Iso→EMIso : ∀ n → Iso (EM G n) (EM H n) Iso.fun (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom e) n diff --git a/Cubical/Homotopy/HiltonMilnor.agda b/Cubical/Homotopy/HiltonMilnor.agda new file mode 100644 index 0000000000..425a5e7977 --- /dev/null +++ b/Cubical/Homotopy/HiltonMilnor.agda @@ -0,0 +1,96 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{- + The finitary Hilton–Milnor splitting + Ω(X ⋁ Y) ≃ ΩX × ΩY × ΩΣ(ΩY ⋀ ΩX) + for pointed types X and Y. +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism using (isoToPath; isoToEquiv) +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Pushout.Flattening +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge + +module Cubical.Homotopy.HiltonMilnor {ℓ} (X∙@(X , x₀) Y∙@(Y , y₀) : Pointed ℓ) where + +fun : X∙ ⋁ Y∙ → typ X∙ × typ Y∙ +fun = ⋁↪ {A = X∙} {B = Y∙} + +-- The required section, given by concatenating the paths +sect : typ (Ω (X∙ ×∙ Y∙)) → typ (Ω (X∙ ⋁∙ₗ Y∙)) +sect p = cong (inl ∘ fst) p ∙ push _ ∙ cong (inr ∘ snd) p ∙ sym (push _) + +rInv : ∀ c → cong fun (sect c) ≡ c +rInv c = congFunct fun _ _ + ∙ cong (cong (fun ∘ inl ∘ fst) c ∙_) + (congFunct fun (push _) _ + ∙ sym (lUnit _) + ∙ lemma₂ -- can't be inlined because lossy unification + ∙ sym (rUnit _)) + ∙ lemma₁ (cong fst c) (cong snd c) + where + lemma₁ : (p : x₀ ≡ x₀) (q : y₀ ≡ y₀) + → cong (_, y₀) p ∙ cong (x₀ ,_) q ≡ cong₂ _,_ p q + lemma₁ p q = cong ΣPathP (ΣPathP + (congFunct fst (cong (_, y₀) p) _ ∙ sym (rUnit _) , + congFunct snd (cong (_, y₀) p) _ ∙ sym (lUnit _))) + + lemma₂ : cong fun + (cong (inr ∘ snd) c ∙ sym (push _)) + ≡ cong (⋁↪ ∘ inr ∘ snd) c ∙ refl + lemma₂ = congFunct fun _ _ + +-- Now the splitting lemma gives us something in terms of the fiber of ⋁↪. +-- So we try to characterize the type. +module _ where + open FlatteningLemma + (λ (_ : Unit) → x₀) (λ _ → y₀) + (λ x → (x , y₀) ≡ (x₀ , y₀)) (λ y → (x₀ , y) ≡ (x₀ , y₀)) + (λ _ → idEquiv _) + + private + E≡ : ∀ u → E u ≡ (⋁↪ u ≡ (x₀ , y₀)) + E≡ (inl x) = refl + E≡ (inr x) = refl + E≡ (push a i) j = uaIdEquiv {A = typ (Ω (X∙ ×∙ Y∙))} j i + + Pushout≃ : Pushout Σf Σg ≃ join (typ (Ω Y∙)) (typ (Ω X∙)) + Pushout≃ = pushoutEquiv _ _ _ _ + (ΣUnit _ ∙ₑ invEquiv ΣPathP≃PathPΣ ∙ₑ Σ-swap-≃) + (Σ-cong-equiv-snd (λ _ → invEquiv ΣPathP≃PathPΣ ∙ₑ Σ-swap-≃) + ∙ₑ invEquiv Σ-assoc-≃ + ∙ₑ invEquiv (fiberProjEquiv X _ x₀)) + (Σ-cong-equiv-snd (λ _ → invEquiv ΣPathP≃PathPΣ) + ∙ₑ invEquiv Σ-assoc-≃ + ∙ₑ invEquiv (fiberProjEquiv Y _ y₀)) + (funExt λ _ → transportRefl _) + (funExt λ _ → transportRefl _) + ∙ₑ joinPushout≃join _ _ + + fiber⋁↪≃ : fiber fun (x₀ , y₀) ≃ join (typ (Ω Y∙)) (typ (Ω X∙)) + fiber⋁↪≃ = Σ-cong-equiv-snd (λ a → pathToEquiv (sym (E≡ a))) + ∙ₑ flatten ∙ₑ Pushout≃ + +HMSplit : typ (Ω (X∙ ⋁∙ₗ Y∙)) ≡ typ (Ω X∙) × typ (Ω Y∙) × typ (Ω (Susp∙ (Ω Y∙ ⋀ Ω X∙))) +HMSplit = splitting + ∙ cong₂ (λ A B∙ → A × typ (Ω B∙)) (sym ΣPathP≡PathPΣ) + (ua∙ fiber⋁↪≃ refl ∙ ΣPathP + (sym (isoToPath SmashJoinIso) , toPathP refl)) + ∙ ua Σ-assoc-≃ + where open Split (inl x₀) ⋁↪ sect rInv diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 36fef7ea1f..2f687d7aba 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -13,8 +13,10 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Function open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence open import Cubical.Functions.Morphism +open import Cubical.Functions.Fibration open import Cubical.Data.Nat open import Cubical.Data.Sigma @@ -460,3 +462,79 @@ EH-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) → π-comp (1 + n) p q ≡ π-comp (1 + n) q p EH-π n = elim2 (λ x y → isOfHLevelPath 2 isSetSetTrunc _ _) λ p q → cong ∣_∣₂ (EH n p q) + +{- + If A -> B -> C is a fiber sequence, and Ω B -> Ω C has a section, + then Ω B is the product Ω A × Ω C. + + https://gist.github.com/ecavallo/5b75313d36977c51ca3563de82506123 +-} + +module _ {ℓ} {B C : Type ℓ} (b₀ : B) (π : B → C) where + private + c₀ = π b₀ + A = fiber π c₀ + a₀ : A + a₀ = b₀ , refl + ι : A → B + ι = fst + π∘ι≡0 : ∀ a → π (ι a) ≡ c₀ + π∘ι≡0 = snd + + A∙ B∙ C∙ : Pointed _ + A∙ = (A , a₀) + B∙ = (B , b₀) + C∙ = (C , c₀) + + TwistedProduct : Pointed _ + TwistedProduct = ((Σ[ p ∈ typ (Ω C∙) ] (b₀ , p) ≡ a₀) , refl , refl) + + -- The middle term of a fiber sequence always splits + -- as a twisted product without any conditions. + presplit : Ω B∙ ≡ TwistedProduct + presplit = cong Ω (ua∙ (totalEquiv π) refl) + ∙ sym (ua∙ ΣPath≃PathΣ refl) + ∙ Σ-cong-equiv-snd∙ (λ p + → PathP≃Path _ _ _ ∙ₑ + compPathlEquiv (sym (fromPathP + (ΣPathP (refl , λ i j → p (i ∧ j)))))) + (rCancel λ j → transp (λ _ → A) (~ j) a₀) + + module Split (s : typ (Ω C∙) → typ (Ω B∙)) + (section : ∀ p → cong π (s p) ≡ p) where + movePoint : ∀ p → a₀ ≡ (b₀ , sym p) + movePoint p = ΣPathP (s p , + λ i j → slideSquare (section p) i (~ j)) + + splitting : typ (Ω B∙) ≡ typ (Ω C∙) × typ (Ω A∙) + splitting = cong typ presplit ∙ ua + (Σ-cong-equiv-snd λ p → compPathlEquiv (movePoint _)) + + -- If the section furthermore respects the base point, then + -- the splitting does too. + module _ (h : s refl ≡ refl) (coh : cong (cong π) h ≡ section refl) where + movePoint-refl : movePoint refl ≡ refl + movePoint-refl i j = h i j , λ k → + slideSquare (cong sym coh ∙ lemma _ _) (~ i) (~ k) j + where + lemma : ∀ {ℓ} {X : Type ℓ} {x : X} (p : x ≡ x) (q : p ≡ refl) + → sym q ≡ flipSquare (slideSquare q) + lemma {x = x} p q = J (λ p (q : refl ≡ p) + → q ≡ flipSquare (slideSquare (sym q))) + (λ i j k → hfill + (slideSquareFaces + {a₀₀ = x} {a₋ = refl} {a₁₋ = refl} {a₋₀ = refl} + j k) (inS x) i) + (sym q) + + twisted∙ : TwistedProduct ≡ Ω C∙ ×∙ Ω A∙ + twisted∙ = Σ-cong-equiv-snd∙ + (λ p → compPathlEquiv (movePoint _)) + (cong (λ t → compPathlEquiv t .fst refl) movePoint-refl ∙ sym (rUnit _)) + + splitting∙ : Ω B∙ ≡ Ω C∙ ×∙ Ω A∙ + splitting∙ = presplit ∙ twisted∙ + + -- splitting and splitting∙ agrees + splitting-typ : cong typ splitting∙ ≡ splitting + splitting-typ = congFunct typ presplit twisted∙ diff --git a/Cubical/README.agda b/Cubical/README.agda index bfe8dee179..af375d0d6e 100644 --- a/Cubical/README.agda +++ b/Cubical/README.agda @@ -4,7 +4,7 @@ module Cubical.README where ------------------------------------------------------------------------ -- An experimental library for Cubical Agda ------------------------------------------------------------------------ +------------------------------------------------------------------------ -- The library comes with a .agda-lib file, for use with the library -- management system. diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda index 07d31701af..b6e658f74b 100644 --- a/Cubical/Relation/Binary/Base.agda +++ b/Cubical/Relation/Binary/Base.agda @@ -77,19 +77,16 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where -- Sum types don't play nicely with props, so we truncate isCotrans : Type (ℓ-max ℓ ℓ') - isCotrans = (a b c : A) → R a b → (R a c ⊔′ R b c) + isCotrans = (a b c : A) → R a b → R a c ⊔′ R b c isWeaklyLinear : Type (ℓ-max ℓ ℓ') isWeaklyLinear = (a b c : A) → R a b → R a c ⊔′ R c b isConnected : Type (ℓ-max ℓ ℓ') - isConnected = (a b : A) → ¬ (a ≡ b) → R a b ⊔′ R b a + isConnected = (a b : A) → (¬ R a b) × (¬ R b a) → a ≡ b - isStronglyConnected : Type (ℓ-max ℓ ℓ') - isStronglyConnected = (a b : A) → R a b ⊔′ R b a - - isStronglyConnected→isConnected : isStronglyConnected → isConnected - isStronglyConnected→isConnected strong a b _ = strong a b + isTotal : Type (ℓ-max ℓ ℓ') + isTotal = (a b : A) → R a b ⊔′ R b a isIrrefl×isTrans→isAsym : isIrrefl × isTrans → isAsym isIrrefl×isTrans→isAsym (irrefl , trans) a₀ a₁ Ra₀a₁ Ra₁a₀ @@ -98,6 +95,9 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where WellFounded→isIrrefl : WellFounded R → isIrrefl WellFounded→isIrrefl well = WFI.induction well λ a f Raa → f a Raa Raa + isAsym→isIrrefl : isAsym → isIrrefl + isAsym→isIrrefl asym a Raa = asym a a Raa Raa + IrreflKernel : Rel A A (ℓ-max ℓ ℓ') IrreflKernel a b = R a b × (¬ a ≡ b) @@ -116,6 +116,9 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where NegationRel : Rel A A ℓ' NegationRel a b = ¬ (R a b) + Dual : Rel A A ℓ' + Dual a b = R b a + module _ {ℓ'' : Level} (P : Embedding A ℓ'') @@ -154,10 +157,15 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where isEffective = (a b : A) → isEquiv (eq/ {R = R} a b) + isDecidable : Type (ℓ-max ℓ ℓ') + isDecidable = (a b : A) → Dec (R a b) impliesIdentity : Type _ impliesIdentity = {a a' : A} → (R a a') → (a ≡ a') + isTight : Type _ + isTight = (a b : A) → ¬ R a b → a ≡ b + inequalityImplies : Type _ inequalityImplies = (a b : A) → ¬ a ≡ b → R a b @@ -241,14 +249,6 @@ isIrreflIrreflKernel _ _ (_ , ¬a≡a) = ¬a≡a refl isReflReflClosure : ∀{ℓ ℓ'} {A : Type ℓ} (R : Rel A A ℓ') → isRefl (ReflClosure R) isReflReflClosure _ _ = inr refl -isConnectedStronglyConnectedIrreflKernel : ∀{ℓ ℓ'} {A : Type ℓ} (R : Rel A A ℓ') - → isStronglyConnected R - → isConnected (IrreflKernel R) -isConnectedStronglyConnectedIrreflKernel R strong a b ¬a≡b - = ∥₁.map (λ x → ⊎.rec (λ Rab → inl (Rab , ¬a≡b)) - (λ Rba → inr (Rba , (λ b≡a → ¬a≡b (sym b≡a)))) x) - (strong a b) - isSymSymKernel : ∀{ℓ ℓ'} {A : Type ℓ} (R : Rel A A ℓ') → isSym (SymKernel R) isSymSymKernel _ _ _ (Rab , Rba) = Rba , Rab diff --git a/Cubical/Relation/Binary/Order.agda b/Cubical/Relation/Binary/Order.agda index 9c15bb7931..7b777b96d8 100644 --- a/Cubical/Relation/Binary/Order.agda +++ b/Cubical/Relation/Binary/Order.agda @@ -2,9 +2,9 @@ module Cubical.Relation.Binary.Order where open import Cubical.Relation.Binary.Order.Apartness public -open import Cubical.Relation.Binary.Order.Preorder public +open import Cubical.Relation.Binary.Order.Proset public open import Cubical.Relation.Binary.Order.Poset public open import Cubical.Relation.Binary.Order.Toset public -open import Cubical.Relation.Binary.Order.StrictPoset public +open import Cubical.Relation.Binary.Order.Quoset public +open import Cubical.Relation.Binary.Order.StrictOrder public open import Cubical.Relation.Binary.Order.Loset public -open import Cubical.Relation.Binary.Order.Properties public diff --git a/Cubical/Relation/Binary/Order/Apartness/Base.agda b/Cubical/Relation/Binary/Order/Apartness/Base.agda index 35dc95528f..5801016a17 100644 --- a/Cubical/Relation/Binary/Order/Apartness/Base.agda +++ b/Cubical/Relation/Binary/Order/Apartness/Base.agda @@ -66,8 +66,8 @@ record ApartnessStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc Apartness : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) Apartness ℓ ℓ' = TypeWithStr ℓ (ApartnessStr ℓ') -apartness : (A : Type ℓ) (_#_ : A → A → Type ℓ') (h : IsApartness _#_) → Apartness ℓ ℓ' -apartness A _#_ h = A , apartnessstr _#_ h +apartness : (A : Type ℓ) → (_#_ : Rel A A ℓ') → IsApartness _#_ → Apartness ℓ ℓ' +apartness A _#_ apart = A , (apartnessstr _#_ apart) record IsApartnessEquiv {A : Type ℓ₀} {B : Type ℓ₁} (M : ApartnessStr ℓ₀' A) (e : A ≃ B) (N : ApartnessStr ℓ₁' B) diff --git a/Cubical/Relation/Binary/Order/Apartness/Properties.agda b/Cubical/Relation/Binary/Order/Apartness/Properties.agda index 77f24ce04f..c82e5d77fc 100644 --- a/Cubical/Relation/Binary/Order/Apartness/Properties.agda +++ b/Cubical/Relation/Binary/Order/Apartness/Properties.agda @@ -1,6 +1,7 @@ {-# OPTIONS --safe #-} module Cubical.Relation.Binary.Order.Apartness.Properties where +open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude open import Cubical.Functions.Embedding @@ -51,3 +52,20 @@ module _ (λ a → IsApartness.is-irrefl apa (f a)) (λ a b c → IsApartness.is-cotrans apa (f a) (f b) (f c)) λ a b → IsApartness.is-sym apa (f a) (f b) + + isApartnessDual : IsApartness R → IsApartness (Dual R) + isApartnessDual apa + = isapartness (IsApartness.is-set apa) + (λ a b → IsApartness.is-prop-valued apa b a) + (IsApartness.is-irrefl apa) + (λ a b c Rba → ∥₁.map (⊎.rec (inl ∘ s c a) + (inr ∘ s c b)) + (IsApartness.is-cotrans apa a b c (s a b Rba))) + s + where s : isSym (Dual R) + s a b = IsApartness.is-sym apa b a + +DualApartness : Apartness ℓ ℓ' → Apartness ℓ ℓ' +DualApartness (_ , apa) + = apartness _ (BinaryRelation.Dual (ApartnessStr._#_ apa)) + (isApartnessDual (ApartnessStr.isApartness apa)) diff --git a/Cubical/Relation/Binary/Order/Loset/Base.agda b/Cubical/Relation/Binary/Order/Loset/Base.agda index 0f2d3bd941..ee0fc5f99d 100644 --- a/Cubical/Relation/Binary/Order/Loset/Base.agda +++ b/Cubical/Relation/Binary/Order/Loset/Base.agda @@ -69,8 +69,8 @@ record LosetStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc ℓ') Loset : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) Loset ℓ ℓ' = TypeWithStr ℓ (LosetStr ℓ') -loset : (A : Type ℓ) (_<_ : A → A → Type ℓ') (h : IsLoset _<_) → Loset ℓ ℓ' -loset A _<_ h = A , losetstr _<_ h +loset : (A : Type ℓ) → (_<_ : Rel A A ℓ') → IsLoset _<_ → Loset ℓ ℓ' +loset A _<_ los = A , (losetstr _<_ los) record IsLosetEquiv {A : Type ℓ₀} {B : Type ℓ₁} (M : LosetStr ℓ₀' A) (e : A ≃ B) (N : LosetStr ℓ₁' B) @@ -98,7 +98,7 @@ isPropIsLoset _<_ = isOfHLevelRetractFromIso 1 IsLosetIsoΣ (isPropΠ5 (λ _ _ _ _ _ → isPropValued< _ _)) (isPropΠ3 (λ x y _ → isProp¬ (y < x))) (isPropΠ4 λ _ _ _ _ → squash₁) - (isPropΠ3 λ _ _ _ → squash₁)) + (isPropΠ3 λ _ _ _ → isSetA _ _)) 𝒮ᴰ-Loset : DUARel (𝒮-Univ ℓ) (LosetStr ℓ') (ℓ-max ℓ ℓ') 𝒮ᴰ-Loset = diff --git a/Cubical/Relation/Binary/Order/Loset/Properties.agda b/Cubical/Relation/Binary/Order/Loset/Properties.agda index 34e0e14e72..943f105a63 100644 --- a/Cubical/Relation/Binary/Order/Loset/Properties.agda +++ b/Cubical/Relation/Binary/Order/Loset/Properties.agda @@ -12,8 +12,9 @@ open import Cubical.HITs.PropositionalTruncation as ∥₁ open import Cubical.Relation.Binary.Base open import Cubical.Relation.Binary.Order.Apartness.Base -open import Cubical.Relation.Binary.Order.Toset.Base -open import Cubical.Relation.Binary.Order.StrictPoset.Base +open import Cubical.Relation.Binary.Order.Toset +open import Cubical.Relation.Binary.Order.Poset.Base +open import Cubical.Relation.Binary.Order.StrictOrder.Base open import Cubical.Relation.Binary.Order.Loset.Base open import Cubical.Relation.Nullary @@ -29,13 +30,14 @@ module _ open BinaryRelation - isLoset→isStrictPoset : IsLoset R → IsStrictPoset R - isLoset→isStrictPoset loset = isstrictposet - (IsLoset.is-set loset) - (IsLoset.is-prop-valued loset) - (IsLoset.is-irrefl loset) - (IsLoset.is-trans loset) - (IsLoset.is-asym loset) + isLoset→isStrictOrder : IsLoset R → IsStrictOrder R + isLoset→isStrictOrder loset = isstrictorder + (IsLoset.is-set loset) + (IsLoset.is-prop-valued loset) + (IsLoset.is-irrefl loset) + (IsLoset.is-trans loset) + (IsLoset.is-asym loset) + (IsLoset.is-weakly-linear loset) private transrefl : isTrans R → isTrans (ReflClosure R) @@ -49,8 +51,8 @@ module _ antisym irr trans a b (inl _) (inr b≡a) = sym b≡a antisym irr trans a b (inr a≡b) _ = a≡b - isLoset→isTosetReflClosure : Discrete A → IsLoset R → IsToset (ReflClosure R) - isLoset→isTosetReflClosure disc loset + isLoset→isTosetReflClosure : IsLoset R → isDecidable R → IsToset (ReflClosure R) + isLoset→isTosetReflClosure loset dec = istoset (IsLoset.is-set loset) (λ a b → isProp⊎ (IsLoset.is-prop-valued loset a b) (IsLoset.is-set loset a b) @@ -61,9 +63,33 @@ module _ (transrefl (IsLoset.is-trans loset)) (antisym (IsLoset.is-irrefl loset) (IsLoset.is-trans loset)) - λ a b → decRec (λ a≡b → ∣ inl (inr a≡b) ∣₁) - (λ ¬a≡b → ∥₁.map (⊎.map (λ Rab → inl Rab) λ Rba → inl Rba) - (IsLoset.is-connected loset a b ¬a≡b)) (disc a b) + λ a b → decRec (λ Rab → ∣ inl (inl Rab) ∣₁) + (λ ¬Rab → decRec (λ Rba → ∣ inr (inl Rba) ∣₁) + (λ ¬Rba → ∣ inl (inr (IsLoset.is-connected loset a b + (¬Rab , ¬Rba))) ∣₁) + (dec b a)) + (dec a b) + + isLosetDecidable→isTosetDecidable : IsLoset R → isDecidable R → isDecidable (ReflClosure R) + isLosetDecidable→isTosetDecidable los dec a b with dec a b + ... | yes Rab = yes (inl Rab) + ... | no ¬Rab with dec b a + ... | yes Rba = no (⊎.rec ¬Rab λ a≡b → IsLoset.is-irrefl los b (subst (R b) a≡b Rba)) + ... | no ¬Rba = yes (inr (IsLoset.is-connected los a b (¬Rab , ¬Rba))) + + isLosetDecidable→Discrete : IsLoset R → isDecidable R → Discrete A + isLosetDecidable→Discrete los dec = isTosetDecidable→Discrete + (isLoset→isTosetReflClosure los dec) + (isLosetDecidable→isTosetDecidable los dec) + + isLoset→isPosetNegationRel : IsLoset R → IsPoset (NegationRel R) + isLoset→isPosetNegationRel loset + = isposet (IsLoset.is-set loset) + (λ a b → isProp¬ (R a b)) + (IsLoset.is-irrefl loset) + (λ a b c ¬Rab ¬Rbc Rac → ∥₁.rec isProp⊥ (⊎.rec ¬Rab ¬Rbc) + (IsLoset.is-weakly-linear loset a c b Rac)) + λ a b ¬Rab ¬Rba → IsLoset.is-connected loset a b (¬Rab , ¬Rba) isLosetInduced : IsLoset R → (B : Type ℓ'') → (f : B ↪ A) → IsLoset (InducedRelation R (B , f)) @@ -74,18 +100,33 @@ module _ (λ a b c → IsLoset.is-trans los (f a) (f b) (f c)) (λ a b → IsLoset.is-asym los (f a) (f b)) (λ a b c → IsLoset.is-weakly-linear los (f a) (f b) (f c)) - λ a b ¬a≡b → IsLoset.is-connected los (f a) (f b) - λ fa≡fb → ¬a≡b (isEmbedding→Inj emb a b fa≡fb) - -Loset→StrictPoset : Loset ℓ ℓ' → StrictPoset ℓ ℓ' -Loset→StrictPoset (_ , los) - = _ , strictposetstr (LosetStr._<_ los) - (isLoset→isStrictPoset (LosetStr.isLoset los)) + λ a b ineq → isEmbedding→Inj emb a b + (IsLoset.is-connected los (f a) (f b) ineq) + + isLosetDual : IsLoset R → IsLoset (Dual R) + isLosetDual los + = isloset (IsLoset.is-set los) + (λ a b → IsLoset.is-prop-valued los b a) + (IsLoset.is-irrefl los) + (λ a b c Rab Rbc → IsLoset.is-trans los c b a Rbc Rab) + (λ a b → IsLoset.is-asym los b a) + (λ a b c Rba → ∥₁.map (⊎.rec inr inl) + (IsLoset.is-weakly-linear los b a c Rba)) + λ { a b (¬Rba , ¬Rab) → IsLoset.is-connected los a b (¬Rab , ¬Rba) } + +Loset→StrictOrder : Loset ℓ ℓ' → StrictOrder ℓ ℓ' +Loset→StrictOrder (_ , los) + = strictorder _ (LosetStr._<_ los) + (isLoset→isStrictOrder (LosetStr.isLoset los)) Loset→Toset : (los : Loset ℓ ℓ') - → Discrete (fst los) + → BinaryRelation.isDecidable (LosetStr._<_ (snd los)) → Toset ℓ (ℓ-max ℓ ℓ') -Loset→Toset (_ , los) disc - = _ , tosetstr (BinaryRelation.ReflClosure (LosetStr._<_ los)) - (isLoset→isTosetReflClosure disc - (LosetStr.isLoset los)) +Loset→Toset (_ , los) dec + = toset _ (BinaryRelation.ReflClosure (LosetStr._<_ los)) + (isLoset→isTosetReflClosure (LosetStr.isLoset los) dec) + +DualLoset : Loset ℓ ℓ' → Loset ℓ ℓ' +DualLoset (_ , los) + = loset _ (BinaryRelation.Dual (LosetStr._<_ los)) + (isLosetDual (LosetStr.isLoset los)) diff --git a/Cubical/Relation/Binary/Order/Poset/Base.agda b/Cubical/Relation/Binary/Order/Poset/Base.agda index 8d6fddd9e4..35ae0fef5c 100644 --- a/Cubical/Relation/Binary/Order/Poset/Base.agda +++ b/Cubical/Relation/Binary/Order/Poset/Base.agda @@ -59,8 +59,8 @@ record PosetStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc ℓ') Poset : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) Poset ℓ ℓ' = TypeWithStr ℓ (PosetStr ℓ') -poset : (A : Type ℓ) (_≤_ : A → A → Type ℓ') (h : IsPoset _≤_) → Poset ℓ ℓ' -poset A _≤_ h = A , posetstr _≤_ h +poset : (A : Type ℓ) → (_≤_ : Rel A A ℓ') → IsPoset _≤_ → Poset ℓ ℓ' +poset A _≤_ pos = A , (posetstr _≤_ pos) record IsPosetEquiv {A : Type ℓ₀} {B : Type ℓ₁} (M : PosetStr ℓ₀' A) (e : A ≃ B) (N : PosetStr ℓ₁' B) @@ -76,6 +76,12 @@ record IsPosetEquiv {A : Type ℓ₀} {B : Type ℓ₁} field pres≤ : (x y : A) → x M.≤ y ≃ equivFun e x N.≤ equivFun e y + -- This also holds in the other direction, which helps a lot in proofs + pres≤⁻ : (x y : B) → x N.≤ y ≃ invEq e x M.≤ invEq e y + pres≤⁻ x y = invEquiv (compEquiv (pres≤ (invEq e x) (invEq e y)) + (subst2Equiv N._≤_ (secEq e x) (secEq e y))) + +unquoteDecl IsPosetEquivIsoΣ = declareRecordIsoΣ IsPosetEquivIsoΣ (quote IsPosetEquiv) PosetEquiv : (M : Poset ℓ₀ ℓ₀') (M : Poset ℓ₁ ℓ₁') → Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) PosetEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsPosetEquiv (M .snd) e (N .snd) @@ -89,6 +95,16 @@ isPropIsPoset _≤_ = isOfHLevelRetractFromIso 1 IsPosetIsoΣ (isPropΠ5 λ _ _ _ _ _ → isPropValued≤ _ _) (isPropΠ4 λ _ _ _ _ → isSetA _ _)) +isPropIsPosetEquiv : {A : Type ℓ₀} {B : Type ℓ₁} + (M : PosetStr ℓ₀' A) + (e : A ≃ B) + (N : PosetStr ℓ₁' B) + → isProp (IsPosetEquiv M e N) +isPropIsPosetEquiv M e N = isOfHLevelRetractFromIso 1 IsPosetEquivIsoΣ + (isPropΠ2 λ _ _ → isOfHLevel≃ 1 + (IsPoset.is-prop-valued (PosetStr.isPoset M) _ _) + (IsPoset.is-prop-valued (PosetStr.isPoset N) _ _)) + 𝒮ᴰ-Poset : DUARel (𝒮-Univ ℓ) (PosetStr ℓ') (ℓ-max ℓ ℓ') 𝒮ᴰ-Poset = 𝒮ᴰ-Record (𝒮-Univ _) IsPosetEquiv diff --git a/Cubical/Relation/Binary/Order/Poset/Instances/Embedding.agda b/Cubical/Relation/Binary/Order/Poset/Instances/Embedding.agda new file mode 100644 index 0000000000..648cc67fd4 --- /dev/null +++ b/Cubical/Relation/Binary/Order/Poset/Instances/Embedding.agda @@ -0,0 +1,145 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Order.Poset.Instances.Embedding where + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Functions.Embedding + +open import Cubical.HITs.PropositionalTruncation as ∥₁ + +open import Cubical.Relation.Binary.Order.Proset.Properties +open import Cubical.Relation.Binary.Order.Poset.Base +open import Cubical.Relation.Binary.Order.Poset.Properties + +private + variable + ℓ ℓ' : Level + +-- The collection of embeddings on a type happens to form a complete lattice + +-- First, we show that it's a poset +isPoset⊆ₑ : {A : Type ℓ} → IsPoset {A = Embedding A ℓ'} _⊆ₑ_ +isPoset⊆ₑ = isposet isSetEmbedding + isProp⊆ₑ + isRefl⊆ₑ + isTrans⊆ₑ + isAntisym⊆ₑ + +EmbeddingPoset : (A : Type ℓ) (ℓ' : Level) → Poset _ _ +EmbeddingPoset A ℓ' + = poset (Embedding A ℓ') _⊆ₑ_ (isPoset⊆ₑ) + +-- Then we describe it as both a lattice and a complete lattice +isMeet∩ₑ : {A : Type ℓ} + (X Y : Embedding A ℓ) + → isMeet (isPoset→isProset isPoset⊆ₑ) X Y (X ∩ₑ Y) +isMeet∩ₑ X Y Z = propBiimpl→Equiv (isProp⊆ₑ Z (X ∩ₑ Y)) + (isProp× (isProp⊆ₑ Z X) + (isProp⊆ₑ Z Y)) + (λ Z⊆X∩Y → + (λ x x∈Z → + subst (_∈ₑ X) + (Z⊆X∩Y x x∈Z .snd) + (Z⊆X∩Y x x∈Z .fst .snd .fst)) , + (λ x x∈Z → + subst (_∈ₑ Y) + (Z⊆X∩Y x x∈Z .snd) + (Z⊆X∩Y x x∈Z .fst .snd .snd))) + λ { (Z⊆X , Z⊆Y) x x∈Z → + (x , ((Z⊆X x x∈Z) , + (Z⊆Y x x∈Z))) , refl } + +isMeetSemipseudolatticeEmbeddingPoset : {A : Type ℓ} + → isMeetSemipseudolattice (EmbeddingPoset A ℓ) +isMeetSemipseudolatticeEmbeddingPoset X Y + = X ∩ₑ Y , isMeet∩ₑ X Y + +isGreatestEmbeddingPosetTotal : {A : Type ℓ} + → isGreatest (isPoset→isProset isPoset⊆ₑ) + (Embedding A ℓ , id↪ (Embedding A ℓ)) + (A , (id↪ A)) +isGreatestEmbeddingPosetTotal _ x _ + = x , refl + +isMeetSemilatticeEmbeddingPoset : {A : Type ℓ} + → isMeetSemilattice (EmbeddingPoset A ℓ) +isMeetSemilatticeEmbeddingPoset {A = A} + = isMeetSemipseudolatticeEmbeddingPoset , + (A , (id↪ A)) , + isGreatestEmbeddingPosetTotal + +isJoin∪ₑ : {A : Type ℓ} + (X Y : Embedding A ℓ) + → isJoin (isPoset→isProset isPoset⊆ₑ) X Y (X ∪ₑ Y) +isJoin∪ₑ X Y Z + = propBiimpl→Equiv (isProp⊆ₑ (X ∪ₑ Y) Z) + (isProp× (isProp⊆ₑ X Z) + (isProp⊆ₑ Y Z)) + (λ X∪Y⊆Z → + (λ x x∈X → X∪Y⊆Z x ((x , ∣ inl x∈X ∣₁) , refl)) , + (λ x x∈Y → X∪Y⊆Z x ((x , ∣ inr x∈Y ∣₁) , refl))) + λ { (X⊆Z , Y⊆Z) x x∈X∪Y → + ∥₁.rec (isProp∈ₑ x Z) + (⊎.rec (λ a∈X → subst (_∈ₑ Z) + (x∈X∪Y .snd) + (X⊆Z _ a∈X)) + (λ a∈Y → subst (_∈ₑ Z) + (x∈X∪Y .snd) + (Y⊆Z _ a∈Y))) + (x∈X∪Y .fst .snd) } + +isJoinSemipseudolatticeEmbeddingPoset : {A : Type ℓ} + → isJoinSemipseudolattice (EmbeddingPoset A ℓ) +isJoinSemipseudolatticeEmbeddingPoset X Y + = X ∪ₑ Y , isJoin∪ₑ X Y + +isLeast∅ : {A : Type ℓ} + → isLeast (isPoset→isProset isPoset⊆ₑ) (Embedding A ℓ , id↪ (Embedding A ℓ)) ((Σ[ x ∈ A ] ⊥) , EmbeddingΣProp λ _ → isProp⊥) +isLeast∅ _ _ ((_ , ()) , _) + +isJoinSemilatticeEmbeddingPoset : {A : Type ℓ} + → isJoinSemilattice (EmbeddingPoset A ℓ) +isJoinSemilatticeEmbeddingPoset {A = A} + = isJoinSemipseudolatticeEmbeddingPoset , + ((Σ[ x ∈ A ] ⊥) , EmbeddingΣProp λ _ → isProp⊥) , + isLeast∅ + +isLatticeEmbeddingPoset : {A : Type ℓ} + → isLattice (EmbeddingPoset A ℓ) +isLatticeEmbeddingPoset = isMeetSemilatticeEmbeddingPoset , + isJoinSemilatticeEmbeddingPoset + +isInfimum⋂ₑ : {A : Type ℓ} + {I : Type ℓ} + (P : I → Embedding A ℓ) + → isInfimum (isPoset→isProset isPoset⊆ₑ) P (⋂ₑ P) +fst (isInfimum⋂ₑ P) i y ((a , ∀i) , a≡y) = subst (_∈ₑ P i) a≡y (∀i i) +snd (isInfimum⋂ₑ P) (X , lwr) y y∈X = (y , λ i → lwr i y y∈X) , refl + +isMeetCompleteSemipseudolatticeEmbeddingPoset : {A : Type ℓ} + → isMeetCompleteSemipseudolattice (EmbeddingPoset A ℓ) +isMeetCompleteSemipseudolatticeEmbeddingPoset P + = (⋂ₑ P) , (isInfimum⋂ₑ P) + +isSupremum⋃ₑ : {A : Type ℓ} + {I : Type ℓ} + (P : I → Embedding A ℓ) + → isSupremum (isPoset→isProset isPoset⊆ₑ) P (⋃ₑ P) +fst (isSupremum⋃ₑ P) i y y∈Pi = (y , ∣ i , y∈Pi ∣₁) , refl +snd (isSupremum⋃ₑ P) (X , upr) y ((a , ∃i) , a≡y) + = ∥₁.rec (isProp∈ₑ y X) (λ (i , a∈Pi) → upr i y (subst (_∈ₑ P i) a≡y a∈Pi)) ∃i + +isJoinCompleteSemipseudolatticeEmbeddingPoset : {A : Type ℓ} + → isJoinCompleteSemipseudolattice (EmbeddingPoset A ℓ) +isJoinCompleteSemipseudolatticeEmbeddingPoset P + = (⋃ₑ P) , (isSupremum⋃ₑ P) + +isCompleteLatticeEmbeddingPoset : {A : Type ℓ} + → isCompleteLattice (EmbeddingPoset A ℓ) +isCompleteLatticeEmbeddingPoset = isMeetCompleteSemipseudolatticeEmbeddingPoset , + isJoinCompleteSemipseudolatticeEmbeddingPoset diff --git a/Cubical/Relation/Binary/Order/Poset/Mappings.agda b/Cubical/Relation/Binary/Order/Poset/Mappings.agda new file mode 100644 index 0000000000..161f366072 --- /dev/null +++ b/Cubical/Relation/Binary/Order/Poset/Mappings.agda @@ -0,0 +1,1228 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Order.Poset.Mappings where + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence + +open import Cubical.Algebra.Semigroup + +open import Cubical.Data.Sigma + +open import Cubical.Functions.Embedding +open import Cubical.Functions.Image +open import Cubical.Functions.Logic using (_⊔′_) +open import Cubical.Functions.Preimage + +open import Cubical.Reflection.RecordEquiv + +open import Cubical.HITs.PropositionalTruncation as ∥₁ +open import Cubical.HITs.SetQuotients + +open import Cubical.Relation.Binary.Order.Poset.Base +open import Cubical.Relation.Binary.Order.Poset.Properties +open import Cubical.Relation.Binary.Order.Poset.Subset +open import Cubical.Relation.Binary.Order.Poset.Instances.Embedding +open import Cubical.Relation.Binary.Order.Proset +open import Cubical.Relation.Binary.Base + + +private + variable + ℓ ℓ' ℓ'' ℓ''' ℓ₀ ℓ₀' ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +record IsIsotone {A : Type ℓ₀} {B : Type ℓ₁} + (M : PosetStr ℓ₀' A) (f : A → B) (N : PosetStr ℓ₁' B) + : Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') ℓ₁') + where + -- Shorter qualified names + private + module M = PosetStr M + module N = PosetStr N + + field + pres≤ : (x y : A) → x M.≤ y → f x N.≤ f y + +unquoteDecl IsIsotoneIsoΣ = declareRecordIsoΣ IsIsotoneIsoΣ (quote IsIsotone) + +isPropIsIsotone : {A : Type ℓ₀} {B : Type ℓ₁} + (M : PosetStr ℓ₀' A) (f : A → B) (N : PosetStr ℓ₁' B) + → isProp (IsIsotone M f N) +isPropIsIsotone M f N = isOfHLevelRetractFromIso 1 IsIsotoneIsoΣ + (isPropΠ3 λ x y _ → IsPoset.is-prop-valued (PosetStr.isPoset N) (f x) (f y)) + +IsIsotone-∘ : {A : Type ℓ₀} {B : Type ℓ₁} {C : Type ℓ₂} + → (M : PosetStr ℓ₀' A) (f : A → B) + (N : PosetStr ℓ₁' B) (g : B → C) (O : PosetStr ℓ₂' C) + → IsIsotone M f N + → IsIsotone N g O + → IsIsotone M (g ∘ f) O +IsIsotone.pres≤ (IsIsotone-∘ M f N g O isf isg) x y x≤y + = IsIsotone.pres≤ isg (f x) (f y) (IsIsotone.pres≤ isf x y x≤y) + +record IsAntitone {A : Type ℓ₀} {B : Type ℓ₁} + (M : PosetStr ℓ₀' A) (f : A → B) (N : PosetStr ℓ₁' B) + : Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') ℓ₁') + where + constructor + isantitone + -- Shorter qualified names + private + module M = PosetStr M + module N = PosetStr N + + field + inv≤ : (x y : A) → x M.≤ y → f y N.≤ f x + +unquoteDecl IsAntitoneIsoΣ = declareRecordIsoΣ IsAntitoneIsoΣ (quote IsAntitone) + +isPropIsAntitone : {A : Type ℓ₀} {B : Type ℓ₁} + (M : PosetStr ℓ₀' A) (f : A → B) (N : PosetStr ℓ₁' B) + → isProp (IsAntitone M f N) +isPropIsAntitone M f N = isOfHLevelRetractFromIso 1 IsAntitoneIsoΣ + (isPropΠ3 λ x y _ → IsPoset.is-prop-valued (PosetStr.isPoset N) (f y) (f x)) + +module _ + (A : Poset ℓ₀ ℓ₀') + (B : Poset ℓ₁ ℓ₁') + (f : ⟨ A ⟩ → ⟨ B ⟩) + where + + DualIsotone : IsIsotone (snd A) f (snd B) + → IsAntitone (snd (DualPoset A)) f (snd B) + IsAntitone.inv≤ (DualIsotone is) x y = IsIsotone.pres≤ is y x + + DualIsotone' : IsIsotone (snd (DualPoset A)) f (snd B) + → IsAntitone (snd A) f (snd B) + IsAntitone.inv≤ (DualIsotone' is) x y = IsIsotone.pres≤ is y x + + IsotoneDual : IsIsotone (snd A) f (snd B) + → IsAntitone (snd A) f (snd (DualPoset B)) + IsAntitone.inv≤ (IsotoneDual is) = IsIsotone.pres≤ is + + IsotoneDual' : IsIsotone (snd A) f (snd (DualPoset B)) + → IsAntitone (snd A) f (snd B) + IsAntitone.inv≤ (IsotoneDual' is) = IsIsotone.pres≤ is + + DualAntitone : IsAntitone (snd A) f (snd B) + → IsIsotone (snd (DualPoset A)) f (snd B) + IsIsotone.pres≤ (DualAntitone is) x y = IsAntitone.inv≤ is y x + + DualAntitone' : IsAntitone (snd (DualPoset A)) f (snd B) + → IsIsotone (snd A) f (snd B) + IsIsotone.pres≤ (DualAntitone' is) x y = IsAntitone.inv≤ is y x + + AntitoneDual : IsAntitone (snd A) f (snd B) + → IsIsotone (snd A) f (snd (DualPoset B)) + IsIsotone.pres≤ (AntitoneDual is) = IsAntitone.inv≤ is + + AntitoneDual' : IsAntitone (snd A) f (snd (DualPoset B)) + → IsIsotone (snd A) f (snd B) + IsIsotone.pres≤ (AntitoneDual' is) = IsAntitone.inv≤ is + + DualIsotoneDual : IsIsotone (snd A) f (snd B) + → IsIsotone (snd (DualPoset A)) f (snd (DualPoset B)) + IsIsotone.pres≤ (DualIsotoneDual is) x y = IsIsotone.pres≤ is y x + + DualIsotoneDual' : IsIsotone (snd (DualPoset A)) f (snd (DualPoset B)) + → IsIsotone (snd A) f (snd B) + IsIsotone.pres≤ (DualIsotoneDual' is) x y = IsIsotone.pres≤ is y x + + DualAntitoneDual : IsAntitone (snd A) f (snd B) + → IsAntitone (snd (DualPoset A)) f (snd (DualPoset B)) + IsAntitone.inv≤ (DualAntitoneDual is) x y = IsAntitone.inv≤ is y x + + DualAntitoneDual' : IsAntitone (snd (DualPoset A)) f (snd (DualPoset B)) + → IsAntitone (snd A) f (snd B) + IsAntitone.inv≤ (DualAntitoneDual' is) x y = IsAntitone.inv≤ is y x + +IsPosetEquiv→IsIsotone : (P S : Poset ℓ ℓ') + (e : ⟨ P ⟩ ≃ ⟨ S ⟩) + → IsPosetEquiv (snd P) e (snd S) + → IsIsotone (snd P) (equivFun e) (snd S) +IsIsotone.pres≤ (IsPosetEquiv→IsIsotone P S e eq) x y = equivFun (IsPosetEquiv.pres≤ eq x y) + +-- Isotone maps are characterized by their actions on down-sets and up-sets +module _ + {P : Poset ℓ₀ ℓ₀'} + {S : Poset ℓ₁ ℓ₁'} + (f : ⟨ P ⟩ → ⟨ S ⟩) + where + private + isP = PosetStr.isPoset (snd P) + isS = PosetStr.isPoset (snd S) + + _≤P_ = PosetStr._≤_ (snd P) + _≤S_ = PosetStr._≤_ (snd S) + + propS = IsPoset.is-prop-valued isS + rflS = IsPoset.is-refl isS + transS = IsPoset.is-trans isS + + IsIsotone→PreimagePrincipalDownsetIsDownset : IsIsotone (snd P) f (snd S) + → ∀ y → isDownset P (f ⃖ (principalDownset S y)) + IsIsotone→PreimagePrincipalDownsetIsDownset is y (x , inPrex) z z≤x + = ∥₁.rec (isEmbedding→hasPropFibers (preimageInclusion f (principalDownset S y) .snd) z) + (λ { ((b , b≤y) , fibb) → + (z , ∣ (f z , transS (f z) (f x) y + (IsIsotone.pres≤ is z x z≤x) + (subst (_≤S y) fibb b≤y)) , + refl ∣₁) , + refl }) inPrex + + IsIsotone→PreimagePrincipalUpsetIsUpset : IsIsotone (snd P) f (snd S) + → ∀ y → isUpset P (f ⃖ (principalUpset S y)) + IsIsotone→PreimagePrincipalUpsetIsUpset is y (x , inPrex) z x≤z + = ∥₁.rec (isEmbedding→hasPropFibers (preimageInclusion f (principalUpset S y) .snd) z) + (λ { ((b , y≤b) , fibb) → + (z , ∣ (f z , transS y (f x) (f z) + (subst (y ≤S_) fibb y≤b) + (IsIsotone.pres≤ is x z x≤z)) , + refl ∣₁) , + refl }) inPrex + + PreimagePrincipalDownsetIsDownset→IsIsotone : (∀ x → isDownset P (f ⃖ (principalDownset S x))) + → IsIsotone (snd P) f (snd S) + IsIsotone.pres≤ (PreimagePrincipalDownsetIsDownset→IsIsotone down) y x y≤x + = ∥₁.rec (propS _ _) (λ { ((b , b≤fx) , fibb) → subst (_≤S f x) (fibb ∙ cong f fiba) b≤fx }) pre + where fib = down (f x) (x , ∣ ((f x) , (rflS (f x))) , refl ∣₁) y y≤x + + pre = fib .fst .snd + fiba = fib .snd + + PreimagePrincipalUpsetIsUpset→IsIsotone : (∀ x → isUpset P (f ⃖ (principalUpset S x))) + → IsIsotone (snd P) f (snd S) + IsIsotone.pres≤ (PreimagePrincipalUpsetIsUpset→IsIsotone up) x y x≤y + = ∥₁.rec (propS _ _) (λ { ((b , fx≤b) , fibb) → subst (f x ≤S_) (fibb ∙ cong f fiba) fx≤b }) pre + where fib = up (f x) (x , ∣ ((f x) , (rflS (f x))) , refl ∣₁) y x≤y + + pre = fib .fst .snd + fiba = fib .snd + +-- The next part requires our posets to operate over the same universes +module _ + (P S : Poset ℓ ℓ') + (f : ⟨ P ⟩ → ⟨ S ⟩) + where + private + isP = PosetStr.isPoset (snd P) + isS = PosetStr.isPoset (snd S) + + _≤P_ = PosetStr._≤_ (snd P) + _≤S_ = PosetStr._≤_ (snd S) + + propP = IsPoset.is-prop-valued isP + rflP = IsPoset.is-refl isP + antiP = IsPoset.is-antisym isP + transP = IsPoset.is-trans isP + + propS = IsPoset.is-prop-valued isS + rflS = IsPoset.is-refl isS + antiS = IsPoset.is-antisym isS + transS = IsPoset.is-trans isS + + -- We can now define the type of residuated maps + isResiduated : Type _ + isResiduated = ∀ y → isPrincipalDownset P (f ⃖ (principalDownset S y)) + + hasResidual : Type (ℓ-max ℓ ℓ') + hasResidual = (IsIsotone (snd P) f (snd S)) × + (Σ[ g ∈ (⟨ S ⟩ → ⟨ P ⟩) ] (IsIsotone (snd S) g (snd P) × + (∀ x → x ≤P (g ∘ f) x) × + (∀ x → (f ∘ g) x ≤S x))) + + isResiduated→hasResidual : isResiduated + → hasResidual + isResiduated→hasResidual down = isotonef , g , isotoneg , g∘f , f∘g + where isotonef : IsIsotone (snd P) f (snd S) + isotonef = PreimagePrincipalDownsetIsDownset→IsIsotone f + λ x → isPrincipalDownset→isDownset P (f ⃖ principalDownset S x) (down x) + + isotonef⃖ : ∀ x y → x ≤S y → (f ⃖ (principalDownset S x)) ⊆ₑ (f ⃖ (principalDownset S y)) + isotonef⃖ x y x≤y z ((a , pre) , fiba) + = ∥₁.rec (isProp∈ₑ z (f ⃖ principalDownset S y)) + (λ { ((b , b≤x) , fibb) → (a , ∣ (b , (transS b x y b≤x x≤y)) , fibb ∣₁) , fiba }) pre + + g : ⟨ S ⟩ → ⟨ P ⟩ + g x = down x .fst + + isotoneg : IsIsotone (snd S) g (snd P) + IsIsotone.pres≤ isotoneg x y x≤y + = invEq + (principalDownsetMembership P (g x) (g y)) + (subst + (g x ∈ₑ_) + (down y .snd) + (isotonef⃖ x y x≤y (g x) + (subst (g x ∈ₑ_) + (sym (down x .snd)) + (equivFun (principalDownsetMembership P (g x) (g x)) (rflP (g x)))))) + + g∘f : ∀ x → x ≤P (g ∘ f) x + g∘f x = invEq (principalDownsetMembership P x (g (f x))) + (subst (x ∈ₑ_) (down (f x) .snd) + ((x , ∣ ((f x) , (rflS (f x))) , refl ∣₁) , refl)) + + f∘g : ∀ y → (f ∘ g) y ≤S y + f∘g y = ∥₁.rec (propS _ _) + (λ { ((a , a≤y) , fib) → + subst (_≤S y) (fib ∙ cong f (gy∈pre .snd)) a≤y }) + (gy∈pre .fst .snd) + where gy∈pre : g y ∈ₑ (f ⃖ (principalDownset S y)) + gy∈pre = subst (g y ∈ₑ_) (sym (down y .snd)) + (equivFun (principalDownsetMembership P (g y) (g y)) (rflP (g y))) + + hasResidual→isResiduated : hasResidual + → isResiduated + hasResidual→isResiduated (isf , g , isg , g∘f , f∘g) y + = (g y) , (equivFun (EmbeddingIP _ _) + ((λ x ((a , pre) , fiba) → + ∥₁.rec (isProp∈ₑ x (principalDownset P (g y))) + (λ { ((b , b≤y) , fibb) → + equivFun (principalDownsetMembership P x (g y)) + (transP x (g (f x)) (g y) (g∘f x) + (IsIsotone.pres≤ isg (f x) y + (subst (_≤S y) + (fibb ∙ cong f fiba) b≤y))) }) pre) , + λ x x∈g → (x , ∣ ((f x) , + (transS (f x) (f (g y)) y + (IsIsotone.pres≤ isf x (g y) + (invEq (principalDownsetMembership P x (g y)) x∈g)) + (f∘g y))) , refl ∣₁) , refl)) + + isPropIsResiduated : isProp isResiduated + isPropIsResiduated = isPropΠ λ _ → isPropIsPrincipalDownset P _ + + residualUnique : (p q : hasResidual) + → p .snd .fst ≡ q .snd .fst + residualUnique (isf₀ , g , isg , g∘f , f∘g) + (isf₁ , g* , isg* , g*∘f , f∘g*) + = funExt λ x → antiP (g x) (g* x) (g≤g* x) (g*≤g x) + where g≤g* : ∀ x → g x ≤P g* x + g≤g* x = transP (g x) ((g* ∘ f) (g x)) (g* x) (g*∘f (g x)) + (IsIsotone.pres≤ isg* (f (g x)) x (f∘g x)) + + g*≤g : ∀ x → g* x ≤P g x + g*≤g x = transP (g* x) ((g ∘ f) (g* x)) (g x) (g∘f (g* x)) + (IsIsotone.pres≤ isg (f (g* x)) x (f∘g* x)) + + isPropHasResidual : isProp hasResidual + isPropHasResidual p q = ≡-× (isPropIsIsotone _ f _ _ _) + (Σ≡Prop (λ g → isProp× (isPropIsIsotone _ g _) + (isProp× (isPropΠ (λ x → propP x (g (f x)))) + (isPropΠ λ x → propS (f (g x)) x))) + (residualUnique p q)) + + residual : hasResidual → ⟨ S ⟩ → ⟨ P ⟩ + residual (_ , g , _) = g + + AbsorbResidual : (res : hasResidual) + → f ∘ (residual res) ∘ f ≡ f + AbsorbResidual (isf , f⁺ , _ , f⁺∘f , f∘f⁺) + = funExt λ x → antiS ((f ∘ f⁺ ∘ f) x) (f x) + (f∘f⁺ (f x)) + (IsIsotone.pres≤ isf x (f⁺ (f x)) (f⁺∘f x)) + + ResidualAbsorb : (res : hasResidual) + → (residual res) ∘ f ∘ (residual res) ≡ (residual res) + ResidualAbsorb (_ , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) + = funExt λ x → antiP ((f⁺ ∘ f ∘ f⁺) x) (f⁺ x) + (IsIsotone.pres≤ isf⁺ (f (f⁺ x)) x (f∘f⁺ x)) + (f⁺∘f (f⁺ x)) + +isResidual : (P S : Poset ℓ ℓ') + → (f⁺ : ⟨ S ⟩ → ⟨ P ⟩) + → Type (ℓ-max ℓ ℓ') +isResidual P S f⁺ = Σ[ f ∈ (⟨ P ⟩ → ⟨ S ⟩) ] (Σ[ res ∈ hasResidual P S f ] f⁺ ≡ residual P S f res) + +isResidualOfUnique : (P S : Poset ℓ ℓ') + → (f⁺ : ⟨ S ⟩ → ⟨ P ⟩) + → (p q : isResidual P S f⁺) + → p .fst ≡ q .fst +isResidualOfUnique P S h (f , (isf , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) , h≡f⁺) + (g , (isg , g⁺ , isg⁺ , g⁺∘g , g∘g⁺) , h≡g⁺) + = funExt λ x → anti (f x) (g x) + (trans (f x) (f (f⁺ (g x))) (g x) + (IsIsotone.pres≤ isf x (f⁺ (g x)) + (subst (x ≤_) (sym (p (g x))) (g⁺∘g x))) + (f∘f⁺ (g x))) + (trans (g x) (g (g⁺ (f x))) (f x) + (IsIsotone.pres≤ isg x (g⁺ (f x)) + (subst (x ≤_) (p (f x)) (f⁺∘f x))) + (g∘g⁺ (f x))) + where _≤_ = PosetStr._≤_ (snd P) + anti = IsPoset.is-antisym (PosetStr.isPoset (snd S)) + trans = IsPoset.is-trans (PosetStr.isPoset (snd S)) + p = funExt⁻ ((sym h≡f⁺) ∙ h≡g⁺) + +isPropIsResidual : (P S : Poset ℓ ℓ') + → (f⁺ : ⟨ S ⟩ → ⟨ P ⟩) + → isProp (isResidual P S f⁺) +isPropIsResidual P S f⁺ p q + = Σ≡Prop (λ f → isPropΣ (isPropHasResidual P S f) + λ _ → isSet→ (IsPoset.is-set (PosetStr.isPoset (snd P))) _ _) + (isResidualOfUnique P S f⁺ p q) + +hasResidual-∘ : (E F G : Poset ℓ ℓ') + → (f : ⟨ E ⟩ → ⟨ F ⟩) + → (g : ⟨ F ⟩ → ⟨ G ⟩) + → hasResidual E F f + → hasResidual F G g + → hasResidual E G (g ∘ f) +hasResidual-∘ E F G f g (isf , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) (isg , g⁺ , isg⁺ , g⁺∘g , g∘g⁺) + = is , f⁺ ∘ g⁺ , is⁺ , ⁺∘ , ∘⁺ + where _≤E_ = PosetStr._≤_ (snd E) + _≤G_ = PosetStr._≤_ (snd G) + + transE = IsPoset.is-trans (PosetStr.isPoset (snd E)) + transG = IsPoset.is-trans (PosetStr.isPoset (snd G)) + + is : IsIsotone (snd E) (g ∘ f) (snd G) + is = IsIsotone-∘ (snd E) f (snd F) g (snd G) isf isg + + is⁺ : IsIsotone (snd G) (f⁺ ∘ g⁺) (snd E) + is⁺ = IsIsotone-∘ (snd G) g⁺ (snd F) f⁺ (snd E) isg⁺ isf⁺ + + ⁺∘ : ∀ x → x ≤E ((f⁺ ∘ g⁺) ∘ (g ∘ f)) x + ⁺∘ x = transE x ((f⁺ ∘ f) x) (((f⁺ ∘ g⁺) ∘ g ∘ f) x) + (f⁺∘f x) + (IsIsotone.pres≤ isf⁺ (f x) (g⁺ (g (f x))) (g⁺∘g (f x))) + + ∘⁺ : ∀ x → ((g ∘ f) ∘ (f⁺ ∘ g⁺)) x ≤G x + ∘⁺ x = transG (((g ∘ f) ∘ f⁺ ∘ g⁺) x) ((g ∘ g⁺) x) x + (IsIsotone.pres≤ isg (f (f⁺ (g⁺ x))) (g⁺ x) (f∘f⁺ (g⁺ x))) + (g∘g⁺ x) + +isResidual-∘ : (E F G : Poset ℓ ℓ') + → (f⁺ : ⟨ F ⟩ → ⟨ E ⟩) + → (g⁺ : ⟨ G ⟩ → ⟨ F ⟩) + → isResidual E F f⁺ + → isResidual F G g⁺ + → isResidual E G (f⁺ ∘ g⁺) +isResidual-∘ E F G f⁺ g⁺ (f , resf , f⁺≡f*) + (g , resg , g⁺≡g*) + = (g ∘ f) , + (hasResidual-∘ E F G f g resf resg) , + (funExt (λ x → cong f⁺ (funExt⁻ g⁺≡g* x) ∙ funExt⁻ f⁺≡f* _)) + +EqualResidual→Involution : (P : Poset ℓ ℓ') + → (f : ⟨ P ⟩ → ⟨ P ⟩) + → (res : hasResidual P P f) + → f ≡ (residual P P f res) + → f ∘ f ≡ idfun ⟨ P ⟩ +EqualResidual→Involution P f (isf , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) f≡f⁺ + = funExt λ x → anti (f (f x)) x + (subst (λ g → (f ∘ g) x ≤ x) (sym f≡f⁺) (f∘f⁺ x)) + (subst (λ g → x ≤ (g ∘ f) x) (sym f≡f⁺) (f⁺∘f x)) + where pos = PosetStr.isPoset (snd P) + _≤_ = PosetStr._≤_ (snd P) + anti = IsPoset.is-antisym pos + +Involution→EqualResidual : (P : Poset ℓ ℓ') + → (f : ⟨ P ⟩ → ⟨ P ⟩) + → (res : hasResidual P P f) + → f ∘ f ≡ idfun ⟨ P ⟩ + → f ≡ (residual P P f res) +Involution→EqualResidual P f res inv + = sym (cong₂ (λ g h → g ∘ f⁺ ∘ h) (sym inv) (sym inv) ∙ + cong (λ g → f ∘ g ∘ f) (AbsorbResidual P P f res) ∙ + cong (f ∘_) inv) + where f⁺ = res .snd .fst + +Res : Poset ℓ ℓ' → Semigroup (ℓ-max ℓ ℓ') +fst (Res E) = Σ[ f ∈ (⟨ E ⟩ → ⟨ E ⟩) ] hasResidual E E f +SemigroupStr._·_ (snd (Res E)) (f , resf) (g , resg) + = (g ∘ f) , (hasResidual-∘ E E E f g resf resg) +IsSemigroup.is-set (SemigroupStr.isSemigroup (snd (Res E))) + = isSetΣ (isSet→ (IsPoset.is-set (PosetStr.isPoset (snd E)))) + λ f → isProp→isSet (isPropHasResidual E E f) +IsSemigroup.·Assoc (SemigroupStr.isSemigroup (snd (Res E))) (f , _) (g , _) (h , _) + = Σ≡Prop (λ f → isPropHasResidual E E f) refl + +Res⁺ : Poset ℓ ℓ' → Semigroup (ℓ-max ℓ ℓ') +fst (Res⁺ E) = Σ[ f⁺ ∈ (⟨ E ⟩ → ⟨ E ⟩) ] isResidual E E f⁺ +SemigroupStr._·_ (snd (Res⁺ E)) (f⁺ , isresf⁺) (g⁺ , isresg⁺) + = (f⁺ ∘ g⁺) , isResidual-∘ E E E f⁺ g⁺ isresf⁺ isresg⁺ +IsSemigroup.is-set (SemigroupStr.isSemigroup (snd (Res⁺ E))) + = isSetΣ (isSet→ (IsPoset.is-set (PosetStr.isPoset (snd E)))) + λ f⁺ → isProp→isSet (isPropIsResidual E E f⁺) +IsSemigroup.·Assoc (SemigroupStr.isSemigroup (snd (Res⁺ E))) (f⁺ , _) (g⁺ , _) (h⁺ , _) + = Σ≡Prop (λ f⁺ → isPropIsResidual E E f⁺) refl + +isClosure : (E : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ E ⟩) + → Type (ℓ-max ℓ ℓ') +isClosure E f = IsIsotone (snd E) f (snd E) × (f ≡ f ∘ f) × (∀ x → x ≤ f x) + where _≤_ = PosetStr._≤_ (snd E) + +isDualClosure : (E : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ E ⟩) + → Type (ℓ-max ℓ ℓ') +isDualClosure E f = IsIsotone (snd E) f (snd E) × (f ≡ f ∘ f) × (∀ x → f x ≤ x) + where _≤_ = PosetStr._≤_ (snd E) + +-- This can be made more succinct +isClosure' : (E : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ E ⟩) + → Type (ℓ-max ℓ ℓ') +isClosure' E f = ∀ x y → x ≤ f y ≃ f x ≤ f y + where _≤_ = PosetStr._≤_ (snd E) + +isDualClosure' : (E : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ E ⟩) + → Type (ℓ-max ℓ ℓ') +isDualClosure' E f = ∀ x y → f x ≤ y ≃ f x ≤ f y + where _≤_ = PosetStr._≤_ (snd E) + +isClosure→isClosure' : (E : Poset ℓ ℓ') + → ∀ f + → isClosure E f + → isClosure' E f +isClosure→isClosure' E f (isf , f≡f∘f , x≤fx) x y + = propBiimpl→Equiv (prop _ _) (prop _ _) + (λ x≤fy → subst (f x ≤_) (sym (funExt⁻ f≡f∘f y)) + (IsIsotone.pres≤ isf x (f y) x≤fy)) + (trans x (f x) (f y) (x≤fx x)) + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + prop = IsPoset.is-prop-valued is + trans = IsPoset.is-trans is + +isDualClosure→isDualClosure' : (E : Poset ℓ ℓ') + → ∀ f + → isDualClosure E f + → isDualClosure' E f +isDualClosure→isDualClosure' E f (isf , f≡f∘f , fx≤x) x y + = propBiimpl→Equiv (prop _ _) (prop _ _) + (λ fx≤y → subst (_≤ f y) (sym (funExt⁻ f≡f∘f x)) + (IsIsotone.pres≤ isf (f x) y fx≤y)) + λ fx≤fy → trans (f x) (f y) y fx≤fy (fx≤x y) + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + prop = IsPoset.is-prop-valued is + trans = IsPoset.is-trans is + +isClosure'→isClosure : (E : Poset ℓ ℓ') + → ∀ f + → isClosure' E f + → isClosure E f +isClosure'→isClosure E f eq + = isf , + (funExt λ x → anti (f x) (f (f x)) + (IsIsotone.pres≤ isf x (f x) (x≤fx x)) + (equivFun (eq (f x) x) (rfl (f x)))) , + x≤fx + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + rfl = IsPoset.is-refl is + anti = IsPoset.is-antisym is + trans = IsPoset.is-trans is + + x≤fx : ∀ x → x ≤ f x + x≤fx x = invEq (eq x x) (rfl (f x)) + + isf : IsIsotone (snd E) f (snd E) + IsIsotone.pres≤ isf x y x≤y + = equivFun (eq x y) + (trans x y (f y) x≤y (x≤fx y)) + +isDualClosure'→isDualClosure : (E : Poset ℓ ℓ') + → ∀ f + → isDualClosure' E f + → isDualClosure E f +isDualClosure'→isDualClosure E f eq + = isf , + (funExt (λ x → anti (f x) (f (f x)) + (equivFun (eq x (f x)) (rfl (f x))) + (IsIsotone.pres≤ isf (f x) x (fx≤x x)))) , + fx≤x + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + rfl = IsPoset.is-refl is + anti = IsPoset.is-antisym is + trans = IsPoset.is-trans is + + fx≤x : ∀ x → f x ≤ x + fx≤x x = invEq (eq x x) (rfl (f x)) + + isf : IsIsotone (snd E) f (snd E) + IsIsotone.pres≤ isf x y x≤y + = equivFun (eq x y) + (trans (f x) x y (fx≤x x) x≤y) + +isClosure→ComposedResidual : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → isClosure E f + → Σ[ F ∈ Poset ℓ ℓ' ] (Σ[ g ∈ (⟨ E ⟩ → ⟨ F ⟩) ] (Σ[ res ∈ hasResidual E F g ] f ≡ (residual E F g res) ∘ g)) +isClosure→ComposedResidual {ℓ} {ℓ'} {E = E} {f = f} (isf , f≡f∘f , x≤fx) = F , ♮ , (is♮ , ♮⁺ , is♮⁺ , x≤fx , ♮∘♮⁺) , refl + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + set = IsPoset.is-set is + prop = IsPoset.is-prop-valued is + rfl = IsPoset.is-refl is + anti = IsPoset.is-antisym is + trans = IsPoset.is-trans is + + kerf : Rel ⟨ E ⟩ ⟨ E ⟩ ℓ + kerf x y = f x ≡ f y + + F' = ⟨ E ⟩ / kerf + + _⊑'_ : F' → F' → hProp _ + _⊑'_ = fun + where + fun₀ : ⟨ E ⟩ → F' → hProp _ + fst (fun₀ x [ y ]) = f x ≤ f y + snd (fun₀ x [ y ]) = prop (f x) (f y) + fun₀ x (eq/ a b fa≡fb i) = record + { fst = cong (f x ≤_) fa≡fb i + ; snd = isProp→PathP (λ i → isPropIsProp {A = cong (f x ≤_) fa≡fb i}) + (prop (f x) (f a)) (prop (f x) (f b)) i + } + fun₀ x (squash/ a b p q i j) = isSet→SquareP (λ _ _ → isSetHProp) + (λ _ → fun₀ x a) + (λ _ → fun₀ x b) + (λ i → fun₀ x (p i)) + (λ i → fun₀ x (q i)) j i + + toPath : ∀ a b (p : kerf a b) (y : F') → fun₀ a y ≡ fun₀ b y + toPath a b fa≡fb = elimProp (λ _ → isSetHProp _ _) λ c → + Σ≡Prop (λ _ → isPropIsProp) (cong (_≤ f c) fa≡fb) + + fun : F' → F' → hProp _ + fun [ a ] y = fun₀ a y + fun (eq/ a b fa≡fb i) y = toPath a b fa≡fb y i + fun (squash/ x y p q i j) z = isSet→SquareP (λ _ _ → isSetHProp) + (λ _ → fun x z) (λ _ → fun y z) (λ i → fun (p i) z) (λ i → fun (q i) z) j i + + _⊑_ : Rel F' F' ℓ' + a ⊑ b = (a ⊑' b) .fst + + open BinaryRelation _⊑_ + + isProp⊑ : isPropValued + isProp⊑ a b = (a ⊑' b) .snd + + isRefl⊑ : isRefl + isRefl⊑ = elimProp (λ x → isProp⊑ x x) + (rfl ∘ f) + + isAntisym⊑ : isAntisym + isAntisym⊑ = elimProp2 (λ x y → isPropΠ2 λ _ _ → squash/ x y) + λ a b fa≤fb fb≤fa → eq/ a b (anti (f a) (f b) fa≤fb fb≤fa) + + isTrans⊑ : isTrans + isTrans⊑ = elimProp3 (λ x _ z → isPropΠ2 λ _ _ → isProp⊑ x z) + λ a b c → trans (f a) (f b) (f c) + + poset⊑ : IsPoset _⊑_ + poset⊑ = isposet squash/ isProp⊑ isRefl⊑ isTrans⊑ isAntisym⊑ + + F : Poset ℓ ℓ' + F = F' , (posetstr _⊑_ poset⊑) + + ♮ : ⟨ E ⟩ → ⟨ F ⟩ + ♮ = [_] + + is♮ : IsIsotone (snd E) ♮ (snd F) + IsIsotone.pres≤ is♮ x y x≤y = IsIsotone.pres≤ isf x y x≤y + + ♮⁺ : ⟨ F ⟩ → ⟨ E ⟩ + ♮⁺ [ x ] = f x + ♮⁺ (eq/ a b fa≡fb i) = fa≡fb i + ♮⁺ (squash/ x y p q i j) = isSet→SquareP (λ _ _ → set) + (λ _ → ♮⁺ x) + (λ _ → ♮⁺ y) + (λ i → ♮⁺ (p i)) + (λ i → ♮⁺ (q i)) j i + + is♮⁺ : IsIsotone (snd F) ♮⁺ (snd E) + IsIsotone.pres≤ is♮⁺ = elimProp2 (λ x y → isPropΠ λ _ → prop (♮⁺ x) (♮⁺ y)) + λ x y fx≤fy → fx≤fy + + ♮∘♮⁺ : ∀ x → (♮ ∘ ♮⁺) x ⊑ x + ♮∘♮⁺ = elimProp (λ x → isProp⊑ ((♮ ∘ ♮⁺) x) x) + λ x → subst (_≤ f x) (funExt⁻ f≡f∘f x) (rfl (f x)) + +isDualClosure→ComposedResidual : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → isDualClosure E f + → Σ[ F ∈ Poset ℓ ℓ' ] (Σ[ g ∈ (⟨ F ⟩ → ⟨ E ⟩) ] (Σ[ res ∈ hasResidual F E g ] f ≡ g ∘ (residual F E g res))) +isDualClosure→ComposedResidual {ℓ} {ℓ'} {E = E} {f = f} (isf , f≡f∘f , fx≤x) = F , ♮ , (is♮ , ♮⁺ , is♮⁺ , ♮⁺∘♮ , fx≤x) , refl + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + set = IsPoset.is-set is + prop = IsPoset.is-prop-valued is + rfl = IsPoset.is-refl is + anti = IsPoset.is-antisym is + trans = IsPoset.is-trans is + + kerf : Rel ⟨ E ⟩ ⟨ E ⟩ ℓ + kerf x y = f x ≡ f y + + F' = ⟨ E ⟩ / kerf + + _⊑'_ : F' → F' → hProp _ + _⊑'_ = fun + where + fun₀ : ⟨ E ⟩ → F' → hProp _ + fst (fun₀ x [ y ]) = f x ≤ f y + snd (fun₀ x [ y ]) = prop (f x) (f y) + fun₀ x (eq/ a b fa≡fb i) = record + { fst = cong (f x ≤_) fa≡fb i + ; snd = isProp→PathP (λ i → isPropIsProp {A = cong (f x ≤_) fa≡fb i}) + (prop (f x) (f a)) (prop (f x) (f b)) i + } + fun₀ x (squash/ a b p q i j) = isSet→SquareP (λ _ _ → isSetHProp) + (λ _ → fun₀ x a) + (λ _ → fun₀ x b) + (λ i → fun₀ x (p i)) + (λ i → fun₀ x (q i)) j i + + toPath : ∀ a b (p : kerf a b) (y : F') → fun₀ a y ≡ fun₀ b y + toPath a b fa≡fb = elimProp (λ _ → isSetHProp _ _) λ c → + Σ≡Prop (λ _ → isPropIsProp) (cong (_≤ f c) fa≡fb) + + fun : F' → F' → hProp _ + fun [ a ] y = fun₀ a y + fun (eq/ a b fa≡fb i) y = toPath a b fa≡fb y i + fun (squash/ x y p q i j) z = isSet→SquareP (λ _ _ → isSetHProp) + (λ _ → fun x z) (λ _ → fun y z) (λ i → fun (p i) z) (λ i → fun (q i) z) j i + + _⊑_ : Rel F' F' ℓ' + a ⊑ b = (a ⊑' b) .fst + + open BinaryRelation _⊑_ + + isProp⊑ : isPropValued + isProp⊑ a b = (a ⊑' b) .snd + + isRefl⊑ : isRefl + isRefl⊑ = elimProp (λ x → isProp⊑ x x) + (rfl ∘ f) + + isAntisym⊑ : isAntisym + isAntisym⊑ = elimProp2 (λ x y → isPropΠ2 λ _ _ → squash/ x y) + λ a b fa≤fb fb≤fa → eq/ a b (anti (f a) (f b) fa≤fb fb≤fa) + + isTrans⊑ : isTrans + isTrans⊑ = elimProp3 (λ x _ z → isPropΠ2 λ _ _ → isProp⊑ x z) + λ a b c → trans (f a) (f b) (f c) + + poset⊑ : IsPoset _⊑_ + poset⊑ = isposet squash/ isProp⊑ isRefl⊑ isTrans⊑ isAntisym⊑ + + F : Poset ℓ ℓ' + F = F' , (posetstr _⊑_ poset⊑) + + ♮⁺ : ⟨ E ⟩ → ⟨ F ⟩ + ♮⁺ = [_] + + is♮⁺ : IsIsotone (snd E) ♮⁺ (snd F) + IsIsotone.pres≤ is♮⁺ x y x≤y = IsIsotone.pres≤ isf x y x≤y + + ♮ : ⟨ F ⟩ → ⟨ E ⟩ + ♮ [ x ] = f x + ♮ (eq/ a b fa≡fb i) = fa≡fb i + ♮ (squash/ x y p q i j) = isSet→SquareP (λ _ _ → set) + (λ _ → ♮ x) + (λ _ → ♮ y) + (λ i → ♮ (p i)) + (λ i → ♮ (q i)) j i + + is♮ : IsIsotone (snd F) ♮ (snd E) + IsIsotone.pres≤ is♮ = elimProp2 (λ x y → isPropΠ λ _ → prop (♮ x) (♮ y)) + λ x y fx≤fy → fx≤fy + + ♮⁺∘♮ : ∀ x → x ⊑ (♮⁺ ∘ ♮) x + ♮⁺∘♮ = elimProp (λ x → isProp⊑ x ((♮⁺ ∘ ♮) x)) + λ x → subst (f x ≤_) (funExt⁻ f≡f∘f x) (rfl (f x)) + +ComposedResidual→isClosure : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → Σ[ F ∈ Poset ℓ ℓ' ] (Σ[ g ∈ (⟨ E ⟩ → ⟨ F ⟩) ] (Σ[ res ∈ hasResidual E F g ] f ≡ (residual E F g res) ∘ g)) + → isClosure E f +ComposedResidual→isClosure {E = E} {f = f} (F , g , (isg , g⁺ , isg⁺ , g⁺∘g , g∘g⁺) , f≡g⁺∘g) + = subst (λ x → IsIsotone (snd E) x (snd E)) (sym f≡g⁺∘g) (IsIsotone-∘ (snd E) g (snd F) g⁺ (snd E) isg isg⁺) , + f≡g⁺∘g ∙ + sym (cong (g⁺ ∘_) + (AbsorbResidual E F g (isg , g⁺ , isg⁺ , g⁺∘g , g∘g⁺))) ∙ + cong (_∘ g⁺ ∘ g) (sym f≡g⁺∘g) ∙ + cong (f ∘_) (sym f≡g⁺∘g) , + λ x → subst (x ≤_) (sym (funExt⁻ f≡g⁺∘g x)) (g⁺∘g x) + where _≤_ = PosetStr._≤_ (snd E) + +ComposedResidual→isDualClosure : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → Σ[ F ∈ Poset ℓ ℓ' ] (Σ[ g ∈ (⟨ F ⟩ → ⟨ E ⟩) ] (Σ[ res ∈ hasResidual F E g ] f ≡ g ∘ (residual F E g res))) + → isDualClosure E f +ComposedResidual→isDualClosure {E = E} {f = f} (F , g , (isg , g⁺ , isg⁺ , g⁺∘g , g∘g⁺) , f≡g∘g⁺) + = subst (λ x → IsIsotone (snd E) x (snd E)) (sym f≡g∘g⁺) (IsIsotone-∘ (snd E) g⁺ (snd F) g (snd E) isg⁺ isg) , + f≡g∘g⁺ ∙ + sym (cong (g ∘_) (ResidualAbsorb F E g (isg , g⁺ , isg⁺ , g⁺∘g , g∘g⁺))) ∙ + cong (_∘ g ∘ g⁺) (sym f≡g∘g⁺) ∙ + cong (f ∘_) (sym f≡g∘g⁺) , + λ x → subst (_≤ x) (sym (funExt⁻ f≡g∘g⁺ x)) (g∘g⁺ x) + where _≤_ = PosetStr._≤_ (snd E) + +isPropIsClosure : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → isProp (isClosure E f) +isPropIsClosure {E = E} {f = f} + = isProp× (isPropIsIsotone (snd E) f (snd E)) + (isProp× (isSet→ (IsPoset.is-set is) _ _) + (isPropΠ λ x → IsPoset.is-prop-valued is x (f x))) + where is = PosetStr.isPoset (snd E) + +isPropIsClosure' : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → isProp (isClosure' E f) +isPropIsClosure' {E = E} {f = f} + = isPropΠ2 λ x y → isOfHLevel≃ 1 (prop x (f y)) (prop (f x) (f y)) + where prop = IsPoset.is-prop-valued (PosetStr.isPoset (snd E)) + +isPropIsDualClosure : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → isProp (isDualClosure E f) +isPropIsDualClosure {E = E} {f = f} + = isProp× (isPropIsIsotone (snd E) f (snd E)) + (isProp× (isSet→ (IsPoset.is-set is) _ _) + (isPropΠ λ x → IsPoset.is-prop-valued is (f x) x)) + where is = PosetStr.isPoset (snd E) + +isPropIsDualClosure' : {E : Poset ℓ ℓ'} + {f : ⟨ E ⟩ → ⟨ E ⟩} + → isProp (isDualClosure' E f) +isPropIsDualClosure' {E = E} {f = f} + = isPropΠ2 λ x y → isOfHLevel≃ 1 (prop (f x) y) (prop (f x) (f y)) + where prop = IsPoset.is-prop-valued (PosetStr.isPoset (snd E)) + +isClosureSubset : (E : Poset ℓ ℓ') + → (F : Embedding ⟨ E ⟩ ℓ) + → Type _ +isClosureSubset E F = Σ[ f ∈ (⟨ E ⟩ → ⟨ E ⟩) ] (isClosure E f × (F ≡ (Image f , imageInclusion f))) + +isDualClosureSubset : (E : Poset ℓ ℓ') + → (F : Embedding ⟨ E ⟩ ℓ) + → Type _ +isDualClosureSubset E F = Σ[ f ∈ (⟨ E ⟩ → ⟨ E ⟩) ] (isDualClosure E f × (F ≡ (Image f , imageInclusion f))) + +ClosureSubsetOperatorUnique : {E : Poset ℓ ℓ'} + {F : Embedding ⟨ E ⟩ ℓ} + → (f g : isClosureSubset E F) + → f .fst ≡ g .fst +ClosureSubsetOperatorUnique {E = E} {F = F} + (f , (isf , f≡f∘f , x≤fx) , F≡Imf) + (g , (isg , g≡g∘g , x≤gx) , F≡Img) + = funExt λ x → anti (f x) (g x) (fx≤gx x) (gx≤fx x) + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + prop = IsPoset.is-prop-valued is + anti = IsPoset.is-antisym is + + Imf⊆Img : (Image f , imageInclusion f) ⊆ₑ (Image g , imageInclusion g) + Imf⊆Img x = subst (x ∈ₑ_) (sym F≡Imf ∙ F≡Img) + + Img⊆Imf : (Image g , imageInclusion g) ⊆ₑ (Image f , imageInclusion f) + Img⊆Imf x = subst (x ∈ₑ_) (sym F≡Img ∙ F≡Imf) + + fx≤gx : ∀ x → f x ≤ g x + fx≤gx x = ∥₁.rec (prop (f x) (g x)) + (λ { (a , fa≡gx) → subst (f x ≤_) + (sym (funExt⁻ f≡f∘f a) ∙ + fa≡gx ∙ + lemma .snd) + (IsIsotone.pres≤ isf x (f a) + (subst (x ≤_) + (sym (fa≡gx ∙ lemma .snd)) + (x≤gx x))) }) + (lemma .fst .snd) + where lemma = Img⊆Imf (g x) (((g x) , ∣ x , refl ∣₁) , refl) + + gx≤fx : ∀ x → g x ≤ f x + gx≤fx x = ∥₁.rec (prop (g x) (f x)) + (λ { (a , ga≡fx) → subst (g x ≤_) + (sym (funExt⁻ g≡g∘g a) ∙ + ga≡fx ∙ + lemma .snd) + (IsIsotone.pres≤ isg x (g a) + (subst (x ≤_) + (sym (ga≡fx ∙ lemma .snd)) + (x≤fx x))) }) + (lemma .fst .snd) + where lemma = Imf⊆Img (f x) (((f x) , ∣ x , refl ∣₁) , refl) + +DualClosureSubsetOperatorUnique : {E : Poset ℓ ℓ'} + {F : Embedding ⟨ E ⟩ ℓ} + → (f g : isDualClosureSubset E F) + → f .fst ≡ g .fst +DualClosureSubsetOperatorUnique {E = E} {F = F} + (f , (isf , f≡f∘f , fx≤x) , F≡Imf) + (g , (isg , g≡g∘g , gx≤x) , F≡Img) + = funExt λ x → anti (f x) (g x) (fx≤gx x) (gx≤fx x) + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + prop = IsPoset.is-prop-valued is + anti = IsPoset.is-antisym is + + Imf⊆Img : (Image f , imageInclusion f) ⊆ₑ (Image g , imageInclusion g) + Imf⊆Img x = subst (x ∈ₑ_) (sym F≡Imf ∙ F≡Img) + + Img⊆Imf : (Image g , imageInclusion g) ⊆ₑ (Image f , imageInclusion f) + Img⊆Imf x = subst (x ∈ₑ_) (sym F≡Img ∙ F≡Imf) + + gx≤fx : ∀ x → g x ≤ f x + gx≤fx x = ∥₁.rec (prop (g x) (f x)) + (λ { (a , fa≡gx) → subst (_≤ f x) (sym (funExt⁻ f≡f∘f a) ∙ + fa≡gx ∙ + lemma .snd) + (IsIsotone.pres≤ isf (f a) x + (subst (_≤ x) + (sym (fa≡gx ∙ lemma .snd)) + (gx≤x x))) }) + (lemma .fst .snd) + where lemma = Img⊆Imf (g x) (((g x) , ∣ x , refl ∣₁) , refl) + + fx≤gx : ∀ x → f x ≤ g x + fx≤gx x = ∥₁.rec (prop (f x) (g x)) + (λ { (a , ga≡fx) → subst (_≤ g x) + (sym (funExt⁻ g≡g∘g a) ∙ + ga≡fx ∙ + lemma .snd) + (IsIsotone.pres≤ isg (g a) x + (subst (_≤ x) + (sym (ga≡fx ∙ lemma .snd)) + (fx≤x x))) }) + (lemma .fst .snd) + where lemma = Imf⊆Img (f x) (((f x) , ∣ x , refl ∣₁) , refl) + +isPropIsClosureSubset : {E : Poset ℓ ℓ'} + {F : Embedding ⟨ E ⟩ ℓ} + → isProp (isClosureSubset E F) +isPropIsClosureSubset p q = Σ≡Prop (λ f → isProp× isPropIsClosure (isSetEmbedding _ _)) + (ClosureSubsetOperatorUnique p q) + +isPropIsDualClosureSubset : {E : Poset ℓ ℓ'} + {F : Embedding ⟨ E ⟩ ℓ} + → isProp (isDualClosureSubset E F) +isPropIsDualClosureSubset p q = Σ≡Prop (λ f → isProp× isPropIsDualClosure (isSetEmbedding _ _)) + (DualClosureSubsetOperatorUnique p q) + +isClosureSubset→IntersectionBottom : (E : Poset ℓ ℓ') + (F : Embedding ⟨ E ⟩ ℓ) + → isClosureSubset E F + → ∀ x + → Least (isPoset→isProset (PosetStr.isPoset (snd E))) (principalUpset E x ∩ₑ F) +isClosureSubset→IntersectionBottom E F (f , (isf , f≡f∘f , x≤fx) , F≡Imf) x + = (f x , fx∈x↑ , fx∈F ) , least + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + prop = IsPoset.is-prop-valued is + + fx∈x↑ : f x ∈ₑ principalUpset E x + fx∈x↑ = equivFun (principalUpsetMembership E x (f x)) (x≤fx x) + + fx∈F : f x ∈ₑ F + fx∈F = subst (f x ∈ₑ_) (sym F≡Imf) ((f x , ∣ x , refl ∣₁) , refl) + + least : isLeast (isPoset→isProset is) (principalUpset E x ∩ₑ F) (f x , fx∈x↑ , fx∈F) + least (y , y∈x↑ , y∈F) = ∥₁.rec (prop _ _) + (λ { (a , fa≡fz) + → subst (f x ≤_) + (sym (funExt⁻ f≡f∘f a ∙ + cong f (fa≡fz ∙ + lemma .snd)) ∙ + fa≡fz ∙ + lemma .snd) + (IsIsotone.pres≤ isf x y + (invEq (principalUpsetMembership E x y) y∈x↑)) }) + (lemma .fst .snd) + where lemma = subst (y ∈ₑ_) F≡Imf y∈F + +IntersectionBottom→isClosureSubset : (E : Poset ℓ ℓ) + (F : Embedding ⟨ E ⟩ ℓ) + → (∀ x → Least (isPoset→isProset (PosetStr.isPoset (snd E))) (principalUpset E x ∩ₑ F)) + → isClosureSubset E F +IntersectionBottom→isClosureSubset E F least + = f , (isf , f≡f∘f , x≤f) , F≡Imf + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + rfl = IsPoset.is-refl is + anti = IsPoset.is-antisym is + + f : ⟨ E ⟩ → ⟨ E ⟩ + f x = least x .fst .fst + + isf : IsIsotone (snd E) f (snd E) + IsIsotone.pres≤ isf x y x≤y = least x .snd (f y , y↑∩F⊆x↑∩F (f y) ((least y .fst) , refl) .fst .snd) + where x↑ = principalUpset E x + y↑ = principalUpset E y + + y↑⊆x↑ = principalUpsetInclusion E x y x≤y + y↑∩F⊆x↑∩F = isMeetIsotone + (isPoset→isProset isPoset⊆ₑ) y↑ x↑ F F + (y↑ ∩ₑ F) + (x↑ ∩ₑ F) + (isMeet∩ₑ y↑ F) + (isMeet∩ₑ x↑ F) + y↑⊆x↑ + (isRefl⊆ₑ F) + + x≤f : ∀ x → x ≤ f x + x≤f x = invEq (principalUpsetMembership E x (f x)) (least x .fst .snd .fst) + + F≡fF : ∀ y → y ∈ₑ F + → y ≡ f y + F≡fF y y∈F = anti y (f y) (x≤f y) + (least y .snd (y , equivFun (principalUpsetMembership E y y) (rfl y) , y∈F)) + + f≡f∘f : f ≡ (f ∘ f) + f≡f∘f = funExt λ x → F≡fF (f x) (least x .fst .snd .snd) + + F⊆Imf : F ⊆ₑ (Image f , imageInclusion f) + F⊆Imf x x∈F = (x , ∣ x , (sym (F≡fF x x∈F)) ∣₁) , refl + + Imf⊆F : (Image f , imageInclusion f) ⊆ₑ F + Imf⊆F x ((a , ima) , fib) + = ∥₁.rec (isProp∈ₑ x F) + (λ { (b , fb≡a) → + subst (_∈ₑ F) + (fb≡a ∙ fib) + (least b .fst .snd .snd) }) ima + + F≡Imf : F ≡ (Image f , imageInclusion f) + F≡Imf = isAntisym⊆ₑ F (Image f , imageInclusion f) F⊆Imf Imf⊆F + +isBicomplete : (E : Poset ℓ ℓ') + (F : Embedding ⟨ E ⟩ ℓ) + → Type _ +isBicomplete E F = isClosureSubset E F × isClosureSubset (DualPoset E) F + +isBicompleteResiduatedClosureImage : (E : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ E ⟩) + → hasResidual E E f + → isClosure E f + → isBicomplete E (Image f , imageInclusion f) +isBicompleteResiduatedClosureImage E f (isf , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) (_ , f≡f∘f , x≤fx) + = (f , clsf , refl) , f⁺ , clsf⁺ , Imf≡Imf⁺ + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + + anti = IsPoset.is-antisym is + + resf = (isf , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) + clsf = (isf , f≡f∘f , x≤fx) + + f≡f⁺∘f : f ≡ f⁺ ∘ f + f≡f⁺∘f = funExt λ x → anti (f x) ((f⁺ ∘ f) x) (f≤f⁺∘f x) (f⁺∘f≤f x) + where f≤f⁺∘f : ∀ x → f x ≤ (f⁺ ∘ f) x + f≤f⁺∘f x = subst (f x ≤_) + (cong f⁺ (sym (funExt⁻ f≡f∘f x))) (f⁺∘f (f x)) + + f⁺∘f≤f : ∀ x → (f⁺ ∘ f) x ≤ f x + f⁺∘f≤f x = subst ((f⁺ ∘ f) x ≤_) + (funExt⁻ (AbsorbResidual E E f resf) x) + (x≤fx ((f⁺ ∘ f) x)) + + f⁺≡f∘f⁺ : f⁺ ≡ f ∘ f⁺ + f⁺≡f∘f⁺ = funExt λ x → sym (funExt⁻ (ResidualAbsorb E E f resf) x) ∙ + funExt⁻ (sym f≡f⁺∘f) (f⁺ x) + + Imf⊆Imf⁺ : (Image f , imageInclusion f) ⊆ₑ (Image f⁺ , imageInclusion f⁺) + Imf⊆Imf⁺ x ((a , ima) , fib) = ∥₁.rec (isProp∈ₑ x (Image f⁺ , imageInclusion f⁺)) + (λ { (b , fb≡a) → (x , ∣ x , (cong f⁺ (sym (fb≡a ∙ fib)) ∙ + funExt⁻ (sym f≡f⁺∘f) b ∙ + fb≡a ∙ + fib) ∣₁) , refl }) + ima + + Imf⁺⊆Imf : (Image f⁺ , imageInclusion f⁺) ⊆ₑ (Image f , imageInclusion f) + Imf⁺⊆Imf x ((a , ima) , fib) = ∥₁.rec (isProp∈ₑ x (Image f , imageInclusion f)) + (λ { (b , f⁺b≡a) → (x , ∣ x , (cong f (sym (f⁺b≡a ∙ fib)) ∙ + funExt⁻ (sym f⁺≡f∘f⁺) b ∙ + f⁺b≡a ∙ + fib) ∣₁) , refl }) + ima + + Imf≡Imf⁺ : (Image f , imageInclusion f) ≡ (Image f⁺ , imageInclusion f⁺) + Imf≡Imf⁺ = isAntisym⊆ₑ _ _ Imf⊆Imf⁺ Imf⁺⊆Imf + + clsf⁺ : isClosure (DualPoset E) f⁺ + clsf⁺ = DualIsotoneDual E E f⁺ isf⁺ , + f⁺≡f∘f⁺ ∙ cong (_∘ f⁺) f≡f⁺∘f ∙ cong (f⁺ ∘_) (sym f⁺≡f∘f⁺) , + λ x → subst (_≤ x) (funExt⁻ (sym f⁺≡f∘f⁺) x) (f∘f⁺ x) + +isBicomplete→ClosureOperatorHasResidual : (E : Poset ℓ ℓ') + (F : Embedding ⟨ E ⟩ ℓ) + → (bi : isBicomplete E F) + → hasResidual E E (bi . fst .fst) +isBicomplete→ClosureOperatorHasResidual E F ((f , (isf , f≡f∘f , x≤fx) , F≡Imf) , + f⁺ , (isf⁺ , f⁺≡f⁺∘f⁺ , f⁺x≤x) , F≡Imf⁺) + = isf , f⁺ , DualIsotoneDual' E E f⁺ isf⁺ , f⁺∘f , f∘f⁺ + where _≤_ = PosetStr._≤_ (snd E) + is = PosetStr.isPoset (snd E) + set = IsPoset.is-set is + + Imf⁺⊆Imf : (Image f⁺ , imageInclusion f⁺) ⊆ₑ (Image f , imageInclusion f) + Imf⁺⊆Imf = subst (Imf⁺ ⊆ₑ_) ((sym F≡Imf⁺) ∙ F≡Imf) (isRefl⊆ₑ Imf⁺) + where Imf⁺ = (Image f⁺ , imageInclusion f⁺) + + Imf⊆Imf⁺ : (Image f , imageInclusion f) ⊆ₑ (Image f⁺ , imageInclusion f⁺) + Imf⊆Imf⁺ = subst (Imf ⊆ₑ_) ((sym F≡Imf) ∙ F≡Imf⁺) (isRefl⊆ₑ Imf) + where Imf = (Image f , imageInclusion f) + + f≡f⁺∘f : ∀ x → f x ≡ (f⁺ ∘ f) x + f≡f⁺∘f x = ∥₁.rec (set _ _) (λ { (b , f⁺b≡a) → (sym (f⁺b≡a ∙ fib) ∙ funExt⁻ f⁺≡f⁺∘f⁺ b) ∙ cong f⁺ (f⁺b≡a ∙ fib) }) ima + where lemma = Imf⊆Imf⁺ (f x) (((f x) , ∣ x , refl ∣₁) , refl) + + ima = lemma .fst .snd + fib = lemma .snd + + f⁺≡f∘f⁺ : ∀ x → f⁺ x ≡ (f ∘ f⁺) x + f⁺≡f∘f⁺ x = ∥₁.rec (set _ _) (λ { (b , fb≡a) → (sym (fb≡a ∙ fib) ∙ funExt⁻ f≡f∘f b) ∙ cong f (fb≡a ∙ fib) }) ima + where lemma = Imf⁺⊆Imf (f⁺ x) (((f⁺ x) , ∣ x , refl ∣₁) , refl) + + ima = lemma .fst .snd + fib = lemma .snd + + f⁺∘f : ∀ x → x ≤ (f⁺ ∘ f) x + f⁺∘f x = subst (x ≤_) (f≡f⁺∘f x) (x≤fx x) + + f∘f⁺ : ∀ x → (f ∘ f⁺) x ≤ x + f∘f⁺ x = subst (_≤ x) (f⁺≡f∘f⁺ x) (f⁺x≤x x) + +IsPosetEquiv→isResiduatedBijection : (P S : Poset ℓ ℓ') + (e : ⟨ P ⟩ ≃ ⟨ S ⟩) + → IsPosetEquiv (snd P) e (snd S) + → hasResidual P S (equivFun e) +IsPosetEquiv→isResiduatedBijection P S e eq + = IsPosetEquiv→IsIsotone P S e eq , invEq e , is⁻ , + (λ x → subst (x ≤P_) (sym (retEq e x)) (rflP x)) , + λ x → subst (_≤S x) (sym (secEq e x)) (rflS x) + where _≤P_ = PosetStr._≤_ (snd P) + _≤S_ = PosetStr._≤_ (snd S) + + rflP = IsPoset.is-refl (PosetStr.isPoset (snd P)) + rflS = IsPoset.is-refl (PosetStr.isPoset (snd S)) + + is⁻ : IsIsotone (snd S) (invEq e) (snd P) + IsIsotone.pres≤ is⁻ x y = equivFun (IsPosetEquiv.pres≤⁻ eq x y) + +isResiduatedBijection→IsPosetEquiv : (P S : Poset ℓ ℓ') + (e : ⟨ P ⟩ ≃ ⟨ S ⟩) + → hasResidual P S (equivFun e) + → IsPosetEquiv (snd P) e (snd S) +IsPosetEquiv.pres≤ (isResiduatedBijection→IsPosetEquiv P S e + (ise , e⁻ , ise⁻ , e⁻∘e , e∘e⁻)) x y + = propBiimpl→Equiv (propP _ _) (propS _ _) (IsIsotone.pres≤ ise x y) (subst2 _≤P_ (lemma x) (lemma y) ∘ IsIsotone.pres≤ ise⁻ _ _) + where _≤P_ = PosetStr._≤_ (snd P) + isP = PosetStr.isPoset (snd P) + propP = IsPoset.is-prop-valued isP + rflP = IsPoset.is-refl isP + + _≤S_ = PosetStr._≤_ (snd S) + isS = PosetStr.isPoset (snd S) + propS = IsPoset.is-prop-valued isS + antiS = IsPoset.is-antisym isS + + e∘e⁻x≡x : ∀ x → equivFun e (e⁻ x) ≡ x + e∘e⁻x≡x x = antiS _ x (e∘e⁻ x) + (subst2 _≤S_ (secEq e x) + (cong (equivFun e ∘ e⁻) (secEq e x)) + (IsIsotone.pres≤ ise _ _ (e⁻∘e (invEq e x)))) + + e⁻≡inv : ∀ x → e⁻ x ≡ invEq e x + e⁻≡inv x = sym (retEq e (e⁻ x)) ∙ cong (invEq e) (e∘e⁻x≡x x) + + lemma : ∀ x → e⁻ (equivFun e x) ≡ x + lemma x = e⁻≡inv (equivFun e x) ∙ retEq e x + +-- We can weaken the equivalence of a poset equivalence to a surjection +isOrderRecovering→isEmbedding : (P S : Poset ℓ ℓ') + (f : ⟨ P ⟩ → ⟨ S ⟩) + → (∀ x y → (PosetStr._≤_ (snd S) (f x) (f y)) + → (PosetStr._≤_ (snd P) x y)) + → isEmbedding f +isOrderRecovering→isEmbedding P S f is = emb + where _≤_ = PosetStr._≤_ (snd S) + + isP = PosetStr.isPoset (snd P) + isS = PosetStr.isPoset (snd S) + + setS = IsPoset.is-set isS + + antiP = IsPoset.is-antisym isP + rflS = IsPoset.is-refl isS + + emb : isEmbedding f + emb = injEmbedding setS λ {w} {x} fw≡fx + → antiP w x (is w x (subst (f w ≤_) fw≡fx (rflS (f w)))) + (is x w (subst (_≤ f w) fw≡fx (rflS (f w)))) + +-- Galois connections work similarly to residuals, but are antitone +module _ + (E F : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ F ⟩) + (g : ⟨ F ⟩ → ⟨ E ⟩) + where + private + _≤E_ = PosetStr._≤_ (snd E) + _≤F_ = PosetStr._≤_ (snd F) + + propE = IsPoset.is-prop-valued (PosetStr.isPoset (snd E)) + propF = IsPoset.is-prop-valued (PosetStr.isPoset (snd F)) + + isGaloisConnection : Type (ℓ-max ℓ ℓ') + isGaloisConnection = IsAntitone (snd E) f (snd F) × + IsAntitone (snd F) g (snd E) × + (∀ x → x ≤F (f ∘ g) x) × + (∀ x → x ≤E (g ∘ f) x) + + isPropIsGaloisConnection : isProp isGaloisConnection + isPropIsGaloisConnection = isProp× (isPropIsAntitone _ _ _) + (isProp× (isPropIsAntitone _ _ _) + (isProp× (isPropΠ λ x → propF x _) + (isPropΠ λ x → propE x _))) + + isGaloisConnection→hasResidualDual : isGaloisConnection + → hasResidual E (DualPoset F) f + isGaloisConnection→hasResidualDual (antif , antig , f∘g , g∘f) + = AntitoneDual E F f antif , g , DualAntitone F E g antig , g∘f , f∘g + + AbsorbGaloisConnection : isGaloisConnection + → f ∘ g ∘ f ≡ f + AbsorbGaloisConnection conn + = AbsorbResidual E (DualPoset F) f (isGaloisConnection→hasResidualDual conn) + + GaloisConnectionAbsorb : isGaloisConnection + → g ∘ f ∘ g ≡ g + GaloisConnectionAbsorb conn + = ResidualAbsorb E (DualPoset F) f (isGaloisConnection→hasResidualDual conn) + + GaloisConnectionClosure : isGaloisConnection + → isClosure E (g ∘ f) + GaloisConnectionClosure conn + = ComposedResidual→isClosure (DualPoset F , f , isGaloisConnection→hasResidualDual conn , refl) + + GaloisConnectionDualClosure : isGaloisConnection + → isDualClosure (DualPoset F) (f ∘ g) + GaloisConnectionDualClosure conn + = ComposedResidual→isDualClosure (E , f , isGaloisConnection→hasResidualDual conn , refl) + +hasResidual→isGaloisConnectionDual : (E F : Poset ℓ ℓ') + (f : ⟨ E ⟩ → ⟨ F ⟩) + → (res : hasResidual E F f) + → isGaloisConnection E (DualPoset F) f (residual E F f res) +hasResidual→isGaloisConnectionDual E F f (isf , f⁺ , isf⁺ , f⁺∘f , f∘f⁺) + = (IsotoneDual E F f isf) , (DualIsotone F E f⁺ isf⁺) , f∘f⁺ , f⁺∘f diff --git a/Cubical/Relation/Binary/Order/Poset/Properties.agda b/Cubical/Relation/Binary/Order/Poset/Properties.agda index a3122fb1ca..323689256e 100644 --- a/Cubical/Relation/Binary/Order/Poset/Properties.agda +++ b/Cubical/Relation/Binary/Order/Poset/Properties.agda @@ -1,8 +1,14 @@ {-# OPTIONS --safe #-} module Cubical.Relation.Binary.Order.Poset.Properties where -open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Data.Sum + +open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Transport open import Cubical.Functions.Embedding @@ -12,14 +18,14 @@ open import Cubical.Data.Sigma open import Cubical.Relation.Binary.Base open import Cubical.Relation.Binary.Order.Poset.Base -open import Cubical.Relation.Binary.Order.Preorder.Base -open import Cubical.Relation.Binary.Order.StrictPoset.Base +open import Cubical.Relation.Binary.Order.Proset +open import Cubical.Relation.Binary.Order.Quoset.Base open import Cubical.Relation.Nullary private variable - ℓ ℓ' ℓ'' : Level + ℓ ℓ' ℓ'' ℓ₀ ℓ₀' ℓ₁ ℓ₁' : Level module _ {A : Type ℓ} @@ -28,22 +34,30 @@ module _ open BinaryRelation - isPoset→isPreorder : IsPoset R → IsPreorder R - isPoset→isPreorder poset = ispreorder + isPoset→isProset : IsPoset R → IsProset R + isPoset→isProset poset = isproset (IsPoset.is-set poset) (IsPoset.is-prop-valued poset) (IsPoset.is-refl poset) (IsPoset.is-trans poset) + isPosetDecidable→Discrete : IsPoset R → isDecidable R → Discrete A + isPosetDecidable→Discrete pos dec a b with dec a b + ... | no ¬Rab = no (λ a≡b → ¬Rab (subst (R a) a≡b (IsPoset.is-refl pos a))) + ... | yes Rab with dec b a + ... | no ¬Rba = no (λ a≡b → ¬Rba (subst (λ x → R x a) a≡b + (IsPoset.is-refl pos a))) + ... | yes Rba = yes (IsPoset.is-antisym pos a b Rab Rba) + private transirrefl : isTrans R → isAntisym R → isTrans (IrreflKernel R) transirrefl trans anti a b c (Rab , ¬a≡b) (Rbc , ¬b≡c) = trans a b c Rab Rbc , λ a≡c → ¬a≡b (anti a b Rab (subst (R b) (sym a≡c) Rbc)) - isPoset→isStrictPosetIrreflKernel : IsPoset R → IsStrictPoset (IrreflKernel R) - isPoset→isStrictPosetIrreflKernel poset - = isstrictposet (IsPoset.is-set poset) + isPoset→isQuosetIrreflKernel : IsPoset R → IsQuoset (IrreflKernel R) + isPoset→isQuosetIrreflKernel poset + = isquoset (IsPoset.is-set poset) (λ a b → isProp× (IsPoset.is-prop-valued poset a b) (isProp¬ (a ≡ b))) (isIrreflIrreflKernel R) @@ -54,6 +68,13 @@ module _ , transirrefl (IsPoset.is-trans poset) (IsPoset.is-antisym poset))) + isPosetDecidable→isQuosetDecidable : IsPoset R → isDecidable R → isDecidable (IrreflKernel R) + isPosetDecidable→isQuosetDecidable pos dec a b with dec a b + ... | no ¬Rab = no λ { (Rab , _) → ¬Rab Rab } + ... | yes Rab with (isPosetDecidable→Discrete pos dec) a b + ... | yes a≡b = no λ { (_ , ¬a≡b) → ¬a≡b a≡b } + ... | no a≢b = yes (Rab , a≢b) + isPosetInduced : IsPoset R → (B : Type ℓ'') → (f : B ↪ A) → IsPoset (InducedRelation R (B , f)) isPosetInduced pos B (f , emb) = isposet (Embedding-into-isSet→isSet (f , emb) (IsPoset.is-set pos)) @@ -63,28 +84,532 @@ module _ λ a b a≤b b≤a → isEmbedding→Inj emb a b (IsPoset.is-antisym pos (f a) (f b) a≤b b≤a) -Poset→Preorder : Poset ℓ ℓ' → Preorder ℓ ℓ' -Poset→Preorder (_ , pos) = _ , preorderstr (PosetStr._≤_ pos) - (isPoset→isPreorder (PosetStr.isPoset pos)) + isPosetDual : IsPoset R → IsPoset (Dual R) + isPosetDual pos + = isposet (IsPoset.is-set pos) + (λ a b → IsPoset.is-prop-valued pos b a) + (IsPoset.is-refl pos) + (λ a b c Rab Rbc → IsPoset.is-trans pos c b a Rbc Rab) + (λ a b Rab Rba → IsPoset.is-antisym pos a b Rba Rab) + +Poset→Proset : Poset ℓ ℓ' → Proset ℓ ℓ' +Poset→Proset (_ , pos) + = proset _ (PosetStr._≤_ pos) + (isPoset→isProset (PosetStr.isPoset pos)) + +Poset→Quoset : Poset ℓ ℓ' → Quoset ℓ (ℓ-max ℓ ℓ') +Poset→Quoset (_ , pos) + = quoset _ (BinaryRelation.IrreflKernel (PosetStr._≤_ pos)) + (isPoset→isQuosetIrreflKernel (PosetStr.isPoset pos)) + +DualPoset : Poset ℓ ℓ' → Poset ℓ ℓ' +DualPoset (_ , pos) + = poset _ (BinaryRelation.Dual (PosetStr._≤_ pos)) + (isPosetDual (PosetStr.isPoset pos)) + +module _ + {A : Type ℓ} + {_≤_ : Rel A A ℓ'} + (pos : IsPoset _≤_) + where + + private + pre = isPoset→isProset pos + + set = IsPoset.is-set pos + + prop = IsPoset.is-prop-valued pos + + rfl = IsPoset.is-refl pos + + anti = IsPoset.is-antisym pos + + trans = IsPoset.is-trans pos + + module _ + {P : Embedding A ℓ''} + where + + private + toA = fst (snd P) + + emb = snd (snd P) + + isLeast→ContrMinimal : ∀ n → isLeast pre P n → ∀ m → isMinimal pre P m → n ≡ m + isLeast→ContrMinimal n isln m ismm + = isEmbedding→Inj emb n m (anti (toA n) (toA m) (isln m) (ismm n (isln m))) + + isGreatest→ContrMaximal : ∀ n → isGreatest pre P n → ∀ m → isMaximal pre P m → n ≡ m + isGreatest→ContrMaximal n isgn m ismm + = isEmbedding→Inj emb n m (anti (toA n) (toA m) (ismm n (isgn m)) (isgn m)) + + isLeastUnique : ∀ n m → isLeast pre P n → isLeast pre P m → n ≡ m + isLeastUnique n m isln islm + = isEmbedding→Inj emb n m (anti (toA n) (toA m) (isln m) (islm n)) + + isGreatestUnique : ∀ n m → isGreatest pre P n → isGreatest pre P m → n ≡ m + isGreatestUnique n m isgn isgm + = isEmbedding→Inj emb n m (anti (toA n) (toA m) (isgm n) (isgn m)) + + LeastUnique : (n m : Least pre P) → n ≡ m + LeastUnique (n , ln) (m , lm) = Σ≡Prop (λ _ → isPropIsLeast pre P _) (isLeastUnique n m ln lm) + + GreatestUnique : (n m : Greatest pre P) → n ≡ m + GreatestUnique (n , gn) (m , gm) = Σ≡Prop (λ _ → isPropIsGreatest pre P _) (isGreatestUnique n m gn gm) + + module _ + {B : Type ℓ''} + {P : B → A} + where + + isInfimum→ContrMaximalLowerBound : ∀ n → isInfimum pre P n + → ∀ m → isMaximalLowerBound pre P m + → n ≡ m + isInfimum→ContrMaximalLowerBound n (isln , isglbn) m (islm , ismlbm) + = anti n m (ismlbm (n , isln) (isglbn (m , islm))) (isglbn (m , islm)) + + isSupremum→ContrMinimalUpperBound : ∀ n → isSupremum pre P n + → ∀ m → isMinimalUpperBound pre P m + → n ≡ m + isSupremum→ContrMinimalUpperBound n (isun , islubn) m (isum , ismubm) + = anti n m (islubn (m , isum)) (ismubm (n , isun) (islubn (m , isum))) + + isInfimumUnique : ∀ n m → isInfimum pre P n → isInfimum pre P m → n ≡ m + isInfimumUnique n m (isln , infn) (islm , infm) + = anti n m (infm (n , isln)) (infn (m , islm)) + + isSupremumUnique : ∀ n m → isSupremum pre P n → isSupremum pre P m → n ≡ m + isSupremumUnique n m (isun , supn) (isum , supm) + = anti n m (supn (m , isum)) (supm (n , isun)) + + InfimumUnique : (n m : Infimum pre P) → n ≡ m + InfimumUnique (n , infn) (m , infm) = Σ≡Prop (λ _ → isPropIsInfimum pre P _) + (isInfimumUnique n m infn infm) + + SupremumUnique : (n m : Supremum pre P) → n ≡ m + SupremumUnique (n , supn) (m , supm) = Σ≡Prop (λ _ → isPropIsSupremum pre P _) + (isSupremumUnique n m supn supm) + + isMeetUnique : ∀ a b x y → isMeet pre a b x → isMeet pre a b y → x ≡ y + isMeetUnique a b x y infx infy = anti x y x≤y y≤x + where x≤y : x ≤ y + x≤y = invEq (infy x) (infx x .fst (rfl x)) + + y≤x : y ≤ x + y≤x = invEq (infx y) (infy y .fst (rfl y)) + + isJoinUnique : ∀ a b x y → isJoin pre a b x → isJoin pre a b y → x ≡ y + isJoinUnique a b x y supx supy = anti x y x≤y y≤x + where x≤y : x ≤ y + x≤y = invEq (supx y) (supy y .fst (rfl y)) + + y≤x : y ≤ x + y≤x = invEq (supy x) (supx x .fst (rfl x)) + + MeetUnique : ∀ a b → (x y : Meet pre a b) → x ≡ y + MeetUnique a b (x , mx) (y , my) = Σ≡Prop (isPropIsMeet pre a b) + (isMeetUnique a b x y mx my) + + JoinUnique : ∀ a b → (x y : Join pre a b) → x ≡ y + JoinUnique a b (x , jx) (y , jy) = Σ≡Prop (isPropIsJoin pre a b) + (isJoinUnique a b x y jx jy) + + module _ + (meet : ∀ a b → Meet pre a b) + where + + meetIdemp : ∀ a → meet a a .fst ≡ a + meetIdemp a + = PathPΣ (MeetUnique a a (meet a a) + (a , (isReflIsMeet pre a))) .fst + + meetComm : ∀ a b → meet a b .fst ≡ meet b a .fst + meetComm a b = sym (PathPΣ (MeetUnique b a (meet b a) + ((meet a b .fst) , + (isMeetComm pre a b + (meet a b .fst) + (meet a b .snd)))) .fst) + + meetAssoc : ∀ a b c + → meet a (meet b c .fst) .fst ≡ meet (meet a b .fst) c .fst + meetAssoc a b c + = sym (PathPΣ (MeetUnique + (meet a b .fst) c + (meet (meet a b .fst) c) + ((meet a (meet b c .fst) .fst) , + (isMeetAssoc pre a b c + (meet a b .fst) + (meet b c .fst) + (meet a (meet b c .fst) .fst) + (meet a b .snd) + (meet b c .snd) + (meet a (meet b c .fst) .snd)))) .fst) + + module _ + (join : ∀ a b → Join pre a b) + where + + joinIdem : ∀ a → join a a .fst ≡ a + joinIdem a + = PathPΣ (JoinUnique a a (join a a) + (a , (isReflIsJoin pre a))) .fst + + joinComm : ∀ a b → join a b .fst ≡ join b a .fst + joinComm a b = sym (PathPΣ (JoinUnique b a (join b a) + ((join a b .fst) , + (isJoinComm pre a b + (join a b .fst) + (join a b .snd)))) .fst) + + joinAssoc : ∀ a b c + → join a (join b c .fst) .fst ≡ join (join a b .fst) c .fst + joinAssoc a b c + = sym (PathPΣ (JoinUnique + (join a b .fst) c + (join (join a b .fst) c) + ((join a (join b c .fst) .fst) , + (isJoinAssoc pre a b c + (join a b .fst) + (join b c .fst) + (join a (join b c .fst) .fst) + (join a b .snd) + (join b c .snd) + (join a (join b c .fst) .snd)))) .fst) + + order≃meet : ∀ a b a∧b + → isMeet pre a b a∧b + → (a ≤ b) ≃ (a ≡ a∧b) + order≃meet a b a∧b funa∧b + = propBiimpl→Equiv (prop a b) + (set a a∧b) + (λ a≤b → anti a a∧b (invEq (funa∧b a) (rfl a , a≤b)) + (funa∧b a∧b .fst (rfl a∧b) .fst)) + λ a≡a∧b → subst (_≤ b) (sym a≡a∧b) + (funa∧b a∧b .fst (rfl a∧b) .snd) + + order≃join : ∀ a b a∨b + → isJoin pre a b a∨b + → (a ≤ b) ≃ (b ≡ a∨b) + order≃join a b a∨b funa∨b + = propBiimpl→Equiv (prop a b) + (set b a∨b) + (λ a≤b → anti b a∨b (funa∨b a∨b .fst (rfl a∨b) .snd) + (invEq (funa∨b b) (a≤b , rfl b))) + λ b≡a∨b → subst (a ≤_) (sym b≡a∨b) + (funa∨b a∨b .fst (rfl a∨b) .fst) + +-- Equivalences of posets respect meets and joins +IsPosetEquivRespectsMeet : {P : Poset ℓ₀ ℓ₀'} {S : Poset ℓ₁ ℓ₁'} (e : PosetEquiv P S) + → ∀ a b a∧b + → isMeet (isPoset→isProset (PosetStr.isPoset (P .snd))) a b a∧b + ≃ isMeet (isPoset→isProset (PosetStr.isPoset (S .snd))) + (equivFun (e .fst) a) + (equivFun (e .fst) b) + (equivFun (e .fst) a∧b) +IsPosetEquivRespectsMeet {P = P} {S = S} (e , posEq) a b a∧b + = propBiimpl→Equiv (isPropIsMeet proP a b a∧b) + (isPropIsMeet proS (equivFun e a) (equivFun e b) (equivFun e a∧b)) + (λ meet x + → IsPosetEquiv.pres≤⁻ posEq x (equivFun e a∧b) ∙ₑ + substEquiv (invEq e x ≤P_) (retEq e a∧b) ∙ₑ + meet (invEq e x) ∙ₑ + ≃-× (IsPosetEquiv.pres≤ posEq (invEq e x) a ∙ₑ + substEquiv (_≤S equivFun e a) (secEq e x)) + (IsPosetEquiv.pres≤ posEq (invEq e x) b ∙ₑ + substEquiv (_≤S equivFun e b) (secEq e x))) + λ meet x + → IsPosetEquiv.pres≤ posEq x a∧b ∙ₑ + meet (equivFun e x) ∙ₑ + ≃-× (IsPosetEquiv.pres≤⁻ posEq (equivFun e x) (equivFun e a) ∙ₑ + subst2Equiv _≤P_ (retEq e x) (retEq e a)) + (IsPosetEquiv.pres≤⁻ posEq (equivFun e x) (equivFun e b) ∙ₑ + subst2Equiv _≤P_ (retEq e x) (retEq e b)) + where _≤P_ = PosetStr._≤_ (P .snd) + _≤S_ = PosetStr._≤_ (S .snd) + + posP = PosetStr.isPoset (P .snd) + posS = PosetStr.isPoset (S .snd) + + proP = isPoset→isProset posP + proS = isPoset→isProset posS + +IsPosetEquivRespectsJoin : {P : Poset ℓ₀ ℓ₀'} {S : Poset ℓ₁ ℓ₁'} (e : PosetEquiv P S) + → ∀ a b a∨b + → isJoin (isPoset→isProset (PosetStr.isPoset (P .snd))) a b a∨b + ≃ isJoin (isPoset→isProset (PosetStr.isPoset (S .snd))) + (equivFun (e .fst) a) + (equivFun (e .fst) b) + (equivFun (e .fst) a∨b) + +IsPosetEquivRespectsJoin {P = P} {S = S} (e , posEq) a b a∨b + = propBiimpl→Equiv (isPropIsJoin proP a b a∨b) + (isPropIsJoin proS (equivFun e a) (equivFun e b) (equivFun e a∨b)) + (λ join x + → IsPosetEquiv.pres≤⁻ posEq (equivFun e a∨b) x ∙ₑ + substEquiv (_≤P invEq e x) (retEq e a∨b) ∙ₑ + join (invEq e x) ∙ₑ + ≃-× (IsPosetEquiv.pres≤ posEq a (invEq e x) ∙ₑ + substEquiv (equivFun e a ≤S_) (secEq e x)) + (IsPosetEquiv.pres≤ posEq b (invEq e x) ∙ₑ + substEquiv (equivFun e b ≤S_) (secEq e x))) + λ join x + → IsPosetEquiv.pres≤ posEq a∨b x ∙ₑ + join (equivFun e x) ∙ₑ + ≃-× (IsPosetEquiv.pres≤⁻ posEq (equivFun e a) (equivFun e x) ∙ₑ + subst2Equiv _≤P_ (retEq e a) (retEq e x)) + (IsPosetEquiv.pres≤⁻ posEq (equivFun e b) (equivFun e x) ∙ₑ + subst2Equiv _≤P_ (retEq e b) (retEq e x)) + where _≤P_ = PosetStr._≤_ (P .snd) + _≤S_ = PosetStr._≤_ (S .snd) + + posP = PosetStr.isPoset (P .snd) + posS = PosetStr.isPoset (S .snd) + + proP = isPoset→isProset posP + proS = isPoset→isProset posS + +module _ + (pos : Poset ℓ ℓ') + where + private + pro = isPoset→isProset (PosetStr.isPoset (snd pos)) + + is = PosetStr.isPoset (snd pos) + + isMeetSemipseudolattice : Type (ℓ-max ℓ ℓ') + isMeetSemipseudolattice = ∀ a b → Meet pro a b + + isMeetCompleteSemipseudolattice : {ℓ'' : Level} → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓ'')) + isMeetCompleteSemipseudolattice {ℓ'' = ℓ''} = {B : Type ℓ''} + → (P : B → ⟨ pos ⟩) + → Infimum pro P + + isJoinSemipseudolattice : Type (ℓ-max ℓ ℓ') + isJoinSemipseudolattice = ∀ a b → Join pro a b + + isJoinCompleteSemipseudolattice : {ℓ'' : Level} → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓ'')) + isJoinCompleteSemipseudolattice {ℓ'' = ℓ''} = {B : Type ℓ''} + → (P : B → ⟨ pos ⟩) + → Supremum pro P + + isMeetSemilattice : Type (ℓ-max ℓ ℓ') + isMeetSemilattice = isMeetSemipseudolattice × Greatest pro (⟨ pos ⟩ , (id↪ _)) + + isJoinSemilattice : Type (ℓ-max ℓ ℓ') + isJoinSemilattice = isJoinSemipseudolattice × Least pro (⟨ pos ⟩ , (id↪ _)) + + isPseudolattice : Type (ℓ-max ℓ ℓ') + isPseudolattice = isMeetSemipseudolattice × isJoinSemipseudolattice + + isLattice : Type (ℓ-max ℓ ℓ') + isLattice = isMeetSemilattice × isJoinSemilattice + + isCompleteLattice : {ℓ'' : Level} → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓ'')) + isCompleteLattice {ℓ'' = ℓ''} = isMeetCompleteSemipseudolattice {ℓ''} × isJoinCompleteSemipseudolattice {ℓ''} + + -- These are all propositonal statements + isPropIsMeetSemipseudolattice : isProp isMeetSemipseudolattice + isPropIsMeetSemipseudolattice + = isPropΠ2 λ a b → MeetUnique is a b + + isPropIsMeetCompleteSemipseudolattice : {ℓ'' : Level} → isProp (isMeetCompleteSemipseudolattice {ℓ''}) + isPropIsMeetCompleteSemipseudolattice + = isPropImplicitΠ λ _ → isPropΠ λ _ → InfimumUnique is + + isPropIsJoinSemipseudolattice : isProp isJoinSemipseudolattice + isPropIsJoinSemipseudolattice + = isPropΠ2 λ a b → JoinUnique is a b + + isPropIsJoinCompleteSemipseudolattice : {ℓ'' : Level} → isProp (isJoinCompleteSemipseudolattice {ℓ''}) + isPropIsJoinCompleteSemipseudolattice + = isPropImplicitΠ λ _ → isPropΠ λ _ → SupremumUnique is + + isPropIsMeetSemilattice : isProp isMeetSemilattice + isPropIsMeetSemilattice = isProp× isPropIsMeetSemipseudolattice + (GreatestUnique is {P = ⟨ pos ⟩ , (id↪ ⟨ pos ⟩)}) + + isPropIsJoinSemilattice : isProp isJoinSemilattice + isPropIsJoinSemilattice = isProp× isPropIsJoinSemipseudolattice + (LeastUnique is {P = ⟨ pos ⟩ , (id↪ ⟨ pos ⟩)}) + + isPropIsPseudolattice : isProp isPseudolattice + isPropIsPseudolattice = isProp× isPropIsMeetSemipseudolattice + isPropIsJoinSemipseudolattice + + isPropIsLattice : isProp isLattice + isPropIsLattice = isProp× isPropIsMeetSemilattice + isPropIsJoinSemilattice + + isPropIsCompleteLattice : {ℓ'' : Level} → isProp (isCompleteLattice {ℓ''}) + isPropIsCompleteLattice = isProp× isPropIsMeetCompleteSemipseudolattice + isPropIsJoinCompleteSemipseudolattice + + -- Complete semipseudolattices are semipseudolattices + isMeetCompleteSemipseudolattice→isMeetSemipseudolattice : isMeetCompleteSemipseudolattice + → isMeetSemipseudolattice + isMeetCompleteSemipseudolattice→isMeetSemipseudolattice meet + = (λ a b → (meet fst .fst) , invEq (isMeet≃isInfimum pro a b _) + (meet fst .snd)) + + isJoinCompleteSemipseudolattice→isJoinSemipseudolattice : isJoinCompleteSemipseudolattice + → isJoinSemipseudolattice + isJoinCompleteSemipseudolattice→isJoinSemipseudolattice join + = (λ a b → (join fst .fst) , invEq (isJoin≃isSupremum pro a b _) + (join fst .snd)) + + -- They also imply each other, though we keep the two distinct for morphism reasons + isMeetCompleteSemipseudolattice≃isJoinCompleteSemipseudolattice : (isMeetCompleteSemipseudolattice {ℓ-max ℓ ℓ'}) + ≃ (isJoinCompleteSemipseudolattice {ℓ-max ℓ ℓ'}) + isMeetCompleteSemipseudolattice≃isJoinCompleteSemipseudolattice + = propBiimpl→Equiv isPropIsMeetCompleteSemipseudolattice + isPropIsJoinCompleteSemipseudolattice to from + where to : isMeetCompleteSemipseudolattice + → isJoinCompleteSemipseudolattice {ℓ-max ℓ ℓ'} + to meet P = (fst lemma) , + isInfimumOfUpperBounds→isSupremum pro P (fst lemma) + (snd lemma) + where P↑ : Type _ + P↑ = UpperBound pro P + + lemma = meet {B = P↑} fst + + from : isJoinCompleteSemipseudolattice + → isMeetCompleteSemipseudolattice {ℓ-max ℓ ℓ'} + from join P = (fst lemma) , + isSupremumOfLowerBounds→isInfimum pro P (fst lemma) + (snd lemma) + where P↓ : Type _ + P↓ = LowerBound pro P + + lemma = join {B = P↓} fst + + isCompleteLattice→isLattice : isCompleteLattice + → isLattice + isCompleteLattice→isLattice (meet , join) + = ((isMeetCompleteSemipseudolattice→isMeetSemipseudolattice meet) , + (join (λ x → x) .fst) , + isUpperBoundOfSelf→isGreatestOfSelf pro (join (λ x → x) .snd .fst)) , + ((isJoinCompleteSemipseudolattice→isJoinSemipseudolattice join) , + meet (λ x → x) .fst , + isLowerBoundOfSelf→isLeastOfSelf pro (meet (λ x → x) .snd .fst)) + + module _ + (lat : isPseudolattice) + where + private + _∧l_ : ⟨ pos ⟩ → ⟨ pos ⟩ → ⟨ pos ⟩ + a ∧l b = (lat .fst a b) .fst + + _∨l_ : ⟨ pos ⟩ → ⟨ pos ⟩ → ⟨ pos ⟩ + a ∨l b = (lat .snd a b) .fst + + set = IsPoset.is-set is + + rfl = IsPoset.is-refl is + + MeetDistLJoin : Type ℓ + MeetDistLJoin = ∀ a b c → a ∧l (b ∨l c) ≡ (a ∧l b) ∨l (a ∧l c) + + MeetDistRJoin : Type ℓ + MeetDistRJoin = ∀ a b c → (a ∨l b) ∧l c ≡ (a ∧l c) ∨l (b ∧l c) + + JoinDistLMeet : Type ℓ + JoinDistLMeet = ∀ a b c → a ∨l (b ∧l c) ≡ (a ∨l b) ∧l (a ∨l c) + + JoinDistRMeet : Type ℓ + JoinDistRMeet = ∀ a b c → (a ∧l b) ∨l c ≡ (a ∨l c) ∧l (b ∨l c) + + MeetAbsorbLJoin : ∀ a b → a ∧l (a ∨l b) ≡ a + MeetAbsorbLJoin a b + = sym (equivFun + (order≃meet + is a (a ∨l b) (a ∧l (a ∨l b)) + (lat .fst a (a ∨l b) .snd)) + (equivFun (lat .snd a b .snd (a ∨l b)) + (rfl (a ∨l b)) .fst)) + + MeetAbsorbRJoin : ∀ a b → (a ∨l b) ∧l a ≡ a + MeetAbsorbRJoin a b + = meetComm is (lat .fst) (a ∨l b) a ∙ + MeetAbsorbLJoin a b + + JoinAbsorbLMeet : ∀ a b → a ∨l (a ∧l b) ≡ a + JoinAbsorbLMeet a b + = sym (equivFun + (order≃join + is (a ∧l b) a (a ∨l (a ∧l b)) + (isJoinComm pro a (a ∧l b) (a ∨l (a ∧l b)) + (lat .snd a (a ∧l b) .snd))) + (equivFun (lat .fst a b .snd (a ∧l b)) (rfl (a ∧l b)) .fst)) + + JoinAbsorbRMeet : ∀ a b → (a ∧l b) ∨l a ≡ a + JoinAbsorbRMeet a b + = joinComm is (lat .snd) (a ∧l b) a ∙ + JoinAbsorbLMeet a b + + isPropMeetDistLJoin : isProp MeetDistLJoin + isPropMeetDistLJoin = isPropΠ3 λ _ _ _ → set _ _ + + isPropMeetDistRJoin : isProp MeetDistRJoin + isPropMeetDistRJoin = isPropΠ3 λ _ _ _ → set _ _ + + isPropJoinDistLMeet : isProp JoinDistLMeet + isPropJoinDistLMeet = isPropΠ3 λ _ _ _ → set _ _ + + isPropJoinDistRMeet : isProp JoinDistRMeet + isPropJoinDistRMeet = isPropΠ3 λ _ _ _ → set _ _ + + MeetDistLJoin≃MeetDistRJoin : MeetDistLJoin ≃ MeetDistRJoin + MeetDistLJoin≃MeetDistRJoin + = propBiimpl→Equiv isPropMeetDistLJoin + isPropMeetDistRJoin + (λ distl a b c → meetComm is (lat .fst) (a ∨l b) c ∙ + distl c a b ∙ + cong₂ _∨l_ (meetComm is (lat .fst) c a) + (meetComm is (lat .fst) c b)) + λ distr a b c → meetComm is (lat .fst) a (b ∨l c) ∙ + distr b c a ∙ + cong₂ _∨l_ (meetComm is (lat .fst) b a) + (meetComm is (lat .fst) c a) -Poset→StrictPoset : Poset ℓ ℓ' → StrictPoset ℓ (ℓ-max ℓ ℓ') -Poset→StrictPoset (_ , pos) - = _ , strictposetstr (BinaryRelation.IrreflKernel (PosetStr._≤_ pos)) - (isPoset→isStrictPosetIrreflKernel (PosetStr.isPoset pos)) + JoinDistLMeet≃JoinDistRMeet : JoinDistLMeet ≃ JoinDistRMeet + JoinDistLMeet≃JoinDistRMeet + = propBiimpl→Equiv isPropJoinDistLMeet + isPropJoinDistRMeet + (λ distl a b c → joinComm is (lat .snd) (a ∧l b) c ∙ + distl c a b ∙ + cong₂ _∧l_ (joinComm is (lat .snd) c a) + (joinComm is (lat .snd) c b)) + λ distr a b c → joinComm is (lat .snd) a (b ∧l c) ∙ + distr b c a ∙ + cong₂ _∧l_ (joinComm is (lat .snd) b a) + (joinComm is (lat .snd) c a) + MeetDistLJoin≃JoinDistLMeet : MeetDistLJoin ≃ JoinDistLMeet + MeetDistLJoin≃JoinDistLMeet + = propBiimpl→Equiv isPropMeetDistLJoin + isPropJoinDistLMeet + (λ dist a b c → (a ∨l (b ∧l c)) ≡⟨ cong (_∨l (b ∧l c)) (sym (JoinAbsorbLMeet a c)) ⟩ + ((a ∨l (a ∧l c)) ∨l (b ∧l c)) ≡⟨ sym (joinAssoc is (lat .snd) a (a ∧l c) (b ∧l c)) ⟩ + (a ∨l ((a ∧l c) ∨l (b ∧l c))) ≡⟨ cong (a ∨l_) (sym (equivFun MeetDistLJoin≃MeetDistRJoin dist a b c)) ⟩ + (a ∨l ((a ∨l b) ∧l c)) ≡⟨ cong (_∨l ((a ∨l b) ∧l c)) (sym (MeetAbsorbRJoin a b)) ⟩ + (((a ∨l b) ∧l a) ∨l ((a ∨l b) ∧l c)) ≡⟨ sym (dist (a ∨l b) a c) ⟩ + ((a ∨l b) ∧l (a ∨l c)) ∎) + λ dist' a b c → (a ∧l (b ∨l c)) ≡⟨ cong (_∧l (b ∨l c)) (sym (MeetAbsorbLJoin a c)) ⟩ + ((a ∧l (a ∨l c)) ∧l (b ∨l c)) ≡⟨ sym (meetAssoc is (lat .fst) a (a ∨l c) (b ∨l c)) ⟩ + (a ∧l ((a ∨l c) ∧l (b ∨l c))) ≡⟨ cong (a ∧l_) (sym (equivFun JoinDistLMeet≃JoinDistRMeet dist' a b c)) ⟩ + (a ∧l ((a ∧l b) ∨l c)) ≡⟨ cong (_∧l ((a ∧l b) ∨l c)) (sym (JoinAbsorbRMeet a b)) ⟩ + (((a ∧l b) ∨l a) ∧l ((a ∧l b) ∨l c)) ≡⟨ sym (dist' (a ∧l b) a c) ⟩ + ((a ∧l b) ∨l (a ∧l c)) ∎ -module PosetDownset (P' : Poset ℓ ℓ') where - private P = fst P' - open PosetStr (snd P') + MeetDistRJoin≃JoinDistRMeet : MeetDistRJoin ≃ JoinDistRMeet + MeetDistRJoin≃JoinDistRMeet = invEquiv MeetDistLJoin≃MeetDistRJoin ∙ₑ + MeetDistLJoin≃JoinDistLMeet ∙ₑ + JoinDistLMeet≃JoinDistRMeet - ↓ : P → Type (ℓ-max ℓ ℓ') - ↓ u = Σ[ v ∈ P ] v ≤ u + -- Since all of those varieties of distributivity are equivalent, we say that MeetDistLJoin is our canonical version of distributivity + isDistributive : Type ℓ + isDistributive = MeetDistLJoin - ↓ᴾ : P → Poset (ℓ-max ℓ ℓ') ℓ' - fst (↓ᴾ u) = ↓ u - PosetStr._≤_ (snd (↓ᴾ u)) v w = v .fst ≤ w .fst - PosetStr.isPoset (snd (↓ᴾ u)) = - isPosetInduced - (PosetStr.isPoset (snd P')) - _ - (EmbeddingΣProp (λ a → is-prop-valued _ _)) + isPropIsDistributive : isProp isDistributive + isPropIsDistributive = isPropMeetDistLJoin diff --git a/Cubical/Relation/Binary/Order/Poset/Subset.agda b/Cubical/Relation/Binary/Order/Poset/Subset.agda new file mode 100644 index 0000000000..9153602a41 --- /dev/null +++ b/Cubical/Relation/Binary/Order/Poset/Subset.agda @@ -0,0 +1,249 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Order.Poset.Subset where + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Transport + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ + +open import Cubical.Functions.Embedding +open import Cubical.Functions.Preimage + +open import Cubical.HITs.PropositionalTruncation as ∥₁ + +open import Cubical.Relation.Binary.Order.Poset.Base +open import Cubical.Relation.Binary.Order.Poset.Properties + + +private + variable + ℓ ℓ' ℓ'' ℓ''' ℓ₀ ℓ₁ : Level + +module _ + (P : Poset ℓ ℓ') + where + private + _≤_ = PosetStr._≤_ (snd P) + + is = PosetStr.isPoset (snd P) + prop = IsPoset.is-prop-valued is + + rfl = IsPoset.is-refl is + anti = IsPoset.is-antisym is + trans = IsPoset.is-trans is + + module _ + (S : Embedding ⟨ P ⟩ ℓ'') + where + private + f = S .snd .fst + isDownset : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isDownset = ∀ x → ∀ y → y ≤ f x → y ∈ₑ S + + isPropIsDownset : isProp isDownset + isPropIsDownset = isPropΠ3 λ _ y _ → isProp∈ₑ y S + + isUpset : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isUpset = ∀ x → ∀ y → f x ≤ y → y ∈ₑ S + + isPropIsUpset : isProp isUpset + isPropIsUpset = isPropΠ3 λ _ y _ → isProp∈ₑ y S + + module _ + (A : Embedding ⟨ P ⟩ ℓ₀) + (B : Embedding ⟨ P ⟩ ℓ₁) + where + module _ + (downA : isDownset A) + (downB : isDownset B) + where + isDownset∩ : isDownset (A ∩ₑ B) + isDownset∩ (x , (a , a≡x) , (b , b≡x)) y y≤x + = (y , (downA a y (subst (y ≤_) (sym a≡x) y≤x) , + downB b y (subst (y ≤_) (sym b≡x) y≤x))) , refl + + isDownset∪ : isDownset (A ∪ₑ B) + isDownset∪ (x , A⊎B) y y≤x + = ∥₁.rec (isProp∈ₑ y (A ∪ₑ B)) + (⊎.rec (λ { (a , a≡x) → + (y , ∣ inl (downA a y (subst (y ≤_) + (sym a≡x) y≤x)) ∣₁) , refl }) + (λ { (b , b≡x) → + (y , ∣ inr (downB b y (subst (y ≤_) + (sym b≡x) y≤x)) ∣₁) , refl })) A⊎B + + module _ + (upA : isUpset A) + (upB : isUpset B) + where + isUpset∩ : isUpset (A ∩ₑ B) + isUpset∩ (x , (a , a≡x) , (b , b≡x)) y x≤y + = (y , (upA a y (subst (_≤ y) (sym a≡x) x≤y) , + upB b y (subst (_≤ y) (sym b≡x) x≤y))) , refl + + isUpset∪ : isUpset (A ∪ₑ B) + isUpset∪ (x , A⊎B) y x≤y + = ∥₁.rec (isProp∈ₑ y (A ∪ₑ B)) + (⊎.rec (λ { (a , a≡x) → + (y , ∣ inl (upA a y (subst (_≤ y) + (sym a≡x) x≤y)) ∣₁) , refl }) + (λ { (b , b≡x) → + (y , ∣ inr (upB b y (subst (_≤ y) + (sym b≡x) x≤y)) ∣₁) , refl })) A⊎B + + module _ + {I : Type ℓ''} + (S : I → Embedding ⟨ P ⟩ ℓ''') + where + module _ + (downS : ∀ i → isDownset (S i)) + where + isDownset⋂ : isDownset (⋂ₑ S) + isDownset⋂ (x , ∀x) y y≤x + = (y , λ i → downS i (∀x i .fst) y (subst (y ≤_) (sym (∀x i .snd)) y≤x)) , refl + + isDownset⋃ : isDownset (⋃ₑ S) + isDownset⋃ (x , ∃i) y y≤x + = (y , (∥₁.map (λ { (i , a , a≡x) → i , downS i a y (subst (y ≤_) (sym a≡x) y≤x) }) ∃i)) , refl + + module _ + (upS : ∀ i → isUpset (S i)) + where + isUpset⋂ : isUpset (⋂ₑ S) + isUpset⋂ (x , ∀x) y x≤y + = (y , λ i → upS i (∀x i .fst) y (subst (_≤ y) (sym (∀x i .snd)) x≤y)) , refl + + isUpset⋃ : isUpset (⋃ₑ S) + isUpset⋃ (x , ∃i) y x≤y + = (y , (∥₁.map (λ { (i , a , a≡x) → i , upS i a y (subst (_≤ y) (sym a≡x) x≤y) }) ∃i)) , refl + + + module _ + (x : ⟨ P ⟩) + where + principalDownset : Embedding ⟨ P ⟩ (ℓ-max ℓ ℓ') + principalDownset = (Σ[ y ∈ ⟨ P ⟩ ] y ≤ x) , EmbeddingΣProp λ _ → prop _ x + + principalUpset : Embedding ⟨ P ⟩ (ℓ-max ℓ ℓ') + principalUpset = (Σ[ y ∈ ⟨ P ⟩ ] x ≤ y) , EmbeddingΣProp λ _ → prop x _ + + isDownsetPrincipalDownset : ∀ x → isDownset (principalDownset x) + isDownsetPrincipalDownset x (y , y≤x) z z≤y = (z , (trans z y x z≤y y≤x)) , refl + + isUpsetPrincipalUpset : ∀ x → isUpset (principalUpset x) + isUpsetPrincipalUpset x (y , x≤y) z y≤z = (z , (trans x y z x≤y y≤z)) , refl + + principalDownsetMembership : ∀ x y + → x ≤ y + ≃ x ∈ₑ principalDownset y + principalDownsetMembership x y + = propBiimpl→Equiv (prop x y) + (isProp∈ₑ x (principalDownset y)) + (λ x≤y → (x , x≤y) , refl) + λ { ((a , a≤y) , fib) → subst (_≤ y) fib a≤y} + + principalUpsetMembership : ∀ x y + → x ≤ y + ≃ y ∈ₑ principalUpset x + principalUpsetMembership x y + = propBiimpl→Equiv (prop x y) + (isProp∈ₑ y (principalUpset x)) + (λ x≤y → (y , x≤y) , refl) + λ { ((a , x≤a) , fib) → subst (x ≤_) fib x≤a } + + principalUpsetInclusion : ∀ x y + → x ≤ y + → principalUpset y ⊆ₑ principalUpset x + principalUpsetInclusion x y x≤y z z∈y↑ + = equivFun (principalUpsetMembership x z) + (trans x y z x≤y (invEq (principalUpsetMembership y z) z∈y↑)) + + principalDownsetInclusion : ∀ x y + → x ≤ y + → principalDownset x ⊆ₑ principalDownset y + principalDownsetInclusion x y x≤y z z∈x↓ + = equivFun (principalDownsetMembership z y) + (trans z x y (invEq (principalDownsetMembership z x) z∈x↓) x≤y) + + module _ + (S : Embedding ⟨ P ⟩ (ℓ-max ℓ ℓ')) + where + isPrincipalDownset : Type _ + isPrincipalDownset = Σ[ x ∈ ⟨ P ⟩ ] S ≡ (principalDownset x) + + isPropIsPrincipalDownset : isProp isPrincipalDownset + isPropIsPrincipalDownset (x , S≡x↓) (y , S≡y↓) + = Σ≡Prop (λ x → isSetEmbedding S (principalDownset x)) + (anti x y x≤y y≤x) + where x≤y : x ≤ y + x≤y = invEq (principalDownsetMembership x y) + (subst (x ∈ₑ_) + (sym S≡x↓ ∙ S≡y↓) + (equivFun (principalDownsetMembership x x) (rfl x))) + + y≤x : y ≤ x + y≤x = invEq (principalDownsetMembership y x) + (subst (y ∈ₑ_) + (sym S≡y↓ ∙ S≡x↓) + (equivFun (principalDownsetMembership y y) (rfl y))) + + isPrincipalUpset : Type _ + isPrincipalUpset = Σ[ x ∈ ⟨ P ⟩ ] S ≡ (principalUpset x) + + isPropIsPrincipalUpset : isProp isPrincipalUpset + isPropIsPrincipalUpset (x , S≡x↑) (y , S≡y↑) + = Σ≡Prop (λ x → isSetEmbedding S (principalUpset x)) + (anti x y x≤y y≤x) + where x≤y : x ≤ y + x≤y = invEq (principalUpsetMembership x y) + (subst (y ∈ₑ_) + (sym S≡y↑ ∙ S≡x↑) + (equivFun (principalUpsetMembership y y) (rfl y))) + + y≤x : y ≤ x + y≤x = invEq (principalUpsetMembership y x) + (subst (x ∈ₑ_) + (sym S≡x↑ ∙ S≡y↑) + (equivFun (principalUpsetMembership x x) (rfl x))) + + isPrincipalDownset→isDownset : isPrincipalDownset → (isDownset S) + isPrincipalDownset→isDownset (x , p) = transport⁻ (cong isDownset p) (isDownsetPrincipalDownset x) + + isPrincipalUpset→isUpset : isPrincipalUpset → (isUpset S) + isPrincipalUpset→isUpset (x , p) = transport⁻ (cong isUpset p) (isUpsetPrincipalUpset x) + +module PosetDownset (P' : Poset ℓ ℓ') where + private P = ⟨ P' ⟩ + open PosetStr (snd P') + + ↓ : P → Type (ℓ-max ℓ ℓ') + ↓ u = principalDownset P' u .fst + + ↓ᴾ : P → Poset (ℓ-max ℓ ℓ') ℓ' + fst (↓ᴾ u) = ↓ u + PosetStr._≤_ (snd (↓ᴾ u)) v w = v .fst ≤ w .fst + PosetStr.isPoset (snd (↓ᴾ u)) = + isPosetInduced + (PosetStr.isPoset (snd P')) + _ + (principalDownset P' u .snd) + +module PosetUpset (P' : Poset ℓ ℓ') where + private P = ⟨ P' ⟩ + open PosetStr (snd P') + + ↑ : P → Type (ℓ-max ℓ ℓ') + ↑ u = principalUpset P' u .fst + + ↑ᴾ : P → Poset (ℓ-max ℓ ℓ') ℓ' + fst (↑ᴾ u) = principalUpset P' u .fst + PosetStr._≤_ (snd (↑ᴾ u)) v w = v .fst ≤ w .fst + PosetStr.isPoset (snd (↑ᴾ u)) = + isPosetInduced + (PosetStr.isPoset (snd P')) + _ + (principalUpset P' u .snd) diff --git a/Cubical/Relation/Binary/Order/Preorder.agda b/Cubical/Relation/Binary/Order/Preorder.agda deleted file mode 100644 index 6b9ce2af2f..0000000000 --- a/Cubical/Relation/Binary/Order/Preorder.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Relation.Binary.Order.Preorder where - -open import Cubical.Relation.Binary.Order.Preorder.Base public -open import Cubical.Relation.Binary.Order.Preorder.Properties public diff --git a/Cubical/Relation/Binary/Order/Preorder/Base.agda b/Cubical/Relation/Binary/Order/Preorder/Base.agda deleted file mode 100644 index 5bf7dfd797..0000000000 --- a/Cubical/Relation/Binary/Order/Preorder/Base.agda +++ /dev/null @@ -1,135 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Relation.Binary.Order.Preorder.Base where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Transport -open import Cubical.Foundations.SIP - -open import Cubical.Data.Sigma - -open import Cubical.Reflection.RecordEquiv -open import Cubical.Reflection.StrictEquiv - -open import Cubical.Displayed.Base -open import Cubical.Displayed.Auto -open import Cubical.Displayed.Record -open import Cubical.Displayed.Universe - -open import Cubical.Relation.Binary.Base - -open Iso -open BinaryRelation - - -private - variable - ℓ ℓ' ℓ'' ℓ₀ ℓ₀' ℓ₁ ℓ₁' : Level - -record IsPreorder {A : Type ℓ} (_≲_ : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where - no-eta-equality - constructor ispreorder - - field - is-set : isSet A - is-prop-valued : isPropValued _≲_ - is-refl : isRefl _≲_ - is-trans : isTrans _≲_ - -unquoteDecl IsPreorderIsoΣ = declareRecordIsoΣ IsPreorderIsoΣ (quote IsPreorder) - - -record PreorderStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc ℓ')) where - - constructor preorderstr - - field - _≲_ : A → A → Type ℓ' - isPreorder : IsPreorder _≲_ - - infixl 7 _≲_ - - open IsPreorder isPreorder public - -Preorder : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) -Preorder ℓ ℓ' = TypeWithStr ℓ (PreorderStr ℓ') - -preorder : (A : Type ℓ) (_≲_ : A → A → Type ℓ') (h : IsPreorder _≲_) → Preorder ℓ ℓ' -preorder A _≲_ h = A , preorderstr _≲_ h - -record IsPreorderEquiv {A : Type ℓ₀} {B : Type ℓ₁} - (M : PreorderStr ℓ₀' A) (e : A ≃ B) (N : PreorderStr ℓ₁' B) - : Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') ℓ₁') - where - constructor - ispreorderequiv - -- Shorter qualified names - private - module M = PreorderStr M - module N = PreorderStr N - - field - pres≲ : (x y : A) → x M.≲ y ≃ equivFun e x N.≲ equivFun e y - - -PreorderEquiv : (M : Preorder ℓ₀ ℓ₀') (M : Preorder ℓ₁ ℓ₁') → Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) -PreorderEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsPreorderEquiv (M .snd) e (N .snd) - -isPropIsPreorder : {A : Type ℓ} (_≲_ : A → A → Type ℓ') → isProp (IsPreorder _≲_) -isPropIsPreorder _≲_ = isOfHLevelRetractFromIso 1 IsPreorderIsoΣ - (isPropΣ isPropIsSet - λ isSetA → isPropΣ (isPropΠ2 (λ _ _ → isPropIsProp)) - λ isPropValued≲ → isProp× - (isPropΠ (λ _ → isPropValued≲ _ _)) - (isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isPropValued≲ _ _)) - -𝒮ᴰ-Preorder : DUARel (𝒮-Univ ℓ) (PreorderStr ℓ') (ℓ-max ℓ ℓ') -𝒮ᴰ-Preorder = - 𝒮ᴰ-Record (𝒮-Univ _) IsPreorderEquiv - (fields: - data[ _≲_ ∣ autoDUARel _ _ ∣ pres≲ ] - prop[ isPreorder ∣ (λ _ _ → isPropIsPreorder _) ]) - where - open PreorderStr - open IsPreorder - open IsPreorderEquiv - -PreorderPath : (M N : Preorder ℓ ℓ') → PreorderEquiv M N ≃ (M ≡ N) -PreorderPath = ∫ 𝒮ᴰ-Preorder .UARel.ua - --- an easier way of establishing an equivalence of preorders -module _ {P : Preorder ℓ₀ ℓ₀'} {S : Preorder ℓ₁ ℓ₁'} (e : ⟨ P ⟩ ≃ ⟨ S ⟩) where - private - module P = PreorderStr (P .snd) - module S = PreorderStr (S .snd) - - module _ (isMon : ∀ x y → x P.≲ y → equivFun e x S.≲ equivFun e y) - (isMonInv : ∀ x y → x S.≲ y → invEq e x P.≲ invEq e y) where - open IsPreorderEquiv - open IsPreorder - - makeIsPreorderEquiv : IsPreorderEquiv (P .snd) e (S .snd) - pres≲ makeIsPreorderEquiv x y = propBiimpl→Equiv (P.isPreorder .is-prop-valued _ _) - (S.isPreorder .is-prop-valued _ _) - (isMon _ _) (isMonInv' _ _) - where - isMonInv' : ∀ x y → equivFun e x S.≲ equivFun e y → x P.≲ y - isMonInv' x y ex≲ey = transport (λ i → retEq e x i P.≲ retEq e y i) (isMonInv _ _ ex≲ey) - - -module PreorderReasoning (P' : Preorder ℓ ℓ') where - private P = fst P' - open PreorderStr (snd P') - open IsPreorder - - _≲⟨_⟩_ : (x : P) {y z : P} → x ≲ y → y ≲ z → x ≲ z - x ≲⟨ p ⟩ q = isPreorder .is-trans x _ _ p q - - _◾ : (x : P) → x ≲ x - x ◾ = isPreorder .is-refl x - - infixr 0 _≲⟨_⟩_ - infix 1 _◾ diff --git a/Cubical/Relation/Binary/Order/Preorder/Properties.agda b/Cubical/Relation/Binary/Order/Preorder/Properties.agda deleted file mode 100644 index deec933ddf..0000000000 --- a/Cubical/Relation/Binary/Order/Preorder/Properties.agda +++ /dev/null @@ -1,57 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Relation.Binary.Order.Preorder.Properties where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels - -open import Cubical.Functions.Embedding - -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.StrictPoset.Base - -open import Cubical.Relation.Nullary - -private - variable - ℓ ℓ' ℓ'' : Level - -module _ - {A : Type ℓ} - {R : Rel A A ℓ'} - where - - open BinaryRelation - - isPreorder→isEquivRelSymKernel : IsPreorder R → isEquivRel (SymKernel R) - isPreorder→isEquivRelSymKernel preorder - = equivRel (λ a → (IsPreorder.is-refl preorder a) - , (IsPreorder.is-refl preorder a)) - (isSymSymKernel R) - (λ a b c (Rab , Rba) (Rbc , Rcb) - → IsPreorder.is-trans preorder a b c Rab Rbc - , IsPreorder.is-trans preorder c b a Rcb Rba) - - isPreorder→isStrictPosetAsymKernel : IsPreorder R → IsStrictPoset (AsymKernel R) - isPreorder→isStrictPosetAsymKernel preorder - = isstrictposet (IsPreorder.is-set preorder) - (λ a b → isProp× (IsPreorder.is-prop-valued preorder a b) (isProp¬ (R b a))) - (λ a (Raa , ¬Raa) → ¬Raa (IsPreorder.is-refl preorder a)) - (λ a b c (Rab , ¬Rba) (Rbc , ¬Rcb) - → IsPreorder.is-trans preorder a b c Rab Rbc - , λ Rca → ¬Rcb (IsPreorder.is-trans preorder c a b Rca Rab)) - (isAsymAsymKernel R) - - isPreorderInduced : IsPreorder R → (B : Type ℓ'') → (f : B ↪ A) → IsPreorder (InducedRelation R (B , f)) - isPreorderInduced pre B (f , emb) - = ispreorder (Embedding-into-isSet→isSet (f , emb) (IsPreorder.is-set pre)) - (λ a b → IsPreorder.is-prop-valued pre (f a) (f b)) - (λ a → IsPreorder.is-refl pre (f a)) - λ a b c → IsPreorder.is-trans pre (f a) (f b) (f c) - -Preorder→StrictPoset : Preorder ℓ ℓ' → StrictPoset ℓ ℓ' -Preorder→StrictPoset (_ , pre) - = _ , strictposetstr (BinaryRelation.AsymKernel (PreorderStr._≲_ pre)) - (isPreorder→isStrictPosetAsymKernel (PreorderStr.isPreorder pre)) diff --git a/Cubical/Relation/Binary/Order/Properties.agda b/Cubical/Relation/Binary/Order/Properties.agda deleted file mode 100644 index 5f1eaaafdb..0000000000 --- a/Cubical/Relation/Binary/Order/Properties.agda +++ /dev/null @@ -1,302 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Relation.Binary.Order.Properties where - -open import Cubical.Data.Sigma -open import Cubical.Data.Sum as ⊎ - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels - -open import Cubical.Functions.Embedding - -open import Cubical.HITs.PropositionalTruncation as ∥₁ - -open import Cubical.Relation.Binary.Base -open import Cubical.Relation.Binary.Order.Poset -open import Cubical.Relation.Binary.Order.Toset -open import Cubical.Relation.Binary.Order.Preorder.Base - -private - variable - ℓ ℓ' ℓ'' : Level - -module _ - {A : Type ℓ} - {_≲_ : Rel A A ℓ'} - (pre : IsPreorder _≲_) - where - - private - prop : ∀ a b → isProp (a ≲ b) - prop = IsPreorder.is-prop-valued pre - - module _ - (P : Embedding A ℓ'') - where - - private - subtype : Type ℓ'' - subtype = fst P - - toA : subtype → A - toA = fst (snd P) - - isMinimal : (n : subtype) → Type (ℓ-max ℓ' ℓ'') - isMinimal n = (x : subtype) → toA x ≲ toA n → toA n ≲ toA x - - isPropIsMinimal : ∀ n → isProp (isMinimal n) - isPropIsMinimal n = isPropΠ2 λ x _ → prop (toA n) (toA x) - - Minimal : Type (ℓ-max ℓ' ℓ'') - Minimal = Σ[ n ∈ subtype ] isMinimal n - - isMaximal : (n : subtype) → Type (ℓ-max ℓ' ℓ'') - isMaximal n = (x : subtype) → toA n ≲ toA x → toA x ≲ toA n - - isPropIsMaximal : ∀ n → isProp (isMaximal n) - isPropIsMaximal n = isPropΠ2 λ x _ → prop (toA x) (toA n) - - Maximal : Type (ℓ-max ℓ' ℓ'') - Maximal = Σ[ n ∈ subtype ] isMaximal n - - isLeast : (n : subtype) → Type (ℓ-max ℓ' ℓ'') - isLeast n = (x : subtype) → toA n ≲ toA x - - isPropIsLeast : ∀ n → isProp (isLeast n) - isPropIsLeast n = isPropΠ λ x → prop (toA n) (toA x) - - Least : Type (ℓ-max ℓ' ℓ'') - Least = Σ[ n ∈ subtype ] isLeast n - - isGreatest : (n : subtype) → Type (ℓ-max ℓ' ℓ'') - isGreatest n = (x : subtype) → toA x ≲ toA n - - isPropIsGreatest : ∀ n → isProp (isGreatest n) - isPropIsGreatest n = isPropΠ λ x → prop (toA x) (toA n) - - Greatest : Type (ℓ-max ℓ' ℓ'') - Greatest = Σ[ n ∈ subtype ] isGreatest n - - module _ - {B : Type ℓ''} - (P : B → A) - where - - isLowerBound : (n : A) → Type (ℓ-max ℓ' ℓ'') - isLowerBound n = (x : B) → n ≲ P x - - isPropIsLowerBound : ∀ n → isProp (isLowerBound n) - isPropIsLowerBound n = isPropΠ λ x → prop n (P x) - - LowerBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - LowerBound = Σ[ n ∈ A ] isLowerBound n - - isUpperBound : (n : A) → Type (ℓ-max ℓ' ℓ'') - isUpperBound n = (x : B) → P x ≲ n - - isPropIsUpperBound : ∀ n → isProp (isUpperBound n) - isPropIsUpperBound n = isPropΠ λ x → prop (P x) n - - UpperBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - UpperBound = Σ[ n ∈ A ] isUpperBound n - - module _ - {B : Type ℓ''} - (P : B → A) - where - - isMaximalLowerBound : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - isMaximalLowerBound n - = Σ[ islb ∈ isLowerBound P n ] - isMaximal (LowerBound P - , (EmbeddingΣProp (isPropIsLowerBound P))) (n , islb) - - isPropIsMaximalLowerBound : ∀ n → isProp (isMaximalLowerBound n) - isPropIsMaximalLowerBound n - = isPropΣ (isPropIsLowerBound P n) - λ islb → isPropIsMaximal (LowerBound P - , EmbeddingΣProp (isPropIsLowerBound P)) (n , islb) - - MaximalLowerBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - MaximalLowerBound = Σ[ n ∈ A ] isMaximalLowerBound n - - isMinimalUpperBound : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - isMinimalUpperBound n - = Σ[ isub ∈ isUpperBound P n ] - isMinimal (UpperBound P - , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) - - isPropIsMinimalUpperBound : ∀ n → isProp (isMinimalUpperBound n) - isPropIsMinimalUpperBound n - = isPropΣ (isPropIsUpperBound P n) - λ isub → isPropIsMinimal (UpperBound P - , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) - - MinimalUpperBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - MinimalUpperBound = Σ[ n ∈ A ] isMinimalUpperBound n - - isInfimum : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - isInfimum n - = Σ[ islb ∈ isLowerBound P n ] - isGreatest (LowerBound P - , EmbeddingΣProp (isPropIsLowerBound P)) (n , islb) - - isPropIsInfimum : ∀ n → isProp (isInfimum n) - isPropIsInfimum n - = isPropΣ (isPropIsLowerBound P n) - λ islb → isPropIsGreatest (LowerBound P - , EmbeddingΣProp (isPropIsLowerBound P)) (n , islb) - - Infimum : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - Infimum = Σ[ n ∈ A ] isInfimum n - - isSupremum : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - isSupremum n - = Σ[ isub ∈ isUpperBound P n ] - isLeast (UpperBound P - , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) - - isPropIsSupremum : ∀ n → isProp (isSupremum n) - isPropIsSupremum n - = isPropΣ (isPropIsUpperBound P n) - λ isub → isPropIsLeast (UpperBound P - , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) - - Supremum : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - Supremum = Σ[ n ∈ A ] isSupremum n - -module _ - {A : Type ℓ} - {_≲_ : Rel A A ℓ'} - (pre : IsPreorder _≲_) - where - - module _ - {P : Embedding A ℓ''} - where - - private - toA : (fst P) → A - toA = fst (snd P) - - isLeast→isMinimal : ∀ n → isLeast pre P n → isMinimal pre P n - isLeast→isMinimal _ isl x _ = isl x - - isGreatest→isMaximal : ∀ n → isGreatest pre P n → isMaximal pre P n - isGreatest→isMaximal _ isg x _ = isg x - - isLeast→isLowerBound : ∀ n → isLeast pre P n → isLowerBound pre toA (toA n) - isLeast→isLowerBound _ isl = isl - - isGreatest→isUpperBound : ∀ n → isGreatest pre P n → isUpperBound pre toA (toA n) - isGreatest→isUpperBound _ isg = isg - - isLeast→isInfimum : ∀ n → isLeast pre P n → isInfimum pre toA (toA n) - isLeast→isInfimum n isl = (isLeast→isLowerBound n isl) , (λ x → snd x n) - - isGreatest→isSupremum : ∀ n → isGreatest pre P n → isSupremum pre toA (toA n) - isGreatest→isSupremum n isg = (isGreatest→isUpperBound n isg) , (λ x → snd x n) - - module _ - {B : Type ℓ''} - {P : B → A} - where - - isInfimum→isLowerBound : ∀ n → isInfimum pre P n → isLowerBound pre P n - isInfimum→isLowerBound _ = fst - - isSupremum→isUpperBound : ∀ n → isSupremum pre P n → isUpperBound pre P n - isSupremum→isUpperBound _ = fst - -module _ - {A : Type ℓ} - {_≤_ : Rel A A ℓ'} - (pos : IsPoset _≤_) - where - - private - pre : IsPreorder _≤_ - pre = isPoset→isPreorder pos - - anti : BinaryRelation.isAntisym _≤_ - anti = IsPoset.is-antisym pos - - module _ - {P : Embedding A ℓ''} - where - - private - toA : (fst P) → A - toA = fst (snd P) - - emb : isEmbedding toA - emb = snd (snd P) - - isLeast→ContrMinimal : ∀ n → isLeast pre P n → ∀ m → isMinimal pre P m → n ≡ m - isLeast→ContrMinimal n isln m ismm - = isEmbedding→Inj emb n m (anti (toA n) (toA m) (isln m) (ismm n (isln m))) - - isGreatest→ContrMaximal : ∀ n → isGreatest pre P n → ∀ m → isMaximal pre P m → n ≡ m - isGreatest→ContrMaximal n isgn m ismm - = isEmbedding→Inj emb n m (anti (toA n) (toA m) (ismm n (isgn m)) (isgn m)) - - isLeastUnique : ∀ n m → isLeast pre P n → isLeast pre P m → n ≡ m - isLeastUnique n m isln islm - = isEmbedding→Inj emb n m (anti (toA n) (toA m) (isln m) (islm n)) - - isGreatestUnique : ∀ n m → isGreatest pre P n → isGreatest pre P m → n ≡ m - isGreatestUnique n m isgn isgm - = isEmbedding→Inj emb n m (anti (toA n) (toA m) (isgm n) (isgn m)) - - module _ - {B : Type ℓ''} - {P : B → A} - where - - isInfimum→ContrMaximalLowerBound : ∀ n → isInfimum pre P n - → ∀ m → isMaximalLowerBound pre P m - → n ≡ m - isInfimum→ContrMaximalLowerBound n (isln , isglbn) m (islm , ismlbm) - = anti n m (ismlbm (n , isln) (isglbn (m , islm))) (isglbn (m , islm)) - - isSupremum→ContrMinimalUpperBound : ∀ n → isSupremum pre P n - → ∀ m → isMinimalUpperBound pre P m - → n ≡ m - isSupremum→ContrMinimalUpperBound n (isun , islubn) m (isum , ismubm) - = anti n m (islubn (m , isum)) (ismubm (n , isun) (islubn (m , isum))) - - isInfimumUnique : ∀ n m → isInfimum pre P n → isInfimum pre P m → n ≡ m - isInfimumUnique n m (isln , infn) (islm , infm) - = anti n m (infm (n , isln)) (infn (m , islm)) - - isSupremumUnique : ∀ n m → isSupremum pre P n → isSupremum pre P m → n ≡ m - isSupremumUnique n m (isun , supn) (isum , supm) - = anti n m (supn (m , isum)) (supm (n , isun)) - -module _ - {A : Type ℓ} - {P : Embedding A ℓ''} - {_≤_ : Rel A A ℓ'} - (tos : IsToset _≤_) - where - - private - prop : ∀ a b → isProp (a ≤ b) - prop = IsToset.is-prop-valued tos - - conn : BinaryRelation.isStronglyConnected _≤_ - conn = IsToset.is-strongly-connected tos - - pre : IsPreorder _≤_ - pre = isPoset→isPreorder (isToset→isPoset tos) - - toA : (fst P) → A - toA = fst (snd P) - - isMinimal→isLeast : ∀ n → isMinimal pre P n → isLeast pre P n - isMinimal→isLeast n ism m - = ∥₁.rec (prop _ _) (⊎.rec (λ n≤m → n≤m) (λ m≤n → ism m m≤n)) (conn (toA n) (toA m)) - - isMaximal→isGreatest : ∀ n → isMaximal pre P n → isGreatest pre P n - isMaximal→isGreatest n ism m - = ∥₁.rec (prop _ _) (⊎.rec (λ n≤m → ism m n≤m) (λ m≤n → m≤n)) (conn (toA n) (toA m)) diff --git a/Cubical/Relation/Binary/Order/Proset.agda b/Cubical/Relation/Binary/Order/Proset.agda new file mode 100644 index 0000000000..c96c4cc15c --- /dev/null +++ b/Cubical/Relation/Binary/Order/Proset.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Order.Proset where + +open import Cubical.Relation.Binary.Order.Proset.Base public +open import Cubical.Relation.Binary.Order.Proset.Properties public diff --git a/Cubical/Relation/Binary/Order/Proset/Base.agda b/Cubical/Relation/Binary/Order/Proset/Base.agda new file mode 100644 index 0000000000..dc80d72b29 --- /dev/null +++ b/Cubical/Relation/Binary/Order/Proset/Base.agda @@ -0,0 +1,138 @@ +{-# OPTIONS --safe #-} +{- + Prosets are preordered sets; a set with a reflexive, transitive relation +-} +module Cubical.Relation.Binary.Order.Proset.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.RecordEquiv +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Relation.Binary.Base + +open Iso +open BinaryRelation + + +private + variable + ℓ ℓ' ℓ'' ℓ₀ ℓ₀' ℓ₁ ℓ₁' : Level + +record IsProset {A : Type ℓ} (_≲_ : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + constructor isproset + + field + is-set : isSet A + is-prop-valued : isPropValued _≲_ + is-refl : isRefl _≲_ + is-trans : isTrans _≲_ + +unquoteDecl IsProsetIsoΣ = declareRecordIsoΣ IsProsetIsoΣ (quote IsProset) + + +record ProsetStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc ℓ')) where + + constructor prosetstr + + field + _≲_ : A → A → Type ℓ' + isProset : IsProset _≲_ + + infixl 7 _≲_ + + open IsProset isProset public + +Proset : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) +Proset ℓ ℓ' = TypeWithStr ℓ (ProsetStr ℓ') + +proset : (A : Type ℓ) → (_≲_ : Rel A A ℓ') → IsProset _≲_ → Proset ℓ ℓ' +proset A _≲_ pros = A , (prosetstr _≲_ pros) + +record IsProsetEquiv {A : Type ℓ₀} {B : Type ℓ₁} + (M : ProsetStr ℓ₀' A) (e : A ≃ B) (N : ProsetStr ℓ₁' B) + : Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') ℓ₁') + where + constructor + isprosetequiv + -- Shorter qualified names + private + module M = ProsetStr M + module N = ProsetStr N + + field + pres≲ : (x y : A) → x M.≲ y ≃ equivFun e x N.≲ equivFun e y + + +ProsetEquiv : (M : Proset ℓ₀ ℓ₀') (M : Proset ℓ₁ ℓ₁') → Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) +ProsetEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsProsetEquiv (M .snd) e (N .snd) + +isPropIsProset : {A : Type ℓ} (_≲_ : A → A → Type ℓ') → isProp (IsProset _≲_) +isPropIsProset _≲_ = isOfHLevelRetractFromIso 1 IsProsetIsoΣ + (isPropΣ isPropIsSet + λ isSetA → isPropΣ (isPropΠ2 (λ _ _ → isPropIsProp)) + λ isPropValued≲ → isProp× + (isPropΠ (λ _ → isPropValued≲ _ _)) + (isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isPropValued≲ _ _)) + +𝒮ᴰ-Proset : DUARel (𝒮-Univ ℓ) (ProsetStr ℓ') (ℓ-max ℓ ℓ') +𝒮ᴰ-Proset = + 𝒮ᴰ-Record (𝒮-Univ _) IsProsetEquiv + (fields: + data[ _≲_ ∣ autoDUARel _ _ ∣ pres≲ ] + prop[ isProset ∣ (λ _ _ → isPropIsProset _) ]) + where + open ProsetStr + open IsProset + open IsProsetEquiv + +ProsetPath : (M N : Proset ℓ ℓ') → ProsetEquiv M N ≃ (M ≡ N) +ProsetPath = ∫ 𝒮ᴰ-Proset .UARel.ua + +-- an easier way of establishing an equivalence of prosets +module _ {P : Proset ℓ₀ ℓ₀'} {S : Proset ℓ₁ ℓ₁'} (e : ⟨ P ⟩ ≃ ⟨ S ⟩) where + private + module P = ProsetStr (P .snd) + module S = ProsetStr (S .snd) + + module _ (isMon : ∀ x y → x P.≲ y → equivFun e x S.≲ equivFun e y) + (isMonInv : ∀ x y → x S.≲ y → invEq e x P.≲ invEq e y) where + open IsProsetEquiv + open IsProset + + makeIsProsetEquiv : IsProsetEquiv (P .snd) e (S .snd) + pres≲ makeIsProsetEquiv x y = propBiimpl→Equiv (P.isProset .is-prop-valued _ _) + (S.isProset .is-prop-valued _ _) + (isMon _ _) (isMonInv' _ _) + where + isMonInv' : ∀ x y → equivFun e x S.≲ equivFun e y → x P.≲ y + isMonInv' x y ex≲ey = transport (λ i → retEq e x i P.≲ retEq e y i) (isMonInv _ _ ex≲ey) + + +module ProsetReasoning (P' : Proset ℓ ℓ') where + private P = fst P' + open ProsetStr (snd P') + open IsProset + + _≲⟨_⟩_ : (x : P) {y z : P} → x ≲ y → y ≲ z → x ≲ z + x ≲⟨ p ⟩ q = isProset .is-trans x _ _ p q + + _◾ : (x : P) → x ≲ x + x ◾ = isProset .is-refl x + + infixr 0 _≲⟨_⟩_ + infix 1 _◾ diff --git a/Cubical/Relation/Binary/Order/Proset/Properties.agda b/Cubical/Relation/Binary/Order/Proset/Properties.agda new file mode 100644 index 0000000000..61343c624a --- /dev/null +++ b/Cubical/Relation/Binary/Order/Proset/Properties.agda @@ -0,0 +1,434 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Order.Proset.Properties where + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude + +open import Cubical.Functions.Embedding + +open import Cubical.HITs.PropositionalTruncation as ∥₁ + +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.Binary.Order.Proset.Base +open import Cubical.Relation.Binary.Order.Quoset.Base + +open import Cubical.Relation.Nullary + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ + {A : Type ℓ} + {R : Rel A A ℓ'} + where + + open BinaryRelation + + isProset→isEquivRelSymKernel : IsProset R → isEquivRel (SymKernel R) + isProset→isEquivRelSymKernel proset + = equivRel (λ a → (IsProset.is-refl proset a) + , (IsProset.is-refl proset a)) + (isSymSymKernel R) + (λ a b c (Rab , Rba) (Rbc , Rcb) + → IsProset.is-trans proset a b c Rab Rbc + , IsProset.is-trans proset c b a Rcb Rba) + + isProset→isQuosetAsymKernel : IsProset R → IsQuoset (AsymKernel R) + isProset→isQuosetAsymKernel proset + = isquoset (IsProset.is-set proset) + (λ a b → isProp× (IsProset.is-prop-valued proset a b) (isProp¬ (R b a))) + (λ a (Raa , ¬Raa) → ¬Raa (IsProset.is-refl proset a)) + (λ a b c (Rab , ¬Rba) (Rbc , ¬Rcb) + → IsProset.is-trans proset a b c Rab Rbc + , λ Rca → ¬Rcb (IsProset.is-trans proset c a b Rca Rab)) + (isAsymAsymKernel R) + + isProsetInduced : IsProset R → (B : Type ℓ'') → (f : B ↪ A) → IsProset (InducedRelation R (B , f)) + isProsetInduced pre B (f , emb) + = isproset (Embedding-into-isSet→isSet (f , emb) (IsProset.is-set pre)) + (λ a b → IsProset.is-prop-valued pre (f a) (f b)) + (λ a → IsProset.is-refl pre (f a)) + λ a b c → IsProset.is-trans pre (f a) (f b) (f c) + + isProsetDual : IsProset R → IsProset (Dual R) + isProsetDual pre + = isproset (IsProset.is-set pre) + (λ a b → IsProset.is-prop-valued pre b a) + (IsProset.is-refl pre) (λ a b c Rab Rbc → IsProset.is-trans pre c b a Rbc Rab) + +Proset→Quoset : Proset ℓ ℓ' → Quoset ℓ ℓ' +Proset→Quoset (_ , pre) + = quoset _ (BinaryRelation.AsymKernel (ProsetStr._≲_ pre)) + (isProset→isQuosetAsymKernel (ProsetStr.isProset pre)) + +DualProset : Proset ℓ ℓ' → Proset ℓ ℓ' +DualProset (_ , pre) + = proset _ (BinaryRelation.Dual (ProsetStr._≲_ pre)) + (isProsetDual (ProsetStr.isProset pre)) + +module _ + {A : Type ℓ} + {_≲_ : Rel A A ℓ'} + (pre : IsProset _≲_) + where + + private + prop = IsProset.is-prop-valued pre + + rfl = IsProset.is-refl pre + + trans = IsProset.is-trans pre + + module _ + (P : Embedding A ℓ'') + where + + private + subtype = fst P + + toA = fst (snd P) + + isMinimal : (n : subtype) → Type (ℓ-max ℓ' ℓ'') + isMinimal n = (x : subtype) → toA x ≲ toA n → toA n ≲ toA x + + isPropIsMinimal : ∀ n → isProp (isMinimal n) + isPropIsMinimal n = isPropΠ2 λ x _ → prop (toA n) (toA x) + + Minimal : Type (ℓ-max ℓ' ℓ'') + Minimal = Σ[ n ∈ subtype ] isMinimal n + + isMaximal : (n : subtype) → Type (ℓ-max ℓ' ℓ'') + isMaximal n = (x : subtype) → toA n ≲ toA x → toA x ≲ toA n + + isPropIsMaximal : ∀ n → isProp (isMaximal n) + isPropIsMaximal n = isPropΠ2 λ x _ → prop (toA x) (toA n) + + Maximal : Type (ℓ-max ℓ' ℓ'') + Maximal = Σ[ n ∈ subtype ] isMaximal n + + isLeast : (n : subtype) → Type (ℓ-max ℓ' ℓ'') + isLeast n = (x : subtype) → toA n ≲ toA x + + isPropIsLeast : ∀ n → isProp (isLeast n) + isPropIsLeast n = isPropΠ λ x → prop (toA n) (toA x) + + Least : Type (ℓ-max ℓ' ℓ'') + Least = Σ[ n ∈ subtype ] isLeast n + + isGreatest : (n : subtype) → Type (ℓ-max ℓ' ℓ'') + isGreatest n = (x : subtype) → toA x ≲ toA n + + isPropIsGreatest : ∀ n → isProp (isGreatest n) + isPropIsGreatest n = isPropΠ λ x → prop (toA x) (toA n) + + Greatest : Type (ℓ-max ℓ' ℓ'') + Greatest = Σ[ n ∈ subtype ] isGreatest n + + module _ + {B : Type ℓ''} + (P : B → A) + where + + isLowerBound : (n : A) → Type (ℓ-max ℓ' ℓ'') + isLowerBound n = (x : B) → n ≲ P x + + isPropIsLowerBound : ∀ n → isProp (isLowerBound n) + isPropIsLowerBound n = isPropΠ λ x → prop n (P x) + + LowerBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + LowerBound = Σ[ n ∈ A ] isLowerBound n + + isUpperBound : (n : A) → Type (ℓ-max ℓ' ℓ'') + isUpperBound n = (x : B) → P x ≲ n + + isPropIsUpperBound : ∀ n → isProp (isUpperBound n) + isPropIsUpperBound n = isPropΠ λ x → prop (P x) n + + UpperBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + UpperBound = Σ[ n ∈ A ] isUpperBound n + + module _ + {B : Type ℓ''} + (P : B → A) + where + + isMaximalLowerBound : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isMaximalLowerBound n + = Σ[ islb ∈ isLowerBound P n ] + isMaximal (LowerBound P + , (EmbeddingΣProp (isPropIsLowerBound P))) (n , islb) + + isPropIsMaximalLowerBound : ∀ n → isProp (isMaximalLowerBound n) + isPropIsMaximalLowerBound n + = isPropΣ (isPropIsLowerBound P n) + λ islb → isPropIsMaximal (LowerBound P + , EmbeddingΣProp (isPropIsLowerBound P)) (n , islb) + + MaximalLowerBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + MaximalLowerBound = Σ[ n ∈ A ] isMaximalLowerBound n + + isMinimalUpperBound : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isMinimalUpperBound n + = Σ[ isub ∈ isUpperBound P n ] + isMinimal (UpperBound P + , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) + + isPropIsMinimalUpperBound : ∀ n → isProp (isMinimalUpperBound n) + isPropIsMinimalUpperBound n + = isPropΣ (isPropIsUpperBound P n) + λ isub → isPropIsMinimal (UpperBound P + , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) + + MinimalUpperBound : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + MinimalUpperBound = Σ[ n ∈ A ] isMinimalUpperBound n + + isInfimum : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isInfimum n + = Σ[ islb ∈ isLowerBound P n ] + isGreatest (LowerBound P + , EmbeddingΣProp (isPropIsLowerBound P)) (n , islb) + + isPropIsInfimum : ∀ n → isProp (isInfimum n) + isPropIsInfimum n + = isPropΣ (isPropIsLowerBound P n) + λ islb → isPropIsGreatest (LowerBound P + , EmbeddingΣProp (isPropIsLowerBound P)) (n , islb) + + Infimum : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + Infimum = Σ[ n ∈ A ] isInfimum n + + isSupremum : (n : A) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isSupremum n + = Σ[ isub ∈ isUpperBound P n ] + isLeast (UpperBound P + , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) + + isPropIsSupremum : ∀ n → isProp (isSupremum n) + isPropIsSupremum n + = isPropΣ (isPropIsUpperBound P n) + λ isub → isPropIsLeast (UpperBound P + , EmbeddingΣProp (isPropIsUpperBound P)) (n , isub) + + Supremum : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + Supremum = Σ[ n ∈ A ] isSupremum n + + isUpperBoundOfSelf→isGreatestOfSelf : ∀{n} → isUpperBound {B = A} (λ x → x) n + → isGreatest (A , (id↪ A)) n + isUpperBoundOfSelf→isGreatestOfSelf ub = ub + + isLowerBoundOfSelf→isLeastOfSelf : ∀{n} → isLowerBound {B = A} (λ x → x) n + → isLeast (A , (id↪ A)) n + isLowerBoundOfSelf→isLeastOfSelf lb = lb + + isInfimumOfUpperBounds→isUpperBound : {B : Type ℓ''} + → (P : B → A) + → ∀ n + → isInfimum {B = UpperBound P} fst n + → isUpperBound P n + isInfimumOfUpperBounds→isUpperBound P _ (_ , gt) x = gt ((P x) , λ { (_ , upy) → upy x }) + + isInfimumOfUpperBounds→isSupremum : {B : Type ℓ''} + → (P : B → A) + → ∀ n + → isInfimum {B = UpperBound P} fst n + → isSupremum P n + isInfimumOfUpperBounds→isSupremum P n (lt , gt) + = isInfimumOfUpperBounds→isUpperBound P n (lt , gt) , lt + + isSupremumOfLowerBounds→isLowerBound : {B : Type ℓ''} + → (P : B → A) + → ∀ n + → isSupremum {B = LowerBound P} fst n + → isLowerBound P n + isSupremumOfLowerBounds→isLowerBound P _ (_ , ls) x = ls ((P x) , λ { (_ , lwy) → lwy x }) + + isSupremumOfLowerBounds→isInfimum : {B : Type ℓ''} + → (P : B → A) + → ∀ n + → isSupremum {B = LowerBound P} fst n + → isInfimum P n + isSupremumOfLowerBounds→isInfimum P n (gs , ls) + = isSupremumOfLowerBounds→isLowerBound P n (gs , ls) , gs + + isMeet : A → A → A → Type (ℓ-max ℓ ℓ') + isMeet a b a∧b = ∀ x → x ≲ a∧b ≃ (x ≲ a × x ≲ b) + + isJoin : A → A → A → Type (ℓ-max ℓ ℓ') + isJoin a b a∨b = ∀ x → a∨b ≲ x ≃ (a ≲ x × b ≲ x) + + Meet : ∀ a b → Type (ℓ-max ℓ ℓ') + Meet a b = Σ[ a∧b ∈ A ] isMeet a b a∧b + + Join : ∀ a b → Type (ℓ-max ℓ ℓ') + Join a b = Σ[ a∨b ∈ A ] isJoin a b a∨b + + isPropIsMeet : ∀ a b a∧b → isProp (isMeet a b a∧b) + isPropIsMeet a b c = isPropΠ λ x → isOfHLevel≃ 1 (prop x c) + (isProp× (prop x a) + (prop x b)) + + isPropIsJoin : ∀ a b a∨b → isProp (isJoin a b a∨b) + isPropIsJoin a b c = isPropΠ λ x → isOfHLevel≃ 1 (prop c x) + (isProp× (prop a x) + (prop b x)) + + isReflIsMeet : ∀ a → isMeet a a a + isReflIsMeet a x = propBiimpl→Equiv (prop x a) + (isProp× (prop x a) (prop x a)) + (λ a → a , a) + fst + + isReflIsJoin : ∀ a → isJoin a a a + isReflIsJoin a x = propBiimpl→Equiv (prop a x) + (isProp× (prop a x) (prop a x)) + (λ a → a , a) + fst + + isMeet≃isInfimum : ∀ a b c + → isMeet a b c + ≃ isInfimum {B = Σ[ x ∈ A ] (x ≡ a) ⊎ (x ≡ b)} fst c + isMeet≃isInfimum a b a∧b + = propBiimpl→Equiv (isPropIsMeet a b a∧b) (isPropIsInfimum fst a∧b) + (λ fun → ⊎.rec (λ x≡a → subst + (a∧b ≲_) + (sym x≡a) + (fun a∧b .fst (rfl a∧b) .fst)) + (λ x≡b → subst + (a∧b ≲_) + (sym x≡b) + (fun a∧b .fst (rfl a∧b) .snd)) + ∘ snd , + λ { (c , uprc) + → invEq (fun c) ((uprc (a , (inl refl))) , + (uprc (b , (inr refl)))) }) + λ { (lb , gt) x + → propBiimpl→Equiv + (prop x a∧b) + (isProp× (prop x a) (prop x b)) + (λ x≤a∧b → (trans x a∧b a x≤a∧b + (lb (a , inl refl))) , + (trans x a∧b b x≤a∧b + (lb (b , inr refl)))) + λ { (x≤a , x≤b) + → gt (x , + (⊎.rec (λ y≡a → subst + (x ≲_) + (sym y≡a) + x≤a) + (λ y≡b → subst + (x ≲_) + (sym y≡b) + x≤b) + ∘ snd)) } } + + isJoin≃isSupremum : ∀ a b c + → isJoin a b c + ≃ isSupremum {B = Σ[ x ∈ A ] (x ≡ a) ⊎ (x ≡ b)} fst c + isJoin≃isSupremum a b a∨b + = propBiimpl→Equiv (isPropIsJoin a b a∨b) (isPropIsSupremum fst a∨b) + (λ fun → ⊎.rec (λ x≡a → subst + (_≲ a∨b) + (sym x≡a) + (fun a∨b .fst (rfl a∨b) .fst)) + (λ x≡b → subst + (_≲ a∨b) + (sym x≡b) + (fun a∨b .fst (rfl a∨b) .snd)) + ∘ snd , + λ { (c , lwrc) + → invEq (fun c) ((lwrc (a , (inl refl))) , + (lwrc (b , (inr refl)))) }) + λ { (ub , lt) x + → propBiimpl→Equiv + (prop a∨b x) + (isProp× (prop a x) (prop b x)) + (λ a∨b≤x → trans a a∨b x + (ub (a , inl refl)) + a∨b≤x , + trans b a∨b x + (ub (b , inr refl)) + a∨b≤x) + λ { (a≤x , b≤x) + → lt (x , + (⊎.rec (λ y≡a → subst + (_≲ x) + (sym y≡a) + a≤x) + (λ y≡b → subst + (_≲ x) + (sym y≡b) + b≤x) + ∘ snd)) } } + + isMeetComm : ∀ a b a∧b → isMeet a b a∧b → isMeet b a a∧b + isMeetComm _ _ _ fun x = compEquiv (fun x) Σ-swap-≃ + + isJoinComm : ∀ a b a∨b → isJoin a b a∨b → isJoin b a a∨b + isJoinComm _ _ _ fun x = compEquiv (fun x) Σ-swap-≃ + + isMeetAssoc : ∀ a b c a∧b b∧c a∧b∧c + → isMeet a b a∧b + → isMeet b c b∧c + → isMeet a b∧c a∧b∧c + → isMeet a∧b c a∧b∧c + isMeetAssoc a b c a∧b b∧c a∧b∧c funa∧b funb∧c funa∧b∧c x + = compEquiv (funa∧b∧c x) + (propBiimpl→Equiv (isProp× (prop x a) (prop x b∧c)) + (isProp× (prop x a∧b) (prop x c)) + (λ { (x≤a , x≤b∧c) + → invEq (funa∧b x) + (x≤a , + (funb∧c x .fst x≤b∧c .fst)) , + funb∧c x .fst x≤b∧c .snd }) + λ { (x≤a∧b , x≤c) + → funa∧b x .fst x≤a∧b .fst , + invEq (funb∧c x) + ((funa∧b x .fst x≤a∧b .snd) , + x≤c) }) + + isJoinAssoc : ∀ a b c a∨b b∨c a∨b∨c + → isJoin a b a∨b + → isJoin b c b∨c + → isJoin a b∨c a∨b∨c + → isJoin a∨b c a∨b∨c + isJoinAssoc a b c a∨b b∨c a∨b∨c funa∨b funb∨c funa∨b∨c x + = compEquiv (funa∨b∨c x) + (propBiimpl→Equiv (isProp× (prop a x) (prop b∨c x)) + (isProp× (prop a∨b x) (prop c x)) + (λ { (a≤x , b∨c≤x) + → invEq (funa∨b x) (a≤x , + (funb∨c x .fst b∨c≤x .fst)) , + funb∨c x .fst b∨c≤x .snd }) + λ { (a∨b≤x , c≤x) + → funa∨b x .fst a∨b≤x .fst , + invEq (funb∨c x) + ((funa∨b x .fst a∨b≤x .snd) , + c≤x) }) + + isMeetIsotone : ∀ a b c d a∧c b∧d + → isMeet a c a∧c + → isMeet b d b∧d + → a ≲ b + → c ≲ d + → a∧c ≲ b∧d + isMeetIsotone a b c d a∧c b∧d meeta∧c meetb∧d a≲b c≲d + = invEq (meetb∧d a∧c) (trans a∧c a b a∧c≲a a≲b , trans a∧c c d a∧c≲c c≲d) + where a∧c≲a = equivFun (meeta∧c a∧c) (rfl a∧c) .fst + a∧c≲c = equivFun (meeta∧c a∧c) (rfl a∧c) .snd + + isJoinIsotone : ∀ a b c d a∨c b∨d + → isJoin a c a∨c + → isJoin b d b∨d + → a ≲ b + → c ≲ d + → a∨c ≲ b∨d + isJoinIsotone a b c d a∨c b∨d joina∨c joinb∨d a≲b c≲d + = invEq (joina∨c b∨d) (trans a b b∨d a≲b b≲b∨d , trans c d b∨d c≲d d≲b∨d) + where b≲b∨d = equivFun (joinb∨d b∨d) (rfl b∨d) .fst + d≲b∨d = equivFun (joinb∨d b∨d) (rfl b∨d) .snd diff --git a/Cubical/Relation/Binary/Order/Quoset.agda b/Cubical/Relation/Binary/Order/Quoset.agda new file mode 100644 index 0000000000..c40af09a6f --- /dev/null +++ b/Cubical/Relation/Binary/Order/Quoset.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Order.Quoset where + +open import Cubical.Relation.Binary.Order.Quoset.Base public +open import Cubical.Relation.Binary.Order.Quoset.Properties public diff --git a/Cubical/Relation/Binary/Order/Quoset/Base.agda b/Cubical/Relation/Binary/Order/Quoset/Base.agda new file mode 100644 index 0000000000..41e78ef8b9 --- /dev/null +++ b/Cubical/Relation/Binary/Order/Quoset/Base.agda @@ -0,0 +1,126 @@ +{-# OPTIONS --safe #-} +{- + Quosets, or quasiordered sets, are posets where the relation is strict, + i.e. irreflexive rather than reflexive +-} +module Cubical.Relation.Binary.Order.Quoset.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.RecordEquiv +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.Nullary.Properties + +open Iso +open BinaryRelation + + +private + variable + ℓ ℓ' ℓ'' ℓ₀ ℓ₀' ℓ₁ ℓ₁' : Level + +record IsQuoset {A : Type ℓ} (_<_ : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + constructor isquoset + + field + is-set : isSet A + is-prop-valued : isPropValued _<_ + is-irrefl : isIrrefl _<_ + is-trans : isTrans _<_ + is-asym : isAsym _<_ + +unquoteDecl IsQuosetIsoΣ = declareRecordIsoΣ IsQuosetIsoΣ (quote IsQuoset) + + +record QuosetStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc ℓ')) where + + constructor quosetstr + + field + _<_ : A → A → Type ℓ' + isQuoset : IsQuoset _<_ + + infixl 7 _<_ + + open IsQuoset isQuoset public + +Quoset : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) +Quoset ℓ ℓ' = TypeWithStr ℓ (QuosetStr ℓ') + +quoset : (A : Type ℓ) → (_<_ : Rel A A ℓ') → IsQuoset _<_ → Quoset ℓ ℓ' +quoset A _<_ quo = A , (quosetstr _<_ quo) + +record IsQuosetEquiv {A : Type ℓ₀} {B : Type ℓ₁} + (M : QuosetStr ℓ₀' A) (e : A ≃ B) (N : QuosetStr ℓ₁' B) + : Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') ℓ₁') + where + constructor + isquosetequiv + -- Shorter qualified names + private + module M = QuosetStr M + module N = QuosetStr N + + field + pres< : (x y : A) → x M.< y ≃ equivFun e x N.< equivFun e y + + +QuosetEquiv : (M : Quoset ℓ₀ ℓ₀') (M : Quoset ℓ₁ ℓ₁') → Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) +QuosetEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsQuosetEquiv (M .snd) e (N .snd) + +isPropIsQuoset : {A : Type ℓ} (_<_ : A → A → Type ℓ') → isProp (IsQuoset _<_) +isPropIsQuoset _<_ = isOfHLevelRetractFromIso 1 IsQuosetIsoΣ + (isPropΣ isPropIsSet + λ isSetA → isPropΣ (isPropΠ2 (λ _ _ → isPropIsProp)) + λ isPropValued< → isProp×2 (isPropΠ (λ x → isProp¬ (x < x))) + (isPropΠ5 (λ _ _ _ _ _ → isPropValued< _ _)) + (isPropΠ3 λ x y _ → isProp¬ (y < x))) + +𝒮ᴰ-Quoset : DUARel (𝒮-Univ ℓ) (QuosetStr ℓ') (ℓ-max ℓ ℓ') +𝒮ᴰ-Quoset = + 𝒮ᴰ-Record (𝒮-Univ _) IsQuosetEquiv + (fields: + data[ _<_ ∣ autoDUARel _ _ ∣ pres< ] + prop[ isQuoset ∣ (λ _ _ → isPropIsQuoset _) ]) + where + open QuosetStr + open IsQuoset + open IsQuosetEquiv + +QuosetPath : (M N : Quoset ℓ ℓ') → QuosetEquiv M N ≃ (M ≡ N) +QuosetPath = ∫ 𝒮ᴰ-Quoset .UARel.ua + +-- an easier way of establishing an equivalence of strict posets +module _ {P : Quoset ℓ₀ ℓ₀'} {S : Quoset ℓ₁ ℓ₁'} (e : ⟨ P ⟩ ≃ ⟨ S ⟩) where + private + module P = QuosetStr (P .snd) + module S = QuosetStr (S .snd) + + module _ (isMon : ∀ x y → x P.< y → equivFun e x S.< equivFun e y) + (isMonInv : ∀ x y → x S.< y → invEq e x P.< invEq e y) where + open IsQuosetEquiv + open IsQuoset + + makeIsQuosetEquiv : IsQuosetEquiv (P .snd) e (S .snd) + pres< makeIsQuosetEquiv x y = propBiimpl→Equiv (P.isQuoset .is-prop-valued _ _) + (S.isQuoset .is-prop-valued _ _) + (isMon _ _) (isMonInv' _ _) + where + isMonInv' : ∀ x y → equivFun e x S.< equivFun e y → x P.< y + isMonInv' x y ex= 1.10 build-type: Simple description: Helper programs. license: MIT -tested-with: GHC == 8.6.5 +tested-with: GHC == 9.10.2 executable Everythings hs-source-dirs: . main-is: Everythings.hs default-language: Haskell2010 - build-depends: base >= 4.9.0.0 && < 4.20 + build-depends: base >= 4.9.0.0 && < 4.22 , directory >= 1.0.0.0 && < 1.4 diff --git a/cubical.agda-lib b/cubical.agda-lib index 064f900d4c..85499ce6c6 100644 --- a/cubical.agda-lib +++ b/cubical.agda-lib @@ -1,4 +1,4 @@ -name: cubical-0.7 +name: cubical-0.8 include: . depend: flags: --cubical --no-import-sorts -WnoUnsupportedIndexedMatch diff --git a/flake.lock b/flake.lock index 6fa40f03e7..b66b947314 100644 --- a/flake.lock +++ b/flake.lock @@ -2,24 +2,22 @@ "nodes": { "agda": { "inputs": { - "flake-utils": [ - "flake-utils" - ], + "flake-parts": "flake-parts", "nixpkgs": [ "nixpkgs" ] }, "locked": { - "lastModified": 1701366566, - "narHash": "sha256-B8Jmjld0gGbkVO08GsovVqrUXCs8VfJ8UdM3sjHnzgM=", + "lastModified": 1726161230, + "narHash": "sha256-IupM/HJk1vpHTRWOakYDDliKn+RLzPjNzGlLTlF0CbY=", "owner": "agda", "repo": "agda", - "rev": "4293e0a94d15acac915ab9088b2ec028f78d14a9", + "rev": "a6fc20c27ae953149b53a8997ba4a1c8b17d628a", "type": "github" }, "original": { "owner": "agda", - "ref": "v2.6.4.1", + "ref": "v2.7.0.1", "repo": "agda", "type": "github" } @@ -27,11 +25,11 @@ "flake-compat": { "flake": false, "locked": { - "lastModified": 1696426674, - "narHash": "sha256-kvjfFW7WAETZlt09AgDn1MrtKzP7t90Vf7vypd3OL1U=", + "lastModified": 1747046372, + "narHash": "sha256-CIVLLkVgvHYbgI2UpXvIIBJ12HWgX+fjA8Xf8PUmqCY=", "owner": "edolstra", "repo": "flake-compat", - "rev": "0f9255e01c2351cc7d116c072cb317785dd33b33", + "rev": "9100a0f413b0c601e0533d1d94ffd501ce2e7885", "type": "github" }, "original": { @@ -40,16 +38,34 @@ "type": "github" } }, + "flake-parts": { + "inputs": { + "nixpkgs-lib": "nixpkgs-lib" + }, + "locked": { + "lastModified": 1701473968, + "narHash": "sha256-YcVE5emp1qQ8ieHUnxt1wCZCC3ZfAS+SRRWZ2TMda7E=", + "owner": "hercules-ci", + "repo": "flake-parts", + "rev": "34fed993f1674c8d06d58b37ce1e0fe5eebcb9f5", + "type": "github" + }, + "original": { + "owner": "hercules-ci", + "repo": "flake-parts", + "type": "github" + } + }, "flake-utils": { "inputs": { "systems": "systems" }, "locked": { - "lastModified": 1705309234, - "narHash": "sha256-uNRRNRKmJyCRC/8y1RqBkqWBLM034y4qN7EprSdmgyA=", + "lastModified": 1731533236, + "narHash": "sha256-l0KFg5HjrsfsO/JpG+r7fRrqm12kzFHyUHqHCVpMMbI=", "owner": "numtide", "repo": "flake-utils", - "rev": "1ef2e671c3b0c19053962c07dbda38332dcebf26", + "rev": "11707dc2f618dd54ca8739b309ec4fc024de578b", "type": "github" }, "original": { @@ -60,11 +76,11 @@ }, "nixpkgs": { "locked": { - "lastModified": 1705883077, - "narHash": "sha256-ByzHHX3KxpU1+V0erFy8jpujTufimh6KaS/Iv3AciHk=", + "lastModified": 1747426788, + "narHash": "sha256-N4cp0asTsJCnRMFZ/k19V9akkxb7J/opG+K+jU57JGc=", "owner": "NixOS", "repo": "nixpkgs", - "rev": "5f5210aa20e343b7e35f40c033000db0ef80d7b9", + "rev": "12a55407652e04dcf2309436eb06fef0d3713ef3", "type": "github" }, "original": { @@ -73,6 +89,24 @@ "type": "indirect" } }, + "nixpkgs-lib": { + "locked": { + "dir": "lib", + "lastModified": 1701253981, + "narHash": "sha256-ztaDIyZ7HrTAfEEUt9AtTDNoCYxUdSd6NrRHaYOIxtk=", + "owner": "NixOS", + "repo": "nixpkgs", + "rev": "e92039b55bcd58469325ded85d4f58dd5a4eaf58", + "type": "github" + }, + "original": { + "dir": "lib", + "owner": "NixOS", + "ref": "nixos-unstable", + "repo": "nixpkgs", + "type": "github" + } + }, "root": { "inputs": { "agda": "agda", diff --git a/flake.nix b/flake.nix index 42f1b1beae..3db1690054 100644 --- a/flake.nix +++ b/flake.nix @@ -8,10 +8,9 @@ flake = false; }; inputs.agda = { - url = "github:agda/agda/v2.6.4.1"; + url = "github:agda/agda/v2.7.0.1"; inputs = { nixpkgs.follows = "nixpkgs"; - flake-utils.follows = "flake-utils"; }; }; @@ -21,7 +20,7 @@ overlay = final: prev: { cubical = final.agdaPackages.mkDerivation rec { pname = "cubical"; - version = "0.7"; + version = "0.8"; src = cleanSourceWith { filter = name: type: @@ -47,7 +46,7 @@ }; agdaWithCubical = final.agdaPackages.agda.withPackages [final.cubical]; }; - overlays = [ agda.overlay overlay ]; + overlays = [ agda.overlays.default overlay ]; in { overlays.default = overlay; } // flake-utils.lib.eachDefaultSystem (system: