├── .gitignore ├── AxiomChoice.agda ├── Basics.agda ├── Everything.agda ├── FinalCoalg ├── InSetoid.agda ├── InTypesAsCoindType.agda └── InTypesAsSetQuot.agda ├── ListRelations.agda ├── ListStuff.agda ├── Pfin ├── AsFreeJoinSemilattice.agda ├── AsSetQuot.agda └── PreservesCountableIntersections.agda ├── README.md ├── SetoidStuff.agda ├── Trees.agda └── Worrell ├── Finality.agda ├── FromInjectivity.agda ├── FromInjectivityOmega.agda ├── Limits.agda ├── NoSurjAlgebra.agda └── ToInjectivity.agda /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | *.agda~ 3 | MAlonzo/** 4 | -------------------------------------------------------------------------------- /AxiomChoice.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module AxiomChoice where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.List renaming (map to mapList; [_] to sing) 8 | open import Cubical.Data.Sigma 9 | open import Cubical.Data.Unit 10 | open import Cubical.Data.Empty renaming (elim to ⊥-elim) 11 | open import Cubical.Data.Nat 12 | open import Cubical.Data.Sum renaming (inl to inj₁; inr to inj₂; map to map⊎) 13 | open import Cubical.Functions.Logic hiding (⊥) 14 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 15 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 16 | renaming (map to ∥map∥; rec to ∥rec∥) 17 | open import Cubical.Relation.Binary hiding (Rel) 18 | open BinaryRelation 19 | 20 | open import Basics 21 | 22 | -- pointwise lifting of a relation to a function space 23 | PW : {X A B : Type} (R : A → B → Type) → (X → A) → (X → B) → Type 24 | PW R f g = ∀ x → R (f x) (g x) 25 | 26 | -- the quotient a function space by the pointwise lifted relation 27 | [_⇒_]/_ : (A B : Type) (R : B → B → Type) → Type 28 | [ A ⇒ B ]/ R = (A → B) / PW R 29 | 30 | -- a function sending equivalence classes of function wrt. pointwise 31 | -- lifted relation and functions into equivalence classes 32 | θ : ∀ A {B} (R : B → B → Type) → [ A ⇒ B ]/ R → A → B / R 33 | θ A R = recQ (isSetΠ (λ _ → squash/)) (λ c x → [ c x ]) 34 | λ c c' r → funExt (λ x → eq/ _ _ (r x)) 35 | 36 | 37 | -- equivalence between two formulation of full axiom of choice; the 38 | -- formulation stating the surjectivity of the map λ g → [_] ∘ g is 39 | -- proved equivalent to the usual presentation of axiom of choice 40 | 41 | -- NB: in the first formulation we do not need to 0-truncate the 42 | -- existence of a section, since the type of sections of θ is a 43 | -- proposition; this follows directly from the lemma SectionIsProp 44 | -- below 45 | 46 | module _ (θInv : ∀ A {B} (R : B → B → Type) → (A → B / R) → [ A ⇒ B ]/ R) 47 | (sectionθ : ∀ A {B} (R : B → B → Type) → section (θ A R) (θInv A R)) where 48 | 49 | ac' : ∀ (A : Type) {B : Type} (R : B → B → Type) 50 | → (f : (A → B) / PW R) → ∃[ g ∈ (A → B) ] [_] ∘ g ≡ θ A R f 51 | ac' A R = elimProp (λ _ → propTruncIsProp) (λ g → ∣ g , refl ∣) 52 | 53 | ac : ∀ (A : Type) {B : Type} (R : B → B → Type) 54 | → (f : A → B / R) → ∃[ g ∈ (A → B) ] [_] ∘ g ≡ f 55 | ac A R f = 56 | ∥map∥ (λ { (g , eq) → g , eq ∙ sectionθ A R f }) (ac' A R (θInv A R f)) 57 | 58 | module _ (ac : ∀ (A : Type) {B : Type} (R : B → B → Type) 59 | → isPropValued R → isEquivRel R 60 | → (f : A → B / R) → ∃[ g ∈ (A → B) ] [_] ∘ g ≡ f) where 61 | 62 | SectionIsProp' : ∀ A {B} (R : B → B → Type) 63 | → isPropValued R → isEquivRel R 64 | → (f : A → B / R) 65 | → (g g' : [ A ⇒ B ]/ R) (eq : θ A R g ≡ f) (eq' : θ A R g' ≡ f) 66 | → g ≡ g' 67 | SectionIsProp' A R Rprop Reqr f = elimProp2 (λ _ _ → isPropΠ (λ _ → isPropΠ λ _ → squash/ _ _)) 68 | λ g g' eq eq' → eq/ _ _ (λ x → effective Rprop Reqr _ _ (λ i → (eq ∙ sym eq') i x)) 69 | 70 | SectionIsProp : ∀ A {B} (R : B → B → Type) 71 | → isPropValued R → isEquivRel R 72 | → (f : A → B / R) 73 | → isProp (Σ ([ A ⇒ B ]/ R) (λ g → θ A R g ≡ f)) 74 | SectionIsProp A R Rprop Reqr f (g , eq) (g' , eq') = 75 | Σ≡Prop (λ _ → isPropFunEq _ _ λ _ → squash/ _ _) 76 | (SectionIsProp' A R Rprop Reqr f g g' eq eq') 77 | 78 | θInvSec : ∀ A {B} (R : B → B → Type) 79 | → isPropValued R → isEquivRel R 80 | → (f : A → B / R) → Σ ([ A ⇒ B ]/ R) (λ g → θ A R g ≡ f) 81 | θInvSec A R Rprop Reqr f = 82 | ∥rec∥ (SectionIsProp A R Rprop Reqr f) 83 | (λ {(g , eq) → [ g ] , eq}) 84 | (ac A R Rprop Reqr f) 85 | 86 | -------------------------------------------------------------------------------- /Basics.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Basics where 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Relation.Binary hiding (Rel) 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sum renaming (rec to rec⊎) 9 | open import Cubical.Data.Empty renaming (rec to ⊥-rec) 10 | open import Cubical.Data.Bool 11 | open import Cubical.Data.Sigma 12 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≟_ to _≟N_) 13 | open BinaryRelation 14 | open import Cubical.Relation.Nullary 15 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 16 | renaming (map to ∥map∥; rec to ∥rec∥) 17 | 18 | -- some basic stuff 19 | hProp₀ = hProp ℓ-zero 20 | 21 | isPropFunEq : ∀{ℓ}{A B : Type ℓ} (f g : A → B) 22 | → (∀ x → isProp (f x ≡ g x)) 23 | → isProp (f ≡ g) 24 | isPropFunEq f g p eq1 eq2 i j x = 25 | p x (λ k → eq1 k x) (λ k → eq2 k x) i j 26 | 27 | EqFiber : {A B : Type} → isSet B 28 | → (f : A → B) (b : B) 29 | → {a a' : A} → a ≡ a' 30 | → (eq : f a ≡ b) (eq' : f a' ≡ b) 31 | → _≡_ {A = fiber f b} (a , eq) (a' , eq') 32 | EqFiber setB f b = 33 | J (λ a _ → (eq : f _ ≡ b) (eq' : f a ≡ b) 34 | → _≡_ {A = fiber f b} (_ , eq) (a , eq')) 35 | λ p q → cong (_ ,_) (setB _ _ p q) 36 | 37 | ≤-suc-cases : ∀ k n → k ≤ suc n 38 | → (k ≤ n) ⊎ (k ≡ suc n) 39 | ≤-suc-cases zero n le = inl zero-≤ 40 | ≤-suc-cases (suc k) zero le = inr (cong suc (≤0→≡0 (pred-≤-pred le))) 41 | ≤-suc-cases (suc k) (suc n) le with ≤-suc-cases k n (pred-≤-pred le) 42 | ... | inl p = inl (suc-≤-suc p) 43 | ... | inr p = inr (cong suc p) 44 | 45 | data isEven : ℕ → Type 46 | data isOdd : ℕ → Type 47 | data isEven where 48 | zero : isEven zero 49 | suc : ∀{n} → isOdd n → isEven (suc n) 50 | data isOdd where 51 | suc : ∀{n} → isEven n → isOdd (suc n) 52 | 53 | isEven? : Bool → ℕ → Type 54 | isEven? false n = isOdd n 55 | isEven? true n = isEven n 56 | 57 | decEven : ∀ n → isEven n ⊎ isOdd n 58 | decEven zero = inl zero 59 | decEven (suc n) = 60 | rec⊎ (λ z → inr (suc z)) (λ z → inl (suc z)) (decEven n) 61 | 62 | even-not-odd : ∀ {n} → isEven n → isOdd n → ⊥ 63 | even-not-odd (suc x₁) (suc x) = even-not-odd x x₁ 64 | 65 | decIsProp : {A : Type} 66 | → isProp A 67 | → isProp (Dec A) 68 | decIsProp propA (yes p) (yes q) = cong yes (propA _ _) 69 | decIsProp propA (yes p) (no ¬p) = ⊥-rec (¬p p) 70 | decIsProp propA (no ¬p) (yes p) = ⊥-rec (¬p p) 71 | decIsProp propA (no ¬p) (no ¬q) = cong no (funExt (λ _ → isProp⊥ _ _)) 72 | 73 | decPropTrunc : {A : Type} → Dec A → Dec ∥ A ∥ 74 | decPropTrunc (yes a) = yes ∣ a ∣ 75 | decPropTrunc (no ¬a) = no (∥rec∥ isProp⊥ ¬a) 76 | 77 | true-before? : (a : ℕ → Bool) (n : ℕ) 78 | → Dec (Σ[ k ∈ ℕ ] (k ≤ n) × (a k ≡ true)) 79 | true-before? a zero with a zero ≟ true 80 | ... | yes p = yes (0 , ≤-refl , p) 81 | ... | no ¬p = no (λ { (k , k≤ , eq) → 82 | ¬p (cong a (sym (≤0→≡0 k≤)) ∙ eq) }) 83 | true-before? a (suc n) with true-before? a n 84 | ... | yes (k , k≤ , eq) = yes (k , ≤-suc k≤ , eq) 85 | ... | no ¬ih with a (suc n) ≟ true 86 | ... | yes p = yes (suc n , ≤-refl , p) 87 | ... | no ¬p = no (λ { (k , k≤ , eq) → rec⊎ (λ r → ¬ih (_ , r , eq)) (λ r → ¬p (cong a (sym r) ∙ eq)) (≤-suc-cases k n k≤) }) 88 | 89 | 90 | 91 | -------------------------------------------------------------------------------- /Everything.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | -- ================================================================= 4 | -- Type Theoretic Constructions of the Final Coalgebra of the Finite 5 | -- Powerset Functor 6 | -- 7 | -- Niccolò Veltri 8 | -- ================================================================= 9 | 10 | module Everything where 11 | 12 | -- some preliminaries 13 | import Basics 14 | 15 | -- basic list stuff 16 | import ListStuff 17 | 18 | -- setoids 19 | import SetoidStuff 20 | 21 | -- alternative formulation of axiom of choice 22 | import AxiomChoice 23 | 24 | -- relation liftings on lists 25 | import ListRelations 26 | 27 | -- non-wellfounded trees with ordered finite branching 28 | import Trees 29 | 30 | -- ================================================================= 31 | 32 | -- finite powerset using HITs 33 | 34 | -- -- as a set quotient of lists 35 | import Pfin.AsSetQuot 36 | 37 | -- -- as the free join semilattice on a type 38 | import Pfin.AsFreeJoinSemilattice 39 | 40 | -- finite powerset preserves countable intersection (assuming 41 | -- countable choice) 42 | import Pfin.PreservesCountableIntersections 43 | 44 | -- ================================================================= 45 | 46 | -- final coalgebra of finite powerset functor in setoids 47 | import FinalCoalg.InSetoid 48 | 49 | -- final coalgebra of finite powerset functor in types 50 | 51 | -- -- as a set quotient of trees (assuming axiom of choice) 52 | import FinalCoalg.InTypesAsSetQuot 53 | 54 | -- -- as a coinductive type 55 | import FinalCoalg.InTypesAsCoindType 56 | 57 | -- ================================================================= 58 | 59 | 60 | -- analysis of Worrell's construction 61 | 62 | -- -- ω-limits 63 | import Worrell.Limits 64 | 65 | -- -- the canonical algebra map of Vω is not surjective (here we use 66 | -- -- PfinQ instead of Pfin, but we know that the two are equivalent, 67 | -- -- see Pfin.AsSetQuot) 68 | import Worrell.NoSurjAlgebra 69 | 70 | -- -- assuming the canonical algebra map of Vω is injective implies 71 | -- -- LLPO 72 | import Worrell.FromInjectivity 73 | 74 | -- -- LLPO and countable choice imply the injectivity of the canonical 75 | -- -- algebra map of Vω 76 | import Worrell.ToInjectivity 77 | 78 | -- -- assuming the canonical projection from Vω2 to Vω is injective 79 | -- -- implies LLPO, showing that Worrell's construction of the final 80 | -- -- coalgebra of Pfin as a subset of Vω cannot be done fully 81 | -- -- constructively 82 | import Worrell.FromInjectivityOmega 83 | 84 | -- -- assuming countable choice and injectivity of the canonical 85 | -- -- algebra map of Vω (or, equivalently, LLPO), we have that Vω2 is 86 | -- -- the final coalgebra of Pfin 87 | import Worrell.Finality 88 | 89 | -------------------------------------------------------------------------------- /FinalCoalg/InSetoid.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module FinalCoalg.InSetoid where 4 | 5 | open import Size 6 | open import Cubical.Core.Everything 7 | open import Cubical.Foundations.Prelude 8 | open import Cubical.Foundations.Everything 9 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 10 | open import Cubical.Relation.Everything 11 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 12 | renaming (map to ∥map∥; rec to ∥rec∥) 13 | open import Cubical.HITs.SetQuotients renaming ([_] to eqCl) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.List renaming (map to mapList) 16 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 17 | open import Cubical.Relation.Binary hiding (Rel) 18 | open BinaryRelation 19 | open isEquivRel 20 | 21 | open import Basics 22 | open import SetoidStuff 23 | open import ListRelations 24 | open import ListStuff 25 | open import Trees 26 | 27 | -- the finite powerset functor on setoids 28 | 29 | PfinS : Setoid₀ → Setoid₀ 30 | PfinS (setoid A R propR eqrR) = 31 | setoid (List A) (Relator R) (isPropRelator R) (isEquivRelRelator eqrR) 32 | 33 | -- the final coalgebra of PfinS 34 | νPfinS : Setoid₀ 35 | νPfinS = setoid (Tree ∞) (ExtEq ∞) isPropExtEq isEquivRelExtEq 36 | 37 | subtreesS : νPfinS →S PfinS νPfinS 38 | subtreesS = (λ x → subtrees x) , (λ r → subtreesE r) 39 | 40 | mapPfinS : {S₁ S₂ : Setoid₀} (f : S₁ →S S₂) 41 | → PfinS S₁ →S PfinS S₂ 42 | mapPfinS {_}{S₂} (f , fr) = mapList f , 43 | (λ { (p , q) → 44 | (λ x mx → 45 | ∥map∥ 46 | (λ { (y , my , r) → f y , ∈mapList my , 47 | S₂ .eqrRel .transitive _ _ _ (subst (S₂ .Rel x) (sym (pre∈mapList mx .snd .snd)) (S₂ .eqrRel .reflexive _)) (fr r) }) 48 | (p _ (pre∈mapList mx .snd .fst))) , 49 | (λ x mx → 50 | ∥map∥ 51 | (λ { (y , my , r) → f y , ∈mapList my , 52 | S₂ .eqrRel .transitive _ _ _ (subst (S₂ .Rel x) (sym (pre∈mapList mx .snd .snd)) (S₂ .eqrRel .reflexive _)) (fr r) }) 53 | (q _ (pre∈mapList mx .snd .fst))) }) 54 | 55 | module _ 56 | (S : Setoid₀) 57 | (cS : S →S PfinS S) 58 | where 59 | 60 | c = cS .mor 61 | cRel = cS .morRel 62 | 63 | -- the function anaTreeS is compatible (respects the relations) 64 | anaTreeRel : ∀ {x y} → S .Rel x y → (j : Size) 65 | → ExtEq j (anaTree c ∞ x) (anaTree c ∞ y) 66 | subtreesE (anaTreeRel r j) {k} = 67 | (λ x mx → 68 | ∥map∥ (λ { (y , my , r') → 69 | anaTree c ∞ y , 70 | ∈mapList my , 71 | subst (λ z → ExtEq k z (anaTree c ∞ y)) 72 | (pre∈mapList mx .snd .snd) 73 | (anaTreeRel r' k)}) 74 | (cRel r .fst _ (pre∈mapList mx .snd .fst))) , 75 | (λ x mx → 76 | ∥map∥ (λ { (y , my , r') → 77 | anaTree c ∞ y , 78 | ∈mapList my , 79 | subst (λ z → ExtEq k z (anaTree c ∞ y)) 80 | (pre∈mapList mx .snd .snd) 81 | (anaTreeRel r' k)}) 82 | (cRel r .snd _ (pre∈mapList mx .snd .fst))) 83 | 84 | -- existence of anamorphism in setoids 85 | anaPfinS : S →S νPfinS 86 | anaPfinS = anaTree c ∞ , λ r → anaTreeRel r ∞ 87 | 88 | anaPfinSEq : subtreesS ∘S anaPfinS ≡S mapPfinS anaPfinS ∘S cS 89 | anaPfinSEq x = reflRelator (reflExtEq ∞) _ 90 | 91 | -- uniqueness 92 | anaPfinUniq' : (fS : S →S νPfinS) 93 | → (∀ x → Relator (ExtEq ∞) (subtrees (fS .mor x)) (mapList (fS .mor) (c x))) 94 | → (j : Size) → ∀ x → ExtEq j (fS .mor x) (anaTree c ∞ x) 95 | subtreesE (anaPfinUniq' fS fSeq j x) {k} = 96 | (λ t mt → 97 | ∥map∥ 98 | (λ { (u , mu , r) → _ , ∈mapList (pre∈mapList mu .snd .fst) , 99 | transExtEq k r 100 | (transExtEq k (subst (ExtEq k u) (sym (pre∈mapList mu .snd .snd)) (reflExtEq ∞ u)) 101 | (anaPfinUniq' fS fSeq k _)) }) 102 | (fSeq x .fst t mt) ) , 103 | λ t mt → 104 | ∥map∥ 105 | (λ { (u , mu , r) → u , mu , subst (λ z → ExtEq k z u) (pre∈mapList mt .snd .snd) 106 | (transExtEq k (symExtEq k (anaPfinUniq' fS fSeq k _)) r) }) 107 | (fSeq x .snd _ (∈mapList (pre∈mapList mt .snd .fst))) 108 | 109 | anaPfinUniq : (fS : S →S νPfinS) 110 | → subtreesS ∘S fS ≡S mapPfinS fS ∘S cS 111 | → fS ≡S anaPfinS 112 | anaPfinUniq fS fSeq = anaPfinUniq' fS fSeq ∞ 113 | 114 | -------------------------------------------------------------------------------- /FinalCoalg/InTypesAsCoindType.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module FinalCoalg.InTypesAsCoindType where 4 | 5 | open import Size 6 | open import Cubical.Core.Everything 7 | open import Cubical.Foundations.Prelude 8 | open import Cubical.Foundations.Everything 9 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 10 | open import Cubical.Relation.Everything 11 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 12 | renaming (map to ∥map∥; rec to ∥rec∥) 13 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 16 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 17 | open import Cubical.Relation.Binary 18 | open import Basics 19 | open import ListRelations 20 | open import Pfin.AsFreeJoinSemilattice 21 | open BinaryRelation 22 | 23 | -- final coalgebra of Pfin as a coinductive type 24 | record νPfin (i : Size) : Type where 25 | constructor buildP 26 | coinductive 27 | field 28 | subtreesP : {j : Size< i} → Pfin (νPfin j) 29 | open νPfin 30 | 31 | -- bisimilarity 32 | record νPfinB (i : Size) (s t : νPfin i) : Type where 33 | coinductive 34 | field 35 | subtreesPB : {j : Size< i} → PfinRel (νPfinB j) (subtreesP s) (subtreesP t) 36 | open νPfinB 37 | 38 | -- bisimilarity is symmetric and reflexive 39 | symνPfinB : ∀ (j : Size) {t t₁} → νPfinB j t t₁ → νPfinB j t₁ t 40 | subtreesPB (symνPfinB j b) = subtreesPB b .snd , subtreesPB b .fst 41 | 42 | reflνPfinB : (i : Size) (t : νPfin i) → νPfinB i t t 43 | subtreesPB (reflνPfinB i t) {j} = 44 | (λ x m → ∣ x , m , reflνPfinB j x ∣) , 45 | (λ x m → ∣ x , m , reflνPfinB j x ∣) 46 | 47 | -- coinduction principle: bisimilarity implies path equality 48 | bisim : (i : Size) (t u : νPfin i) → νPfinB i t u → t ≡ u 49 | subtreesP (bisim s t u b i) {s'} = 50 | PfinEq→Eq {s = subtreesP t}{subtreesP u} 51 | ((λ x m → ∥map∥ (λ { (y , my , b') → y , my , bisim s' x y b' }) (subtreesPB b .fst x m)) , 52 | (λ x m → ∥map∥ (λ { (y , my , b') → y , my , bisim s' x y b' }) (subtreesPB b .snd x m))) 53 | i 54 | 55 | -- anamorphism 56 | anaPfin : {X : Type} (c : X → Pfin X) (j : Size) → X → νPfin j 57 | subtreesP (anaPfin c j x) {k} = mapPfin (anaPfin c k) (c x) 58 | 59 | anaPfinEq : {X : Type} (c : X → Pfin X) (x : X) 60 | → subtreesP (anaPfin c ∞ x) ≡ mapPfin (anaPfin c ∞) (c x) 61 | anaPfinEq c x = refl 62 | 63 | -- uniqueness 64 | anaPfinUniqB : {X : Type} (c : X → Pfin X) 65 | → (f : (s : Size) → X → νPfin s) 66 | → (eq : ∀ (s : Size) (s' : Size< s) x 67 | → subtreesP (f s x) {s'} ≡ mapPfin (f s') (c x)) 68 | → (s : Size) → ∀ x → νPfinB s (f s x) (anaPfin c s x) 69 | anaPfinUniqB' : {X : Type} (c : X → Pfin X) 70 | → (f : (s : Size) → X → νPfin s) 71 | → (eq : ∀ (s : Size) (s' : Size< s) x 72 | → subtreesP (f s x) {s'} ≡ mapPfin (f s') (c x)) 73 | → (s : Size) 74 | → ∀ xs → PfinRel (νPfinB s) (mapPfin (f s) xs) (mapPfin (anaPfin c s) xs) 75 | subtreesPB (anaPfinUniqB c f eq s x) {s'} = 76 | subst (λ z → PfinRel (νPfinB s') z (mapPfin (anaPfin c s') (c x))) (sym (eq s s' x)) 77 | (anaPfinUniqB' c f eq s' (c x)) 78 | 79 | anaPfinUniqB' c f eq s = 80 | elimPfinProp 81 | (λ xs → _ , isPropPfinRel _ (mapPfin (f s) xs) (mapPfin (anaPfin c s) xs)) 82 | ((λ { _ () }) , λ { _ () }) 83 | (λ x → 84 | (λ y → ∥map∥ λ p → _ , ∣ refl ∣ , subst (λ z → νPfinB s z _) (sym p) (anaPfinUniqB c f eq s x)) , 85 | (λ y → ∥map∥ λ p → _ , ∣ refl ∣ , subst (λ z → νPfinB s z _) (sym p) (symνPfinB s (anaPfinUniqB c f eq s x)))) 86 | (λ {x}{y} → PfinRel∪ (νPfinB s) (mapPfin (f s) x) (mapPfin (f s) y) (mapPfin (anaPfin c s) x) (mapPfin (anaPfin c s) y)) 87 | 88 | anaPfinUniq : {X : Type} (c : X → Pfin X) 89 | → (f : X → νPfin ∞) 90 | → (eq : ∀ (s : Size) x → subtreesP (f x) {s} ≡ mapPfin f (c x)) 91 | → f ≡ anaPfin c ∞ 92 | anaPfinUniq c f eq = 93 | funExt λ x → bisim ∞ _ _ (anaPfinUniqB c (λ _ → f) (λ { _ _ → eq _ }) ∞ x) 94 | 95 | -- νPfin is a set 96 | 97 | νPfinIsSet : isSet (νPfin ∞) 98 | subtreesP (νPfinIsSet x y p q j k) = 99 | trunc (subtreesP x) (subtreesP y) 100 | (λ i → subtreesP (p i)) (λ i → subtreesP (q i)) 101 | j k 102 | 103 | 104 | -------------------------------------------------------------------------------- /FinalCoalg/InTypesAsSetQuot.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module FinalCoalg.InTypesAsSetQuot where 4 | 5 | open import Size 6 | open import Cubical.Core.Everything 7 | open import Cubical.Foundations.Prelude 8 | open import Cubical.Foundations.Everything 9 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 10 | open import Cubical.Relation.Everything 11 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 12 | renaming (map to ∥map∥; rec to ∥rec∥) 13 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 16 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 17 | open import Cubical.Relation.Binary 18 | open import Basics 19 | open import ListRelations 20 | open import ListStuff 21 | open import SetoidStuff 22 | open import Trees 23 | open import Pfin.AsSetQuot 24 | open import FinalCoalg.InSetoid 25 | open BinaryRelation 26 | 27 | -- the final coalgebra of PfinQ as quotient of trees 28 | νPfinQ : Type 29 | νPfinQ = Tree ∞ / ExtEq ∞ 30 | 31 | -- νPfinQ is a coalgebra for the functor PfinQ 32 | ξRel : ∀ ts ts' → DRelator (ExtEq ∞) ts ts' 33 | → DRelator _≡_ (mapList {B = νPfinQ} [_] ts) (mapList [_] ts') 34 | ξRel ts ts' p c mc with pre∈mapList mc 35 | ... | (t , mt , eq) = 36 | ∥map∥ (λ { (u , mu , r) → _ , ∈mapList mu , sym eq ∙ eq/ _ _ r}) (p t mt) 37 | 38 | ξ : νPfinQ → PfinQ νPfinQ 39 | ξ = recQ squash/ (λ t → [ mapList [_] (subtrees t) ]) 40 | λ t t' p → eq/ _ _ (ξRel _ _ (subtreesE p .fst) , ξRel _ _ (subtreesE p .snd)) 41 | 42 | -- two quotients of function spaces 43 | [_⇒PfinQ_] : (A B : Type) → Type 44 | [ A ⇒PfinQ B ] = [ A ⇒ (List B) ]/ SameEls 45 | 46 | [_⇒νPfinQ] : (A : Type) → Type 47 | [ A ⇒νPfinQ] = [ A ⇒ Tree ∞ ]/ ExtEq ∞ 48 | 49 | -- towards the construction of the anamorphism: there exists a map 50 | -- from X to νPfinQ, provided that X comes with a "coalgebra" 51 | -- c : [ X ⇒PfinQ X ] 52 | 53 | anaTreeRelCoalg : ∀{X}(c c' : X → List X) 54 | → (∀ x → SameEls (c x) (c' x)) → (j : Size) (x : X) 55 | → ExtEq j (anaTree c ∞ x) (anaTree c' ∞ x) 56 | anaTreeRelCoalg' : ∀{X}(c c' : X → List X) 57 | → (∀ x → SameEls (c x) (c' x)) → (j : Size) (x : X) 58 | → DRelator (ExtEq j) (mapList (anaTree c ∞) (c x)) 59 | (mapList (anaTree c' ∞) (c' x)) 60 | 61 | subtreesE (anaTreeRelCoalg c c' rc j x) {k} = 62 | anaTreeRelCoalg' c c' rc k x , 63 | anaTreeRelCoalg' c' c (λ z → symRelator (rc z)) k x 64 | 65 | anaTreeRelCoalg' c c' rc j x t mt with pre∈mapList mt 66 | ... | y , my , eq = 67 | ∥map∥ 68 | (λ { (z , mz , eqz) → 69 | _ , ∈mapList mz , 70 | subst (λ a → ExtEq j a (anaTree c' ∞ z)) eq 71 | (subst (λ a → ExtEq j (anaTree c ∞ y) (anaTree c' ∞ a)) eqz 72 | (anaTreeRelCoalg c c' rc j y)) } ) 73 | (rc x .fst y my) 74 | 75 | anaPfinQ' : {X : Type} (c : [ X ⇒PfinQ X ]) 76 | → X → νPfinQ 77 | anaPfinQ' = 78 | recQ (isSetΠ (λ _ → squash/)) (λ c x → [ anaTree c ∞ x ]) 79 | λ c c' rc → funExt (λ x → eq/ _ _ (anaTreeRelCoalg c c' rc ∞ x)) 80 | 81 | -- the construction of the anamorphism; 82 | -- for this to work, we assume that θ has a section, i.e. it is a 83 | -- split epimorphism; this is equivalent to full axiom of choice (the 84 | -- equivalence is proved in the end of the file AxiomChoice.agda) 85 | 86 | module _ (θInv : ∀ A {B} (R : B → B → Type) → (A → B / R) → [ A ⇒ B ]/ R) 87 | (sectionθ : ∀ A {B} (R : B → B → Type) → section (θ A R) (θInv A R)) where 88 | 89 | θ1 : ∀{X} → [ X ⇒PfinQ X ] → X → PfinQ X 90 | θ1 = θ _ _ 91 | 92 | θ1Inv : ∀ {X} → (X → PfinQ X) → [ X ⇒PfinQ X ] 93 | θ1Inv = θInv _ _ 94 | 95 | sectionθ1 : ∀{X} (f : X → PfinQ X) → θ1 (θ1Inv f) ≡ f 96 | sectionθ1 = sectionθ _ _ 97 | 98 | θ2 : ∀{X} → [ X ⇒νPfinQ] → X → νPfinQ 99 | θ2 = θ _ _ 100 | 101 | θ2-ns : ∀{X} → [ X ⇒νPfinQ] → X → νPfinQ 102 | θ2-ns = θ _ _ 103 | 104 | θ2Inv : ∀ {X} → (X → νPfinQ) → [ X ⇒νPfinQ] 105 | θ2Inv = θInv _ _ 106 | 107 | θ2Inv-ns : ∀ {X} → (X → νPfinQ) → [ X ⇒νPfinQ] 108 | θ2Inv-ns = θInv _ _ 109 | 110 | sectionθ2 : ∀{X} (f : X → νPfinQ) → θ2 (θ2Inv f) ≡ f 111 | sectionθ2 = sectionθ _ _ 112 | 113 | sectionθ2-ns : ∀{X} (f : X → νPfinQ) → θ2-ns (θ2Inv-ns f) ≡ f 114 | sectionθ2-ns = sectionθ _ _ 115 | 116 | anaPfinQ : {X : Type} (c : X → PfinQ X) 117 | → X → νPfinQ 118 | anaPfinQ c = anaPfinQ' (θ1Inv c) 119 | 120 | 121 | -- the anamorphism is a coalgebra morphism 122 | anaPfinQEq' : {X : Type} (c : [ X ⇒PfinQ X ]) 123 | → ∀ x → ξ (anaPfinQ' c x) ≡ mapPfinQ (anaPfinQ' c) (θ1 c x) 124 | anaPfinQEq' = 125 | elimProp (λ _ → isPropΠ (λ _ → squash/ _ _)) 126 | (λ c x → cong [_] (mapListComp (c x))) 127 | 128 | anaPfinQEq : {X : Type} (c : X → PfinQ X) 129 | → ∀ x → ξ (anaPfinQ c x) ≡ mapPfinQ (anaPfinQ c) (c x) 130 | anaPfinQEq c x = 131 | anaPfinQEq' (θ1Inv c) x ∙ cong (λ f → mapPfinQ (anaPfinQ c) (f x)) (sectionθ1 c) 132 | 133 | 134 | -- uniqueness 135 | 136 | anaPfinQUniq'' : {X : Type} (Xset : isSet X) (c : X → List X) 137 | → (f : [ X ⇒νPfinQ]) (feq : ∀ x → ξ (θ2 f x) ≡ mapPfinQ (θ2 f) [ c x ]) 138 | → ∀ x → θ2 f x ≡ anaPfinQ' [ c ] x 139 | anaPfinQUniq'' Xset c = 140 | elimProp (λ _ → isPropΠ (λ _ → isPropΠ (λ _ → squash/ _ _))) 141 | (λ f feq x → eq/ _ _ 142 | (anaPfinUniq' (Set→Setoid (_ , Xset)) 143 | (c , λ eq → subst (λ z → Relator _≡_ (c _) (c z)) eq (reflRelator (λ _ → refl) _)) 144 | (f , λ eq → subst (λ z → ExtEq ∞ (f _) (f z)) eq (reflExtEq ∞ _)) 145 | (λ y → 146 | (λ t mt → 147 | ∥rec∥ propTruncIsProp (uncurry (elimProp (λ _ → isPropΠ λ _ → propTruncIsProp) 148 | λ { t' (mt' , eq') → ∣ _ , ∈mapList (pre∈mapList mt' .snd .fst) , 149 | effective isPropExtEq isEquivRelExtEq _ _ (eq' ∙ sym (pre∈mapList mt' .snd .snd)) ∣ } )) 150 | (effective isPropSameEls isEquivRelSameEls _ _ (feq y) .fst [ t ] (∈mapList mt))) , 151 | (λ t mt → ∥rec∥ propTruncIsProp (uncurry (elimProp (λ _ → isPropΠ λ _ → propTruncIsProp) 152 | λ {t' (mt' , eq') → ∣ _ , pre∈mapList mt' .snd .fst , 153 | effective isPropExtEq isEquivRelExtEq _ _ (eq' ∙ sym (pre∈mapList mt' .snd .snd)) ∣ })) 154 | (effective isPropSameEls isEquivRelSameEls _ _ (feq y) .snd [ t ] (subst ([ t ] ∈_) (mapListComp (c y)) (∈mapList mt))))) 155 | ∞ 156 | x)) 157 | 158 | 159 | anaPfinQUniq' : {X : Type} (Xset : isSet X) (c : [ X ⇒PfinQ X ]) 160 | → (f : X → νPfinQ) (feq : ∀ x → ξ (f x) ≡ mapPfinQ f (θ1 c x)) 161 | → ∀ x → f x ≡ anaPfinQ' c x 162 | anaPfinQUniq' Xset = 163 | elimProp 164 | (λ _ → isPropΠ (λ _ → isPropΠ (λ _ → isPropΠ (λ _ → squash/ _ _)))) 165 | (λ c f feq x → (λ i → sym (sectionθ2 f) i x) 166 | ∙ anaPfinQUniq'' Xset c (θ2Inv f) 167 | (λ y → (λ i → ξ (sectionθ2 f i y) ) 168 | ∙ feq y 169 | ∙ cong (λ g → [ mapList g (c y) ]) (sym (sectionθ2 f))) 170 | x) 171 | 172 | anaPfinQUniq : {X : Type} (Xset : isSet X) (c : X → PfinQ X) 173 | → (f : X → νPfinQ) (feq : ∀ x → ξ (f x) ≡ mapPfinQ f (c x)) 174 | → f ≡ anaPfinQ c 175 | anaPfinQUniq Xset c f feq = funExt λ x → 176 | anaPfinQUniq' Xset (θ1Inv c) f 177 | (λ y → feq y ∙ λ i → mapPfinQ f (sectionθ1 c (~ i) y)) 178 | x 179 | -------------------------------------------------------------------------------- /ListRelations.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module ListRelations where 4 | 5 | open import Size 6 | open import Cubical.Core.Everything 7 | open import Cubical.Foundations.Prelude 8 | open import Cubical.Foundations.Everything 9 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 10 | open import Cubical.Relation.Everything 11 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 12 | renaming (rec to ∥rec∥; map to ∥map∥) 13 | open import Cubical.HITs.SetQuotients renaming ([_] to eqCl) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.Nat 16 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_; _≟_ to _≟N_) 17 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂) 18 | open import Cubical.Data.List renaming (map to mapList) 19 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 20 | open import Cubical.Relation.Binary 21 | open import Cubical.Relation.Nullary 22 | open BinaryRelation 23 | 24 | open import ListStuff 25 | 26 | -- canonical relation lifting on lists 27 | data ListRel {ℓ}{X Y : Type ℓ} (R : X → Y → Type ℓ) 28 | : List X → List Y → Type ℓ where 29 | [] : ListRel R [] [] 30 | _∷_ : ∀{x y xs ys} → R x y → ListRel R xs ys 31 | → ListRel R (x ∷ xs) (y ∷ ys) 32 | 33 | -- another relation lifting: two lists are (Relator R)-related if for 34 | -- each element in a list there merely exists an R-related element in 35 | -- the other list 36 | DRelator : ∀{ℓ}{X Y : Type ℓ} (R : X → Y → Type ℓ) 37 | → List X → List Y → Type ℓ 38 | DRelator R xs ys = 39 | ∀ x → x ∈ xs → ∃[ y ∈ _ ] y ∈ ys × R x y 40 | 41 | Relator : ∀{ℓ}{X : Type ℓ} (R : X → X → Type ℓ) 42 | → List X → List X → Type ℓ 43 | Relator R xs ys = 44 | DRelator R xs ys × DRelator R ys xs 45 | 46 | isPropRelator : ∀{ℓ}{X : Type ℓ} (R : X → X → Type ℓ) 47 | → ∀ xs ys → isProp (Relator R xs ys) 48 | isPropRelator R _ _ = 49 | isProp× (isPropΠ (λ _ → isPropΠ (λ _ → propTruncIsProp))) 50 | (isPropΠ (λ _ → isPropΠ (λ _ → propTruncIsProp))) 51 | 52 | Relatorₚ : ∀{ℓ}{X : Type ℓ} (R : X → X → hProp ℓ) 53 | → List X → List X → hProp ℓ 54 | Relatorₚ R xs ys = 55 | Relator (λ x y → ⟨ R x y ⟩) xs ys , 56 | isPropRelator _ xs ys 57 | 58 | -- reflexivity, symmetry and transitivity of the relator 59 | reflDRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 60 | → (∀ x → R x x) 61 | → ∀ xs → DRelator R xs xs 62 | reflDRelator reflR xs x mx = ∣ x , mx , reflR x ∣ 63 | 64 | reflRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 65 | → (∀ x → R x x) 66 | → ∀ xs → Relator R xs xs 67 | reflRelator reflR xs = (reflDRelator reflR xs) , (reflDRelator reflR xs) 68 | 69 | symRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 70 | → ∀ {xs ys} → Relator R xs ys → Relator R ys xs 71 | symRelator (p , q) = q , p 72 | 73 | transDRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 74 | → (∀ {x y z} → R x y → R y z → R x z) 75 | → ∀ {xs ys zs} → DRelator R xs ys → DRelator R ys zs → DRelator R xs zs 76 | transDRelator transR p q x mx = 77 | ∥rec∥ propTruncIsProp 78 | (λ { (y , my , ry) → ∥map∥ (λ { (z , mz , rz) → z , mz , transR ry rz }) (q y my)}) 79 | (p x mx) 80 | 81 | transRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 82 | → (∀ {x y z} → R x y → R y z → R x z) 83 | → ∀ {xs ys zs} → Relator R xs ys → Relator R ys zs → Relator R xs zs 84 | transRelator transR (p , q) (p' , q') = 85 | transDRelator transR p p' , transDRelator transR q' q 86 | 87 | isEquivRelRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 88 | → isEquivRel R → isEquivRel (Relator R) 89 | isEquivRelRelator (equivRel reflR _ transR) = 90 | equivRel (reflRelator reflR) 91 | (λ _ _ r → r .snd , r .fst) 92 | λ _ _ _ → transRelator (transR _ _ _) 93 | 94 | -- the relation Relator R is decidable if the relation R is decidable 95 | decDRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 96 | → (∀ x y → Dec (R x y)) → ∀ xs ys → Dec (DRelator R xs ys) 97 | decDRelator decR [] ys = yes (λ _ ()) 98 | decDRelator decR (x ∷ xs) ys with decDRelator decR xs ys 99 | ... | no ¬p = no (λ q → ¬p (λ y my → q y (there my))) 100 | ... | yes p with decMem decR x ys 101 | ... | yes q = yes (λ { ._ here → ∣ q ∣ 102 | ; y (there m) → p y m }) 103 | ... | no ¬q = no (λ q → ∥rec∥ isProp⊥ ¬q (q _ here)) 104 | 105 | decRelator : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 106 | → (∀ x y → Dec (R x y)) → ∀ xs ys → Dec (Relator R xs ys) 107 | decRelator decR xs ys with decDRelator decR xs ys 108 | ... | no ¬p = no (λ x → ¬p (fst x)) 109 | ... | yes p with decDRelator decR ys xs 110 | ... | yes q = yes (p , q) 111 | ... | no ¬q = no (λ z → ¬q (snd z)) 112 | 113 | -- interaction between relator and mapList 114 | DRelatorFunMapList : {X Y : Type}(f g : X → Y) 115 | → {R : Y → Y → Type} → (∀ x → R (f x) (g x)) 116 | → ∀ xs → DRelator R (mapList f xs) (mapList g xs) 117 | DRelatorFunMapList f g {R = R} r xs x mx with pre∈mapList mx 118 | ... | (y , my , eq) = ∣ g y , ∈mapList my , subst (λ z → R z (g y)) eq (r y) ∣ 119 | 120 | RelatorFunMapList : {X Y : Type}(f g : X → Y) 121 | → {R : Y → Y → Type} → (∀ x → R (f x) (g x)) 122 | → (∀ {x x'} → R x x' → R x' x) 123 | → ∀ xs → Relator R (mapList f xs) (mapList g xs) 124 | RelatorFunMapList f g r Rsym xs = 125 | DRelatorFunMapList f g r xs , DRelatorFunMapList g f (λ x → Rsym (r x)) xs 126 | 127 | -- properties of DRelator _≡_ 128 | DRelatorEq++₁ : {A : Type}{xs ys zs : List A} 129 | → DRelator _≡_ xs ys → DRelator _≡_ (xs ++ zs) (ys ++ zs) 130 | DRelatorEq++₁ {xs = xs}{ys} p x mx with ++∈ {xs = xs} mx 131 | ... | inj₁ mx' = ∥map∥ (λ { (y , my , eq) → y , ∈++₁ my , eq}) (p x mx') 132 | ... | inj₂ mx' = ∣ x , ∈++₂ {xs = ys} mx' , refl ∣ 133 | 134 | DRelatorEq++₂ : {A : Type}{xs ys zs : List A} 135 | → DRelator _≡_ ys zs → DRelator _≡_ (xs ++ ys) (xs ++ zs) 136 | DRelatorEq++₂ {xs = xs} p x mx with ++∈ {xs = xs} mx 137 | ... | inj₁ mx' = ∣ x , ∈++₁ mx' , refl ∣ 138 | ... | inj₂ mx' = 139 | ∥map∥ (λ { (y , my , eq) → y , ∈++₂ {xs = xs} my , eq }) (p x mx') 140 | 141 | DRelatorEqCom : {A : Type}(xs ys : List A) 142 | → DRelator _≡_ (xs ++ ys) (ys ++ xs) 143 | DRelatorEqCom xs ys x mx with ++∈ {xs = xs} mx 144 | ... | inj₁ mx' = ∣ x , ∈++₂ {xs = ys} mx' , refl ∣ 145 | ... | inj₂ mx' = ∣ x , ∈++₁ mx' , refl ∣ 146 | 147 | DRelatorEqAss₁ : {A : Type}(xs ys zs : List A) 148 | → DRelator _≡_ (xs ++ ys ++ zs) ((xs ++ ys) ++ zs) 149 | DRelatorEqAss₁ xs ys zs x mx with ++∈ {xs = xs} mx 150 | ... | inj₁ mx' = ∣ x , ∈++₁ (∈++₁ mx') , refl ∣ 151 | ... | inj₂ mx' with ++∈ {xs = ys} mx' 152 | ... | inj₁ mx'' = ∣ x , ∈++₁ (∈++₂ {xs = xs} mx'') , refl ∣ 153 | ... | inj₂ mx'' = ∣ x , ∈++₂ {xs = xs ++ ys} mx'' , refl ∣ 154 | 155 | DRelatorEqAss₂ : {A : Type}(xs ys zs : List A) 156 | → DRelator _≡_ ((xs ++ ys) ++ zs) (xs ++ ys ++ zs) 157 | DRelatorEqAss₂ xs ys zs x mx with ++∈ {xs = xs ++ ys} mx 158 | ... | inj₂ mx' = ∣ x , ∈++₂ {xs = xs} (∈++₂ {xs = ys} mx') , refl ∣ 159 | ... | inj₁ mx' with ++∈ {xs = xs} mx' 160 | ... | inj₁ mx'' = ∣ x , ∈++₁ mx'' , refl ∣ 161 | ... | inj₂ mx'' = ∣ x , ∈++₂ {xs = xs} (∈++₁ mx'') , refl ∣ 162 | 163 | DRelatorEqIdem : {A : Type}(xs : List A) 164 | → DRelator _≡_ (xs ++ xs) xs 165 | DRelatorEqIdem xs x mx with ++∈ {xs = xs} mx 166 | ... | inj₁ mx' = ∣ x , mx' , refl ∣ 167 | ... | inj₂ mx' = ∣ x , mx' , refl ∣ 168 | 169 | DRelatorEqNr : {A : Type}(xs : List A) 170 | → DRelator _≡_ (xs ++ []) xs 171 | DRelatorEqNr xs x mx with ++∈ {xs = xs} mx 172 | ... | inj₁ mx' = ∣ x , mx' , refl ∣ 173 | ... | inj₂ () 174 | 175 | -- given two duplicate-free lists xs and ys such that each element of 176 | -- xs is also ys, then xs is shorter than ys 177 | lengthDupFree : {A : Type} 178 | → (xs ys : List A) → dupFree xs → dupFree ys 179 | → DRelator _≡_ xs ys → length xs ≤N length ys 180 | lengthDupFree [] ys dxs dys r = zero-≤ 181 | lengthDupFree (x ∷ xs) ys (p , dxs) dys r = 182 | ∥rec∥ m≤n-isProp 183 | (λ { (y , m , eq) → 184 | ≤-trans (suc-≤-suc (lengthDupFree xs (remove ys m) dxs 185 | (dupFreeRemove m dys) 186 | λ z m' → ∥map∥ (λ { (a , ma , eqa) → 187 | a , 188 | ∈removeRel {R = _≡_} (λ _ → refl) m ma 189 | (λ eq' → p (subst (_∈ xs) (eqa ∙ eq' ∙ sym eq) m')) , 190 | eqa }) 191 | (r z (there m')))) 192 | (subst (suc (length (remove ys m)) ≤N_) (sym (lengthRemove m)) ≤-refl) }) 193 | (r _ here) 194 | -------------------------------------------------------------------------------- /ListStuff.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module ListStuff where 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.List renaming (map to mapList; [_] to sing) 7 | open import Cubical.Data.Sigma 8 | open import Cubical.Data.Empty renaming (rec to ⊥-rec) 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Nat.Order hiding (eq) 11 | open import Cubical.Data.Sum renaming (inl to inj₁; inr to inj₂; map to map⊎) 12 | open import Cubical.Relation.Nullary 13 | 14 | 15 | --open import Cubical.Core.Everything 16 | --open import Cubical.Data.Unit 17 | --open import Cubical.Functions.Logic hiding (⊥) 18 | --open import Cubical.HITs.SetQuotients renaming (rec to recQ) 19 | --open import Cubical.HITs.PropositionalTruncation as PropTrunc 20 | -- renaming (map to ∥map∥; rec to ∥rec∥) 21 | --open import Cubical.Relation.Binary hiding (Rel) 22 | --open BinaryRelation 23 | 24 | -- mapList preserves composition 25 | mapListComp : ∀{ℓ}{X Y Z : Type ℓ}{g : Y → Z}{f : X → Y} (xs : List X) 26 | → mapList g (mapList f xs) ≡ mapList (g ∘ f) xs 27 | mapListComp [] = refl 28 | mapListComp (x ∷ xs) = cong (_ ∷_) (mapListComp xs) 29 | 30 | 31 | -- list membership 32 | infix 21 _∈_ 33 | data _∈_ {ℓ}{X : Type ℓ} (x : X) : List X → Type ℓ where 34 | here : ∀{xs} → x ∈ (x ∷ xs) 35 | there : ∀{y xs} → x ∈ xs → x ∈ (y ∷ xs) 36 | 37 | -- properties of list membership 38 | ∈sing : ∀{ℓ}{X : Type ℓ}{x y : X} → x ∈ sing y → x ≡ y 39 | ∈sing here = refl 40 | 41 | ++∈ : ∀{ℓ}{X : Type ℓ}{x : X}{xs ys} → x ∈ (xs ++ ys) → x ∈ xs ⊎ x ∈ ys 42 | ++∈ {xs = []} m = inj₂ m 43 | ++∈ {xs = x ∷ xs} here = inj₁ here 44 | ++∈ {xs = x ∷ xs} (there m) = map⊎ there (λ z → z) (++∈ {xs = xs} m) 45 | 46 | ∈++₁ : ∀{ℓ}{X : Type ℓ}{x : X}{xs ys} → x ∈ xs → x ∈ (xs ++ ys) 47 | ∈++₁ here = here 48 | ∈++₁ (there p) = there (∈++₁ p) 49 | 50 | ∈++₂ : ∀{ℓ}{X : Type ℓ}{x : X}{xs ys} → x ∈ ys → x ∈ (xs ++ ys) 51 | ∈++₂ {xs = []} m = m 52 | ∈++₂ {xs = x ∷ xs} m = there (∈++₂ m) 53 | 54 | hereEq : ∀{ℓ}{X : Type ℓ}{x y : X}{xs} → x ≡ y → x ∈ (y ∷ xs) 55 | hereEq = J (λ z _ → _ ∈ (z ∷ _)) here 56 | 57 | -- length of a mapList 58 | lengthMapList : ∀{ℓ}{A B : Type ℓ}{f : A → B}(xs : List A) 59 | → length (mapList f xs) ≡ length xs 60 | lengthMapList [] = refl 61 | lengthMapList (x ∷ xs) = cong suc (lengthMapList xs) 62 | 63 | -- mapList is a monoid homomorphism (commutes with ++) 64 | mapList++ : ∀{ℓ}{A B : Type ℓ}{f : A → B}(xs ys : List A) 65 | → mapList f (xs ++ ys) ≡ mapList f xs ++ mapList f ys 66 | mapList++ [] ys = refl 67 | mapList++ (x ∷ xs) ys = cong (_ ∷_) (mapList++ xs ys) 68 | 69 | -- removing an element from a list 70 | remove : ∀{ℓ}{A : Type ℓ} {x : A} xs → x ∈ xs → List A 71 | remove (x ∷ xs) here = xs 72 | remove (y ∷ xs) (there m) = y ∷ remove xs m 73 | 74 | -- interaction between ∈ and remove 75 | remove∈ : ∀{ℓ}{A : Type ℓ} {x y : A} {xs} (m : x ∈ xs) 76 | → y ∈ remove xs m → y ∈ xs 77 | remove∈ here m' = there m' 78 | remove∈ (there m) here = here 79 | remove∈ (there m) (there m') = there (remove∈ m m') 80 | 81 | ∈removeRel : ∀{ℓ}{A : Type ℓ} 82 | → {R : A → A → Type ℓ} → (∀ x → R x x) 83 | → {xs : List A} {x w : A} (m : x ∈ xs) 84 | → w ∈ xs → (R w x → ⊥) 85 | → w ∈ remove xs m 86 | ∈removeRel reflR here here ¬r = ⊥-rec (¬r (reflR _)) 87 | ∈removeRel reflR here (there mw) ¬r = mw 88 | ∈removeRel reflR (there mx) here ¬r = here 89 | ∈removeRel reflR (there mx) (there mw) ¬r = there (∈removeRel reflR mx mw ¬r) 90 | 91 | -- length of remove 92 | lengthRemove : ∀{ℓ}{A : Type ℓ} {x : A} {xs} (m : x ∈ xs) 93 | → length xs ≡ suc (length (remove xs m)) 94 | lengthRemove here = refl 95 | lengthRemove (there m) = cong suc (lengthRemove m) 96 | 97 | -- properties of membership in the image of a list 98 | ∈mapList : {A B : Type} {f : A → B} {a : A} {xs : List A} 99 | → a ∈ xs → f a ∈ mapList f xs 100 | ∈mapList here = here 101 | ∈mapList (there ma) = there (∈mapList ma) 102 | 103 | pre∈mapList : {A B : Type} {f : A → B} {b : B} {xs : List A} 104 | → b ∈ mapList f xs → Σ[ a ∈ A ] (a ∈ xs) × (f a ≡ b) 105 | pre∈mapList {xs = x ∷ xs} here = _ , here , refl 106 | pre∈mapList {xs = x ∷ xs} (there mb) with pre∈mapList mb 107 | ... | a , ma , eq = a , there ma , eq 108 | 109 | -- list membership is decidable if equality on the carrier is 110 | -- decidable 111 | decMem : ∀{ℓ}{X : Type ℓ} {R : X → X → Type ℓ} 112 | → (∀ x y → Dec (R x y)) → ∀ x ys 113 | → Dec (Σ[ y ∈ _ ] (y ∈ ys) × R x y) 114 | decMem decR x [] = no (λ { () }) 115 | decMem decR x (y ∷ ys) with decR x y 116 | ... | yes p = yes (y , here , p) 117 | ... | no ¬p with decMem decR x ys 118 | ... | yes (z , m , r) = yes (z , there m , r) 119 | ... | no ¬q = no (λ { (._ , here , r') → ¬p r' 120 | ; (w , there m' , r') → ¬q (w , m' , r') }) 121 | 122 | -- removing all duplicates from a list (assuming decidable equality) 123 | removeDups : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 124 | → (xs : List A) → List A 125 | removeDups decA [] = [] 126 | removeDups decA (x ∷ xs) with decMem decA x xs 127 | ... | yes p = removeDups decA xs 128 | ... | no ¬p = x ∷ removeDups decA xs 129 | 130 | -- by removing all duplicates we obtain a generally shorter list 131 | lengthRemoveDups : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 132 | → (xs : List A) → length (removeDups decA xs) ≤ length xs 133 | lengthRemoveDups decA [] = ≤-refl 134 | lengthRemoveDups decA (x ∷ xs) with decMem decA x xs 135 | ... | yes p = ≤-trans (lengthRemoveDups decA xs) (≤-suc ≤-refl) 136 | ... | no ¬p = suc-≤-suc (lengthRemoveDups decA xs) 137 | 138 | -- lemmata about membership in removeDups 139 | removeDups∈ : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 140 | → (xs : List A) {x : A} 141 | → x ∈ removeDups decA xs → x ∈ xs 142 | removeDups∈ decA (x ∷ xs) m with decMem decA x xs 143 | ... | yes p = there (removeDups∈ decA xs m) 144 | removeDups∈ decA (x ∷ xs) here | no ¬p = here 145 | removeDups∈ decA (x ∷ xs) (there m) | no ¬p = 146 | there (removeDups∈ decA xs m) 147 | 148 | ∈removeDups : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 149 | → (xs : List A) {x : A} 150 | → x ∈ xs → x ∈ removeDups decA xs 151 | ∈removeDups decA (x ∷ xs) m with decMem decA x xs 152 | ∈removeDups decA (x ∷ xs) here | yes (y , m , eq) = subst (_∈ removeDups decA xs) (sym eq) (∈removeDups decA xs m) 153 | ∈removeDups decA (x ∷ xs) (there m) | yes p = ∈removeDups decA xs m 154 | ∈removeDups decA (x ∷ xs) here | no ¬p = here 155 | ∈removeDups decA (x ∷ xs) (there m) | no ¬p = there (∈removeDups decA xs m) 156 | 157 | -- interaction between mapList and removeDups 158 | removeDupMapList : {A B : Type} (decA : (x y : A) → Dec (x ≡ y)) 159 | → (decB : (x y : B) → Dec (x ≡ y)) 160 | → {f : A → B} (injf : ∀ x y → f x ≡ f y → x ≡ y) (xs : List A) 161 | → removeDups decB (mapList f xs) ≡ mapList f (removeDups decA xs) 162 | removeDupMapList decA decB injf [] = refl 163 | removeDupMapList decA decB {f} injf (x ∷ xs) with decMem decA x xs | decMem decB (f x) (mapList f xs) 164 | ... | yes p | yes q = removeDupMapList decA decB injf xs 165 | ... | yes (y , m , eq) | no ¬q = ⊥-rec (¬q (f y , ∈mapList m , cong f eq)) 166 | ... | no ¬p | no ¬q = cong (f x ∷_) (removeDupMapList decA decB injf xs) 167 | ... | no ¬p | yes (y , m , eq) with pre∈mapList m 168 | ... | (x' , m' , eq') = ⊥-rec (¬p (x' , m' , injf _ _ (eq ∙ sym eq'))) 169 | 170 | -- duplicate-free lists 171 | dupFree : {A : Type} → List A → Type 172 | dupFree [] = Unit 173 | dupFree (x ∷ xs) = (x ∈ xs → ⊥) × dupFree xs 174 | 175 | -- duplicate-freeness is preserved under removal of elements 176 | dupFreeRemove : {A : Type} 177 | → {xs : List A} {x : A} (m : x ∈ xs) 178 | → dupFree xs 179 | → dupFree (remove xs m) 180 | dupFreeRemove here (p , dxs) = dxs 181 | dupFreeRemove (there m) (p , dxs) = 182 | (λ m' → p (remove∈ m m')) , dupFreeRemove m dxs 183 | 184 | dupFreeRemoveDups : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 185 | → (xs : List A) → dupFree (removeDups decA xs) 186 | dupFreeRemoveDups decA [] = tt 187 | dupFreeRemoveDups decA (x ∷ xs) with decMem decA x xs 188 | ... | yes p = dupFreeRemoveDups decA xs 189 | ... | no ¬p = (λ q → ¬p (x , removeDups∈ decA xs q , refl)) , (dupFreeRemoveDups decA xs) 190 | -------------------------------------------------------------------------------- /Pfin/AsFreeJoinSemilattice.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Pfin.AsFreeJoinSemilattice where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (rec to ∥rec∥; map to ∥map∥) 12 | open import Cubical.HITs.SetQuotients 13 | renaming ([_] to eqCl; rec to recQ; rec2 to recQ2) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.List 16 | open import Cubical.Data.Nat renaming (elim to elimNat) 17 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂) 18 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 19 | open import Cubical.Relation.Binary 20 | open import Basics 21 | open import ListStuff 22 | open import Cubical.Relation.Nullary 23 | 24 | -- finite powerset as a HIT (the free join semilattice on A) 25 | data Pfin {ℓ} (A : Type ℓ) : Type ℓ where 26 | ø : Pfin A 27 | η : A → Pfin A 28 | _∪_ : Pfin A → Pfin A → Pfin A 29 | com : ∀ x y → x ∪ y ≡ y ∪ x 30 | ass : ∀ x y z → x ∪ (y ∪ z) ≡ (x ∪ y) ∪ z 31 | idem : ∀ x → x ∪ x ≡ x 32 | nr : ∀ x → x ∪ ø ≡ x 33 | trunc : isSet (Pfin A) 34 | 35 | -- recursion principle of Pfin 36 | module _ {ℓ}{A B : Type ℓ} (Bset : isSet B) 37 | (bø : B) (bη : A → B) 38 | (_b∪_ : B → B → B) 39 | (bcom : ∀ x y → x b∪ y ≡ y b∪ x) 40 | (bass : ∀ x y z → x b∪ (y b∪ z) ≡ (x b∪ y) b∪ z) 41 | (bidem : ∀ x → x b∪ x ≡ x) 42 | (bnr : ∀ x → x b∪ bø ≡ x) where 43 | 44 | recPfin : Pfin A → B 45 | recPfin ø = bø 46 | recPfin (η x) = bη x 47 | recPfin (s ∪ s₁) = (recPfin s) b∪ (recPfin s₁) 48 | recPfin (com s s₁ i) = bcom (recPfin s) (recPfin s₁) i 49 | recPfin (ass s s₁ s₂ i) = bass (recPfin s) (recPfin s₁) (recPfin s₂) i 50 | recPfin (idem s i) = bidem (recPfin s) i 51 | recPfin (nr s i) = bnr (recPfin s) i 52 | recPfin (trunc s s₁ x y i i₁) = 53 | Bset (recPfin s) (recPfin s₁) 54 | (\ j → recPfin (x j)) (\ j → recPfin (y j)) 55 | i i₁ 56 | 57 | -- finite subset membership 58 | _∈ₛ_ : ∀{A : Type} → A → Pfin A → hProp₀ 59 | x ∈ₛ ø = ⊥ₚ 60 | x ∈ₛ η y = x ≡ₚ y 61 | x ∈ₛ (s₁ ∪ s₂) = (x ∈ₛ s₁) ⊔ (x ∈ₛ s₂) 62 | x ∈ₛ com s₁ s₂ i = 63 | ⇔toPath {_} {x ∈ₛ s₁ ⊔ x ∈ₛ s₂} {x ∈ₛ s₂ ⊔ x ∈ₛ s₁} 64 | (∥map∥ λ { (inj₁ m) → inj₂ m ; (inj₂ m) → inj₁ m}) 65 | (∥map∥ λ { (inj₁ m) → inj₂ m ; (inj₂ m) → inj₁ m}) 66 | i 67 | x ∈ₛ ass s₁ s₂ s₃ i = 68 | ⇔toPath {_} {x ∈ₛ s₁ ⊔ x ∈ₛ s₂ ⊔ x ∈ₛ s₃} {(x ∈ₛ s₁ ⊔ x ∈ₛ s₂) ⊔ x ∈ₛ s₃} 69 | (∥rec∥ propTruncIsProp λ { (inj₁ m) → inl (inl m) 70 | ; (inj₂ m) → ∥map∥ (map⊎ inr (λ y → y)) m}) 71 | (∥rec∥ propTruncIsProp λ { (inj₁ m) → ∥map∥ (map⊎ (λ y → y) inl) m 72 | ; (inj₂ m) → inr (inr m)}) 73 | i 74 | x ∈ₛ idem s i = 75 | ⇔toPath {_} {x ∈ₛ s ⊔ x ∈ₛ s} {x ∈ₛ s} 76 | (∥rec∥ (isProp⟨⟩ (x ∈ₛ s)) (λ { (inj₁ m) → m ; (inj₂ m) → m })) 77 | inl 78 | i 79 | x ∈ₛ nr s i = 80 | ⇔toPath {_} {x ∈ₛ s ⊔ ⊥ₚ} {x ∈ₛ s} 81 | (∥rec∥ (isProp⟨⟩ (x ∈ₛ s)) (λ { (inj₁ m) → m ; (inj₂ ()) })) 82 | inl 83 | i 84 | x ∈ₛ trunc s₁ s₂ p q i j = 85 | isSetHProp (x ∈ₛ s₁) (x ∈ₛ s₂) (cong (x ∈ₛ_) p) (cong (x ∈ₛ_) q) i j 86 | 87 | -- Pfin action on functions 88 | mapPfin : ∀ {ℓ}{A B : Type ℓ} → (A → B) → Pfin A → Pfin B 89 | mapPfin f ø = ø 90 | mapPfin f (η x) = η (f x) 91 | mapPfin f (x ∪ y) = (mapPfin f x) ∪ (mapPfin f y) 92 | mapPfin f (com x y i) = com (mapPfin f x) (mapPfin f y) i 93 | mapPfin f (ass p p₁ p₂ i) = ass (mapPfin f p) (mapPfin f p₁) (mapPfin f p₂) i 94 | mapPfin f (idem p i) = idem (mapPfin f p) i 95 | mapPfin f (nr p i) = nr (mapPfin f p) i 96 | mapPfin f (trunc p q x y i j) = 97 | trunc _ _ (cong (mapPfin f) x) (cong (mapPfin f) y) i j 98 | 99 | 100 | -- elimination principle into a mere proposition 101 | module _ {ℓ}{A : Type ℓ} 102 | (P : Pfin A → hProp ℓ) (pø : ⟨ P ø ⟩) (pη : ∀ a → ⟨ P (η a) ⟩) 103 | (p∪ : ∀ {s₁ s₂} → ⟨ P s₁ ⟩ → ⟨ P s₂ ⟩ → ⟨ P (s₁ ∪ s₂) ⟩) where 104 | 105 | elimPfinProp : ∀ x → ⟨ P x ⟩ 106 | elimPfinProp ø = pø 107 | elimPfinProp (η x) = pη x 108 | elimPfinProp (s ∪ s') = p∪ (elimPfinProp s) (elimPfinProp s') 109 | elimPfinProp (com s s' i) = 110 | isProp→PathP (λ j → snd (P (com s s' j))) 111 | (p∪ (elimPfinProp s) (elimPfinProp s')) 112 | (p∪ (elimPfinProp s') (elimPfinProp s)) 113 | i 114 | elimPfinProp (ass s s₁ s₂ i) = 115 | isProp→PathP (λ j → snd (P (ass s s₁ s₂ j))) 116 | (p∪ (elimPfinProp s) (p∪ (elimPfinProp s₁) (elimPfinProp s₂))) 117 | (p∪ (p∪ (elimPfinProp s) (elimPfinProp s₁)) (elimPfinProp s₂)) 118 | i 119 | elimPfinProp (idem s i) = 120 | isProp→PathP (λ j → snd (P (idem s j))) 121 | (p∪ (elimPfinProp s) (elimPfinProp s)) 122 | (elimPfinProp s) 123 | i 124 | elimPfinProp (nr s i) = 125 | isProp→PathP (λ j → snd (P (nr s j))) 126 | (p∪ (elimPfinProp s) pø) 127 | (elimPfinProp s) 128 | i 129 | elimPfinProp (trunc s s' p q i j) = 130 | isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (snd (P x))) 131 | (elimPfinProp s) (elimPfinProp s') 132 | (cong elimPfinProp p) (cong elimPfinProp q) 133 | (trunc s s' p q) 134 | i j 135 | 136 | -- functor laws of mapPfin 137 | mapPfinId : ∀{ℓ} {A : Type ℓ} (s : Pfin A) 138 | → mapPfin (λ x → x) s ≡ s 139 | mapPfinId = 140 | elimPfinProp (λ _ → _ , trunc _ _) 141 | refl 142 | (λ _ → refl) 143 | λ eq1 eq2 → cong₂ _∪_ eq1 eq2 144 | 145 | mapPfinComp : ∀{ℓ} {A B C : Type ℓ} {g : B → C} {f : A → B} (s : Pfin A) 146 | → mapPfin g (mapPfin f s) ≡ mapPfin (g ∘ f) s 147 | mapPfinComp = 148 | elimPfinProp (λ _ → _ , trunc _ _) 149 | refl 150 | (λ _ → refl) 151 | λ eq1 eq2 → cong₂ _∪_ eq1 eq2 152 | 153 | -- an algebraic order on subsets , given by the presence of joins 154 | _≤_ : ∀{A : Type} → Pfin A → Pfin A → Type₀ 155 | s ≤ t = (s ∪ t) ≡ t 156 | 157 | antisym≤ : ∀{A : Type}{s t : Pfin A} → s ≤ t → t ≤ s → s ≡ t 158 | antisym≤ p q = sym q ∙ com _ _ ∙ p 159 | 160 | isProp≤ : ∀{A : Type}{s t : Pfin A} → isProp (s ≤ t) 161 | isProp≤ = trunc _ _ 162 | 163 | -- joins are least upper bounds wrt. ≤ 164 | ∪isLub : ∀{A : Type}{s t : Pfin A} (u : Pfin A) 165 | → s ≤ u → t ≤ u → (s ∪ t) ≤ u 166 | ∪isLub {s = s}{t} u ls lt = 167 | sym (ass _ _ _) 168 | ∙ cong (s ∪_) lt 169 | ∙ ls 170 | 171 | -- subset relation expressed in terms of membership 172 | _⊆_ : ∀{A : Type} → Pfin A → Pfin A → Type₀ 173 | s ⊆ t = ∀ x → ⟨ x ∈ₛ s ⟩ → ⟨ x ∈ₛ t ⟩ 174 | 175 | trans⊆ : ∀{A : Type} {xs ys zs : Pfin A} 176 | → xs ⊆ ys → ys ⊆ zs → xs ⊆ zs 177 | trans⊆ p q x m = q x (p x m) 178 | 179 | -- ⊆ implies ≤ 180 | ⊂2≤-η : ∀{A : Type}(a : A) (s : Pfin A) → ⟨ a ∈ₛ s ⟩ → η a ≤ s 181 | ⊂2≤-η a = elimPfinProp (λ _ → _ , isPropΠ λ x → isProp≤) 182 | (λ ()) 183 | (λ b → ∥rec∥ isProp≤ λ eq → cong (_∪ η b) (cong η eq) ∙ idem _) 184 | (λ {s₁}{s₂} p₁ p₂ → ∥rec∥ isProp≤ 185 | λ { (inj₁ m) → ass _ _ _ ∙ cong (_∪ _) (p₁ m) 186 | ; (inj₂ m) → ass _ _ _ ∙ cong (_∪ s₂) (com _ _) ∙ sym (ass _ _ _) ∙ cong (_ ∪_) (p₂ m)}) 187 | 188 | ⊂2≤ : ∀{A : Type}(s t : Pfin A) → t ⊆ s → t ≤ s 189 | ⊂2≤ s = elimPfinProp (λ _ → _ , isPropΠ λ x → isProp≤) 190 | (λ p → com ø s ∙ nr s) 191 | (λ a m → ⊂2≤-η a s (m a ∣ refl ∣)) 192 | (λ p₁ p₂ q → 193 | ∪isLub s (p₁ (λ x m → q x (inl m))) (p₂ (λ x m → q x (inr m)))) 194 | 195 | -- canonical directed relation lifting on Pfin 196 | PfinDRel : ∀{X Y} (R : X → Y → Type₀) 197 | → Pfin X → Pfin Y → Type₀ 198 | PfinDRel R s₁ s₂ = 199 | ∀ x → ⟨ x ∈ₛ s₁ ⟩ → ∃[ y ∈ _ ] ⟨ y ∈ₛ s₂ ⟩ × R x y 200 | 201 | -- symmetric relation lifting 202 | PfinRel : ∀{X} (R : X → X → Type₀) 203 | → Pfin X → Pfin X → Type₀ 204 | PfinRel R s₁ s₂ = 205 | PfinDRel R s₁ s₂ × PfinDRel R s₂ s₁ 206 | 207 | isPropPfinRel : ∀{X} (R : X → X → Type₀) 208 | → ∀ s t → isProp (PfinRel R s t) 209 | isPropPfinRel R s t = 210 | isProp× (isPropΠ (λ _ → isPropΠ (λ _ → propTruncIsProp))) 211 | (isPropΠ (λ _ → isPropΠ (λ _ → propTruncIsProp))) 212 | 213 | PfinRelₚ : ∀{X} (R : X → X → hProp₀) 214 | → Pfin X → Pfin X → hProp₀ 215 | PfinRelₚ R s t = 216 | PfinRel (λ x y → ⟨ R x y ⟩) s t , isPropPfinRel _ s t 217 | 218 | -- relation lifting is a congruence 219 | PfinRel∪ : ∀{X} (R : X → X → Type₀) 220 | → (s s' t t' : Pfin X) 221 | → PfinRel R s t → PfinRel R s' t' 222 | → PfinRel R (s ∪ s') (t ∪ t') 223 | PfinRel∪ R s s' t t' (p , p') (q , q') = 224 | (λ x → 225 | ∥rec∥ propTruncIsProp 226 | λ { (inj₁ m) → ∥map∥ (λ { (y , my , r) → y , inl my , r}) (p _ m) 227 | ; (inj₂ m) → ∥map∥ (λ { (y , my , r) → y , inr my , r}) (q _ m) }) , 228 | (λ x → 229 | ∥rec∥ propTruncIsProp 230 | λ { (inj₁ m) → ∥map∥ (λ { (y , my , r) → y , inl my , r}) (p' _ m) 231 | ; (inj₂ m) → ∥map∥ (λ { (y , my , r) → y , inr my , r}) (q' _ m) }) 232 | 233 | -- extensional equality of finite subsets: they are equal if they 234 | -- contain the same elements 235 | PfinEq : ∀{X} → Pfin X → Pfin X → Type₀ 236 | PfinEq = PfinRel _≡_ 237 | 238 | -- extensional equality of finite subsets is equivalent to path 239 | -- equality 240 | PfinDRel⊆ : ∀{X}(s t : Pfin X) → PfinDRel _≡_ s t → s ⊆ t 241 | PfinDRel⊆ s t p x mx = 242 | ∥rec∥ (snd (x ∈ₛ t)) 243 | (λ { (y , my , eq) → subst (λ z → ⟨ z ∈ₛ t ⟩) (sym eq) my }) (p x mx) 244 | 245 | PfinDRel⊆2 : ∀{X}(s t : Pfin X) → PfinDRel (λ y x → x ≡ y) t s → t ⊆ s 246 | PfinDRel⊆2 s t p x mx = 247 | ∥rec∥ (snd (x ∈ₛ s)) 248 | (λ { (y , my , eq) → subst (λ z → ⟨ z ∈ₛ s ⟩) eq my }) (p x mx) 249 | 250 | PfinEq→Eq : ∀{X} {s t : Pfin X} → PfinEq s t → s ≡ t 251 | PfinEq→Eq {s = s}{t} (p₁ , p₂) = 252 | antisym≤ (⊂2≤ _ _ (PfinDRel⊆ s t p₁)) (⊂2≤ _ _ (PfinDRel⊆ t s p₂)) 253 | 254 | -- properties of membership in the image of a finite subset 255 | ∈ₛmapPfin : ∀{A B} (f : A → B) (a : A) (s : Pfin A) 256 | → ⟨ a ∈ₛ s ⟩ → ⟨ f a ∈ₛ mapPfin f s ⟩ 257 | ∈ₛmapPfin f a = 258 | elimPfinProp (λ x → _ , isPropΠ (λ _ → snd (f a ∈ₛ mapPfin f x))) 259 | (λ ()) 260 | (λ b → ∥map∥ (cong f)) 261 | λ p₁ p₂ → ∥map∥ (map⊎ p₁ p₂) 262 | 263 | pre∈ₛmapPfin : ∀{A B} (f : A → B) (b : B) (s : Pfin A) 264 | → ⟨ b ∈ₛ mapPfin f s ⟩ → ∃[ a ∈ A ] ⟨ a ∈ₛ s ⟩ × (f a ≡ b) 265 | pre∈ₛmapPfin f b = 266 | elimPfinProp (λ x → _ , isPropΠ (λ _ → propTruncIsProp)) 267 | (λ ()) 268 | (λ a → ∥map∥ (λ eq → a , ∣ refl ∣ , sym eq)) 269 | λ p₁ p₂ → ∥rec∥ propTruncIsProp 270 | (λ { (inj₁ m) → ∥map∥ (λ {(a , m , eq) → a , inl m , eq}) (p₁ m) 271 | ; (inj₂ m) → ∥map∥ (λ {(a , m , eq) → a , inr m , eq}) (p₂ m) }) 272 | 273 | 274 | -- some properties of ⊆ 275 | ∪⊆ : ∀{A : Type} (s1 s2 t : Pfin A) → s1 ⊆ t → s2 ⊆ t → (s1 ∪ s2) ⊆ t 276 | ∪⊆ s1 s2 t p q x = 277 | ∥rec∥ (snd (x ∈ₛ t)) λ { (inj₁ m) → p x m ; (inj₂ m) → q x m } 278 | 279 | ∪⊆1 : ∀{A : Type} (s1 s2 t : Pfin A) → (s1 ∪ s2) ⊆ t → s1 ⊆ t 280 | ∪⊆1 s1 s2 t p x m = p x (inl m) 281 | 282 | ∪⊆2 : ∀{A : Type} (s1 s2 t : Pfin A) → (s1 ∪ s2) ⊆ t → s2 ⊆ t 283 | ∪⊆2 s1 s2 t p x m = p x (inr m) 284 | 285 | map∪⊆ : ∀{A : Type} (s1 s2 t1 t2 : Pfin A) → s1 ⊆ t1 → s2 ⊆ t2 → (s1 ∪ s2) ⊆ (t1 ∪ t2) 286 | map∪⊆ s1 s2 t1 t2 p q x = 287 | ∥map∥ λ { (inj₁ m) → inj₁ (p x m) ; (inj₂ m) → inj₂ (q x m) } 288 | 289 | ⊆∪ : ∀{A : Type} (s1 s2 t : Pfin A) 290 | → t ⊆ (s1 ∪ s2) → ∃[ t1 ∈ Pfin A ] Σ[ t2 ∈ Pfin A ] (t1 ⊆ s1) × (t2 ⊆ s2) × (t ≡ t1 ∪ t2) 291 | ⊆∪ s1 s2 = 292 | elimPfinProp (λ _ → _ , isPropΠ (λ _ → propTruncIsProp)) 293 | (λ x → ∣ ø , ø , (λ { _ () }) , (λ { _ () }) , sym (idem _) ∣) 294 | (λ a m → 295 | ∥map∥ 296 | (λ { (inj₁ p) → η a , ø , 297 | (λ x → ∥rec∥ (snd (x ∈ₛ s1)) λ eq → subst (λ z → ⟨ z ∈ₛ s1 ⟩) (sym eq) p) , 298 | (λ { _ () }) , 299 | sym (nr _) ; 300 | (inj₂ p) → ø , η a , 301 | (λ { _ () }) , 302 | (λ x → ∥rec∥ (snd (x ∈ₛ s2)) λ eq → subst (λ z → ⟨ z ∈ₛ s2 ⟩) (sym eq) p) , 303 | sym (com _ _ ∙ nr _) }) 304 | (m a ∣ refl ∣)) 305 | λ ih1 ih2 p → 306 | ∥rec∥ propTruncIsProp 307 | (λ { (u1 , u2 , m1 , m2 , eq1) → 308 | ∥map∥ 309 | (λ { (v1 , v2 , n1 , n2 , eq2) → 310 | (u1 ∪ v1) , (u2 ∪ v2) , 311 | ∪⊆ u1 v1 s1 m1 n1 , 312 | ∪⊆ u2 v2 s2 m2 n2 , 313 | cong₂ _∪_ eq1 eq2 314 | ∙ ass _ _ _ 315 | ∙ cong (_∪ v2) (sym (ass _ _ _) 316 | ∙ cong (u1 ∪_) (com _ _) 317 | ∙ ass _ _ _) 318 | ∙ sym (ass _ _ _) }) 319 | (ih2 (λ x m → p x (inr m))) }) 320 | (ih1 λ x m → p x (inl m)) 321 | 322 | 323 | pre⊆mapPfin : ∀{A B} (f : A → B) (s : Pfin A) (t : Pfin B) 324 | → t ⊆ mapPfin f s → ∃[ s' ∈ Pfin A ] (s' ⊆ s) × (mapPfin f s' ≡ t) 325 | pre⊆mapPfin f s = 326 | elimPfinProp (λ x → _ , isPropΠ (λ _ → propTruncIsProp)) 327 | (λ x → ∣ ø , (λ { _ () }) , refl ∣) 328 | (λ b p → 329 | ∥map∥ 330 | (λ { (a , m , eq) → 331 | η a , 332 | (λ x → ∥rec∥ (snd (x ∈ₛ s)) λ r → subst (λ y → ⟨ y ∈ₛ s ⟩) (sym r) m) , 333 | cong η eq}) 334 | (pre∈ₛmapPfin f b s (p b ∣ refl ∣))) 335 | λ {t1} {t2} ih1 ih2 p → 336 | ∥rec∥ propTruncIsProp 337 | (λ { (u1 , m1 , eq1) → 338 | ∥map∥ 339 | (λ { (u2 , m2 , eq2) → (u1 ∪ u2) , ∪⊆ u1 u2 s m1 m2 , cong₂ _∪_ eq1 eq2 }) 340 | (ih2 (∪⊆2 t1 t2 (mapPfin f s) p)) }) 341 | (ih1 (∪⊆1 t1 t2 (mapPfin f s) p)) 342 | 343 | antisym⊆ : ∀{A : Type}{s t : Pfin A} → s ⊆ t → t ⊆ s → s ≡ t 344 | antisym⊆ p q = antisym≤ (⊂2≤ _ _ p) (⊂2≤ _ _ q) 345 | 346 | -- an equivalent definition of extensional equality of finite subsets 347 | _≡ₛ_ : ∀{A : Type} → Pfin A → Pfin A → Type 348 | s ≡ₛ t = (s ⊆ t) × (t ⊆ s) 349 | 350 | -- injectivity of η 351 | 352 | ηisInjective' : {A : Type} (setA : isSet A) 353 | → {a b : A} 354 | → η a ⊆ η b → a ≡ b 355 | ηisInjective' setA {a} p = 356 | ∥rec∥ (setA _ _) 357 | (λ x → x) 358 | (p a ∣ refl ∣) 359 | 360 | ηisInjective : {A : Type} (setA : isSet A) 361 | → {a b : A} 362 | → η a ≡ η b → a ≡ b 363 | ηisInjective setA {a} eq = 364 | ηisInjective' setA (subst (η a ⊆_) eq (λ _ m → m)) 365 | 366 | -- ø dijoint from η 367 | 368 | ødisjη' : {A : Type} {a : A} → η a ⊆ ø → ⊥ 369 | ødisjη' {a = a} p = p a ∣ refl ∣ 370 | 371 | ødisjη : {A : Type} {a : A} → η a ≡ ø → ⊥ 372 | ødisjη {a = a} eq = ødisjη' (subst (η a ⊆_) eq (λ _ m → m)) 373 | 374 | -- if a function g is injective, then mapPfin g is injective as well 375 | mapPfinInj' : ∀{A B} (g : A → B) → (∀ x y → g x ≡ g y → x ≡ y) 376 | → (s t : Pfin A) → mapPfin g s ⊆ mapPfin g t → s ⊆ t 377 | mapPfinInj' g injg s t p x m = 378 | ∥rec∥ (snd (x ∈ₛ t)) 379 | (λ { (x' , m' , eq) → subst (λ z → ⟨ z ∈ₛ t ⟩) (injg _ _ eq) m' }) 380 | (pre∈ₛmapPfin g _ t (p (g x) (∈ₛmapPfin g x s m))) 381 | 382 | mapPfinInj : ∀{A B} (g : A → B) → (∀ x y → g x ≡ g y → x ≡ y) 383 | → (s t : Pfin A) → mapPfin g s ≡ mapPfin g t → s ≡ t 384 | mapPfinInj g injg s t p = 385 | antisym⊆ 386 | (mapPfinInj' g injg s t (subst (mapPfin g s ⊆_) p (λ _ m → m))) 387 | (mapPfinInj' g injg t s (subst (_⊆ mapPfin g s) p (λ _ m → m))) 388 | 389 | -- mapPfin f s is empty iff s is empty 390 | mapPfinø' : ∀{A B} (f : A → B) (s : Pfin A) 391 | → mapPfin f s ⊆ ø → s ⊆ ø 392 | mapPfinø' f s p x m = p (f x) (∈ₛmapPfin f x s m) 393 | 394 | mapPfinø : ∀{A B} (f : A → B) (s : Pfin A) 395 | → mapPfin f s ≡ ø → s ≡ ø 396 | mapPfinø f s eq = 397 | antisym⊆ 398 | (mapPfinø' f s (subst (mapPfin f s ⊆_) eq (λ _ m → m))) 399 | λ { _ ()} 400 | 401 | -- if mapPfin f s is a singleton η b, then s is a singleton in case f 402 | -- is injective 403 | mapPfinη' : ∀{A B} (setA : isSet A) 404 | → (f : A → B) → (∀ x y → f x ≡ f y → x ≡ y) 405 | → (s : Pfin A) (b : B) 406 | → mapPfin f s ⊆ η b → η b ⊆ mapPfin f s 407 | → Σ[ a ∈ A ] s ≡ η a 408 | mapPfinη' setA f injf s b p q = 409 | ∥rec∥ (λ { (x , r1) (y , r2) → Σ≡Prop (λ _ → trunc _ _) (ηisInjective setA (sym r1 ∙ r2)) }) 410 | (λ { (a , ma , eq) → a , 411 | (antisym⊆ (λ x mx → ∥map∥ (λ eq' → injf x a (eq' ∙ sym eq)) (p (f x) (∈ₛmapPfin f x s mx))) 412 | (λ x → ∥rec∥ (snd (x ∈ₛ s)) (λ eqx → subst (λ z → ⟨ z ∈ₛ s ⟩) (sym eqx) ma))) }) 413 | (pre∈ₛmapPfin f b s (q b ∣ refl ∣)) 414 | 415 | mapPfinη : ∀{A B} (setA : isSet A) 416 | → (f : A → B) → (∀ x y → f x ≡ f y → x ≡ y) 417 | → (s : Pfin A) (b : B) 418 | → mapPfin f s ≡ η b → Σ[ a ∈ A ] s ≡ η a 419 | mapPfinη {A} setA f injf s b eq = 420 | mapPfinη' setA f injf s b 421 | (subst (mapPfin f s ⊆_) eq (λ _ m → m)) (subst (η b ⊆_) (sym eq) (λ _ m → m)) 422 | 423 | -- if mapPfin f s is a binary union t1 ∪ t2, then s is a binary union 424 | -- s1 ∪ s2 with mapPfin f s1 ≡ t1 and mapPfin f s2 ≡ t2 in case f is 425 | -- injective 426 | ∪⊆mapPfin : ∀{A B} (f : A → B) 427 | → (s : Pfin A) (t1 t2 : Pfin B) 428 | → (t1 ∪ t2) ⊆ mapPfin f s 429 | → ∃[ s1 ∈ Pfin A ] Σ[ s2 ∈ Pfin A ] ((s1 ∪ s2) ⊆ s) × (t1 ≡ mapPfin f s1) × (t2 ≡ mapPfin f s2) 430 | ∪⊆mapPfin f s t1 t2 mt = 431 | ∥rec∥ propTruncIsProp 432 | (λ { (u1 , m1 , eq1) → ∥map∥ 433 | (λ { (u2 , m2 , eq2) → u1 , u2 , ∪⊆ u1 u2 s m1 m2 , sym eq1 , sym eq2 }) 434 | (pre⊆mapPfin f s t2 λ x mx → mt x (inr mx)) }) 435 | (pre⊆mapPfin f s t1 λ x mx → mt x (inl mx)) 436 | 437 | ∪≡mapPfin : ∀{A B} (f : A → B) → (∀ x y → f x ≡ f y → x ≡ y) 438 | → (s : Pfin A) (t1 t2 : Pfin B) 439 | → (t1 ∪ t2) ≡ mapPfin f s 440 | → ∃[ s1 ∈ Pfin A ] Σ[ s2 ∈ Pfin A ] (s1 ∪ s2 ≡ s) × (t1 ≡ mapPfin f s1) × (t2 ≡ mapPfin f s2) 441 | ∪≡mapPfin f injf s t1 t2 eq = 442 | ∥rec∥ propTruncIsProp 443 | (λ { (u1 , m1 , eq1) → ∥map∥ 444 | (λ { (u2 , m2 , eq2) → u1 , u2 , mapPfinInj f injf _ _ (cong₂ _∪_ eq1 eq2 ∙ eq) , sym eq1 , sym eq2 }) 445 | (pre⊆mapPfin f s t2 (subst (t2 ⊆_) eq (λ _ → inr))) }) 446 | (pre⊆mapPfin f s t1 (subst (t1 ⊆_) eq (λ _ → inl))) 447 | 448 | 449 | -- if path equality on A is decidable, then also the membership 450 | -- relation between A and Pfin A is decidable 451 | dec∈ₛ : {A : Type} 452 | → ((x y : A) → Dec (x ≡ y)) 453 | → (x : A) (s : Pfin A) → Dec ⟨ x ∈ₛ s ⟩ 454 | dec∈ₛ decEq x = elimPfinProp 455 | (λ s → _ , decIsProp (snd (x ∈ₛ s))) 456 | (no λ x → x) 457 | (λ a → decPropTrunc (decEq x a)) 458 | λ {s1}{s2} → lem {s1}{s2} 459 | where 460 | lem : ∀{s1 s2} → Dec ⟨ x ∈ₛ s1 ⟩ → Dec ⟨ x ∈ₛ s2 ⟩ 461 | → Dec ⟨ x ∈ₛ (s1 ∪ s2) ⟩ 462 | lem (yes p) d2 = yes (inl p) 463 | lem (no ¬p) (yes p) = yes (inr p) 464 | lem (no ¬p) (no ¬q) = no (∥rec∥ isProp⊥ 465 | (λ { (inj₁ x) → ¬p x ; (inj₂ x) → ¬q x })) 466 | 467 | -- if path equality on A is decidable, then also the subset relation 468 | -- on Pfin A is decidable 469 | dec⊆ : {A : Type} 470 | → ((x y : A) → Dec (x ≡ y)) 471 | → (s t : Pfin A) → Dec (s ⊆ t) 472 | dec⊆ {A} decEq = 473 | elimPfinProp (λ _ → _ , isPropΠ λ x → isPropDec (isPropΠ (λ y → isPropΠ (λ _ → snd (y ∈ₛ x))))) 474 | (λ t → yes (λ { _ () })) 475 | (λ a → lem' {a}) 476 | λ {s1}{s2} d1 d2 t → lem {s1} {s2} t (d1 t) (d2 t) 477 | where 478 | lem' : ∀ {a : A} t → Dec (η a ⊆ t) 479 | lem' {a} t with dec∈ₛ decEq a t 480 | ... | yes p = yes (λ y → ∥rec∥ (snd (y ∈ₛ t)) (λ eq → subst (λ z → ⟨ z ∈ₛ t ⟩) (sym eq) p)) 481 | ... | no ¬p = no (λ q → ¬p (q _ ∣ refl ∣)) 482 | 483 | lem : ∀ {s1 s2 : Pfin A} t 484 | → Dec (s1 ⊆ t) → Dec (s2 ⊆ t) 485 | → Dec ((s1 ∪ s2) ⊆ t) 486 | lem {s1}{s2} t (yes p) (yes q) = yes (∪⊆ s1 s2 t p q) 487 | lem t (yes p) (no ¬p) = no (λ q → ¬p λ x mx → q x (inr mx)) 488 | lem t (no ¬p) d2 = no (λ q → ¬p (λ x mx → q x (inl mx))) 489 | 490 | -- if path equality on A is decidable, then also path equality on Pfin A 491 | -- is decidable 492 | PfinDecEq : {A : Type} 493 | → ((x y : A) → Dec (x ≡ y)) 494 | → (x y : Pfin A) → Dec (x ≡ y) 495 | PfinDecEq decEq x y with dec⊆ decEq x y | dec⊆ decEq y x 496 | ... | yes p | yes q = yes (antisym⊆ p q) 497 | ... | yes p | no ¬q = no (λ eq → ¬q (λ a → subst (λ z → ⟨ a ∈ₛ z ⟩) (sym eq))) 498 | ... | no ¬p | _ = no (λ eq → ¬p (λ a → subst (λ z → ⟨ a ∈ₛ z ⟩) eq)) 499 | 500 | -- Pfin preserves isomorphisms and equivalences 501 | Pfin-iso : {A B : Type} → Iso A B → Iso (Pfin A) (Pfin B) 502 | Pfin-iso AisoB = 503 | iso (mapPfin (Iso.fun AisoB)) 504 | (mapPfin (Iso.inv AisoB)) 505 | (λ x → mapPfinComp x ∙ (λ i → mapPfin (λ y → Iso.rightInv AisoB y i) x) ∙ mapPfinId x) 506 | λ x → mapPfinComp x ∙ (λ i → mapPfin (λ y → Iso.leftInv AisoB y i) x) ∙ mapPfinId x 507 | 508 | Pfin≃ : {A B : Type} → A ≃ B → Pfin A ≃ Pfin B 509 | Pfin≃ eq = isoToEquiv (Pfin-iso (equivToIso eq)) 510 | -------------------------------------------------------------------------------- /Pfin/AsSetQuot.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Pfin.AsSetQuot where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (rec to ∥rec∥; map to ∥map∥) 12 | open import Cubical.HITs.SetQuotients 13 | renaming (rec to recQ; rec2 to recQ2; elim to elimQ) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.Nat 16 | open import Cubical.Data.Bool 17 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 18 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂) 19 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 20 | open import Cubical.Relation.Binary 21 | open import Cubical.Relation.Nullary 22 | open import Basics 23 | open import ListStuff 24 | open import ListRelations 25 | open import Trees 26 | open import Pfin.AsFreeJoinSemilattice 27 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_) 28 | 29 | -- the relation relating lists with the same elements 30 | SameEls : {A : Type} → List A → List A → Type 31 | SameEls = Relator _≡_ 32 | 33 | isPropSameEls : ∀{A} (xs ys : List A) → isProp (SameEls xs ys) 34 | isPropSameEls = isPropRelator _ 35 | 36 | isEquivRelSameEls : ∀{A} → BinaryRelation.isEquivRel (SameEls {A}) 37 | isEquivRelSameEls = 38 | BinaryRelation.equivRel (reflRelator (λ _ → refl)) 39 | (λ _ _ r → r .snd , r .fst) 40 | (λ _ _ _ → transRelator _∙_) 41 | 42 | -- finite powerset as a quotient of lists 43 | PfinQ : Type → Type 44 | PfinQ A = List A / SameEls 45 | 46 | -- adding an element to a finite subset 47 | _∷Q_ : ∀ {A} → A → PfinQ A → PfinQ A 48 | _∷Q_ a = recQ squash/ (λ xs → [ a ∷ xs ]) 49 | λ { xs ys (p , q) → eq/ _ _ 50 | ((λ { x here → ∣ x , here , refl ∣ ; x (there m) → ∥map∥ (λ { (y , m' , eq) → y , there m' , eq }) (p x m) }) , 51 | λ { x here → ∣ x , here , refl ∣ ; x (there m) → ∥map∥ (λ { (y , m' , eq) → y , there m' , eq }) (q x m) }) } 52 | 53 | -- membership relation 54 | _∈Q_ : ∀{A} → A → PfinQ A → hProp₀ 55 | _∈Q_ a = elimQ (λ _ → isSetHProp) 56 | (λ xs → ∥ a ∈ xs ∥ , propTruncIsProp) 57 | λ xs ys r → Σ≡Prop (λ _ → isPropIsProp) 58 | (hPropExt propTruncIsProp propTruncIsProp 59 | (∥rec∥ propTruncIsProp (λ m → ∥map∥ (λ { (x , m , eq) → subst (_∈ ys) (sym eq) m }) (r .fst _ m))) 60 | (∥rec∥ propTruncIsProp (λ m → ∥map∥ (λ { (x , m , eq) → subst (_∈ xs) (sym eq) m }) (r .snd _ m)))) 61 | 62 | 63 | -- turning a list into a finite subset 64 | List→Pfin : ∀{A : Type} → List A → Pfin A 65 | List→Pfin [] = ø 66 | List→Pfin (x ∷ xs) = η x ∪ List→Pfin xs 67 | 68 | -- this functions is a monoid morphism (turns ++ into ∪) 69 | List→Pfin++ : {A : Type}(xs ys : List A) 70 | → List→Pfin (xs ++ ys) ≡ List→Pfin xs ∪ List→Pfin ys 71 | List→Pfin++ [] ys = sym (com _ _ ∙ nr _) 72 | List→Pfin++ (x ∷ xs) ys = cong (η x ∪_) (List→Pfin++ xs ys) ∙ ass _ _ _ 73 | 74 | 75 | -- properties of membership in the finite subset associated to a list 76 | ∈ₛList→Pfin : ∀{A : Type} (xs : List A){a : A} 77 | → ⟨ a ∈ₛ List→Pfin xs ⟩ → ∥ a ∈ xs ∥ 78 | ∈ₛList→Pfin (x ∷ xs) = ∥rec∥ propTruncIsProp 79 | λ { (inj₁ p) → ∥map∥ (λ eq → subst (_∈ _) (sym eq) here) p 80 | ; (inj₂ p) → ∥map∥ there (∈ₛList→Pfin xs p)} 81 | 82 | List→Pfin∈ : ∀{A : Type} (xs : List A){a : A} 83 | → a ∈ xs → ⟨ a ∈ₛ List→Pfin xs ⟩ 84 | List→Pfin∈ (x ∷ xs) here = inl ∣ refl ∣ 85 | List→Pfin∈ (x ∷ xs) (there p) = inr (List→Pfin∈ xs p) 86 | 87 | 88 | -- the two presentation of finite powerset (as a quotient and as the 89 | -- free join semilattice) are equivalent 90 | List→PfinRel : ∀{A}{xs ys : List A} 91 | → DRelator _≡_ xs ys → PfinDRel _≡_ (List→Pfin xs) (List→Pfin ys) 92 | List→PfinRel p x mx = 93 | ∥rec∥ propTruncIsProp 94 | (λ mx' → 95 | ∥map∥ (λ { (y , my , eq) → (y , List→Pfin∈ _ my , eq) }) (p x mx')) 96 | (∈ₛList→Pfin _ mx) 97 | 98 | PfinQ→Pfin : ∀{A} → PfinQ A → Pfin A 99 | PfinQ→Pfin = recQ trunc List→Pfin 100 | λ xs ys p → PfinEq→Eq (List→PfinRel (p .fst) , List→PfinRel (p .snd)) 101 | 102 | Pfin→PfinQ : ∀{A} → Pfin A → PfinQ A 103 | Pfin→PfinQ = 104 | recPfin squash/ 105 | [ [] ] 106 | (λ a → [ a ∷ [] ]) 107 | (recQ2 squash/ (λ xs ys → [ xs ++ ys ]) 108 | (λ xs ys zs p → eq/ _ _ (DRelatorEq++₁ (p .fst) , DRelatorEq++₁ (p .snd))) 109 | (λ xs ys zs p → eq/ _ _ (DRelatorEq++₂ (p .fst) , DRelatorEq++₂ (p .snd)))) 110 | (elimProp2 (λ _ _ → squash/ _ _) 111 | λ xs ys → eq/ _ _ (DRelatorEqCom xs ys , DRelatorEqCom ys xs)) 112 | (elimProp3 (λ _ _ _ → squash/ _ _) 113 | λ xs ys zs → eq/ _ _ (DRelatorEqAss₁ xs ys zs , DRelatorEqAss₂ xs ys zs)) 114 | (elimProp (λ _ → squash/ _ _) 115 | λ xs → eq/ _ _ (DRelatorEqIdem xs , λ x mx → ∣ x , ∈++₁ mx , refl ∣)) 116 | (elimProp (λ _ → squash/ _ _) 117 | λ xs → eq/ _ _ (DRelatorEqNr xs , λ x mx → ∣ x , ∈++₁ mx , refl ∣)) 118 | 119 | PfinQ→Pfin→PfinQ' : ∀{A} (xs : List A) 120 | → Pfin→PfinQ (List→Pfin xs) ≡ [ xs ] 121 | PfinQ→Pfin→PfinQ' [] = refl 122 | PfinQ→Pfin→PfinQ' (x ∷ xs) = 123 | cong (recQ squash/ (λ ys → [ x ∷ ys ]) _) (PfinQ→Pfin→PfinQ' xs) 124 | 125 | PfinQ→Pfin→PfinQ : ∀{A} (s : PfinQ A) → Pfin→PfinQ (PfinQ→Pfin s) ≡ s 126 | PfinQ→Pfin→PfinQ = elimProp (λ _ → squash/ _ _) PfinQ→Pfin→PfinQ' 127 | 128 | Pfin→PfinQ→Pfin : ∀{A} (s : Pfin A) → PfinQ→Pfin (Pfin→PfinQ s) ≡ s 129 | Pfin→PfinQ→Pfin {A} = 130 | elimPfinProp (λ _ → _ , trunc _ _) refl (λ _ → nr _) 131 | λ {s₁}{s₂} p q → 132 | lem (Pfin→PfinQ s₁) (Pfin→PfinQ s₂) ∙ cong₂ _∪_ p q 133 | where 134 | lem : (s₁ s₂ : PfinQ A) 135 | → PfinQ→Pfin (recQ2 squash/ (λ xs ys → [ xs ++ ys ]) _ _ s₁ s₂) ≡ 136 | PfinQ→Pfin s₁ ∪ PfinQ→Pfin s₂ 137 | lem = elimProp2 (λ _ _ → trunc _ _) List→Pfin++ 138 | 139 | Pfin≡PfinQ : ∀{A} → Pfin A ≡ PfinQ A 140 | Pfin≡PfinQ = 141 | isoToPath (iso Pfin→PfinQ PfinQ→Pfin PfinQ→Pfin→PfinQ Pfin→PfinQ→Pfin) 142 | 143 | -- PfinQ action on functions 144 | DRelatorMapList : {A B : Type} (f : A → B) {xs ys : List A} 145 | → DRelator _≡_ xs ys → DRelator _≡_ (mapList f xs) (mapList f ys) 146 | DRelatorMapList f p x mx with pre∈mapList mx 147 | ... | y , my , eq = 148 | ∥map∥ (λ { (z , mz , eq') → _ , ∈mapList mz , sym eq ∙ cong f eq'}) (p y my) 149 | 150 | mapPfinQ : ∀{A B} (f : A → B) → PfinQ A → PfinQ B 151 | mapPfinQ f = recQ squash/ (λ xs → [ mapList f xs ]) 152 | λ xs ys p → eq/ _ _ (DRelatorMapList f (p .fst) , DRelatorMapList f (p .snd)) 153 | 154 | mapPfinQComp : ∀{A B C} {g : B → C} {f : A → B} (x : PfinQ A) 155 | → mapPfinQ g (mapPfinQ f x) ≡ mapPfinQ (g ∘ f) x 156 | mapPfinQComp = elimProp (λ _ → squash/ _ _) 157 | (λ xs → cong [_] (mapListComp xs)) 158 | 159 | pre∈mapPfinQ : {A B : Type} {f : A → B} {b : B} (s : PfinQ A) 160 | → ⟨ b ∈Q mapPfinQ f s ⟩ → ∃[ a ∈ A ] ⟨ a ∈Q s ⟩ × (f a ≡ b) 161 | pre∈mapPfinQ = elimProp (λ _ → isPropΠ (λ _ → propTruncIsProp)) 162 | λ xs → ∥map∥ (λ m → _ , ∣ pre∈mapList m .snd .fst ∣ , pre∈mapList m .snd .snd) 163 | 164 | -- the size of a finite subset, which we can define if the carrier 165 | -- type has decidable equality 166 | size : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 167 | → PfinQ A → ℕ 168 | size decA = recQ isSetℕ 169 | (λ xs → length (removeDups decA xs)) 170 | λ xs ys r → ≤-antisym (lengthDupFree (removeDups decA xs) (removeDups decA ys) 171 | (dupFreeRemoveDups decA xs) (dupFreeRemoveDups decA ys) 172 | λ x m → ∥map∥ (λ { (y , m , eq) → y , ∈removeDups decA ys m , eq}) (r .fst _ (removeDups∈ decA xs m))) 173 | (lengthDupFree (removeDups decA ys) (removeDups decA xs) 174 | (dupFreeRemoveDups decA ys) (dupFreeRemoveDups decA xs) 175 | λ x m → ∥map∥ (λ { (y , m , eq) → y , ∈removeDups decA xs m , eq}) (r .snd _ (removeDups∈ decA ys m))) 176 | 177 | -- the size of (x ∷Q s) 178 | size∷Q' : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 179 | → {x : A} (xs : List A) → (∥ x ∈ xs ∥ → ⊥) 180 | → size decA (x ∷Q [ xs ]) ≡ suc (size decA [ xs ]) 181 | size∷Q' decA {x} xs m with decMem decA x xs 182 | ... | yes (x' , m' , eq) = ⊥-rec (m ∣ subst (_∈ xs) (sym eq) m' ∣) 183 | ... | no ¬p = refl 184 | 185 | size∷Q : {A : Type} (decA : (x y : A) → Dec (x ≡ y)) 186 | → {x : A} (s : PfinQ A) → (⟨ x ∈Q s ⟩ → ⊥) 187 | → size decA (x ∷Q s) ≡ suc (size decA s) 188 | size∷Q decA = elimProp (λ _ → isPropΠ (λ _ → isSetℕ _ _)) (size∷Q' decA) 189 | 190 | -- the size of (mapPfinQ f s) 191 | sizeMapPfinQ : {A B : Type} (decA : (x y : A) → Dec (x ≡ y)) 192 | → (decB : (x y : B) → Dec (x ≡ y)) 193 | → {f : A → B} (injf : ∀ x y → f x ≡ f y → x ≡ y) (s : PfinQ A) 194 | → size decB (mapPfinQ f s) ≡ size decA s 195 | sizeMapPfinQ decA decB injf = elimProp (λ _ → isSetℕ _ _) 196 | (λ xs → cong length (removeDupMapList decA decB injf xs) 197 | ∙ lengthMapList (removeDups decA xs)) 198 | 199 | 200 | 201 | -- if path equality on A is decidable, then also path equality on PfinQ A 202 | -- is decidable 203 | PfinQDecEq' : {A : Type} 204 | → ((x y : A) → Dec (x ≡ y)) 205 | → (xs ys : List A) → Dec (_≡_ {A = PfinQ A} [ xs ] [ ys ]) 206 | PfinQDecEq' decA xs ys with decRelator decA xs ys 207 | ... | yes p = yes (eq/ _ _ p) 208 | ... | no ¬p = no (λ p → ¬p (effective isPropSameEls isEquivRelSameEls _ _ p)) 209 | 210 | PfinQDecEq : {A : Type} 211 | → ((x y : A) → Dec (x ≡ y)) 212 | → (x y : PfinQ A) → Dec (x ≡ y) 213 | PfinQDecEq decA = 214 | elimProp2 (λ _ _ → isPropDec (squash/ _ _)) 215 | (PfinQDecEq' decA) 216 | 217 | 218 | 219 | 220 | -------------------------------------------------------------------------------- /Pfin/PreservesCountableIntersections.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Pfin.PreservesCountableIntersections where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (rec to ∥rec∥; map to ∥map∥) 12 | open import Cubical.HITs.SetQuotients 13 | renaming ([_] to eqCl; rec to recQ; rec2 to recQ2) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.List 16 | open import Cubical.Data.Nat renaming (elim to elimNat) 17 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂) 18 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 19 | open import Cubical.Relation.Binary 20 | open import Basics 21 | open import ListStuff 22 | open import Cubical.Relation.Nullary 23 | 24 | open import Pfin.AsFreeJoinSemilattice 25 | 26 | -- -- Pfin preserves intersections 27 | 28 | -- (countable) wide pullpacks 29 | ×pℕ : {A : ℕ → Type} {C : Type} 30 | → (f : ∀ n → A n → C) → Type 31 | ×pℕ {A} f = Σ[ a ∈ ((n : ℕ) → A n) ] ∀ n → f (suc n) (a (suc n)) ≡ f 0 (a 0) 32 | 33 | -- the wide pullback is a set if the components of the wedge are sets 34 | isSet×pℕ : {A : ℕ → Type} {C : Type} 35 | → (∀ n → isSet (A n)) → isSet C 36 | → (f : ∀ n → A n → C) → isSet (×pℕ f) 37 | isSet×pℕ sA sC f = isSetΣ (isSetΠ sA) λ _ → isProp→isSet (isPropΠ (λ _ → sC _ _)) 38 | 39 | -- equality in the wide pullback 40 | Eq×pℕ : {A : ℕ → Type} {C : Type} → isSet C 41 | → (f : ∀ n → A n → C) 42 | → {a a' : ∀ n → A n} → (∀ n → a n ≡ a' n) 43 | → {eq : ∀ n → f (suc n) (a (suc n)) ≡ f 0 (a 0)} 44 | → {eq' : ∀ n → f (suc n) (a' (suc n)) ≡ f 0 (a' 0)} 45 | → _≡_ {A = ×pℕ f} (a , eq) (a' , eq') 46 | Eq×pℕ setC f p = 47 | Σ≡Prop (λ a → isPropΠ (λ _ → setC _ _)) 48 | λ i n → p n i 49 | 50 | -- the construction of a function from the type of finite subsets of a 51 | -- wide pullback to the wide pullback of finite subsets 52 | to×pℕ : {A : ℕ → Type}{C : Type} (f : ∀ n → A n → C) 53 | → Pfin (×pℕ {A}{C} f) → ×pℕ {Pfin ∘ A}{Pfin C} (mapPfin ∘ f) 54 | to×pℕ f s = 55 | (λ n → mapPfin (λ x → x .fst n) s) , 56 | λ n → 57 | mapPfinComp s 58 | ∙ cong (λ z → mapPfin z s) (λ i x → x .snd n i) 59 | ∙ sym (mapPfinComp s) 60 | 61 | -- some auxiliary constructions 62 | funs : {A : ℕ → Type} {C : Type} 63 | → (f0 : A 0 → C) 64 | → (f : ∀ n → A (suc n) → C) 65 | → ∀ n → A n → C 66 | funs f0 f zero = f0 67 | funs f0 f (suc n) = f n 68 | 69 | funsEq : {A : ℕ → Type} {C : Type} 70 | → (f : ∀ n → A n → C) 71 | → ∀ n → funs {A} (f 0) (f ∘ suc) n ≡ f n 72 | funsEq f zero = refl 73 | funsEq f (suc n) = refl 74 | 75 | args : {A : ℕ → Type} 76 | → (a0 : Pfin (A 0)) 77 | → (as : ∀ n → Pfin (A (suc n))) 78 | → ∀ n → Pfin (A n) 79 | args a0 as zero = a0 80 | args a0 as (suc n) = as n 81 | 82 | argsEq : {A : ℕ → Type} 83 | → (a : ∀ n → Pfin (A n)) 84 | → ∀ n → args {A} (a 0) (a ∘ suc) n ≡ a n 85 | argsEq a zero = refl 86 | argsEq a (suc n) = refl 87 | 88 | args∪ : {A : ℕ → Type} 89 | → {a0 b0 : Pfin (A 0)} 90 | → {as bs : ∀ n → Pfin (A (suc n))} 91 | → ∀ n → args {A} (a0 ∪ b0) (λ k → as k ∪ bs k) n 92 | ≡ (args {A} a0 as n ∪ args {A} b0 bs n) 93 | args∪ zero = refl 94 | args∪ (suc n) = refl 95 | 96 | -- to prove that a family (a : ∀ n → A n) in the wide pullback (×pℕ f) 97 | -- is in a certain subset s, in case all functions in (f ∘ suc) are 98 | -- injective, then it sufficient to show that a 0 is among the 99 | -- 0-projections in s 100 | ∈Pfin×pℕ : {A : ℕ → Type} {C : Type} → isSet C 101 | → (f : ∀ n → A n → C) 102 | → (∀ n (x y : A (suc n)) → f (suc n) x ≡ f (suc n) y → x ≡ y) 103 | → {a : ∀ n → A n} (eq : ∀ n → f (suc n) (a (suc n)) ≡ f 0 (a 0)) 104 | → (s : Pfin (×pℕ f)) 105 | → ⟨ a 0 ∈ₛ mapPfin (λ x → x .fst 0) s ⟩ 106 | → ⟨ (a , eq) ∈ₛ s ⟩ 107 | ∈Pfin×pℕ setC f injf {a} eq s ma0 = 108 | ∥rec∥ (snd ((a , eq) ∈ₛ s)) 109 | (λ { ((a' , eq') , m , r) → 110 | J {x = a'} (λ y _ → (eq : ∀ n → f (suc n) (y (suc n)) ≡ f 0 (y 0)) → ⟨ (y , eq) ∈ₛ s ⟩) 111 | (λ eq → subst (λ z → ⟨ (a' , z) ∈ₛ s ⟩) (funExt (λ n → setC _ _ _ _)) m) 112 | (λ { i zero → r i ; i (suc n) → injf n _ _ ((eq' n ∙ cong (f 0) r ∙ sym (eq n))) i }) 113 | eq}) 114 | (pre∈ₛmapPfin (λ x → x .fst 0) _ s ma0) 115 | 116 | -- We show that Pfin preserves a wide pullback ×pℕ in case the family 117 | -- f is made of injective functions, which is to say that Pfin 118 | -- preserves intersections. The construction of an inverse function of 119 | -- to×pℕ requires the assumption of countable choice. We actually prove 120 | -- in one go that to×pℕ is an equivalence. 121 | module _ (cc : (P : ℕ → Type) → (∀ n → ∥ P n ∥) → ∥ (∀ n → P n) ∥) where 122 | 123 | to×pℕEquiv : {A : ℕ → Type} {C : Type} 124 | → (setA : ∀ n → isSet (A (suc n))) (setC : isSet C) 125 | → (f0 : A 0 → C) 126 | → (f : ∀ n → A (suc n) → C) 127 | → (∀ n (x y : A (suc n)) → f n x ≡ f n y → x ≡ y) 128 | → (a0 : Pfin (A 0)) 129 | → (as : ∀ n → Pfin (A (suc n))) 130 | → (eq : ∀ n → mapPfin (f n) (as n) ≡ mapPfin f0 a0) 131 | → isContr (fiber (to×pℕ (funs {A} f0 f)) (args a0 as , eq)) 132 | to×pℕEquiv {A}{C} setA setC f0 f injf = 133 | elimPfinProp (λ _ → _ , isPropΠ (λ _ → isPropΠ (λ _ → isPropIsContr))) 134 | (λ as eq → 135 | (ø , 136 | Eq×pℕ trunc (λ n → mapPfin (funs {A} f0 f n)) (λ { zero → refl ; (suc n) → sym (mapPfinø (f n) (as n) (eq n)) })) , 137 | λ { (w , eqw) → 138 | EqFiber (isSetΣ (isSetΠ (λ _ → trunc)) λ _ → isSetΠ (λ _ → isProp→isSet (trunc _ _))) _ _ 139 | (antisym⊆ (λ { _ () }) 140 | λ x m → subst (λ z → ⟨ x .fst 0 ∈ₛ z ⟩) (funExt⁻ (cong fst eqw) 0) (∈ₛmapPfin (λ z → z .fst 0) x w m)) 141 | _ _ }) 142 | (λ a as eq → 143 | let a' : ∀ n → Σ[ x ∈ A (suc n) ] as n ≡ η x 144 | a' n = mapPfinη (setA n) (f n) (injf n) (as n) (f0 a) (eq n) 145 | in (η ((λ { zero → a ; (suc n) → a' n .fst }) , 146 | λ n → ηisInjective setC (cong (mapPfin (f n)) (sym (a' n .snd)) ∙ eq n)) , 147 | Eq×pℕ trunc (λ n → mapPfin (funs {A} f0 f n)) (λ { zero → refl ; (suc n) → sym (a' n .snd) })) , 148 | λ { (w , eqw) → 149 | EqFiber (isSetΣ (isSetΠ (λ _ → trunc)) λ _ → isSetΠ (λ _ → isProp→isSet (trunc _ _))) _ _ 150 | (antisym⊆ 151 | (λ { x@(a' , fa'≡gb') → ∥rec∥ (snd (x ∈ₛ w)) 152 | λ eqx → ∈Pfin×pℕ setC (funs {A} f0 f) injf fa'≡gb' w 153 | (subst (λ z → ⟨ a' 0 ∈ₛ z 0 ⟩) (cong fst (sym eqw)) ∣ funExt⁻ (cong fst eqx) 0 ∣) }) 154 | λ { x@(a' , fa'≡gb') mx → ∈Pfin×pℕ setC (funs {A} f0 f) injf fa'≡gb' (η (_ , _)) 155 | (subst (λ z → ⟨ a' 0 ∈ₛ z ⟩) (funExt⁻ (cong fst eqw) 0) (∈ₛmapPfin (λ z → z .fst 0) _ w mx)) }) 156 | _ _ }) 157 | λ {s1} {s2} ih1 ih2 as eq → 158 | let p : ∥ (∀ n → Σ[ a1 ∈ Pfin (A (suc n)) ] Σ[ a2 ∈ Pfin (A (suc n)) ] 159 | (a1 ∪ a2 ≡ as n) × (mapPfin f0 s1 ≡ mapPfin (f n) a1) × (mapPfin f0 s2 ≡ mapPfin (f n) a2)) ∥ 160 | p = cc _ (λ n → ∪≡mapPfin (f n) (injf n) (as n) (mapPfin f0 s1) (mapPfin f0 s2) (sym (eq n))) 161 | in ∥rec∥ isPropIsContr 162 | (λ { p → 163 | let u1 : ∀ n → Pfin (A (suc n)) 164 | u1 n = p n .fst 165 | u2 : ∀ n → Pfin (A (suc n)) 166 | u2 n = p n .snd .fst 167 | eqt : ∀ n → u1 n ∪ u2 n ≡ as n 168 | eqt n = p n .snd .snd .fst 169 | eq1 : ∀ n → mapPfin (f n) (u1 n) ≡ mapPfin f0 s1 170 | eq1 n = sym (p n .snd .snd .snd .fst) 171 | eq2 : ∀ n → mapPfin (f n) (u2 n) ≡ mapPfin f0 s2 172 | eq2 n = sym (p n .snd .snd .snd .snd) 173 | ((v1 , q1) , r1) = ih1 u1 eq1 174 | ((v2 , q2) , r2) = ih2 u2 eq2 175 | in ((v1 ∪ v2) , 176 | Eq×pℕ trunc (λ n → mapPfin (funs {A} f0 f n)) 177 | (λ n → cong₂ _∪_ (funExt⁻ (cong fst q1) n) (funExt⁻ (cong fst q2) n) 178 | ∙ sym (args∪ n) 179 | ∙ cong (λ k → args {A} (s1 ∪ s2) k n) 180 | (funExt (λ n → mapPfinInj (f n) (injf n) _ _ (cong₂ _∪_ (eq1 n) (eq2 n) ∙ sym (eq n)))))) , 181 | λ { (w , eqw) → 182 | EqFiber (isSetΣ (isSetΠ (λ _ → trunc)) λ _ → isSetΠ (λ _ → isProp→isSet (trunc _ _))) _ _ 183 | (cong₂ _∪_ (cong fst (r1 (v1 , q1))) (cong fst (r2 (v2 , q2))) 184 | ∙ antisym⊆ 185 | (λ x@(a , fa≡gb) → ∥rec∥ (snd (x ∈ₛ w)) 186 | (λ { (inj₁ mx) → ∈Pfin×pℕ setC (funs {A} f0 f) injf fa≡gb w 187 | (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) 188 | (funExt⁻ (cong fst (sym eqw)) 0) 189 | (inl (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) 190 | (funExt⁻ (cong fst q1) 0) 191 | (∈ₛmapPfin (λ y → y .fst 0) x v1 mx)))) 192 | ; (inj₂ mx) → ∈Pfin×pℕ setC (funs {A} f0 f) injf fa≡gb w 193 | (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) 194 | (funExt⁻ (cong fst (sym eqw)) 0) 195 | (inr (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) 196 | (funExt⁻ (cong fst q2) 0) 197 | (∈ₛmapPfin (λ y → y .fst 0) x v2 mx)))) })) 198 | λ { x@(a , fa≡gb) mx → ∥rec∥ propTruncIsProp 199 | (λ { (inj₁ ma) → inl (∈Pfin×pℕ setC (funs {A} f0 f) injf fa≡gb v1 200 | (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) (funExt⁻ (cong fst (sym q1)) 0) ma)) 201 | ; (inj₂ ma) → inr (∈Pfin×pℕ setC (funs {A} f0 f) injf fa≡gb v2 202 | (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) (funExt⁻ (cong fst (sym q2)) 0) ma)) }) 203 | (subst (λ z → ⟨ a 0 ∈ₛ z ⟩) (funExt⁻ (cong fst eqw) 0) (∈ₛmapPfin (λ y → y .fst 0) x w mx)) }) 204 | _ _ } }) 205 | p 206 | 207 | 208 | Pfin×pℕ : {A : ℕ → Type} {C : Type} 209 | → (setA : ∀ n → isSet (A (suc n))) (setC : isSet C) 210 | → (f : ∀ n → A n → C) 211 | → (injf : ∀ n (x y : A (suc n)) → f (suc n) x ≡ f (suc n) y → x ≡ y) 212 | → Pfin (×pℕ f) ≃ ×pℕ (mapPfin ∘ f) 213 | Pfin×pℕ {A} setA setC f injf = (to×pℕ f) , 214 | (record { equiv-proof = 215 | subst (λ f → ∀ x → isContr (fiber (to×pℕ f) x)) 216 | (funExt (funsEq {A} f)) 217 | (λ x@(a , eq) → subst (isContr ∘ fiber (to×pℕ (funs {A} (f 0) (f ∘ suc)))) 218 | (λ i → (λ n → argsEq {A} a n i) , eq) 219 | (to×pℕEquiv setA setC (f 0) (f ∘ suc) injf (a 0) (a ∘ suc) eq)) }) 220 | 221 | 222 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Type-Theoretic Constructions of the Final Coalgebra of the Finite Powerset Functor 2 | ## Niccolò Veltri 3 | 4 | The finite powerset functor is a construct frequently employed for the specification of nondeterministic 5 | transition systems as coalgebras. The final coalgebra of the finite powerset functor, whose elements 6 | characterize the dynamical behavior of transition systems, is a well-understood object which enjoys 7 | many equivalent presentations in set theoretic foundations based on classical logic. 8 | 9 | In this paper, we discuss various constructions of the final coalgebra of the finite powerset functor 10 | in constructive type theory, and we formalize our results in the Cubical Agda proof assistant. Using 11 | setoids, the final coalgebra of the finite powerset functor can be defined from the final coalgebra of 12 | the list functor. Using types instead of setoids, as it is common in homotopy type theory, one can 13 | specify the finite powerset datatype as a higher inductive type and define its final coalgebra as a 14 | coinductive type. Another construction is obtained by quotienting the final coalgebra of the list 15 | functor, but the proof of finality requires the assumption of the axiom of choice. We conclude the 16 | paper with an analysis of a classical construction by James Worrell, and show that its adaptation to 17 | our constructive setting requires the presence of classical axioms such as countable choice and the 18 | lesser limited principle of omniscience. 19 | 20 | The file [Everything.agda](https://github.com/niccoloveltri/final-pfin/blob/main/Everything.agda) contains the whole Agda development. 21 | -------------------------------------------------------------------------------- /SetoidStuff.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module SetoidStuff where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.List renaming (map to mapList; [_] to sing) 8 | open import Cubical.Data.Sigma 9 | open import Cubical.Data.Unit 10 | open import Cubical.Data.Empty renaming (elim to ⊥-elim) 11 | open import Cubical.Data.Nat 12 | open import Cubical.Data.Sum renaming (inl to inj₁; inr to inj₂; map to map⊎) 13 | open import Cubical.Functions.Logic hiding (⊥) 14 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 15 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 16 | renaming (map to ∥map∥; rec to ∥rec∥) 17 | open import Cubical.Relation.Binary hiding (Rel) 18 | open BinaryRelation 19 | 20 | open import Basics 21 | 22 | -- type of setoids 23 | record Setoid ℓ : Type (ℓ-suc ℓ) where 24 | constructor setoid 25 | field 26 | Carr : Type ℓ 27 | Rel : Carr → Carr → Type ℓ 28 | propRel : isPropValued Rel 29 | eqrRel : isEquivRel Rel 30 | open Setoid public 31 | 32 | Setoid₀ = Setoid ℓ-zero 33 | 34 | -- the unit setoid 35 | UnitS : Setoid₀ 36 | UnitS = setoid Unit (λ _ _ → Unit) 37 | (λ _ _ → isPropUnit) 38 | (equivRel (λ _ → tt) (λ _ _ _ → tt) (λ _ _ _ _ _ → tt)) 39 | 40 | -- type of setoid morphisms 41 | record _→S_ {ℓ} (S₁ S₂ : Setoid ℓ) : Type ℓ where 42 | constructor _,_ 43 | field 44 | mor : S₁ .Carr → S₂ .Carr 45 | morRel : ∀{x y} → S₁ .Rel x y → S₂ .Rel (mor x) (mor y) 46 | open _→S_ public 47 | 48 | -- the unique setoid morphism into UnitS 49 | bangS : {S : Setoid₀} → S →S UnitS 50 | bangS = (λ _ → tt) , (λ _ → tt) 51 | 52 | -- equality of setoid morphisms 53 | _≡S_ : ∀{ℓ} {S₁ S₂ : Setoid ℓ} (f g : S₁ →S S₂) → Type ℓ 54 | _≡S_ {S₂ = S₂} f g = ∀ x → S₂ .Rel (f .mor x) (g .mor x) 55 | 56 | -- the identity setoid morphism 57 | idS : ∀{ℓ} {S : Setoid ℓ} → S →S S 58 | idS = (λ x → x) , λ r → r 59 | 60 | -- composition of setoid morphisms 61 | infix 21 _∘S_ 62 | _∘S_ : ∀{ℓ} {S₁ S₂ S₃ : Setoid ℓ} 63 | → (g : S₂ →S S₃) (f : S₁ →S S₂) 64 | → S₁ →S S₃ 65 | (g , gr) ∘S (f , fr) = g ∘ f , gr ∘ fr 66 | 67 | -- from sets to setoids 68 | Set→Setoid : ∀{ℓ} → hSet ℓ → Setoid ℓ 69 | Set→Setoid (X , Xset) = 70 | setoid X _≡_ Xset (equivRel (λ _ → refl) (λ _ _ → sym) λ _ _ _ → _∙_) 71 | 72 | -- surjective setoid morphisms 73 | isSurjectionS : ∀{ℓ}{S T : Setoid ℓ} → S →S T → Type ℓ 74 | isSurjectionS {T = T} (f , _) = ∀ y → ∃[ x ∈ _ ] T .Rel (f x) y 75 | 76 | 77 | -- pointwise lifting of a relation to a function space 78 | PW : {X A B : Type} (R : A → B → Type) → (X → A) → (X → B) → Type 79 | PW R f g = ∀ x → R (f x) (g x) 80 | 81 | -- the quotient a function space by the pointwise lifted relation 82 | [_⇒_]/_ : (A B : Type) (R : B → B → Type) → Type 83 | [ A ⇒ B ]/ R = (A → B) / PW R 84 | 85 | -- a function sending equivalence classes of function wrt. pointwise 86 | -- lifted relation and functions into equivalence classes 87 | θ : ∀ A {B} (R : B → B → Type) → [ A ⇒ B ]/ R → A → B / R 88 | θ A R = recQ (isSetΠ (λ _ → squash/)) (λ c x → [ c x ]) 89 | λ c c' r → funExt (λ x → eq/ _ _ (r x)) 90 | 91 | 92 | -- equivalence between two formulation of full axiom of choice; the 93 | -- formulation stating the surjectivity of the map λ g → [_] ∘ g is 94 | -- proved equivalent to the usual presentation of axiom of choice 95 | 96 | -- NB: in the first formulation we do not need to 0-truncate the 97 | -- existence of a section, since the type of sections of θ is a 98 | -- proposition; this follows directly from the lemma SectionIsProp 99 | -- below 100 | 101 | module _ (θInv : ∀ A {B} (R : B → B → Type) → (A → B / R) → [ A ⇒ B ]/ R) 102 | (sectionθ : ∀ A {B} (R : B → B → Type) → section (θ A R) (θInv A R)) where 103 | 104 | ac' : ∀ (A : Type) {B : Type} (R : B → B → Type) 105 | → (f : (A → B) / PW R) → ∃[ g ∈ (A → B) ] [_] ∘ g ≡ θ A R f 106 | ac' A R = elimProp (λ _ → propTruncIsProp) (λ g → ∣ g , refl ∣) 107 | 108 | ac : ∀ (A : Type) {B : Type} (R : B → B → Type) 109 | → (f : A → B / R) → ∃[ g ∈ (A → B) ] [_] ∘ g ≡ f 110 | ac A R f = 111 | ∥map∥ (λ { (g , eq) → g , eq ∙ sectionθ A R f }) (ac' A R (θInv A R f)) 112 | 113 | module _ (ac : ∀ (A : Type) {B : Type} (R : B → B → Type) 114 | → isPropValued R → isEquivRel R 115 | → (f : A → B / R) → ∃[ g ∈ (A → B) ] [_] ∘ g ≡ f) where 116 | 117 | SectionIsProp' : ∀ A {B} (R : B → B → Type) 118 | → isPropValued R → isEquivRel R 119 | → (f : A → B / R) 120 | → (g g' : [ A ⇒ B ]/ R) (eq : θ A R g ≡ f) (eq' : θ A R g' ≡ f) 121 | → g ≡ g' 122 | SectionIsProp' A R Rprop Reqr f = elimProp2 (λ _ _ → isPropΠ (λ _ → isPropΠ λ _ → squash/ _ _)) 123 | λ g g' eq eq' → eq/ _ _ (λ x → effective Rprop Reqr _ _ (λ i → (eq ∙ sym eq') i x)) 124 | 125 | SectionIsProp : ∀ A {B} (R : B → B → Type) 126 | → isPropValued R → isEquivRel R 127 | → (f : A → B / R) 128 | → isProp (Σ ([ A ⇒ B ]/ R) (λ g → θ A R g ≡ f)) 129 | SectionIsProp A R Rprop Reqr f (g , eq) (g' , eq') = 130 | Σ≡Prop (λ _ → isPropFunEq _ _ λ _ → squash/ _ _) 131 | (SectionIsProp' A R Rprop Reqr f g g' eq eq') 132 | 133 | θInvSec : ∀ A {B} (R : B → B → Type) 134 | → isPropValued R → isEquivRel R 135 | → (f : A → B / R) → Σ ([ A ⇒ B ]/ R) (λ g → θ A R g ≡ f) 136 | θInvSec A R Rprop Reqr f = 137 | ∥rec∥ (SectionIsProp A R Rprop Reqr f) 138 | (λ {(g , eq) → [ g ] , eq}) 139 | (ac A R Rprop Reqr f) 140 | 141 | -------------------------------------------------------------------------------- /Trees.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Trees where 4 | 5 | open import Size 6 | open import Cubical.Core.Everything 7 | open import Cubical.Foundations.Prelude 8 | open import Cubical.Foundations.Everything 9 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 10 | open import Cubical.Relation.Everything 11 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 12 | renaming (rec to ∥rec∥; map to ∥map∥) 13 | open import Cubical.HITs.SetQuotients renaming ([_] to eqCl) 14 | open import Cubical.Data.Sigma 15 | open import Cubical.Data.List renaming (map to mapList) 16 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 17 | open import Cubical.Relation.Binary 18 | open import Cubical.Relation.Nullary 19 | open BinaryRelation 20 | 21 | open import Basics 22 | open import ListRelations 23 | 24 | -- non-wellfounded trees with finite ordered branching 25 | record Tree (i : Size) : Type where 26 | coinductive 27 | field 28 | subtrees : {j : Size< i} → List (Tree j) 29 | open Tree public 30 | 31 | -- tree bisimilarity 32 | record Bisim (i : Size) (xs ys : Tree i) : Type where 33 | constructor thunkEq 34 | coinductive 35 | field 36 | subtreesB : {j : Size< i} 37 | → ListRel (Bisim j) (subtrees xs) (subtrees ys) 38 | open Bisim public 39 | 40 | -- coinduction principle: tree bisimilarity implies path equality 41 | bisim : (i : Size) {t₁ t₂ : Tree i} → Bisim i t₁ t₂ → t₁ ≡ t₂ 42 | bisimL : (i : Size) {t₁ t₂ : List (Tree i)} 43 | → ListRel (Bisim i) t₁ t₂ → t₁ ≡ t₂ 44 | 45 | subtrees (bisim s b i) {s'} = bisimL s' (subtreesB b) i 46 | 47 | bisimL s [] i = [] 48 | bisimL s (b ∷ bs) i = bisim s b i ∷ bisimL s bs i 49 | 50 | -- existence of anamorphism 51 | anaTree : {X : Type} (c : X → List X) → (i : Size) → X → Tree i 52 | subtrees (anaTree c s x) {s'} = mapList (anaTree c s') (c x) 53 | 54 | -- the anamorphism is a coalgebra morphism (the corresponding square 55 | -- commutes on the nose) 56 | anaTreeEq : {X : Type} (c : X → List X) (x : X) 57 | → subtrees (anaTree c ∞ x) ≡ mapList (anaTree c ∞) (c x) 58 | anaTreeEq c x = refl 59 | 60 | -- uniqueness of universal property of final List-coalgebra 61 | anaTreeUniqB : {X : Type} (c : X → List X) 62 | → (f : (s : Size) → X → Tree s) 63 | → (eq : ∀ (s : Size) (s' : Size< s) x 64 | → subtrees (f s x) {s'} ≡ mapList (f s') (c x)) 65 | → ∀ (s : Size) x → Bisim s (f s x) (anaTree c s x) 66 | anaTreeUniqB' : {X : Type} (c : X → List X) 67 | → (f : (s : Size) → X → Tree s) 68 | → (eq : ∀ (s : Size) (s' : Size< s) x 69 | → subtrees (f s x) {s'} ≡ mapList (f s') (c x)) 70 | → (s : Size) 71 | → ∀ xs → ListRel (Bisim s) (mapList (f s) xs) (mapList (anaTree c s) xs) 72 | 73 | subtreesB (anaTreeUniqB c f eq s x) {s'} = 74 | subst (λ z → ListRel (Bisim s') z (mapList (anaTree c s') (c x))) 75 | (sym (eq s s' x)) 76 | (anaTreeUniqB' c f eq s' (c x)) 77 | 78 | anaTreeUniqB' c f eq s [] = [] 79 | anaTreeUniqB' c f eq s (x ∷ xs) = 80 | anaTreeUniqB c f eq s x ∷ anaTreeUniqB' c f eq s xs 81 | 82 | anaTreeUniq : {X : Type} (c : X → List X) 83 | → (f : X → Tree ∞) 84 | → (eq : ∀ (s : Size) x → subtrees (f x) {s} ≡ mapList f (c x)) 85 | → f ≡ anaTree c ∞ 86 | anaTreeUniq c f eq = 87 | funExt (λ x → 88 | bisim ∞ (anaTreeUniqB c (λ _ → f) (λ {_ _ → eq _}) ∞ x)) 89 | 90 | {- ================================================================ -} 91 | 92 | -- coinductive closure of the relator, which gives a notion of 93 | -- extensional equality of trees 94 | record ExtEq (i : Size) (t₁ t₂ : Tree ∞) : Type where 95 | coinductive 96 | field 97 | subtreesE : {j : Size< i} → Relator (ExtEq j) (subtrees t₁) (subtrees t₂) 98 | open ExtEq public 99 | 100 | isPropExtEq : ∀ t₁ t₂ → isProp (ExtEq ∞ t₁ t₂) 101 | subtreesE (isPropExtEq t₁ t₂ p q i) = 102 | isPropRelator (ExtEq _) (subtrees t₁) (subtrees t₂) 103 | (subtreesE p) (subtreesE q) i 104 | 105 | ExtEqₚ : Tree ∞ → Tree ∞ → hProp₀ 106 | ExtEqₚ t₁ t₂ = ExtEq ∞ t₁ t₂ , isPropExtEq t₁ t₂ 107 | 108 | -- reflexivity, symmetry and transitivity of ExtEqS 109 | reflExtEq : ∀ j t → ExtEq j t t 110 | subtreesE (reflExtEq j t) {k} = 111 | reflRelator (reflExtEq k) (subtrees t) 112 | 113 | symExtEq : ∀{t t'} (j : Size) → ExtEq j t t' → ExtEq j t' t 114 | subtreesE (symExtEq j p) = subtreesE p .snd , subtreesE p .fst 115 | 116 | transExtEq : ∀{t t₁ t₂}(j : Size) 117 | → ExtEq j t t₁ → ExtEq j t₁ t₂ → ExtEq j t t₂ 118 | subtreesE (transExtEq j p q) {k} = 119 | transRelator (transExtEq k) (subtreesE p) (subtreesE q) 120 | 121 | isEquivRelExtEq : isEquivRel (ExtEq ∞) 122 | isEquivRelExtEq = 123 | equivRel (reflExtEq ∞) 124 | (λ _ _ → symExtEq ∞) 125 | (λ _ _ _ → transExtEq ∞) 126 | 127 | {- 128 | -- reflexivity of bisimilarity 129 | -- refl-Bisim : ∀ t → Bisim t t 130 | -- refl-Bisim' : ∀ t → ListRel Bisim t t 131 | -- subtreesB (refl-Bisim ts) = refl-Bisim' (subtrees ts) 132 | -- refl-Bisim' [] = [] 133 | -- refl-Bisim' (t ∷ ts) = (refl-Bisim t) ∷ refl-Bisim' ts 134 | 135 | refl-BisimS : ∀ t → BisimS ∞ t t 136 | refl-BisimS' : ∀ t → ListRel (BisimS ∞) t t 137 | forceEq (refl-BisimS ts) = refl-BisimS' (force ts) 138 | refl-BisimS' [] = [] 139 | refl-BisimS' (t ∷ ts) = (refl-BisimS t) ∷ refl-BisimS' ts 140 | 141 | -- transitivity of bisimilarity 142 | transBisimS : ∀{t t₁ t₂} → BisimS ∞ t t₁ → BisimS ∞ t₁ t₂ → BisimS ∞ t t₂ 143 | transBisimS' : ∀{t t₁ t₂} → ListRel (BisimS ∞) t t₁ → ListRel (BisimS ∞) t₁ t₂ 144 | → ListRel (BisimS ∞) t t₂ 145 | forceEq (transBisimS b b') = transBisimS' (forceEq b) (forceEq b') 146 | transBisimS' [] [] = [] 147 | transBisimS' (b ∷ p) (b' ∷ p') = transBisimS b b' ∷ transBisimS' p p' 148 | 149 | -- equality implies bisimilarity 150 | misibS : {t₁ t₂ : TreeS ∞} → t₁ ≡ t₂ → BisimS ∞ t₁ t₂ 151 | misibS {t₁} = J (λ x p → BisimS ∞ t₁ x) (refl-BisimS t₁) 152 | 153 | misib : {t₁ t₂ : Tree} → t₁ ≡ t₂ → Bisim t₁ t₂ 154 | misib {t₁} = J (λ x p → Bisim t₁ x) (refl-Bisim t₁) 155 | 156 | -} 157 | -------------------------------------------------------------------------------- /Worrell/Finality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Worrell.Finality where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (map to ∥map∥; rec to ∥rec∥) 12 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 13 | open import Cubical.Data.Sigma 14 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 15 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 16 | open import Cubical.Data.Nat renaming (elim to elimNat) 17 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_; _≟_ to _≟N_) 18 | open import Cubical.Data.Bool 19 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂; rec to rec⊎; elim to elim⊎) 20 | open import Cubical.Relation.Binary 21 | open import Cubical.Relation.Nullary 22 | open import Basics 23 | open import ListRelations 24 | open import ListStuff 25 | open import Pfin.AsSetQuot 26 | open import Pfin.AsFreeJoinSemilattice 27 | open import Pfin.PreservesCountableIntersections 28 | open BinaryRelation 29 | 30 | open import Worrell.Limits 31 | 32 | -- Worrell's proof: Vω2 is the final Pfin-coalgebra. This works 33 | -- classically, and constructively requires countable choice and the 34 | -- injectivity of (m : Pfin Vω → Vω) (or analogously LLPO) 35 | module Finality (minj : ∀ s t → m s ≡ m t → s ≡ t) 36 | (cc : (P : ℕ → Type) → (∀ n → ∥ P n ∥) → ∥ (∀ n → P n) ∥) where 37 | 38 | -- since m is injective, so are all the restriction maps iMapPfin2 n 39 | iMapPfin2inj : ∀ n (x y : iPfin2 (suc n)) 40 | → iMapPfin2 n x ≡ iMapPfin2 n y → x ≡ y 41 | iMapPfin2inj zero x y eq = minj x y eq 42 | iMapPfin2inj (suc n) x y eq = 43 | mapPfinInj (iMapPfin2 n) (iMapPfin2inj n) x y eq 44 | 45 | -- injective maps projecting an element in (iPin2 n) to Vω 46 | u : ∀ n → iPfin2 n → Vω 47 | u zero x = x 48 | u (suc n) x = u n (iMapPfin2 n x) 49 | 50 | uinj : ∀ n (x y : iPfin2 n) 51 | → u n x ≡ u n y → x ≡ y 52 | uinj zero x y eq = eq 53 | uinj (suc n) x y eq = 54 | iMapPfin2inj n _ _ (uinj n _ _ eq) 55 | 56 | uLem : ∀ (x : Vω2) n 57 | → u n (iMapPfin2 n (x .fst (suc n))) ≡ m (x .fst 1) 58 | uLem x zero = refl 59 | uLem x (suc n) = cong (λ z → u n (iMapPfin2 n z)) (x .snd (suc n)) ∙ uLem x n 60 | 61 | uLem2 : ∀ (x : ×pℕ u) n 62 | → u (suc n) (x .fst (suc n)) ≡ u n (x .fst n) 63 | uLem2 x zero = x .snd 0 64 | uLem2 (x , p) (suc n) = uLem2 ((λ n → iMapPfin2 n (x (suc n))) , λ n → p (suc n) ∙ sym (p 0)) n 65 | 66 | -- Vω2 is equivalent to the intersection of the family of maps u 67 | Vω2-iso-×pℕ : Iso Vω2 (×pℕ u) 68 | Vω2-iso-×pℕ = Σ-cong-iso-snd 69 | λ x → iso (λ p n → uLem (x , p) n ∙ p 0) 70 | (λ q n → uinj n _ _ (uLem2 (x , q) n)) 71 | (λ _ → isPropΠ (λ _ → isSetVω _ _) _ _) 72 | λ _ → isPropΠ (λ _ → isSetiPfin2 _ _ _) _ _ 73 | 74 | Vω2≃×pℕ : Vω2 ≃ ×pℕ u 75 | Vω2≃×pℕ = isoToEquiv Vω2-iso-×pℕ 76 | 77 | -- The limit of the shifted chain Vω2Sh is equivalent to the 78 | -- intersection of the family of maps (mapPfin ∘ u) 79 | ×pℕSh-iso-Vω2Sh : Iso (×pℕ (mapPfin ∘ u)) Vω2Sh 80 | ×pℕSh-iso-Vω2Sh = Σ-cong-iso-snd 81 | λ x → iso (λ p n → mapPfinInj (u n) (uinj n) _ _ (mapPfinComp (x (suc n)) ∙ lem x p n)) 82 | (λ p n → lem2 x (λ n → sym (mapPfinComp (x (suc n))) ∙ cong (mapPfin (u n)) (p n)) n) 83 | (λ _ → isPropΠ (λ _ → trunc _ _) _ _) 84 | λ _ → isPropΠ (λ _ → trunc _ _) _ _ 85 | where 86 | lem : (x : ∀ n → Pfin (iPfin2 n)) 87 | → (∀ n → mapPfin (u (suc n)) (x (suc n)) ≡ mapPfin (u 0) (x 0)) 88 | → ∀ n → mapPfin (u (suc n)) (x (suc n)) ≡ mapPfin (u n) (x n) 89 | lem x p zero = p 0 90 | lem x p (suc n) = p (suc n) ∙ sym (p n) 91 | 92 | lem2 : (x : ∀ n → Pfin (iPfin2 n)) 93 | → (∀ n → mapPfin (u (suc n)) (x (suc n)) ≡ mapPfin (u n) (x n)) 94 | → ∀ n → mapPfin (u (suc n)) (x (suc n)) ≡ mapPfin (u 0) (x 0) 95 | lem2 x p zero = p 0 96 | lem2 x p (suc n) = p (suc n) ∙ lem2 x p n 97 | 98 | 99 | ×pℕSh≃Vω2Sh : ×pℕ (mapPfin ∘ u) ≃ Vω2Sh 100 | ×pℕSh≃Vω2Sh = isoToEquiv ×pℕSh-iso-Vω2Sh 101 | 102 | -- The shifted chain has equivalent limit 103 | Vω22Sh≃Vω2 : Vω2Sh ≃ Vω2 104 | Vω22Sh≃Vω2 = invEquiv (shift≃ iPfin2-ch isSetiPfin2) 105 | 106 | -- Composing the above equivalence shows that Vω2 is a fixpoint of 107 | -- Pfin 108 | τ-equiv : Pfin Vω2 ≃ Vω2 109 | τ-equiv = 110 | compEquiv (Pfin≃ Vω2≃×pℕ) 111 | (compEquiv (Pfin×pℕ cc (λ _ → trunc) isSetVω u (λ n → uinj (suc n))) 112 | (compEquiv ×pℕSh≃Vω2Sh Vω22Sh≃Vω2)) 113 | 114 | 115 | -- In particular, Vω is a Pfin-coalgebra 116 | τ : Vω2 → Pfin Vω2 117 | τ = invEq τ-equiv 118 | 119 | τ-1≡ : τ-1 ≡ equivFun τ-equiv 120 | τ-1≡ = funExt (λ x → Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin2 _ _ _)) 121 | (funExt (λ n → cong (iMapPfin2 n) (sym (mapPfinComp x))))) 122 | 123 | -- Proof that Vω is the final Pfin-coalgebra 124 | module _ (A : Type) (α : A → Pfin A) 125 | (αsim : A → Vω2) 126 | (αsim-mor : ∀ a → τ (αsim a) ≡ mapPfin αsim (α a)) where 127 | 128 | pα : ∀ n (a : A) → iPfin n 129 | pα zero a = tt 130 | pα (suc n) a = mapPfin (pα n) (α a) 131 | 132 | pα-res : ∀ n (a : A) → iMapPfin n (mapPfin (pα n) (α a)) ≡ pα n a 133 | pα-res zero a = refl 134 | pα-res (suc n) a = mapPfinComp (α a) ∙ cong (λ f → mapPfin f (α a)) (funExt (pα-res n)) 135 | 136 | pα2 : ∀ n (a : A) → iPfin2 n 137 | pα2 zero a = (λ n → pα n a) , λ n → pα-res n a 138 | pα2 (suc n) a = mapPfin (pα2 n) (α a) 139 | 140 | pα2-res : ∀ n (a : A) → iMapPfin2 n (mapPfin (pα2 n) (α a)) ≡ pα2 n a 141 | pα2-res zero a = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 142 | (funExt (λ n → cong (iMapPfin n) (mapPfinComp (α a)) ∙ pα-res n a)) 143 | pα2-res (suc n) a = mapPfinComp (α a) ∙ cong (λ f → mapPfin f (α a)) (funExt (pα2-res n)) 144 | 145 | coneA : Cone iPfin2-ch A 146 | coneA = pα2 , pα2-res 147 | 148 | αbar : A → Vω2 149 | αbar = Iso.inv (AdjLim iPfin2-ch _) coneA 150 | 151 | αbar-mor' : ∀ a → αbar a ≡ τ-1 (mapPfin αbar (α a)) 152 | αbar-mor' a = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin2 _ _ _)) 153 | (funExt (λ n → sym (cong (iMapPfin2 n) (mapPfinComp (α a)) ∙ pα2-res n a))) 154 | 155 | αbar-mor : ∀ a → τ (αbar a) ≡ mapPfin αbar (α a) 156 | αbar-mor a = 157 | τ (αbar a) 158 | ≡⟨ (λ i → τ (αbar-mor' a i)) ⟩ 159 | τ (τ-1 (mapPfin αbar (α a))) 160 | ≡⟨ (λ i → τ (τ-1≡ i (mapPfin αbar (α a)))) ⟩ 161 | τ (equivFun τ-equiv (mapPfin αbar (α a))) 162 | ≡⟨ (λ i → Iso.leftInv (equivToIso τ-equiv) (mapPfin αbar (α a)) i) ⟩ 163 | mapPfin αbar (α a) 164 | ∎ 165 | 166 | αsim-mor' : ∀ a → αsim a ≡ τ-1 (mapPfin αsim (α a)) 167 | αsim-mor' a = 168 | αsim a 169 | ≡⟨ sym (Iso.rightInv (equivToIso τ-equiv) (αsim a)) ⟩ 170 | equivFun τ-equiv (τ (αsim a)) 171 | ≡⟨ (λ i → equivFun τ-equiv (αsim-mor a i) ) ⟩ 172 | equivFun τ-equiv (mapPfin αsim (α a)) 173 | ≡⟨ (λ i → τ-1≡ (~ i) (mapPfin αsim (α a))) ⟩ 174 | τ-1 (mapPfin αsim (α a)) 175 | ∎ 176 | 177 | αbar-eq : ∀ a n → αsim a .fst 0 .fst n ≡ pα n a 178 | αbar-eq a zero = refl 179 | αbar-eq a (suc n) = 180 | funExt⁻ (cong fst (funExt⁻ (cong fst (αsim-mor' a)) 0)) (suc n) 181 | ∙ mapPfinComp ((mapPfin (proj iPfin2-ch 0) (mapPfin αsim (α a)))) 182 | ∙ mapPfinComp (mapPfin αsim (α a)) 183 | ∙ mapPfinComp (α a) 184 | ∙ cong (λ f → mapPfin f (α a)) (funExt (λ x → αsim x .fst 0 .snd n ∙ αbar-eq x n)) 185 | 186 | αbar-eq2 : ∀ a n → αsim a .fst n ≡ pα2 n a 187 | αbar-eq2 a zero = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 188 | (funExt (αbar-eq a)) 189 | αbar-eq2 a (suc n) = 190 | funExt⁻ (cong fst (αsim-mor' a)) (suc n) 191 | ∙ mapPfinComp (mapPfin αsim (α a)) 192 | ∙ mapPfinComp (α a) 193 | ∙ cong (λ f → mapPfin f (α a)) (funExt (λ x → αsim x .snd n ∙ αbar-eq2 x n)) 194 | 195 | αbar-uniq : ∀ a → αsim a ≡ αbar a 196 | αbar-uniq a = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin2 _ _ _)) 197 | (funExt (αbar-eq2 a)) 198 | 199 | 200 | 201 | -------------------------------------------------------------------------------- /Worrell/FromInjectivity.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Worrell.FromInjectivity where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (map to ∥map∥; rec to ∥rec∥) 12 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 13 | open import Cubical.Data.Sigma 14 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 15 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 16 | open import Cubical.Data.Nat renaming (elim to elimNat) 17 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_; _≟_ to _≟N_) 18 | open import Cubical.Data.Bool 19 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂; rec to rec⊎; elim to elim⊎) 20 | open import Cubical.Relation.Binary 21 | open import Cubical.Relation.Nullary 22 | open import Basics 23 | open import ListRelations 24 | open import ListStuff 25 | open import Pfin.AsSetQuot 26 | open import Pfin.AsFreeJoinSemilattice 27 | open import Pfin.PreservesCountableIntersections 28 | open BinaryRelation 29 | 30 | open import Worrell.Limits 31 | 32 | -- We prove that injectivity of (m : Pfin Vω → Vω) implies LLPO 33 | 34 | -- The sequence corresponding to the infinite tree in which each node 35 | -- has exactly one subtree. 36 | long : ∀ n → iPfin n 37 | long zero = tt 38 | long (suc n) = η (long n) 39 | 40 | long-res : ∀ n → iMapPfin n (long (suc n)) ≡ long n 41 | long-res zero = refl 42 | long-res (suc n) = cong η (long-res n) 43 | 44 | long-ch : Vω 45 | long-ch = long , long-res 46 | 47 | -- Given a sequence a : ℕ → Bool, we a variant (long? a) of long, 48 | -- which is equal to long if the sequence a contains only value false, 49 | -- but its height stop growing if there is n : ℕ such that (a n) is 50 | -- the first true in a. 51 | long? : (a : ℕ → Bool) → ∀ n → iPfin n 52 | long? a zero = tt 53 | long? a (suc n) = 54 | if a 0 then ø else η (long? (a ∘ suc) n) 55 | 56 | long?-res : (a : ℕ → Bool) 57 | → ∀ n → iMapPfin n (long? a (suc n)) ≡ long? a n 58 | long?-res a zero = refl 59 | long?-res a (suc n) with a 0 60 | ... | false = cong η (long?-res (a ∘ suc) n) 61 | ... | true = refl 62 | 63 | long?-ch : (a : ℕ → Bool) → Vω 64 | long?-ch a = long? a , long?-res a 65 | 66 | -- connections between long and (long? a) 67 | long?-res' : (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) 68 | → ∀ n → a (suc n) ≡ true → iMapPfin n (long? a (suc n)) ≡ long n 69 | long?-res' a aP zero eq = refl 70 | long?-res' a aP (suc n) eq with dichotomyBool (a 0) 71 | ... | inj₁ p = ⊥-rec (znots (cong fst (aP (_ , p) (_ , eq)))) 72 | ... | inj₂ p = 73 | cong (λ z → mapPfin (iMapPfin n) (if z then ø else η (if a 1 then ø else η (long? (λ x → a (suc (suc x))) n)))) p 74 | ∙ cong η (long?-res' (a ∘ suc) (λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) }) n eq) 75 | 76 | long?≠long : (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) 77 | → ∀ n → long? a (suc n) ≡ long (suc n) → a n ≡ false 78 | long?≠long a aP zero eq with a 0 79 | ... | false = refl 80 | ... | true = ⊥-rec (ødisjη (sym eq)) 81 | long?≠long a aP (suc n) eq with a 0 82 | ... | false = long?≠long (a ∘ suc) (λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) }) n (ηisInjective trunc eq) 83 | ... | true = ⊥-rec (ødisjη (sym eq)) 84 | 85 | -- given a sequence (a : ℕ → Bool) and two values x,y : A, the 86 | -- sequence (seq-ch a x y true) is defined as follows: it returns y if 87 | -- there exists an even number (k ≤ n) such that (a k = true) and 88 | -- (a j = false) for all j < k; in all other cases seq-ch a x y true n 89 | -- returns x. 90 | seq-ch : {A : Type} (a : ℕ → Bool) (x y : A) → Bool → ℕ → A 91 | seq-ch a x y b zero = if a 0 and b then y else x 92 | seq-ch a x y b (suc n) = if a 0 and b then y else seq-ch (a ∘ suc) x y (not b) n 93 | 94 | -- lemmata about the behaviour of seq-ch 95 | seq-ch-lem1 : {A : Type} (a : ℕ → Bool) (x y : A) (b : Bool) (n : ℕ) 96 | → (∀ k → k ≤N n → a k ≡ false) → seq-ch a x y b n ≡ x 97 | seq-ch-lem1 a x y b zero p with p 0 zero-≤ 98 | ... | r with a 0 | b 99 | ... | false | false = refl 100 | ... | false | true = refl 101 | ... | true | false = refl 102 | ... | true | true = ⊥-rec (true≢false r) 103 | seq-ch-lem1 a x y b (suc n) p with p 0 zero-≤ 104 | ... | r with a 0 | b 105 | ... | false | false = seq-ch-lem1 (a ∘ suc) x y true n λ k le → p (suc k) (suc-≤-suc le) 106 | ... | false | true = seq-ch-lem1 (a ∘ suc) x y false n λ k le → p (suc k) (suc-≤-suc le) 107 | ... | true | false = seq-ch-lem1 (a ∘ suc) x y true n λ k le → p (suc k) (suc-≤-suc le) 108 | ... | true | true = ⊥-rec (true≢false r) 109 | 110 | 111 | seq-ch-lem2 : {A : Type} (a : ℕ → Bool) (x y : A) (b : Bool) (n : ℕ) 112 | → ∀ k → k ≤N n → a k ≡ true → isEven? b k → seq-ch a x y b n ≡ y 113 | seq-ch-lem2 a x y b zero zero le eq ev with a 0 114 | ... | false = ⊥-rec (false≢true eq) 115 | seq-ch-lem2 a x y true zero zero le eq ev | true = refl 116 | seq-ch-lem2 a x y b (suc n) zero le eq ev with a 0 117 | ... | false = ⊥-rec (false≢true eq) 118 | seq-ch-lem2 a x y true (suc n) zero le eq ev | true = refl 119 | seq-ch-lem2 a x y b zero (suc k) le eq ev = ⊥-rec (¬-<-zero le) 120 | seq-ch-lem2 a x y b (suc n) (suc k) le eq ev with a 0 121 | seq-ch-lem2 a x y false (suc n) (suc k) le eq (suc ev) | false = seq-ch-lem2 (a ∘ suc) x y true n k (pred-≤-pred le) eq ev 122 | seq-ch-lem2 a x y true (suc n) (suc k) le eq (suc ev) | false = seq-ch-lem2 (a ∘ suc) x y false n k (pred-≤-pred le) eq ev 123 | seq-ch-lem2 a x y false (suc n) (suc k) le eq (suc ev) | true = seq-ch-lem2 (a ∘ suc) x y true n k (pred-≤-pred le) eq ev 124 | seq-ch-lem2 a x y true (suc n) (suc k) le eq ev | true = refl 125 | 126 | seq-ch-lem3 : {A : Type} (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) 127 | → (x y : A) (b : Bool) (n : ℕ) 128 | → ∀ k → k ≤N n → a k ≡ true → isEven? b k → seq-ch a x y (not b) n ≡ x 129 | seq-ch-lem3 a aP x y true zero zero le eq ev with a 0 130 | ... | false = ⊥-rec (false≢true eq) 131 | ... | true = refl 132 | seq-ch-lem3 a aP x y true (suc n) zero le eq ev with dichotomyBool (a 0) 133 | ... | inj₂ q = ⊥-rec (false≢true (sym q ∙ eq)) 134 | ... | inj₁ q = 135 | cong (λ z → if z and false then y else seq-ch (a ∘ suc) x y true n) q 136 | ∙ seq-ch-lem1 (a ∘ suc) x y true n 137 | (λ k le' → rec⊎ (λ p → ⊥-rec (snotz (cong fst (aP (_ , p) (_ , eq))) )) (λ p → p) (dichotomyBool (a (suc k)))) 138 | seq-ch-lem3 a aP x y b zero (suc k) le eq ev = ⊥-rec (¬-<-zero le) 139 | seq-ch-lem3 a aP x y false (suc n) (suc k) le eq (suc ev) with dichotomyBool (a 0) 140 | ... | inj₂ p = 141 | cong (λ z → if z and true then y else seq-ch (a ∘ suc) x y false n) p 142 | ∙ seq-ch-lem3 (a ∘ suc) (λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) }) x y true n k (pred-≤-pred le) eq ev 143 | ... | inj₁ p = ⊥-rec (snotz (cong fst (aP (_ , eq) (_ , p)))) 144 | seq-ch-lem3 a aP x y true (suc n) (suc k) le eq (suc ev) with dichotomyBool (a 0) 145 | ... | inj₂ p = 146 | cong (λ z → if z and false then y else seq-ch (a ∘ suc) x y true n) p 147 | ∙ seq-ch-lem3 (a ∘ suc) ((λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) })) x y false n k (pred-≤-pred le) eq ev 148 | ... | inj₁ p = ⊥-rec (snotz (cong fst (aP (_ , eq) (_ , p)))) 149 | 150 | diag-seq-ch : (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) (n : ℕ) → 151 | res iPfin-ch n (seq-ch a long-ch (long?-ch a) true (suc n) .fst (suc n)) ≡ seq-ch a long-ch (long?-ch a) true n .fst n 152 | diag-seq-ch a aP n with true-before? a n 153 | ... | yes (k , le , eq) with decEven k 154 | ... | inj₁ p = 155 | cong (λ z → res iPfin-ch n (z .fst (suc n))) (seq-ch-lem2 a long-ch (long?-ch a) true (suc n) k (≤-suc le) eq p) 156 | ∙ long?-res a n 157 | ∙ cong (λ z → z .fst n) (sym (seq-ch-lem2 a long-ch (long?-ch a) true n k le eq p)) 158 | ... | inj₂ p = 159 | cong (λ z → res iPfin-ch n (z .fst (suc n))) (seq-ch-lem3 a aP long-ch (long?-ch a) false (suc n) k (≤-suc le) eq p) 160 | ∙ long-res n 161 | ∙ cong (λ z → z .fst n) (sym (seq-ch-lem3 a aP long-ch (long?-ch a) false n k le eq p)) 162 | diag-seq-ch a aP n | no ¬p with dichotomyBool (a (suc n)) 163 | ... | inj₂ q = 164 | cong (λ z → res iPfin-ch n (z .fst (suc n))) (seq-ch-lem1 a long-ch (long?-ch a) true (suc n) 165 | (λ k le → rec⊎ (λ r → ⊥-rec (rec⊎ (λ p → ¬p (k , p , r)) (λ p → false≢true (sym q ∙ cong a (sym p) ∙ r)) (≤-suc-cases _ _ le))) 166 | (λ r → r) 167 | (dichotomyBool (a k)))) 168 | ∙ long-res n 169 | ∙ cong (λ z → z .fst n) (sym (seq-ch-lem1 a long-ch (long?-ch a) true n 170 | λ k le → rec⊎ (λ r → ⊥-rec (¬p (k , le , r))) (λ r → r) (dichotomyBool (a k)))) 171 | ... | inj₁ q with decEven (suc n) 172 | ... | inj₁ ev = 173 | cong (λ z → res iPfin-ch n (z .fst (suc n))) (seq-ch-lem2 a long-ch (long?-ch a) true (suc n) (suc n) ≤-refl q ev) 174 | ∙ long?-res' a aP n q 175 | ∙ cong (λ z → z .fst n) (sym (seq-ch-lem1 a long-ch (long?-ch a) true n 176 | λ k le → rec⊎ (λ r → ⊥-rec (¬p (k , le , r))) (λ r → r) (dichotomyBool (a k)))) 177 | ... | inj₂ odd = 178 | cong (λ z → res iPfin-ch n (z .fst (suc n))) (seq-ch-lem3 a aP long-ch (long?-ch a) false (suc n) (suc n) ≤-refl q odd) 179 | ∙ long-res n 180 | ∙ cong (λ z → z .fst n) (sym (seq-ch-lem1 a long-ch (long?-ch a) true n 181 | λ k le → rec⊎ (λ r → ⊥-rec (¬p (k , le , r))) (λ r → r) (dichotomyBool (a k)))) 182 | 183 | seq-ch-cases : {A : Type} (a : ℕ → Bool) 184 | → (x y : A) (b : Bool) (n : ℕ) 185 | → (seq-ch a x y b n ≡ x) ⊎ (seq-ch a x y b n ≡ y) 186 | seq-ch-cases a x y false zero with a 0 187 | ... | false = inj₁ refl 188 | ... | true = inj₁ refl 189 | seq-ch-cases a x y true zero with a 0 190 | ... | false = inj₁ refl 191 | ... | true = inj₂ refl 192 | seq-ch-cases a x y false (suc n) with a 0 193 | ... | false = seq-ch-cases (a ∘ suc) x y true n 194 | ... | true = seq-ch-cases (a ∘ suc) x y true n 195 | seq-ch-cases a x y true (suc n) with a 0 196 | ... | false = seq-ch-cases (a ∘ suc) x y false n 197 | ... | true = inj₂ refl 198 | 199 | -- Assuming injectivity of m we prove LLPO. 200 | module FromInjectivity (minj : ∀ s t → m s ≡ m t → s ≡ t) where 201 | 202 | -- The implication factors through the fact that two-element subsets 203 | -- of Vω are complete. 204 | 205 | complete : (x y1 y2 : Vω) (z : ℕ → Vω) 206 | → (∀ n → (z n ≡ y1) ⊎ (z n ≡ y2)) 207 | → (∀ n → z n .fst n ≡ x .fst n) 208 | → ⟨ x ∈ₛ (η y1 ∪ η y2) ⟩ 209 | complete x y1 y2 z p q = subst (λ z → ⟨ x ∈ₛ z ⟩) (minj s t (meq-lem s t eq)) (inl ∣ refl ∣) 210 | where 211 | t : Pfin Vω 212 | t = η y1 ∪ η y2 213 | 214 | s : Pfin Vω 215 | s = η x ∪ t 216 | 217 | sub : ∀ n → mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) s 218 | ⊆ mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) t 219 | sub n a = ∥rec∥ propTruncIsProp 220 | (λ { (inj₁ r) → 221 | ∥map∥ (λ eq → map⊎ (λ eq' → ∣ eq ∙ sym (q n) ∙ cong (λ w → w .fst n) eq' ∣) 222 | (λ eq' → ∣ eq ∙ sym (q n) ∙ cong (λ w → w .fst n) eq' ∣) 223 | (p n)) 224 | r ; 225 | (inj₂ r) → r }) 226 | 227 | eq : ∀ n → mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) s 228 | ≡ mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) t 229 | eq n = antisym⊆ (sub n) (λ a → inr) 230 | 231 | llpo : (a : ℕ → Bool) → isProp (Σ[ n ∈ ℕ ] a n ≡ true) 232 | → ∥ (∀ n → isEven n → a n ≡ false) ⊎ (∀ n → isOdd n → a n ≡ false) ∥ 233 | llpo a aP = 234 | ∥rec∥ propTruncIsProp 235 | (λ { (inj₁ p) → 236 | ∥map∥ (λ eq → inj₁ (λ n ev → rec⊎ (λ q → ⊥-rec (case1 eq n ev q)) (λ q → q) (dichotomyBool (a n)))) 237 | p 238 | ; (inj₂ p) → 239 | ∥map∥ (λ eq → inj₂ (λ n odd → rec⊎ (λ q → ⊥-rec (case2 eq n odd q)) (λ q → q) (dichotomyBool (a n)))) 240 | p }) 241 | (complete x y1 y2 z lem1 lem2) 242 | where 243 | y1 : Vω 244 | y1 = long-ch 245 | 246 | y2 : Vω 247 | y2 = long?-ch a 248 | 249 | z : ℕ → Vω 250 | z = seq-ch a y1 y2 true 251 | 252 | x : Vω 253 | x = (λ n → z n .fst n) , 254 | diag-seq-ch a aP 255 | 256 | lem1 : ∀ n → (z n ≡ y1) ⊎ (z n ≡ y2) 257 | lem1 = seq-ch-cases a y1 y2 true 258 | 259 | lem2 : ∀ n → z n .fst n ≡ x .fst n 260 | lem2 n = refl 261 | 262 | case1 : x ≡ y1 → ∀ n → isEven n → a n ≡ true → ⊥ 263 | case1 eqx n ev eq = false≢true (sym (long?≠long a aP n bad) ∙ eq) 264 | where 265 | bad : long? a (suc n) ≡ long (suc n) 266 | bad = 267 | sym (funExt⁻ (cong fst (seq-ch-lem2 a long-ch (long?-ch a) true (suc n) n (≤-suc ≤-refl) eq ev)) (suc n)) 268 | ∙ funExt⁻ (cong fst eqx) (suc n) 269 | 270 | case2 : x ≡ y2 → ∀ n → isOdd n → a n ≡ true → ⊥ 271 | case2 eqx n ev eq = false≢true (sym (long?≠long a aP n (sym bad)) ∙ eq) 272 | where 273 | bad : long (suc n) ≡ long? a (suc n) 274 | bad = 275 | sym (funExt⁻ (cong fst (seq-ch-lem3 a aP long-ch (long?-ch a) false (suc n) n (≤-suc ≤-refl) eq ev)) (suc n)) 276 | ∙ funExt⁻ (cong fst eqx) (suc n) 277 | 278 | -------------------------------------------------------------------------------- /Worrell/FromInjectivityOmega.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Worrell.FromInjectivityOmega where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (map to ∥map∥; rec to ∥rec∥) 12 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 13 | open import Cubical.Data.Sigma 14 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 15 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 16 | open import Cubical.Data.Nat renaming (elim to elimNat) 17 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_; _≟_ to _≟N_) 18 | open import Cubical.Data.Bool 19 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂; rec to rec⊎; elim to elim⊎) 20 | open import Cubical.Relation.Binary 21 | open import Cubical.Relation.Nullary 22 | open import Basics 23 | open import ListRelations 24 | open import ListStuff 25 | open import Pfin.AsSetQuot 26 | open import Pfin.AsFreeJoinSemilattice 27 | open import Pfin.PreservesCountableIntersections 28 | open BinaryRelation 29 | 30 | open import Worrell.Limits 31 | open import Worrell.FromInjectivity 32 | 33 | ℓω : Vω2 → Vω 34 | ℓω x = x .fst zero 35 | 36 | -- We prove that the injectivity of ℓω implies LLPO 37 | 38 | -- The sequences long and (long? a) can be extended to elements of Vω2 39 | long2 : ∀ n → iPfin2 n 40 | long2 zero = long-ch 41 | long2 (suc n) = η (long2 n) 42 | 43 | long2-res : ∀ n → iMapPfin2 n (long2 (suc n)) ≡ long2 n 44 | long2-res zero = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) (funExt long-res) 45 | long2-res (suc n) = cong η (long2-res n) 46 | 47 | long2-ch : Vω2 48 | long2-ch = long2 , long2-res 49 | 50 | long2? : (a : ℕ → Bool) → ∀ n → iPfin2 n 51 | long2? a zero = long?-ch a 52 | long2? a (suc n) = if a 0 then ø else η (long2? (a ∘ suc) n) 53 | 54 | long2?-res : (a : ℕ → Bool) 55 | → ∀ n → iMapPfin2 n (long2? a (suc n)) ≡ long2? a n 56 | long2?-res a zero = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 57 | (funExt (λ n → cong (iMapPfin n) (lem n (a 0)) ∙ long?-res a n)) 58 | where 59 | lem : ∀ n → (b : Bool) 60 | → mapPfin (proj iPfin-ch n) (if b then ø else η (long?-ch (a ∘ suc))) 61 | ≡ (if b then ø else η (long? (a ∘ suc) n)) 62 | lem n false = refl 63 | lem n true = refl 64 | long2?-res a (suc n) with a 0 65 | ... | false = cong η (long2?-res (a ∘ suc) n) 66 | ... | true = refl 67 | 68 | long2?-ch : (a : ℕ → Bool) → Vω2 69 | long2?-ch a = long2? a , long2?-res a 70 | 71 | -- another element of Vω, a variation of (long? a) that takes an 72 | -- additional Boolean value b in input. 73 | 74 | -- (long?? a true) is equal to long if the sequence a contains only 75 | -- value false, but its height stop growing if there is *even* n : ℕ 76 | -- such that (a n) is the first true in a. 77 | long?? : (a : ℕ → Bool) → Bool → ∀ n → iPfin n 78 | long?? a b zero = tt 79 | long?? a b (suc n) = 80 | if a 0 and b then ø else η (long?? (a ∘ suc) (not b) n) 81 | 82 | long??-res : (a : ℕ → Bool) (b : Bool) 83 | → ∀ n → iMapPfin n (long?? a b (suc n)) ≡ long?? a b n 84 | long??-res a b zero = refl 85 | long??-res a b (suc n) with a 0 | b 86 | ... | false | false = cong η (long??-res (a ∘ suc) true n) 87 | ... | false | true = cong η (long??-res (a ∘ suc) false n) 88 | ... | true | false = cong η (long??-res (a ∘ suc) true n) 89 | ... | true | true = refl 90 | 91 | long??-ch : (a : ℕ → Bool) → Bool → Vω 92 | long??-ch a b = long?? a b , long??-res a b 93 | 94 | -- lemmata about the behaviour of long?? 95 | long??-lem1 : (a : ℕ → Bool) (b : Bool) (n : ℕ) 96 | → (∀ k → k ≤N n → a k ≡ false) → long?? a b n ≡ long n 97 | long??-lem1 a b zero p = refl 98 | long??-lem1 a b (suc n) p with dichotomyBool (a 0) 99 | ... | inj₁ eq = ⊥-rec (true≢false (sym eq ∙ p 0 zero-≤)) 100 | long??-lem1 a true (suc n) p | inj₂ eq = 101 | cong (λ z → if z and true then ø else η (long?? (a ∘ suc) false n)) eq 102 | ∙ cong η (long??-lem1 (a ∘ suc) false n (λ k le → p (suc k) (suc-≤-suc le))) 103 | long??-lem1 a false (suc n) p | inj₂ eq = 104 | cong (λ z → if z and false then ø else η (long?? (a ∘ suc) true n)) eq 105 | ∙ cong η (long??-lem1 (a ∘ suc) true n (λ k le → p (suc k) (suc-≤-suc le))) 106 | 107 | long??-lem2 : (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) 108 | → (b : Bool) (n : ℕ) 109 | → ∀ k → k ≤N n → a k ≡ true → isEven? b k → long?? a b n ≡ long? a n 110 | long??-lem2 a aP b zero k lek eqk evk = refl 111 | long??-lem2 a aP true (suc n) zero lek eqk evk = 112 | cong (λ z → if z and true then ø else η (long?? (a ∘ suc) false n)) eqk 113 | ∙ cong (λ z → if z then ø else η (long? (a ∘ suc) n)) (sym eqk) 114 | long??-lem2 a aP b (suc n) (suc k) lek eqk evk with dichotomyBool (a 0) 115 | long??-lem2 a aP b (suc n) (suc k) lek eqk evk | inj₁ eq0 = 116 | ⊥-rec (snotz (cong fst (aP (_ , eqk) (_ , eq0)))) 117 | long??-lem2 a aP false (suc n) (suc k) lek eqk (suc ev) | inj₂ eq0 = 118 | cong (λ z → if z and false then ø else η (long?? (a ∘ suc) true n)) eq0 119 | ∙ cong η (long??-lem2 (a ∘ suc) (λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) }) true n k (pred-≤-pred lek) eqk ev) 120 | ∙ cong (λ z → if z then ø else η (long? (a ∘ suc) n)) (sym eq0) 121 | long??-lem2 a aP true (suc n) (suc k) lek eqk (suc odd) | inj₂ eq0 = 122 | cong (λ z → if z and true then ø else η (long?? (a ∘ suc) false n)) eq0 123 | ∙ cong η (long??-lem2 (a ∘ suc) (λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) }) false n k (pred-≤-pred lek) eqk odd) 124 | ∙ cong (λ z → if z then ø else η (long? (a ∘ suc) n)) (sym eq0) 125 | 126 | long??-lem3 : (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) 127 | → (b : Bool) (n : ℕ) 128 | → ∀ k → k ≤N n → a k ≡ true → isEven? b k → long?? a (not b) n ≡ long n 129 | long??-lem3 a aP b zero k le eqk evk = refl 130 | long??-lem3 a aP true (suc n) zero le eq0 evk = 131 | cong (λ z → if z and false then ø else η (long?? (a ∘ suc) true n)) eq0 132 | ∙ cong η (long??-lem1 (a ∘ suc) true n 133 | (λ k le → rec⊎ (λ eqk → ⊥-rec (snotz (cong fst (aP (_ , eqk) (_ , eq0))))) (λ x → x) (dichotomyBool (a (suc k))))) 134 | long??-lem3 a aP b (suc n) (suc k) le eqk evk with dichotomyBool (a 0) 135 | ... | inj₁ eq0 = ⊥-rec (snotz (cong fst (aP (_ , eqk) (_ , eq0)))) 136 | long??-lem3 a aP false (suc n) (suc k) le eqk (suc ev) | inj₂ eq0 = 137 | cong (λ z → if z and true then ø else η (long?? (a ∘ suc) false n)) eq0 138 | ∙ cong η (long??-lem3 (a ∘ suc) ((λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) })) true n k (pred-≤-pred le) eqk ev) 139 | long??-lem3 a aP true (suc n) (suc k) le eqk (suc odd) | inj₂ eq0 = 140 | cong (λ z → if z and false then ø else η (long?? (a ∘ suc) true n)) eq0 141 | ∙ cong η (long??-lem3 (a ∘ suc) ((λ { (x , q) (y , r) → Σ≡Prop (λ _ → isSetBool _ _) (injSuc (cong fst (aP (_ , q) (_ , r)))) })) false n k (pred-≤-pred le) eqk odd) 142 | 143 | -- (long?? a true) is the same as the diagonal of 144 | -- (seq-ch a long-ch (long?-ch a) true) 145 | seq-ch-long?? : (a : ℕ → Bool) (aP : isProp (Σ[ n ∈ ℕ ] a n ≡ true)) (n : ℕ) 146 | → seq-ch a long-ch (long?-ch a) true n .fst n ≡ long??-ch a true .fst n 147 | seq-ch-long?? a aP n with true-before? a n 148 | ... | yes (k , le , eq) with decEven k 149 | ... | inj₁ ev = 150 | cong (λ z → z .fst n) (seq-ch-lem2 a long-ch (long?-ch a) true n k le eq ev) 151 | ∙ sym (long??-lem2 a aP true n k le eq ev) 152 | ... | inj₂ odd = 153 | cong (λ z → z .fst n) (seq-ch-lem3 a aP long-ch (long?-ch a) false n k le eq odd) 154 | ∙ sym (long??-lem3 a aP false n k le eq odd ) 155 | seq-ch-long?? a aP n | no ¬p = 156 | cong (λ z → z .fst n) (seq-ch-lem1 a long-ch (long?-ch a) true n 157 | (λ k le → rec⊎ (λ eq → ⊥-rec (¬p (k , le , eq))) (λ x → x) (dichotomyBool (a k)))) 158 | ∙ sym (long??-lem1 a true n (λ k le → rec⊎ (λ eq → ⊥-rec (¬p (k , le , eq))) (λ x → x) (dichotomyBool (a k)))) 159 | 160 | -- long?? also extends to an element of Vω2 161 | long2?? : (a : ℕ → Bool) → Bool → ∀ n → iPfin2 n 162 | long2?? a b zero = long??-ch a b 163 | long2?? a b (suc n) = 164 | if a 0 and b then ø else η (long2?? (a ∘ suc) (not b) n) 165 | 166 | long2??-res : (a : ℕ → Bool) (b : Bool) 167 | → ∀ n → iMapPfin2 n (long2?? a b (suc n)) ≡ long2?? a b n 168 | long2??-res a b zero = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 169 | (funExt (λ n → cong (iMapPfin n) (lem n b (a 0)) ∙ long??-res a b n)) 170 | where 171 | lem : ∀ n → (b b' : Bool) 172 | → mapPfin (proj iPfin-ch n) (if b' and b then ø else η (long??-ch (a ∘ suc) (not b))) 173 | ≡ (if b' and b then ø else η (long?? (a ∘ suc) (not b) n)) 174 | lem n false false = refl 175 | lem n true false = refl 176 | lem n false true = refl 177 | lem n true true = refl 178 | long2??-res a b (suc n) with a 0 | b 179 | ... | false | false = cong η (long2??-res (a ∘ suc) true n) 180 | ... | false | true = cong η (long2??-res (a ∘ suc) false n) 181 | ... | true | false = cong η (long2??-res (a ∘ suc) true n) 182 | ... | true | true = refl 183 | 184 | long2??-ch : (a : ℕ → Bool) → Bool → Vω2 185 | long2??-ch a b = long2?? a b , long2??-res a b 186 | 187 | -- injectivity of ℓω implies LLPO 188 | 189 | -- again the implication factors through some form of completeness 190 | -- theorem 191 | module FromInjectivityOmega (ℓωinj : (s t : Vω2) → ℓω s ≡ ℓω t → s ≡ t) where 192 | 193 | ℓω+1inj : (s t : Vω2) → m (s .fst 1) ≡ m (t .fst 1) → s ≡ t 194 | ℓω+1inj s t eq = ℓωinj s t (sym (s .snd 0) ∙ eq ∙ t .snd 0) 195 | 196 | joinω : (x y : ∀ n → iPfin n) → ∀ n → iPfin n 197 | joinω x y zero = tt 198 | joinω x y (suc n) = η (x n) ∪ η (y n) 199 | 200 | joinω-res : (x y : Vω) 201 | → ∀ n → iMapPfin n (joinω (x .fst) (y .fst) (suc n)) ≡ joinω (x .fst) (y .fst) n 202 | joinω-res x y zero = refl 203 | joinω-res x y (suc n) = cong₂ _∪_ (cong η (x .snd n)) (cong η (y .snd n)) 204 | 205 | joinω-ch : (x y : Vω) → Vω 206 | joinω-ch x y = joinω (x .fst) (y .fst) , joinω-res x y 207 | 208 | joinω2 : (x y : ∀ n → iPfin2 n) → ∀ n → iPfin2 n 209 | joinω2 x y zero = joinω-ch (x 0) (y 0) 210 | joinω2 x y (suc n) = η (x n) ∪ η (y n) 211 | 212 | joinω2-res : (x y : Vω2) 213 | → ∀ n → iMapPfin2 n (joinω2 (x .fst) (y .fst) (suc n)) ≡ joinω2 (x .fst) (y .fst) n 214 | joinω2-res x y zero = 215 | Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 216 | (funExt (joinω-res (x .fst 0) (y .fst 0))) 217 | joinω2-res x y (suc n) = cong₂ _∪_ (cong η (x .snd n)) (cong η (y .snd n)) 218 | 219 | joinω2-ch : (x y : Vω2) → Vω2 220 | joinω2-ch x y = joinω2 (x .fst) (y .fst) , joinω2-res x y 221 | 222 | 3joinω : (x y z : ∀ n → iPfin n) → ∀ n → iPfin n 223 | 3joinω x y z zero = tt 224 | 3joinω x y z (suc n) = η (x n) ∪ (η (y n) ∪ η (z n)) 225 | 226 | 3joinω-res : (x y z : Vω) 227 | → ∀ n → iMapPfin n (3joinω (x .fst) (y .fst) (z .fst) (suc n)) ≡ 3joinω (x .fst) (y .fst) (z .fst) n 228 | 3joinω-res x y z zero = refl 229 | 3joinω-res x y z (suc n) = cong₂ _∪_ (cong η (x .snd n)) (cong₂ _∪_ (cong η (y .snd n)) (cong η (z .snd n))) 230 | 231 | 3joinω-ch : (x y z : Vω) → Vω 232 | 3joinω-ch x y z = 3joinω (x .fst) (y .fst) (z .fst) , 3joinω-res x y z 233 | 234 | 3joinω2 : (x y z : ∀ n → iPfin2 n) → ∀ n → iPfin2 n 235 | 3joinω2 x y z zero = 3joinω-ch (x 0) (y 0) (z 0) 236 | 3joinω2 x y z (suc n) = η (x n) ∪ (η (y n) ∪ η (z n)) 237 | 238 | 3joinω2-res : (x y z : Vω2) 239 | → ∀ n → iMapPfin2 n (3joinω2 (x .fst) (y .fst) (z .fst) (suc n)) ≡ 3joinω2 (x .fst) (y .fst) (z .fst) n 240 | 3joinω2-res x y z zero = 241 | Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 242 | (funExt (3joinω-res (x .fst 0) (y .fst 0) (z .fst 0))) 243 | 3joinω2-res x y z (suc n) = cong₂ _∪_ (cong η (x .snd n)) (cong₂ _∪_ (cong η (y .snd n)) (cong η (z .snd n))) 244 | 245 | 3joinω2-ch : (x y z : Vω2) → Vω2 246 | 3joinω2-ch x y z = 3joinω2 (x .fst) (y .fst) (z .fst) , 3joinω2-res x y z 247 | 248 | complete : (x : Vω2) (y1 y2 : Vω2) (z : ℕ → Vω) 249 | → (∀ n → (z n ≡ y1 .fst 0) ⊎ (z n ≡ y2 .fst 0)) 250 | → (∀ n → z n .fst n ≡ x .fst 0 .fst n) 251 | → ⟨ x .fst 0 ∈ₛ (η (y1 .fst 0) ∪ η (y2 .fst 0)) ⟩ 252 | complete x y1 y2 z p q = 253 | subst (λ z → ⟨ x .fst 0 ∈ₛ z ⟩) (funExt⁻ (cong fst (ℓω+1inj s t (meq-lem (s .fst 1) (t .fst 1) eq))) 1) (inl ∣ refl ∣) 254 | where 255 | t : Vω2 256 | t = joinω2-ch y1 y2 257 | 258 | s : Vω2 259 | s = 3joinω2-ch x y1 y2 260 | 261 | sub : ∀ n → mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) (s .fst 1) 262 | ⊆ mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) (t .fst 1) 263 | sub zero a _ = ∣ inj₁ ∣ refl ∣ ∣ 264 | sub (suc n) a = ∥rec∥ propTruncIsProp 265 | (λ { (inj₁ r) → 266 | ∥map∥ (λ eq → map⊎ (λ eq' → ∣ eq ∙ sym (q (suc n)) ∙ cong (λ w → w .fst (suc n)) eq' ∣) 267 | (λ eq' → ∣ eq ∙ sym (q (suc n)) ∙ cong (λ w → w .fst (suc n)) eq' ∣) 268 | (p (suc n))) 269 | r ; 270 | (inj₂ r) → r }) 271 | 272 | eq : ∀ n → mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) (s .fst 1) 273 | ≡ mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) (t .fst 1) 274 | eq n = antisym⊆ (sub n) (λ a → inr) 275 | 276 | 277 | llpo : (a : ℕ → Bool) → isProp (Σ[ n ∈ ℕ ] a n ≡ true) 278 | → ∥ (∀ n → isEven n → a n ≡ false) ⊎ (∀ n → isOdd n → a n ≡ false) ∥ 279 | llpo a aP = 280 | ∥rec∥ propTruncIsProp 281 | (λ { (inj₁ p) → 282 | ∥map∥ (λ eq → inj₁ (λ n ev → rec⊎ (λ q → ⊥-rec (case1 eq n ev q)) (λ q → q) (dichotomyBool (a n)))) 283 | p 284 | ; (inj₂ p) → 285 | ∥map∥ (λ eq → inj₂ (λ n odd → rec⊎ (λ q → ⊥-rec (case2 eq n odd q)) (λ q → q) (dichotomyBool (a n)))) 286 | p }) 287 | (complete x y1 y2 z lem1 lem2) 288 | where 289 | y1 : Vω2 290 | y1 = long2-ch 291 | 292 | y2 : Vω2 293 | y2 = long2?-ch a 294 | 295 | z : ℕ → Vω 296 | z = seq-ch a long-ch (long?-ch a) true 297 | 298 | x : Vω2 299 | x = long2??-ch a true 300 | 301 | lem1 : ∀ n → (z n ≡ y1 .fst 0) ⊎ (z n ≡ y2 .fst 0) 302 | lem1 = seq-ch-cases a _ _ true 303 | 304 | lem2 : ∀ n → z n .fst n ≡ x .fst 0 .fst n 305 | lem2 n = seq-ch-long?? a aP n 306 | 307 | case1 : x .fst 0 ≡ y1 .fst 0 → ∀ n → isEven n → a n ≡ true → ⊥ 308 | case1 eqx n ev eq = false≢true (sym (long?≠long a aP n bad) ∙ eq) 309 | where 310 | bad : long? a (suc n) ≡ long (suc n) 311 | bad = 312 | sym (funExt⁻ (cong fst (seq-ch-lem2 a long-ch (long?-ch a) true (suc n) n (≤-suc ≤-refl) eq ev)) (suc n)) 313 | ∙ seq-ch-long?? a aP (suc n) 314 | ∙ funExt⁻ (cong fst eqx) (suc n) 315 | 316 | case2 : x .fst 0 ≡ y2 .fst 0 → ∀ n → isOdd n → a n ≡ true → ⊥ 317 | case2 eqx n ev eq = false≢true (sym (long?≠long a aP n (sym bad)) ∙ eq) 318 | where 319 | bad : long (suc n) ≡ long? a (suc n) 320 | bad = 321 | sym (funExt⁻ (cong fst (seq-ch-lem3 a aP long-ch (long?-ch a) false (suc n) n (≤-suc ≤-refl) eq ev)) (suc n)) 322 | ∙ seq-ch-long?? a aP (suc n) 323 | ∙ funExt⁻ (cong fst eqx) (suc n) 324 | 325 | -------------------------------------------------------------------------------- /Worrell/Limits.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Worrell.Limits where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (map to ∥map∥; rec to ∥rec∥) 12 | open import Cubical.HITs.SetQuotients renaming (rec to recQ) 13 | open import Cubical.Data.Sigma 14 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 15 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 16 | open import Cubical.Data.Nat renaming (elim to elimNat) 17 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_; _≟_ to _≟N_) 18 | open import Cubical.Data.Bool 19 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂; rec to rec⊎; elim to elim⊎) 20 | open import Cubical.Relation.Binary 21 | open import Cubical.Relation.Nullary 22 | open import Basics 23 | open import ListRelations 24 | open import ListStuff 25 | open import Pfin.AsSetQuot 26 | open import Pfin.AsFreeJoinSemilattice 27 | open import Pfin.PreservesCountableIntersections 28 | open BinaryRelation 29 | 30 | -- Stuff about ω-limits 31 | 32 | -- ω-chains 33 | record ωChain ℓ : Type (ℓ-suc ℓ) where 34 | constructor _,_ 35 | field 36 | Ch : ℕ → Type ℓ 37 | res : (n : ℕ) → Ch (suc n) → Ch n 38 | open ωChain public 39 | 40 | ωChain₀ = ωChain ℓ-zero 41 | 42 | -- the limit of a ω-chain 43 | ωLimit : ∀{ℓ} → ωChain ℓ → Type ℓ 44 | ωLimit (A , r) = 45 | Σ[ x ∈ ((n : ℕ) → A n) ] ∀ n → r n (x (suc n)) ≡ x n 46 | 47 | -- definition of cone for a ω-chain 48 | Cone : ∀{ℓ} → ωChain ℓ → Type ℓ → Type ℓ 49 | Cone (A , r) C = 50 | Σ[ f ∈ (∀ n → C → A n) ] ∀ n x → r n (f (suc n) x) ≡ f n x 51 | 52 | -- the limit is a cone 53 | proj : ∀{ℓ} (c@(A , r) : ωChain ℓ) n → ωLimit c → A n 54 | proj (A , r) n (x , eq) = x n 55 | 56 | ConeωLimit : ∀{ℓ} (c : ωChain ℓ) → Cone c (ωLimit c) 57 | ConeωLimit c@(A , r) = proj c , λ n x → x .snd n 58 | 59 | -- isomorphism between cones with vertex C and functions from C to the 60 | -- ω-limit 61 | AdjLim : ∀{ℓ} (c : ωChain ℓ) (C : Type ℓ) 62 | → Iso (C → ωLimit c) (Cone c C) 63 | AdjLim (A , r) C = iso 64 | (λ g → (λ n x → g x .fst n) , λ n x → g x .snd n) 65 | (λ h@(f , eq) x → (λ n → f n x) , (λ n → eq n x)) 66 | (λ _ → refl) 67 | (λ g → funExt (λ _ → refl)) 68 | 69 | -- lemma for proving injectivity of functions into a ω-limit 70 | lem-injective-lim : ∀{ℓ} (c : ωChain ℓ) (C : Type ℓ) 71 | → (g : C → ωLimit c) 72 | → (∀ x y → (∀ n → Iso.fun (AdjLim c C) g .fst n x ≡ Iso.fun (AdjLim c C) g .fst n y) → x ≡ y) 73 | → ∀ x y → g x ≡ g y → x ≡ y 74 | lem-injective-lim c C g p x y eq = p x y (funExt⁻ (cong fst eq)) 75 | 76 | -- existence-part of the universal property of the limit: given a type 77 | -- L which is a cone, then there exists a map from L to the limit 78 | ∃upωLimit : ∀{ℓ} (c : ωChain ℓ) (L : Type ℓ) 79 | → Cone c L 80 | → L → ωLimit c 81 | ∃upωLimit c L = Iso.inv (AdjLim c L) 82 | 83 | -- shifted chain, which has the same (isomorphic) limit 84 | shift : ∀{ℓ} → ωChain ℓ → ωChain ℓ 85 | shift (A , r) = (λ n → A (suc n)) , λ n → r (suc n) 86 | 87 | shift-iso : ∀{ℓ} (c : ωChain ℓ) 88 | → (∀ n → isSet (c .Ch n)) 89 | → Iso (ωLimit c) (ωLimit (shift c)) 90 | shift-iso c@(A , r) cS = 91 | iso (Iso.inv (AdjLim (shift c) _) cone-sh) 92 | (Iso.inv (AdjLim c _) cone) 93 | (λ { (x , p) → Σ≡Prop (λ _ → isPropΠ (λ n → cS (suc n) _ _)) (λ i n → p n i) }) 94 | λ { (x , p) → Σ≡Prop (λ _ → isPropΠ (λ n → cS n _ _)) (λ i n → p n i) } 95 | where 96 | cone-sh : Cone (shift c) (ωLimit c) 97 | cone-sh = (λ n x → x .fst (suc n)) , λ n x → x .snd (suc n) 98 | 99 | cone : Cone c (ωLimit (shift c)) 100 | cone = (λ n x → r n (x .fst n)) , λ n x → cong (r n) (x .snd n) 101 | 102 | shift≃ : ∀{ℓ} (c : ωChain ℓ) 103 | → (∀ n → isSet (c .Ch n)) 104 | → ωLimit c ≃ ωLimit (shift c) 105 | shift≃ c cS = isoToEquiv (shift-iso c cS) 106 | 107 | -- restrictions between types in a ω-chain 108 | ires : ∀{ℓ} (c : ωChain ℓ) 109 | → ∀ n k → k ≤N n 110 | → c .Ch n → c .Ch k 111 | ires c zero k le y = J (λ k _ → c .Ch k) y (sym (≤0→≡0 le)) 112 | ires c (suc n) k le y with ≤-suc-cases k n le 113 | ... | inj₁ p = ires c n k p (c .res n y) 114 | ... | inj₂ p = J (λ k _ → c .Ch k) y (sym p) 115 | 116 | ires-eq : ∀{ℓ} (c : ωChain ℓ) (x : ωLimit c) 117 | → ∀ n k (le : k ≤N n) 118 | → ires c n k le (x .fst n) ≡ x .fst k 119 | ires-eq c x zero k le = 120 | J (λ k eq → J (λ k _ → c .Ch k) (x .fst 0) eq ≡ x .fst k) (JRefl (λ k _ → c .Ch k) (x .fst 0)) (sym (≤0→≡0 le)) 121 | ires-eq c x (suc n) k le with ≤-suc-cases k n le 122 | ... | inj₁ p = cong (ires c n k p) (x .snd n) ∙ ires-eq c x n k p 123 | ... | inj₂ p = J (λ k eq → J (λ k _ → c .Ch k) (x .fst (suc n)) eq ≡ x .fst k) (JRefl (λ k _ → c .Ch k) (x .fst (suc n))) (sym p) 124 | 125 | -- the ω-chain of iterated applications of Pfin to the unit type 126 | iPfin : ℕ → Type 127 | iPfin zero = Unit 128 | iPfin (suc n) = Pfin (iPfin n) 129 | 130 | isSetiPfin : ∀ n → isSet (iPfin n) 131 | isSetiPfin zero = λ _ _ _ _ → refl 132 | isSetiPfin (suc n) = trunc 133 | 134 | iMapPfin : ∀ n → iPfin (suc n) → iPfin n 135 | iMapPfin zero = λ _ → tt 136 | iMapPfin (suc n) = mapPfin (iMapPfin n) 137 | 138 | iPfin-ch : ωChain₀ 139 | iPfin-ch = iPfin , iMapPfin 140 | 141 | -- the limit of the ω-chain iPfin-ch 142 | Vω : Type 143 | Vω = ωLimit iPfin-ch 144 | 145 | -- Vω is a set 146 | isSetVω : isSet Vω 147 | isSetVω = 148 | isSetΣ (isSetΠ (λ _ → isSetiPfin _)) 149 | (λ _ → isProp→isSet (isPropΠ (λ _ → isSetiPfin _ _ _))) 150 | 151 | -- the limit of the shifted ω-chain 152 | VωSh : Type 153 | VωSh = ωLimit (shift iPfin-ch) 154 | 155 | -- path equality in (iPfin n) is decidable 156 | iPfinDecEq : ∀ n (x y : iPfin n) → Dec (x ≡ y) 157 | iPfinDecEq zero x y = yes refl 158 | iPfinDecEq (suc n) = PfinDecEq (iPfinDecEq n) 159 | 160 | -- this limit is an algebra for Pfin, which is given by the universal 161 | -- property of the limit (i.e. Pfin Vω is a cone) 162 | ConePfinVω : Cone iPfin-ch (Pfin Vω) 163 | ConePfinVω = 164 | (λ n x → iMapPfin n (mapPfin (proj iPfin-ch n) x)) , 165 | λ n x → cong (iMapPfin n) 166 | (mapPfinComp x ∙ cong (λ f → mapPfin f x) 167 | (funExt (λ y → y .snd n))) 168 | 169 | m : Pfin Vω → Vω 170 | m = Iso.inv (AdjLim iPfin-ch _) ConePfinVω 171 | 172 | -- a characterization of the equality (m s ≡ m t) 173 | meq-lem : ∀ s t 174 | → (∀ n → mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) s 175 | ≡ mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) t) 176 | → m s ≡ m t 177 | meq-lem s t p = Σ≡Prop (λ _ → isPropΠ (λ _ → isSetiPfin _ _ _)) 178 | (funExt (λ { zero → refl ; 179 | (suc n) → antisym⊆ 180 | (λ x mx → ∥rec∥ (snd ((x ∈ₛ mapPfin (iMapPfin n) (mapPfin (proj iPfin-ch (suc n)) t)))) 181 | (λ { (y , my , eq) → subst (λ z → ⟨ x ∈ₛ z ⟩) 182 | (sym (mapPfinComp t 183 | ∙ cong (λ f → mapPfin f t) (funExt (λ z → z .snd n)) 184 | ∙ sym (p n) 185 | ∙ cong (λ f → mapPfin f s) (sym (funExt (λ z → z .snd n))) 186 | ∙ sym (mapPfinComp s))) 187 | (subst (λ z → ⟨ z ∈ₛ mapPfin (iMapPfin n) (mapPfin (proj iPfin-ch (suc n)) s) ⟩) 188 | eq 189 | (∈ₛmapPfin (iMapPfin n) y (mapPfin (proj iPfin-ch (suc n)) s) my)) }) 190 | (pre∈ₛmapPfin (iMapPfin n) x (mapPfin (proj iPfin-ch (suc n)) s) mx)) 191 | (λ x mx → ∥rec∥ (snd ((x ∈ₛ mapPfin (iMapPfin n) (mapPfin (proj iPfin-ch (suc n)) s)))) 192 | (λ { (y , my , eq) → subst (λ z → ⟨ x ∈ₛ z ⟩) 193 | (sym (mapPfinComp s 194 | ∙ cong (λ f → mapPfin f s) (funExt (λ z → z .snd n)) 195 | ∙ p n 196 | ∙ cong (λ f → mapPfin f t) (sym (funExt (λ z → z .snd n))) 197 | ∙ sym (mapPfinComp t))) 198 | (subst (λ z → ⟨ z ∈ₛ mapPfin (iMapPfin n) (mapPfin (proj iPfin-ch (suc n)) t) ⟩) 199 | eq 200 | (∈ₛmapPfin (iMapPfin n) y (mapPfin (proj iPfin-ch (suc n)) t) my)) }) 201 | (pre∈ₛmapPfin (iMapPfin n) x (mapPfin (proj iPfin-ch (suc n)) t) mx))})) 202 | 203 | meq-lem2 : ∀ s t 204 | → m s ≡ m t 205 | → ∀ n → mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) s 206 | ≡ mapPfin (λ (x : ωLimit iPfin-ch) → x .fst n) t 207 | meq-lem2 s t eq n = 208 | cong (λ f → mapPfin f s) (funExt (λ x → sym (x .snd n))) 209 | ∙ sym (mapPfinComp s) 210 | ∙ funExt⁻ (cong fst eq) (suc n) 211 | ∙ mapPfinComp t 212 | ∙ cong (λ f → mapPfin f t) (funExt (λ x → x .snd n)) 213 | 214 | 215 | -- an alternative but equivalent construction of the limit Vω using 216 | -- PfinQ (set quotient of lists) instead of Pfin. 217 | 218 | iPfinQ : ℕ → Type 219 | iPfinQ zero = Unit 220 | iPfinQ (suc n) = PfinQ (iPfinQ n) 221 | 222 | isSetiPfinQ : ∀ n → isSet (iPfinQ n) 223 | isSetiPfinQ zero = λ _ _ _ _ → refl 224 | isSetiPfinQ (suc n) = squash/ 225 | 226 | iMapPfinQ : ∀ n → iPfinQ (suc n) → iPfinQ n 227 | iMapPfinQ zero = λ _ → tt 228 | iMapPfinQ (suc n) = mapPfinQ (iMapPfinQ n) 229 | 230 | iPfinQ-ch : ωChain₀ 231 | iPfinQ-ch = iPfinQ , iMapPfinQ 232 | 233 | VωQ : Type 234 | VωQ = ωLimit iPfinQ-ch 235 | 236 | -- path equality in (iPfinQ n) is decidable 237 | iPfinQDecEq : ∀ n (x y : iPfinQ n) → Dec (x ≡ y) 238 | iPfinQDecEq zero x y = yes refl 239 | iPfinQDecEq (suc n) = PfinQDecEq (iPfinQDecEq n) 240 | 241 | -- the PfinQ-algebra structure on VωQ, analogous to the function m 242 | ConePfinQVωQ : Cone iPfinQ-ch (PfinQ VωQ) 243 | ConePfinQVωQ = 244 | (λ n x → iMapPfinQ n (mapPfinQ (proj iPfinQ-ch n) x)) , 245 | λ n x → cong (iMapPfinQ n) 246 | (mapPfinQComp x ∙ cong (λ f → mapPfinQ f x) 247 | (funExt (λ y → y .snd n))) 248 | 249 | mQ : PfinQ VωQ → VωQ 250 | mQ = Iso.inv (AdjLim iPfinQ-ch _) ConePfinQVωQ 251 | 252 | 253 | -- the ω-chain of iterated applications of Pfin to Vω 254 | iPfin2 : ℕ → Type 255 | iPfin2 zero = Vω 256 | iPfin2 (suc n) = Pfin (iPfin2 n) 257 | 258 | isSetiPfin2 : ∀ n → isSet (iPfin2 n) 259 | isSetiPfin2 zero = isSetVω 260 | isSetiPfin2 (suc n) = trunc 261 | 262 | iMapPfin2 : ∀ n → iPfin2 (suc n) → iPfin2 n 263 | iMapPfin2 zero = m 264 | iMapPfin2 (suc n) = mapPfin (iMapPfin2 n) 265 | 266 | iPfin2-ch : ωChain₀ 267 | iPfin2-ch = iPfin2 , iMapPfin2 268 | 269 | -- the limit of the ω-chain iPfin2-ch 270 | Vω2 : Type 271 | Vω2 = ωLimit iPfin2-ch 272 | 273 | -- algebra map on Vω2 274 | ConePfinVω2 : Cone iPfin2-ch (Pfin Vω2) 275 | ConePfinVω2 = 276 | (λ n x → iMapPfin2 n (mapPfin (proj iPfin2-ch n) x)) , 277 | λ n x → cong (iMapPfin2 n) 278 | (mapPfinComp x ∙ cong (λ f → mapPfin f x) 279 | (funExt (λ y → y .snd n))) 280 | 281 | τ-1 : Pfin Vω2 → Vω2 282 | τ-1 = Iso.inv (AdjLim iPfin2-ch _) ConePfinVω2 283 | 284 | -- the limit of the shifted (ω+ω)-chain 285 | Vω2Sh : Type 286 | Vω2Sh = ωLimit (shift iPfin2-ch) 287 | -------------------------------------------------------------------------------- /Worrell/NoSurjAlgebra.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --no-import-sorts #-} 2 | 3 | module Worrell.NoSurjAlgebra where 4 | 5 | open import Cubical.Core.Everything 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Everything 8 | open import Cubical.Functions.Logic renaming (⊥ to ⊥ₚ) 9 | open import Cubical.Relation.Everything 10 | open import Cubical.HITs.PropositionalTruncation as PropTrunc 11 | renaming (map to ∥map∥; rec to ∥rec∥) 12 | open import Cubical.HITs.SetQuotients renaming (rec to recQ; elimProp to elimQProp; elim to elimQ; elimProp2 to elimQProp2) 13 | open import Cubical.Data.Sigma 14 | open import Cubical.Data.List renaming (map to mapList) hiding ([_]) 15 | open import Cubical.Data.Empty renaming (elim to ⊥-elim; rec to ⊥-rec) 16 | open import Cubical.Data.Unit 17 | open import Cubical.Data.Nat renaming (elim to elimNat) 18 | open import Cubical.Data.Nat.Order hiding (eq) renaming (_≤_ to _≤N_; _≟_ to _≟N_) 19 | open import Cubical.Data.Bool 20 | open import Cubical.Data.Sum renaming (map to map⊎; inl to inj₁; inr to inj₂; rec to rec⊎; elim to elim⊎) 21 | open import Cubical.Relation.Binary 22 | open import Cubical.Relation.Nullary 23 | open import Basics 24 | open import ListRelations 25 | open import ListStuff 26 | open import Pfin.AsSetQuot 27 | open BinaryRelation 28 | open import Worrell.Limits 29 | 30 | -- A counterexample to the surjectivity of the PfinQ-algebra function 31 | -- mQ : PfinQ VωQ → VωQ. This also shows that m : Pfin Vω → Vω is not 32 | -- surjective. 33 | 34 | -- a sequence of finite trees in which top-level branching is strictly 35 | -- increasing 36 | growing : ∀ n → iPfinQ n 37 | growing zero = tt 38 | growing (suc zero) = [ tt ∷ [] ] 39 | growing (suc (suc n)) = [ [] ] ∷Q mapPfinQ (λ x → [ x ∷ [] ]) (growing (suc n)) 40 | 41 | -- an alternative but equivalent definition growing', for which is 42 | -- easier to prove that it is in fact an element of VωQ 43 | growing'-aux : ∀ n → iPfinQ (suc n) → iPfinQ (suc (suc n)) 44 | growing'-aux n = 45 | recQ squash/ 46 | (λ xs → [ [ [] ] ∷ mapList (λ x → [ x ∷ [] ]) xs ]) 47 | (λ xs ys r → eq/ _ _ 48 | ((λ { .([ [] ]) here → ∣ _ , here , refl ∣ ; 49 | x (there m) → ∥map∥ (λ { (y , m' , eq) → _ , there (∈mapList m') , sym (pre∈mapList m .snd .snd) ∙ cong (λ z → [ z ∷ [] ]) eq }) (r .fst _ (pre∈mapList m .snd .fst)) }) , 50 | λ { .([ [] ]) here → ∣ _ , here , refl ∣ ; 51 | x (there m) → ∥map∥ (λ { (y , m' , eq) → _ , there (∈mapList m') , sym (pre∈mapList m .snd .snd) ∙ cong (λ z → [ z ∷ [] ]) eq }) (r .snd _ (pre∈mapList m .snd .fst)) })) 52 | 53 | 54 | growing' : ∀ n → iPfinQ n 55 | growing' zero = tt 56 | growing' (suc zero) = [ tt ∷ [] ] 57 | growing' (suc (suc n)) = growing'-aux n (growing' (suc n)) 58 | 59 | growing≡growing' : ∀ n → growing n ≡ growing' n 60 | growing≡growing' zero = refl 61 | growing≡growing' (suc zero) = refl 62 | growing≡growing' (suc (suc n)) = 63 | elimQProp {B = λ x → ([ [] ] ∷Q mapPfinQ (λ x → [ x ∷ [] ]) x) ≡ growing'-aux n x} 64 | (λ _ → squash/ _ _) 65 | (λ _ → refl) 66 | (growing (suc n)) 67 | ∙ cong (growing'-aux n) (growing≡growing' (suc n)) 68 | 69 | growingEq' : ∀ n → iMapPfinQ n (growing' (suc n)) ≡ growing' n 70 | growingEq' zero = refl 71 | growingEq' (suc zero) = eq/ _ _ 72 | ((λ { _ here → ∣ tt , here , refl ∣ ; _ (there m) → ∣ tt , m , refl ∣ }) , 73 | λ { _ here → ∣ tt , here , refl ∣ }) 74 | growingEq' (suc (suc n)) = 75 | elimQProp {B = λ x → mapPfinQ (mapPfinQ (iMapPfinQ n)) ((growing'-aux (suc n) (growing'-aux n x))) ≡ growing'-aux n (mapPfinQ (iMapPfinQ n) (growing'-aux n x))} 76 | (λ _ → squash/ _ _) 77 | (λ xs → 78 | cong [_] (cong (λ z → [ [] ] ∷ [ iMapPfinQ n [ [] ] ∷ [] ] ∷ z) (mapListComp (mapList (λ x → [ x ∷ [] ]) xs))) 79 | ∙ cong [_] (cong (λ z → [ [] ] ∷ [ iMapPfinQ n [ [] ] ∷ [] ] ∷ z) (mapListComp xs)) 80 | ∙ cong [_] (cong (λ z → [ [] ] ∷ [ iMapPfinQ n [ [] ] ∷ [] ] ∷ z) (sym (mapListComp xs))) 81 | ∙ cong [_] (cong (λ z → [ [] ] ∷ [ iMapPfinQ n [ [] ] ∷ [] ] ∷ z) (sym (mapListComp (mapList (λ x → [ x ∷ [] ]) xs))))) 82 | (growing' (suc n)) 83 | ∙ cong (growing'-aux n) (growingEq' (suc n)) 84 | 85 | growingEq : ∀ n → iMapPfinQ n (growing (suc n)) ≡ growing n 86 | growingEq n = 87 | cong (iMapPfinQ n) (growing≡growing' (suc n)) 88 | ∙ growingEq' n 89 | ∙ sym (growing≡growing' n) 90 | 91 | -- growing is an element of VωQ 92 | growing-ch : VωQ 93 | growing-ch = growing , growingEq 94 | 95 | -- the growing sequence is in fact growing: 96 | -- the size of (growing (suc n)) is (suc n) 97 | sizeGrowing : ∀ n → size (iPfinQDecEq _) (growing (suc n)) ≡ suc n 98 | sizeGrowing zero = refl 99 | sizeGrowing (suc n) = 100 | size∷Q (iPfinQDecEq _) (mapPfinQ (λ x → [ x ∷ [] ]) (growing (suc n))) 101 | (λ p → ∥rec∥ isProp⊥ 102 | (λ { (s , m , eq) → ∥rec∥ isProp⊥ 103 | (λ { (_ , () , _) }) 104 | (effective isPropSameEls isEquivRelSameEls _ _ eq .fst _ here) }) 105 | (pre∈mapPfinQ (growing (suc n)) p)) 106 | ∙ cong suc (sizeMapPfinQ (iPfinQDecEq _) (iPfinQDecEq _) (inj[x∷[]] n) (growing (suc n))) 107 | ∙ cong suc (sizeGrowing n) 108 | where 109 | inj[x∷[]] : ∀ n (x y : iPfinQ n) → _≡_ {A = PfinQ (iPfinQ n)} [ x ∷ [] ] [ y ∷ [] ] → x ≡ y 110 | inj[x∷[]] zero _ _ _ = refl 111 | inj[x∷[]] (suc n) = elimQProp2 (λ _ _ → isPropΠ (λ _ → squash/ _ _)) 112 | λ xs ys r → eq/ _ _ 113 | (∥rec∥ (isPropSameEls _ _) 114 | (λ { (.([ ys ]) , here , eq) → effective isPropSameEls isEquivRelSameEls _ _ eq }) 115 | (effective isPropSameEls isEquivRelSameEls _ _ r .fst _ here)) 116 | 117 | -- growing is not in the image of mQ: 118 | -- there is no element s : PfinQ VωQ such that (mQ s ≡ growing-ch). 119 | -- Therefore mQ is not surjective 120 | noSurjAlg' : (xs : List VωQ) → mQ [ xs ] ≡ growing-ch → ⊥ 121 | noSurjAlg' xs eq = ¬m