diff --git a/formalizations/guarded-cubical/Common/Monotone.agda b/formalizations/guarded-cubical/Common/Monotone.agda new file mode 100644 index 0000000000000000000000000000000000000000..6d44ff8ec01c21f6927fe8d1ac920d9de76ec8ed --- /dev/null +++ b/formalizations/guarded-cubical/Common/Monotone.agda @@ -0,0 +1,229 @@ +{-# OPTIONS --safe #-} + +module Common.Monotone where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Relation.Binary.Base +open import Cubical.Foundations.Function +open import Cubical.Foundations.Structure + +open import Cubical.Foundations.HLevels +open import Cubical.Reflection.Base hiding (_$_) +open import Cubical.Reflection.RecordEquiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Sigma + + +open import Common.Preorder + +open BinaryRelation + + +private + variable + â„“ â„“' : Level + + +-- Helper functions + +-- The relation associated to a preorder X +rel : (X : Preorder â„“ â„“') -> (⟨ X ⟩ -> ⟨ X ⟩ -> Type â„“') +rel X = PreorderStr._≤_ (X .snd) + +reflexive : (X : Preorder â„“ â„“') -> (x : ⟨ X ⟩) -> (rel X x x) +reflexive X x = IsPreorder.is-refl (PreorderStr.isPreorder (str X)) x + +transitive : (X : Preorder â„“ â„“') -> (x y z : ⟨ X ⟩) -> + rel X x y -> rel X y z -> rel X x z +transitive X x y z x≤y y≤z = + IsPreorder.is-trans (PreorderStr.isPreorder (str X)) x y z x≤y y≤z + +isSet-preorder : (X : Preorder â„“ â„“') -> isSet ⟨ X ⟩ +isSet-preorder X = IsPreorder.is-set (PreorderStr.isPreorder (str X)) + +isPropValued-preorder : (X : Preorder â„“ â„“') -> + isPropValued (PreorderStr._≤_ (str X)) +isPropValued-preorder X = IsPreorder.is-prop-valued + (PreorderStr.isPreorder (str X)) + + +-- If Y is a set, then functions into Y form a set +isSetArrow : {X Y : Type} -> isSet Y -> isSet (X -> Y) +isSetArrow Hy = λ f g p q i j x → + Hy (f x) (g x) (λ k -> p k x) (λ k -> q k x) i j + + +module _ {â„“ â„“' : Level} where + + -- Because of a bug with Cubical Agda's reflection, we need to declare + -- a separate version of MonFun where the arguments to the isMon + -- constructor are explicit. + -- See https://github.com/agda/cubical/issues/995. + module _ {X Y : Preorder â„“ â„“'} where + + module X = PreorderStr (X .snd) + module Y = PreorderStr (Y .snd) + _≤X_ = X._≤_ + _≤Y_ = Y._≤_ + + monotone' : (⟨ X ⟩ -> ⟨ Y ⟩) -> Type (â„“-max â„“ â„“') + monotone' f = (x y : ⟨ X ⟩) -> x ≤X y → f x ≤Y f y + + monotone : (⟨ X ⟩ -> ⟨ Y ⟩) -> Type (â„“-max â„“ â„“') + monotone f = {x y : ⟨ X ⟩} -> x ≤X y → f x ≤Y f y + + -- Monotone functions from X to Y + -- This definition works with Cubical Agda's reflection. + record MonFun' (X Y : Preorder â„“ â„“') : Type ((â„“-max â„“ â„“')) where + field + f : (X .fst) → (Y .fst) + isMon : monotone' {X} {Y} f + + -- This is the definition we want, where the first two arguments to isMon + -- are implicit. + record MonFun (X Y : Preorder â„“ â„“') : Type ((â„“-max â„“ â„“')) where + field + f : (X .fst) → (Y .fst) + isMon : monotone {X} {Y} f + + open MonFun + + isoMonFunMonFun' : {X Y : Preorder â„“ â„“'} -> Iso (MonFun X Y) (MonFun' X Y) + isoMonFunMonFun' = iso + (λ g → record { f = MonFun.f g ; isMon = λ x y x≤y → isMon g x≤y }) + (λ g → record { f = MonFun'.f g ; isMon = λ {x} {y} x≤y -> MonFun'.isMon g x y x≤y }) + (λ g → refl) + (λ g → refl) + + {- + isPropIsMon : {X Y : Preorder â„“ â„“'} -> {f : MonFun X Y} -> + isProp (∀ x y -> rel X x y -> rel Y (MonFun.f f x) (MonFun.f f y)) + isPropIsMon {X} {Y} {f} = + isPropÎ 3 λ x y x≤y -> isPropValued-preorder Y (MonFun.f f x) (MonFun.f f y) + -} + + + isPropIsMon : {X Y : Preorder â„“ â„“'} -> (f : MonFun X Y) -> + isProp (monotone {X} {Y} (MonFun.f f)) + isPropIsMon {X} {Y} f = + isPropImplicitÎ 2 (λ x y -> + isPropÎ (λ x≤y -> isPropValued-preorder Y (MonFun.f f x) (MonFun.f f y))) + + isPropIsMon' : {X Y : Preorder â„“ â„“'} -> (f : ⟨ X ⟩ -> ⟨ Y ⟩) -> + isProp (monotone' {X} {Y} f) + isPropIsMon' {X} {Y} f = + isPropÎ 3 (λ x y x≤y -> isPropValued-preorder Y (f x) (f y)) + + + -- Equivalence between MonFun' record and a sigma type + unquoteDecl MonFun'IsoΣ = declareRecordIsoΣ MonFun'IsoΣ (quote (MonFun')) + + Sigma : (X Y : Preorder â„“ â„“') -> Type (â„“-max â„“ â„“') + Sigma X Y = + (Σ (X .fst → Y .fst) + (λ z → (x y : ⟨ X ⟩) → _≤X_ {X} {Y} x y → _≤Y_ {X} {Y} (z x) (z y))) + + test : {X Y : Preorder â„“ â„“'} -> Iso (MonFun' X Y) (Sigma X Y) + test = MonFun'IsoΣ + + MonFun≡Sigma : {X Y : Preorder â„“ â„“'} -> MonFun' X Y ≡ Sigma X Y + MonFun≡Sigma = isoToPath MonFun'IsoΣ + + Sigma≡ : {X Y : Preorder â„“ â„“'} -> {s1 s2 : Sigma X Y} -> + s1 .fst ≡ s2 .fst -> s1 ≡ s2 + Sigma≡ {X} {Y} = Σ≡Prop (λ f → isPropÎ 3 + (λ x y x≤y -> isPropValued-preorder Y (f x) (f y))) + + -- Equality of monotone functions is equivalent to equality of the + -- underlying functions. + EqMon' : {X Y : Preorder â„“ â„“'} -> (f g : MonFun' X Y) -> + MonFun'.f f ≡ MonFun'.f g -> f ≡ g + EqMon' {X} {Y} f g p = isoFunInjective MonFun'IsoΣ f g + (Σ≡Prop (λ h → isPropÎ 3 (λ y z q → isPropValued-preorder Y (h y) (h z))) p) + + EqMon : {X Y : Preorder â„“ â„“'} -> (f g : MonFun X Y) -> + MonFun.f f ≡ MonFun.f g -> f ≡ g + EqMon {X} {Y} f g p = isoFunInjective isoMonFunMonFun' f g (EqMon' _ _ p) + + -- isSet for Sigma + isSetSigma : {X Y : Preorder â„“ â„“'} -> isSet (Sigma X Y) + isSetSigma {X} {Y} = isSetΣSndProp + (isSet→ (isSet-preorder Y)) + λ f -> isPropIsMon' {X} {Y} f + + -- isSet for monotone functions + MonFunIsSet : {X Y : Preorder â„“ â„“'} -> isSet (MonFun X Y) + MonFunIsSet {X} {Y} = let composedIso = (compIso isoMonFunMonFun' MonFun'IsoΣ) in + isSetRetract + (Iso.fun composedIso) (Iso.inv composedIso) (Iso.leftInv composedIso) + (isSetΣSndProp + (isSet→ (isSet-preorder Y)) + (isPropIsMon' {X} {Y})) + + + + -- Ordering on monotone functions + module _ {X Y : Preorder â„“ â„“'} where + + _≤mon_ : + MonFun X Y → MonFun X Y → Type (â„“-max â„“ â„“') + _≤mon_ f g = ∀ x -> rel Y (MonFun.f f x) (MonFun.f g x) + + ≤mon-prop : isPropValued _≤mon_ + ≤mon-prop f g = + isPropÎ (λ x -> isPropValued-preorder Y (MonFun.f f x) (MonFun.f g x)) + + ≤mon-refl : isRefl _≤mon_ + ≤mon-refl = λ f x → reflexive Y (MonFun.f f x) + + ≤mon-trans : isTrans _≤mon_ + ≤mon-trans = λ f g h f≤g g≤h x → + transitive Y (MonFun.f f x) (MonFun.f g x) (MonFun.f h x) + (f≤g x) (g≤h x) + + -- Alternate definition of ordering on monotone functions, where we allow for the + -- arguments to be distinct + _≤mon-het_ : MonFun X Y -> MonFun X Y -> Type (â„“-max â„“ â„“') + _≤mon-het_ f f' = ∀ x x' -> rel X x x' -> rel Y (MonFun.f f x) (MonFun.f f' x') + + ≤mon→≤mon-het : (f f' : MonFun X Y) -> f ≤mon f' -> f ≤mon-het f' + ≤mon→≤mon-het f f' f≤f' = λ x x' x≤x' → + MonFun.f f x ≤⟨ MonFun.isMon f x≤x' ⟩ + MonFun.f f x' ≤⟨ f≤f' x' ⟩ + MonFun.f f' x' â—¾ + where + open PreorderReasoning Y + + + -- Predomain of monotone functions between two predomains + IntHom : Preorder â„“ â„“' -> Preorder â„“ â„“' -> + Preorder (â„“-max â„“ â„“') (â„“-max â„“ â„“') -- (â„“-max (â„“-suc â„“) (â„“-suc â„“')) (â„“-max â„“ â„“') + IntHom X Y = + MonFun X Y , + (preorderstr + (_≤mon_) + (ispreorder MonFunIsSet ≤mon-prop ≤mon-refl ≤mon-trans)) + + -- Notation + _==>_ : Preorder â„“ â„“' -> Preorder â„“ â„“' -> + Preorder (â„“-max â„“ â„“') (â„“-max â„“ â„“') -- (â„“-max (â„“-suc â„“) (â„“-suc â„“')) (â„“-max â„“ â„“') + X ==> Y = IntHom X Y -- IntHom X Y + + + + -- Some basic combinators/utility functions on monotone functions + + MonId : {X : Preorder â„“ â„“'} -> MonFun X X + MonId = record { f = λ x -> x ; isMon = λ x≤y → x≤y } + + _$_ : {X Y : Preorder â„“ â„“'} -> MonFun X Y -> ⟨ X ⟩ -> ⟨ Y ⟩ + f $ x = MonFun.f f x + + MonComp : {X Y Z : Preorder â„“ â„“'} -> + MonFun X Y -> MonFun Y Z -> MonFun X Z + MonComp f g = record { + f = λ x -> g $ (f $ x) ; + isMon = λ {x1} {x2} x1≤x2 → isMon g (isMon f x1≤x2) } diff --git a/formalizations/guarded-cubical/Common/Preorder.agda b/formalizations/guarded-cubical/Common/Preorder.agda new file mode 100644 index 0000000000000000000000000000000000000000..f83768128b8a5c7e307afd38b74368d58271adf5 --- /dev/null +++ b/formalizations/guarded-cubical/Common/Preorder.agda @@ -0,0 +1,145 @@ +{-# OPTIONS --safe #-} +{-# OPTIONS --cubical #-} + +module Common.Preorder where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +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Î 5 (λ _ _ _ _ _ -> isPropValued≤ _ _))) +{- + (isPropΣ isPropIsSet + λ isSetA → isPropΣ (isPropÎ 2 (λ _ _ → isPropIsProp)) + λ isPropValued≤ → isProp×2 + (isPropÎ (λ _ → isPropValued≤ _ _)) + (isPropÎ 5 λ _ _ _ _ _ → isPropValued≤ _ _) + (isPropÎ 4 λ _ _ _ _ → isSetA _ _)) +-} + +ð’®á´°-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 _â—¾