├── FiniteSets ├── Kuratowski.agda ├── Kuratowski │ ├── Base.agda │ ├── Decidable.agda │ ├── Decidable │ │ ├── Base.agda │ │ └── Properties.agda │ └── Properties.agda ├── List.agda ├── List │ ├── Base.agda │ ├── Decidable │ │ ├── Base.agda │ │ └── Properties.agda │ └── Properties.agda └── Semilattice.agda ├── README.md └── finiteSets.agda-lib /FiniteSets/Kuratowski.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.Kuratowski where 4 | 5 | open import FiniteSets.Kuratowski.Base public 6 | open import FiniteSets.Kuratowski.Properties public 7 | -------------------------------------------------------------------------------- /FiniteSets/Kuratowski/Base.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.Kuratowski.Base where 4 | 5 | open import Cubical.Foundations.Prelude 6 | import Cubical.Foundations.Logic as L 7 | 8 | open import Cubical.Foundations.HLevels 9 | 10 | private 11 | variable 12 | ℓ : Level 13 | A B C : Set ℓ 14 | P : A → Set ℓ 15 | x y : A 16 | p q : x ≡ y 17 | 18 | infix 7 _∈′_ 19 | infix 7 _∈_ 20 | infix 7 _∉_ 21 | --infix 6 _∉′_ 22 | infix 7 _⊆_ 23 | 24 | data K (A : Set ℓ) : Set ℓ where 25 | ∅ : K A 26 | [_] : A → K A 27 | _∪_ : K A → K A → K A 28 | nl : ∀ x → ∅ ∪ x ≡ x 29 | nr : ∀ x → x ∪ ∅ ≡ x 30 | idem : ∀ a → [ a ] ∪ [ a ] ≡ [ a ] 31 | assoc : ∀ x y z → x ∪ (y ∪ z) ≡ (x ∪ y) ∪ z 32 | com : ∀ x y → x ∪ y ≡ y ∪ x 33 | trunc : isSet (K A) 34 | infixr 10 _∪_ 35 | 36 | elimK : (PSet : {x : K A} → isSet (P x)) 37 | → (z : P ∅) 38 | → (s : (a : A) → P [ a ]) 39 | → (f : (x y : K A) → P x → P y → P (x ∪ y)) 40 | → (∀ x px → PathP (λ i → P (nl x i)) (f ∅ x z px) px) 41 | → (∀ x px → PathP (λ i → P (nr x i)) (f x ∅ px z) px) 42 | → (∀ a → PathP (λ i → P (idem a i)) (f [ a ] [ a ] (s a) (s a)) (s a)) 43 | → (∀ x y z Px Py Pz → 44 | PathP (λ i → P (assoc x y z i)) 45 | (f x (y ∪ z) Px (f y z Py Pz)) (f (x ∪ y) z (f x y Px Py) Pz)) 46 | → (∀ x y Px Py → PathP (λ i → P (com x y i)) (f x y Px Py) (f y x Py Px)) 47 | → (x : K A) → P x 48 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ ∅ = z 49 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ [ x ] = s x 50 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (x ∪ y) = f x y (g x) (g y) 51 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 52 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (nl x i) = nlᴾ x (g x) i 53 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 54 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (nr x i) = nrᴾ x (g x) i 55 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 56 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (idem a i) = idemᴾ a i 57 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 58 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (assoc x y k i) = assocᴾ x y k (g x) (g y) (g k) i 59 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 60 | elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (com x y i) = comᴾ x y (g x) (g y) i 61 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 62 | elimK {A = A} PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ (trunc x y p q i j) = 63 | isOfHLevel→isOfHLevelDep {n = 2} (λ _ → PSet) (g x) (g y) (cong g p) (cong g q) (trunc x y p q) i j 64 | where g = elimK PSet z s f nlᴾ nrᴾ idemᴾ assocᴾ comᴾ 65 | 66 | elimKprop : (PProp : {x : K A} → isProp (P x)) 67 | → (z : P ∅) 68 | → (s : (a : A) → P [ a ]) 69 | → (f : (x y : K A) → P x → P y → P (x ∪ y)) 70 | → (x : K A) → P x 71 | elimKprop {P = P} PProp z s f = elimK (isProp→isSet PProp) z s f 72 | (λ x px → toPathP (PProp (transp (λ i → P (nl x i)) i0 (f ∅ x z px)) px)) 73 | (λ x px → toPathP (PProp (transp (λ i → P (nr x i)) i0 (f x ∅ px z)) px)) 74 | (λ a → toPathP (PProp (transp (λ i → P (idem a i)) i0 (f [ a ] [ a ] (s a) (s a))) (s a))) 75 | (λ x y z Px Py Pz → toPathP (PProp (transp (λ i → P (assoc x y z i)) i0 76 | (f x (y ∪ z) Px (f y z Py Pz))) 77 | (f (x ∪ y) z (f x y Px Py) Pz))) 78 | λ x y Px Py → toPathP (PProp (transp (λ i → P (com x y i)) i0 (f x y Px Py)) (f y x Py Px)) 79 | 80 | recK : isSet B 81 | → (z : B) 82 | → (ins : A → B) 83 | → (f : B → B → B) 84 | → (∀ x → f z x ≡ x) 85 | → (∀ x → f x z ≡ x) 86 | → (∀ a → f (ins a) (ins a) ≡ ins a) 87 | → (∀ x y z → f x (f y z) ≡ f (f x y) z) 88 | → (∀ x y → f x y ≡ f y x) 89 | → K A → B 90 | recK Bset z ins f nlᴮ nrᴮ idemᴮ assocᴮ comᴮ = elimK Bset z ins 91 | (λ _ _ → f) (λ _ → nlᴮ) (λ _ → nrᴮ) idemᴮ (λ _ _ _ → assocᴮ) λ _ _ → comᴮ 92 | 93 | mapK : (A → B) → K A → K B 94 | mapK f = recK trunc ∅ (λ a → [ f a ]) _∪_ nl nr (λ a → idem (f a)) assoc com 95 | -------------------------------------------------------------------------------- 96 | -- 97 | module _ where 98 | open import FiniteSets.Semilattice 99 | 100 | KIsFree : (L : Semilattice ℓ) → (A → Semilattice.A L) → K A → (Semilattice.A L) 101 | KIsFree L f = recK AisSet (Semilattice.⊥ L) f _⊔_ 102 | ⊔-identityˡ ⊔-identityʳ (λ _ → ⊔-idem _) ⊔-assoc ⊔-comm 103 | where open Semilattice L 104 | 105 | KPropRec : (A → L.hProp) → (a : K A) → L.hProp 106 | KPropRec f = KIsFree hProp-Semilattice f 107 | where open hProp 108 | 109 | -------------------------------------------------------------------------------- 110 | -- Properties (Mere propositions) 111 | open L hiding ([_]) 112 | -- A mere identity 113 | infix 8 _≡ₖ_ 114 | 115 | _≡ₖ_ : K A → K A → hProp 116 | x ≡ₖ y = (x ≡ y) , trunc x y 117 | 118 | _≢ₖ_ : K A → K A → hProp 119 | x ≢ₖ y = ¬ (x ≡ₖ y) 120 | 121 | -------------------------------------------------------------------------------- 122 | -- Membership relation 123 | 124 | data _∈′_ {A : Set ℓ} (a : A) : (x : K A) → Set ℓ where 125 | here : a ∈′ [ a ] 126 | left : ∀ {x y} → (a∈x : a ∈′ x) → a ∈′ x ∪ y 127 | right : ∀ {x y} → (a∈y : a ∈′ y) → a ∈′ x ∪ y 128 | sq : ∀ {x} → isProp (a ∈′ x) 129 | 130 | _∈_ : (a : A) → K A → hProp 131 | a ∈ x = a ∈′ x , sq 132 | 133 | _∉_ : (a : A) → K A → hProp 134 | a ∉ x = ¬ (a ∈ x) 135 | 136 | _⊆_ : K A → K A → hProp 137 | x ⊆ y = ∀[ a ] a ∈ x ⇒ a ∈ y 138 | 139 | elim∈prop : {a : A} → ∀ {P : ∀ {x} → a ∈′ x → Set ℓ} 140 | → (PProp : ∀ {x} {p : a ∈′ x} → isProp (P p)) 141 | → (P here) 142 | → (∀ x y → (a∈′x : a ∈′ x) → P a∈′x → P (left {y = y} a∈′x)) 143 | → (∀ x y → (a∈′y : a ∈′ y) → P a∈′y → P (right {x = x} a∈′y)) 144 | → ∀ x (a∈′x : a ∈′ x) → P a∈′x 145 | elim∈prop PProp hereᴾ leftᴾ rightᴾ ._ here = hereᴾ 146 | elim∈prop PProp hereᴾ leftᴾ rightᴾ .(_ ∪ _) (left {x} {y} a∈x) = 147 | leftᴾ x y a∈x (g x a∈x) 148 | where g = elim∈prop PProp hereᴾ leftᴾ rightᴾ 149 | elim∈prop PProp hereᴾ leftᴾ rightᴾ .(_ ∪ _) (right {x} {y} a∈y) = 150 | rightᴾ x y a∈y (g y a∈y) 151 | where g = elim∈prop PProp hereᴾ leftᴾ rightᴾ 152 | elim∈prop {P = P} PProp hereᴾ leftᴾ rightᴾ x (sq a∈x a∈x′ i) = toPathP {A = λ i → P (sq a∈x a∈x′ i)} 153 | (PProp (transp (λ i → P (sq a∈x a∈x′ i)) i0 (g x a∈x)) (g x a∈x′)) i 154 | where g = elim∈prop PProp hereᴾ leftᴾ rightᴾ 155 | 156 | -- An alternative definition `a ∈ x` of membership relation defined by 157 | -- recursion on the finite set x. 158 | _∈ₚ_ : {A : Set (ℓ-zero)} (a : A) → K A → hProp {ℓ-zero} 159 | a ∈ₚ x = KPropRec (λ c → a ≡ₚ c) x 160 | -------------------------------------------------------------------------------- /FiniteSets/Kuratowski/Decidable.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | 4 | module FiniteSets.Kuratowski.Decidable where 5 | 6 | open import FiniteSets.Kuratowski.Decidable.Base public 7 | open import FiniteSets.Kuratowski.Decidable.Properties public 8 | 9 | -------------------------------------------------------------------------------- /FiniteSets/Kuratowski/Decidable/Base.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --safe #-} 2 | 3 | open import Cubical.Foundations.Logic hiding (⊔-idem) 4 | 5 | module FiniteSets.Kuratowski.Decidable.Base where 6 | 7 | open import Cubical.Foundations.Prelude 8 | open import Cubical.HITs.PropositionalTruncation 9 | open import Cubical.Data.Sum 10 | open import Cubical.Relation.Nullary 11 | 12 | open import FiniteSets.Kuratowski renaming ([_] to K[_]) 13 | open import FiniteSets.Kuratowski.Properties 14 | 15 | open import FiniteSets.Semilattice 16 | 17 | module _ {A : Set} 18 | (_≟_ : [ ∀[ x ∶ A ] ∀[ y ∶ A ] Decₚ (x ≡ₚ y) ]) where 19 | 20 | _∈?_ : [ ∀[ a ] ∀[ x ∶ K A ] Decₚ (a ∈ x) ] 21 | _∈?_ a = elimKprop {P = λ y → [ Decₚ (a ∈ y) ]} (isPropDec sq) (no (a∉∅ a)) f g 22 | where 23 | f : [ ∀[ b ] Decₚ (a ∈ K[ b ]) ] 24 | f b with a ≟ b 25 | ... | yes a≡b = yes (substₚ (λ b → a ∈ K[ b ]) a≡b here) 26 | ... | no ¬a≡b = no λ a∈[b] → ¬a≡b (∈⇒∈ₚ a∈[b]) 27 | 28 | g : [ ∀[ x ] ∀[ y ] Decₚ (a ∈ x) ⇒ Decₚ (a ∈ y) ⇒ Decₚ (a ∈ x ∪ y) ] 29 | g x y (yes p) (yes q) = yes (left p) 30 | g x y (yes p) (no ¬q) = yes (left p) 31 | g x y (no ¬p) (yes q) = yes (right q) 32 | g x y (no ¬p) (no ¬q) = 33 | no λ a∈x∪y → elimPropTrunc (λ _ → λ ()) 34 | (elim-⊎ (λ a∈ₚx → ¬p (∈ₚ⇒∈ a∈ₚx)) λ a∈ₚy → ¬q (∈ₚ⇒∈ a∈ₚy)) (∈⇒∈ₚ a∈x∪y) 35 | 36 | delete : A → K A → K A 37 | delete x = recK trunc ∅ eq? _∪_ nl nr (λ a → ⊔-idem (eq? a)) assoc com 38 | where 39 | open Semilattice (KSemilattice A) hiding (A) 40 | 41 | eq? : A → K A 42 | eq? y with x ≟ y 43 | ... | yes _ = ∅ 44 | ... | no _ = K[ y ] 45 | 46 | module _ {A : Set} 47 | (_∈?_ : [ ∀[ a ∶ A ] ∀[ x ] Decₚ (a ∈ x) ]) where 48 | _≟_ : [ ∀[ x ∶ A ] ∀[ y ∶ A ] Decₚ (x ≡ₚ y) ] 49 | a ≟ b with a ∈? K[ b ] 50 | ... | yes p = yes (∈⇒∈ₚ p) 51 | ... | no ¬p = no λ a≡b → ¬p (∈ₚ⇒∈ a≡b) 52 | -------------------------------------------------------------------------------- /FiniteSets/Kuratowski/Decidable/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.Kuratowski.Decidable.Properties where 4 | 5 | open import Cubical.Foundations.Prelude 6 | 7 | open import FiniteSets.Kuratowski 8 | open import FiniteSets.Kuratowski.Decidable.Base 9 | -------------------------------------------------------------------------------- /FiniteSets/Kuratowski/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --safe #-} 2 | 3 | module FiniteSets.Kuratowski.Properties where 4 | 5 | open import Cubical.Foundations.Prelude 6 | 7 | open import Cubical.Foundations.Logic as L hiding (⊔-idem) 8 | open import Cubical.Foundations.Function 9 | open import Cubical.Foundations.Isomorphism 10 | 11 | open import Cubical.HITs.SetTruncation 12 | open import Cubical.HITs.PropositionalTruncation 13 | 14 | open import Cubical.Data.Prod 15 | open import Cubical.Data.Sum 16 | 17 | open import FiniteSets.Kuratowski.Base renaming ([_] to K[_]) 18 | open import FiniteSets.Semilattice 19 | 20 | private 21 | variable 22 | ℓ : Level 23 | -------------------------------------------------------------------------------- 24 | -- K A is a ∪-semilattice. 25 | 26 | module Lattice where 27 | variable 28 | A : Set ℓ 29 | ∪-idem : (x : K A) → x ∪ x ≡ x 30 | ∪-idem = elimKprop (trunc _ _) (nl ∅) idem λ x y x∪x=x y∪y=y → 31 | (x ∪ y) ∪ x ∪ y ≡⟨ assoc (x ∪ y) x y ⟩ 32 | ((x ∪ y) ∪ x) ∪ y ≡⟨ cong (_∪ y) (cong (_∪ x) (com x y)) ⟩ 33 | ((y ∪ x) ∪ x) ∪ y ≡⟨ cong (_∪ y) (sym (assoc y x x)) ⟩ 34 | (y ∪ x ∪ x) ∪ y ≡⟨ cong (_∪ y) (cong (y ∪_) x∪x=x) ⟩ 35 | (y ∪ x) ∪ y ≡⟨ cong (_∪ y) (com y x) ⟩ 36 | (x ∪ y) ∪ y ≡⟨ sym (assoc x y y) ⟩ 37 | x ∪ y ∪ y ≡⟨ cong (x ∪_) y∪y=y ⟩ 38 | x ∪ y ∎ 39 | 40 | KSemilattice : Set ℓ → Semilattice ℓ 41 | KSemilattice A = record 42 | { A = K A 43 | ; _⊔_ = _∪_ 44 | ; ⊥ = ∅ 45 | ; isSemilattice = record 46 | { AisSet = trunc 47 | ; ⊔-identityˡ = nl 48 | ; ⊔-identityʳ = nr 49 | ; ⊔-idem = ∪-idem 50 | ; ⊔-assoc = assoc 51 | ; ⊔-comm = com } 52 | } 53 | open Lattice using (KSemilattice) public 54 | 55 | module _ {A : Set}where 56 | open Properties (KSemilattice A) 57 | 58 | -- Two definitions of membership relations are the same. 59 | ∈⇒∈ₚ : ∀ {a : A} {x} → [ a ∈ x ⇒ a ∈ₚ x ] 60 | ∈⇒∈ₚ {a = a} {x = x} a∈x = elim∈prop {P = λ {y} _ → [ a ∈ₚ y ]} 61 | (λ {y} → snd (a ∈ₚ y)) ∣ refl ∣ 62 | (λ _ _ _ → L.inl) (λ _ _ _ → L.inr) x a∈x 63 | 64 | ∈ₚ⇒∈ : ∀ {a} {x : K A} → [ a ∈ₚ x ⇒ a ∈ x ] 65 | ∈ₚ⇒∈ {a} {x} a∈x = elimKprop {P = λ y → [ (a ∈ₚ y) ⇒ a ∈ y ]} (propPi λ _ → snd (a ∈ _)) 66 | (λ ()) (λ b a≡b → substₚ (λ c → a ∈ K[ c ]) a≡b here) 67 | (λ y z py pz pyz → elimPropTrunc (λ _ → sq) 68 | (elim-⊎ (λ a∈y → left (py a∈y)) λ a∈z → right (pz a∈z)) pyz) x a∈x 69 | 70 | ∈≡∈ₚ : ∀ {a : A}{x : K A} → a ∈ x ≡ a ∈ₚ x 71 | ∈≡∈ₚ = ⇔toPath ∈⇒∈ₚ ∈ₚ⇒∈ 72 | 73 | [a]≢∅ : {a : A} → [ K.[ a ] ≢ₖ ∅ ] 74 | [a]≢∅ p = subst (fst ∘ KPropRec (λ a → (K.[ a ] ≡ ∅) , trunc _ _)) p p 75 | 76 | -- not used 77 | []-injective : {A : Set} {a b : A} → K.[ a ] ≡ₖ K.[ b ] ≡ a ≡ₚ b 78 | []-injective {a = a} {b} = 79 | ⇒∶ (λ eq → subst (fst ∘ a ∈ₚ_) eq ∣ refl ∣) 80 | ⇐∶ elimPropTrunc (λ _ → trunc K[ a ] K[ b ]) (cong K[_]) 81 | -------------------------------------------------------------------------------- 82 | -- a ∉ ∅ 83 | 84 | a∉∅ : (a : A) → [ a ∉ ∅ ] 85 | a∉∅ a = lem refl 86 | where 87 | lem : ∀ {a x} → x ≡ ∅ → [ a ∉ x ] 88 | lem {a} {x} x≡∅ a∈x = elim∈prop {P = λ {y} b∈y → y ≡ ∅ → [ a ∉ y ]} 89 | (propPi (λ _ → snd (a ∉ _))) (λ [a]≡∅ _ → [a]≢∅ [a]≡∅) 90 | (λ _ _ a∈y h y∪z≡∅ _ → h (⊔-conicalˡ _ _ y∪z≡∅) a∈y) 91 | (λ _ _ a∈z h y∪z≡∅ _ → h (⊔-conicalʳ _ _ y∪z≡∅) a∈z) x a∈x x≡∅ a∈x 92 | 93 | a∈x⇒[a]∪x≡x : ∀ (a : A) x → [ a ∈ x ] → K.[ a ] ∪ x ≡ x 94 | a∈x⇒[a]∪x≡x a = elim∈prop (trunc _ _) (idem a) 95 | (λ x y a∈x a∪x≡x → 96 | K.[ a ] ∪ x ∪ y ≡⟨ assoc _ _ _ ⟩ 97 | (K.[ a ] ∪ x) ∪ y ≡⟨ cong (_∪ y) a∪x≡x ⟩ 98 | x ∪ y ∎) 99 | (λ x y a∈y a∪y≡y → 100 | K.[ a ] ∪ x ∪ y ≡⟨ cong (K.[ a ] ∪_) (com _ _) ⟩ 101 | K.[ a ] ∪ y ∪ x ≡⟨ assoc _ _ _ ⟩ 102 | (K.[ a ] ∪ y) ∪ x ≡⟨ cong (_∪ x) a∪y≡y ⟩ 103 | y ∪ x ≡⟨ com _ _ ⟩ 104 | x ∪ y ∎) 105 | 106 | y⊆x⇒y∪x≡x : ∀ {x y : K A} → (y ⊆ x) ≡ (y ∪ x) ≡ₖ x 107 | y⊆x⇒y∪x≡x {x = x} {y} = 108 | ⇒∶ elimKprop {P = λ z → [ z ⊆ x ⇒ (z ∪ x ≡ₖ x) ]} 109 | (λ p q → funExt λ f → trunc _ _ (p f ) (q f)) 110 | (λ _ → nl _) 111 | (λ a p → a∈x⇒[a]∪x≡x a x (p a here)) 112 | (λ u v pu pv f → 113 | (u ∪ v) ∪ x ≡⟨ sym (assoc _ _ _) ⟩ 114 | u ∪ (v ∪ x) ≡⟨ cong (u ∪_) (pv (λ x z → f x (right z))) ⟩ 115 | u ∪ x ≡⟨ pu (λ x z → f x (left z)) ⟩ 116 | x ∎) y 117 | ⇐∶ λ y∪x≡x a a∈y → subst (λ z → [ a ∈ z ]) y∪x≡x (left a∈y) 118 | 119 | y∪x≡x∧x∪y≡y : {x y : K A} → (y ∪ x ≡ₖ x) ⊓ (x ∪ y ≡ₖ y) ≡ x ≡ₖ y 120 | y∪x≡x∧x∪y≡y {x} {y} = 121 | ⇒∶ (λ { (y≤x , x≤y) → ≤-antisym x≤y y≤x }) 122 | ⇐∶ λ x≡y → sym (cong (_∪ x) x≡y) ∙ ⊔-idem _ , cong (_∪ y) x≡y ∙ ⊔-idem _ 123 | where open Semilattice (KSemilattice A) 124 | 125 | setExt : {x y : K A} → (x ≡ₖ y) ≡ (∀[ a ∶ A ] a ∈ y ⇔ a ∈ x) 126 | setExt {x} {y} = 127 | x ≡ₖ y 128 | ≡⟨ sym (y∪x≡x∧x∪y≡y) ⟩ 129 | (y ∪ x ≡ₖ x) ⊓ (x ∪ y ≡ₖ y) 130 | ≡⟨ cong₂ (_⊓_) (sym y⊆x⇒y∪x≡x) (sym y⊆x⇒y∪x≡x) ⟩ 131 | (y ⊆ x) ⊓ (x ⊆ y) 132 | ≡⟨ cong₂ {y = ∀[ a ∶ A ] a ∈ y ⇒ a ∈ x} _⊓_ refl {v = ∀[ a ∶ A ] a ∈ x ⇒ a ∈ y} refl ⟩ 133 | (∀[ a ∶ A ] a ∈ y ⇒ a ∈ x) ⊓ (∀[ a ∶ A ] a ∈ x ⇒ a ∈ y) 134 | ≡⟨ ⊓-∀-distrib (λ a → a ∈ y ⇒ a ∈ x) (λ a → a ∈ x ⇒ a ∈ y) ⟩ 135 | (∀[ a ∶ A ] (a ∈ y ⇔ a ∈ x)) ∎ 136 | 137 | -------------------------------------------------------------------------------- /FiniteSets/List.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.List where 4 | 5 | open import FiniteSets.List.Base public 6 | -------------------------------------------------------------------------------- /FiniteSets/List/Base.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.List.Base where 4 | 5 | open import Cubical.Foundations.Prelude 6 | open import Cubical.Foundations.HLevels 7 | 8 | infixr 8 _++_ 9 | infixr 10 _∷_ 10 | 11 | private 12 | variable 13 | ℓ : Level 14 | A B : Set ℓ 15 | 16 | data L (A : Set ℓ) : Set ℓ where 17 | [] : L A 18 | _∷_ : A → L A → L A 19 | dup : ∀ a xs → a ∷ a ∷ xs ≡ a ∷ xs 20 | com : ∀ a b xs → a ∷ b ∷ xs ≡ b ∷ a ∷ xs 21 | trunc : isSet (L A) 22 | 23 | pattern [_] a = a ∷ [] 24 | 25 | elimL : ∀ {P : L A → Set ℓ} 26 | → (PSet : {xs : L A} → isSet (P xs)) 27 | → (z : P []) 28 | → (f : ∀ a xs → P xs → P (a ∷ xs)) 29 | → (∀ a xs pxs → PathP (λ i → P (dup a xs i)) 30 | (f a (a ∷ xs) (f a xs pxs)) (f a xs pxs)) 31 | → (∀ a b xs pxs → PathP (λ i → P (com a b xs i)) 32 | (f a (b ∷ xs) (f b xs pxs)) (f b (a ∷ xs) (f a xs pxs))) 33 | → (xs : L A) → P xs 34 | elimL PSet z f fdup fcom [] = z 35 | elimL PSet z f fdup fcom (x ∷ xs) = f x xs (elimL PSet z f fdup fcom xs) 36 | elimL {P = P} PSet z f fdup fcom (dup a xs i) = fdup a xs (elimL PSet z f fdup fcom xs) i 37 | elimL PSet z f fdup fcom (com a b xs i) = fcom a b xs (elimL PSet z f fdup fcom xs) i 38 | elimL {A = A} PSet z f fdup fcom (trunc xs ys p q i j) = 39 | isOfHLevel→isOfHLevelDep {n = 2} (\_ → PSet) (g xs) (g ys) (cong g p) (cong g q) (trunc xs ys p q) i j 40 | where g = elimL PSet z f fdup fcom 41 | 42 | elimLprop : ∀ {P : L A → Set ℓ} 43 | → (PSet : {xs : L A} → isProp (P xs)) 44 | → (z : P []) 45 | → (f : ∀ a xs → P xs → P (a ∷ xs)) 46 | → (xs : L A) → P xs 47 | elimLprop {P = P} Pprop z f = elimL (isProp→isSet Pprop) z f 48 | (λ a xs pxs → toPathP (Pprop (transp (λ i → P (dup a xs i)) i0 (f a (a ∷ xs) (f a xs pxs))) (f a xs pxs))) 49 | λ a b xs pxs → toPathP (Pprop (transp (λ i → P (com a b xs i)) i0 50 | (f a (b ∷ xs) (f b xs pxs))) (f b (a ∷ xs) (f a xs pxs))) 51 | 52 | recL : isSet B 53 | → (z : B) 54 | → (f : A → B → B) 55 | → (∀ a p → f a (f a p) ≡ f a p) 56 | → (∀ a b p → f a (f b p) ≡ f b (f a p)) 57 | → L A → B 58 | recL BSet z f dupᴮ comᴮ = 59 | elimL BSet z (λ a _ b → f a b) 60 | (λ a _ pxs → dupᴮ a pxs) (λ a b _ pb → comᴮ a b pb) 61 | 62 | _++_ : L A → L A → L A 63 | _++_ xs ys = recL trunc ys (λ x xs++ys → x ∷ xs++ys) dup com xs 64 | -------------------------------------------------------------------------------- /FiniteSets/List/Decidable/Base.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical --safe #-} 2 | 3 | module FiniteSets.List.Decidable.Base where 4 | 5 | open import Cubical.Foundations.Prelude 6 | open import Cubical.Foundations.Logic 7 | open import Cubical.Relation.Nullary hiding (¬_) 8 | 9 | open import FiniteSets.List renaming ([_] to L[_]) 10 | open import FiniteSets.List.Properties 11 | import FiniteSets.Kuratowski as K 12 | 13 | module _ {A : Set} (_≟_ : [ ∀[ x ∶ A ] ∀[ y ∶ A ] Decₚ (x ≡ₚ y) ]) where 14 | delete : A → L A → L A 15 | delete x = recL trunc [] eq? dupEq? comEq? 16 | where 17 | eq? : A → L A → L A 18 | eq? y xs with x ≟ y 19 | ... | yes p = xs 20 | ... | no ¬p = y ∷ xs 21 | 22 | dupEq? : ∀ y xs → eq? y (eq? y xs) ≡ eq? y xs 23 | dupEq? y xs with x ≟ y 24 | ... | yes p = refl 25 | ... | no ¬p = dup y xs 26 | 27 | comEq? : ∀ y z xs → eq? y (eq? z xs) ≡ eq? z (eq? y xs) 28 | comEq? y z xs with x ≟ y | x ≟ z 29 | comEq? y z xs | yes p | yes q = refl 30 | comEq? y z xs | yes p | no ¬q = refl 31 | comEq? y z xs | no ¬p | yes q = refl 32 | comEq? y z xs | no ¬p | no ¬q = com y z xs 33 | -------------------------------------------------------------------------------- /FiniteSets/List/Decidable/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.List.Decidable.Properties where 4 | 5 | open import Cubical.Foundations.Prelude 6 | open import Cubical.Data.Sum 7 | open import Cubical.Data.Empty 8 | open import Cubical.Relation.Nullary 9 | 10 | open import FiniteSets.List 11 | open import FiniteSets.List.Decidable.Base 12 | -------------------------------------------------------------------------------- /FiniteSets/List/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.List.Properties where 4 | 5 | open import Cubical.Foundations.Prelude 6 | 7 | open import Cubical.Foundations.HLevels 8 | open import Cubical.Foundations.Isomorphism 9 | open import Cubical.Foundations.Function 10 | 11 | open import FiniteSets.Semilattice 12 | open import FiniteSets.List as L renaming ([_] to L[_]) 13 | open import FiniteSets.Kuratowski as K 14 | renaming ( [_] to K[_] 15 | ; trunc to trunck 16 | ; com to comk 17 | ; _∈_ to _K∈_ 18 | ; _∈′_ to _K∈′_) 19 | hiding (_∉_; _⊆_) 20 | 21 | private 22 | variable 23 | ℓ : Level 24 | A : Set ℓ 25 | 26 | infix 7 _∈′_ 27 | infix 7 _∈_ 28 | infix 7 _∉_ 29 | infix 7 _⊆_ 30 | 31 | ++-identityʳ : (xs : L A) → xs ++ [] ≡ xs 32 | ++-identityʳ = elimLprop (trunc _ _) refl (λ a _ p → (cong (a ∷_) p)) 33 | 34 | ++-assoc : (ys zs xs : L A) 35 | → xs ++ ys ++ zs ≡ (xs ++ ys) ++ zs 36 | ++-assoc ys zs = elimLprop (trunc _ _) refl λ a xs p → cong (a ∷_) p 37 | 38 | [a]++xs≡xs++[a] : ∀ xs (a : A) → a ∷ xs ≡ xs ++ L[ a ] 39 | [a]++xs≡xs++[a] = elimLprop (propPi (λ _ → trunc _ _)) (λ _ → refl) λ a xs pxs b → 40 | b ∷ a ∷ xs ≡⟨ com _ _ _ ⟩ 41 | a ∷ b ∷ xs ≡⟨ cong (a ∷_) (pxs b) ⟩ 42 | a ∷ (xs ++ L[ b ]) ≡⟨ refl ⟩ 43 | a ∷ xs ++ L[ b ] ∎ 44 | 45 | ++-comm : ∀ xs → (ys : L A) → xs ++ ys ≡ ys ++ xs 46 | ++-comm = elimLprop (propPi (λ _ → trunc _ _)) (λ ys → sym (++-identityʳ ys)) λ a xs pxs ys → 47 | a ∷ (xs ++ ys) ≡⟨ cong (a ∷_) (pxs ys) ⟩ 48 | a ∷ (ys ++ xs) ≡⟨ [a]++xs≡xs++[a] (ys ++ xs) a ⟩ 49 | (ys ++ xs) ++ L[ a ] ≡⟨ sym (++-assoc _ _ ys) ⟩ 50 | ys ++ xs ++ L[ a ] ≡⟨ cong (ys ++_) (sym ([a]++xs≡xs++[a] _ _)) ⟩ 51 | ys ++ a ∷ xs ∎ 52 | 53 | ++-idem : (xs : L A) → xs ++ xs ≡ xs 54 | ++-idem = elimLprop (trunc _ _) refl λ a xs pxs → 55 | (a ∷ xs) ++ a ∷ xs ≡⟨ refl ⟩ 56 | a ∷ (xs ++ a ∷ xs) ≡⟨ cong (a ∷_) (++-comm xs (a ∷ xs)) ⟩ 57 | a ∷ a ∷ (xs ++ xs) ≡⟨ dup a _ ⟩ 58 | a ∷ xs ++ xs ≡⟨ cong (a ∷_) pxs ⟩ 59 | a ∷ xs ∎ 60 | 61 | L-Semilattice : Set ℓ → Semilattice ℓ 62 | L-Semilattice A = record 63 | { A = L A 64 | ; _⊔_ = _++_ 65 | ; ⊥ = [] 66 | ; isSemilattice = record 67 | { AisSet = trunc 68 | ; ⊔-identityˡ = λ _ → refl 69 | ; ⊔-identityʳ = ++-identityʳ 70 | ; ⊔-idem = ++-idem 71 | ; ⊔-assoc = λ xs ys zs → ++-assoc ys zs xs 72 | ; ⊔-comm = ++-comm 73 | } } 74 | 75 | KtoL : K A → L A 76 | KtoL {A = A} = recK trunc [] L[_] _⊔_ ⊔-identityˡ ⊔-identityʳ (⊔-idem ∘ L[_]) ⊔-assoc ⊔-comm 77 | where open Semilattice (L-Semilattice A) 78 | 79 | LtoK : L A → K A 80 | LtoK = recL trunck ∅ (λ a x → K[ a ] ∪ x) 81 | (λ a x → assoc _ _ _ ∙ cong (_∪ x) (idem a)) 82 | (λ a b x → assoc _ _ _ ∙ cong (_∪ x) (comk _ _) ∙ sym (assoc _ _ _)) 83 | 84 | K≡L : K A ≡ L A 85 | K≡L = isoToPath (iso KtoL LtoK f∘g=id g∘f=id) 86 | where 87 | f = KtoL 88 | g = LtoK 89 | 90 | g-homo : ∀ xs ys → g (xs ++ ys) ≡ g xs ∪ g ys 91 | g-homo = elimLprop (propPi λ _ → trunck _ _) (λ ys → sym (nl (g ys))) 92 | λ a xs f ys → 93 | K[ a ] ∪ g (xs ++ ys) ≡⟨ cong (K[ a ] ∪_) (f ys) ⟩ 94 | K[ a ] ∪ (g xs ∪ g ys) ≡⟨ assoc _ _ _ ⟩ 95 | g (a ∷ xs) ∪ g ys ∎ 96 | 97 | f∘g=id : section f g 98 | f∘g=id = elimLprop (trunc _ _) refl λ a xs → cong (a ∷_) 99 | 100 | g∘f=id : retract f g 101 | g∘f=id = elimKprop (trunck _ _) refl (λ _ → nr _) 102 | λ x y g∘fx≡x g∘fy≡y → 103 | g (f (x ∪ y)) ≡⟨ refl ⟩ 104 | g (f x ++ f y) ≡⟨ g-homo (f x) (f y) ⟩ 105 | g (f x) ∪ g (f y) ≡⟨ cong₂ (_∪_) g∘fx≡x g∘fy≡y ⟩ 106 | x ∪ y ∎ 107 | 108 | open import Cubical.Foundations.Logic 109 | 110 | _∈′_ : (a : A) → L A → Set _ 111 | a ∈′ xs = a K∈′ LtoK xs 112 | 113 | _∈_ : (a : A) → L A → hProp 114 | a ∈ xs = a ∈′ xs , sq 115 | 116 | _∉_ : (a : A) → L A → hProp 117 | a ∉ x = ¬ (a ∈ x) 118 | 119 | _⊆_ : L A → L A → hProp 120 | x ⊆ y = ∀[ a ] a ∈ x ⇒ a ∈ y 121 | -------------------------------------------------------------------------------- /FiniteSets/Semilattice.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module FiniteSets.Semilattice where 4 | open import Cubical.Foundations.Prelude 5 | 6 | private 7 | variable 8 | ℓ : Level 9 | 10 | record IsSemilattice (A : Set ℓ) (_⊔_ : A → A → A) (⊥ : A) : Set ℓ where 11 | field 12 | AisSet : isSet A 13 | ⊔-identityˡ : ∀ x → ⊥ ⊔ x ≡ x 14 | ⊔-identityʳ : ∀ x → x ⊔ ⊥ ≡ x 15 | ⊔-idem : ∀ x → x ⊔ x ≡ x 16 | ⊔-assoc : ∀ x y z → x ⊔ (y ⊔ z) ≡ (x ⊔ y) ⊔ z 17 | ⊔-comm : ∀ x y → x ⊔ y ≡ y ⊔ x 18 | 19 | record Semilattice (ℓ : Level) : Set (ℓ-suc ℓ) where 20 | field 21 | A : Set ℓ 22 | _⊔_ : A → A → A 23 | ⊥ : A 24 | isSemilattice : IsSemilattice A _⊔_ ⊥ 25 | open IsSemilattice isSemilattice public 26 | 27 | _≤_ : A → A → Set ℓ 28 | x ≤ y = x ⊔ y ≡ y 29 | infixr 7 _⊔_ 30 | infix 5 _≤_ 31 | 32 | module Bool where 33 | open import Cubical.Data.Bool 34 | 35 | Or-Semilattice : Semilattice ℓ-zero 36 | Or-Semilattice = record 37 | { A = Bool 38 | ; _⊔_ = _or_ 39 | ; ⊥ = false 40 | ; isSemilattice = record 41 | { AisSet = isSetBool 42 | ; ⊔-identityˡ = or-identityˡ 43 | ; ⊔-identityʳ = or-identityʳ 44 | ; ⊔-idem = or-idem 45 | ; ⊔-assoc = or-assoc 46 | ; ⊔-comm = or-comm 47 | } 48 | } 49 | 50 | module hProp where 51 | import Cubical.Foundations.Logic as L 52 | import Cubical.Foundations.HLevels as L 53 | 54 | hProp-Semilattice : Semilattice _ 55 | hProp-Semilattice = record 56 | { A = L.hProp 57 | ; _⊔_ = L._⊔_ 58 | ; ⊥ = L.⊥ 59 | ; isSemilattice = record 60 | { AisSet = L.isSetHProp 61 | ; ⊔-identityˡ = L.⊔-identityˡ 62 | ; ⊔-identityʳ = L.⊔-identityʳ 63 | ; ⊔-idem = L.⊔-idem 64 | ; ⊔-assoc = L.⊔-assoc 65 | ; ⊔-comm = L.⊔-comm 66 | } 67 | } 68 | 69 | module Properties {ℓ} (L : Semilattice ℓ) where 70 | open Semilattice L 71 | 72 | ≤-refl : (x : A) → x ≤ x 73 | ≤-refl = ⊔-idem 74 | 75 | ≤-antisym : {x y : A} → x ≤ y → y ≤ x → x ≡ y 76 | ≤-antisym {x = x} {y} x≤y y≤x = 77 | x ≡⟨ sym y≤x ⟩ 78 | y ⊔ x ≡⟨ ⊔-comm _ _ ⟩ 79 | x ⊔ y ≡⟨ x≤y ⟩ 80 | y ∎ 81 | 82 | ≤-trans : (x y z : A) → x ≤ y → y ≤ z → x ≤ z 83 | ≤-trans x y z x≤y y≤z = 84 | x ⊔ z ≡⟨ cong (x ⊔_) (sym y≤z) ⟩ 85 | x ⊔ y ⊔ z ≡⟨ ⊔-assoc _ _ _ ⟩ 86 | (x ⊔ y) ⊔ z ≡⟨ cong (_⊔ z) x≤y ⟩ 87 | y ⊔ z ≡⟨ y≤z ⟩ 88 | z ∎ 89 | 90 | ⊥-minimum : (x : A) → ⊥ ≤ x 91 | ⊥-minimum = ⊔-identityˡ 92 | 93 | ≤-isAlgOrder : (x y z : A) → x ⊔ z ≡ y → x ≤ y 94 | ≤-isAlgOrder x y z p = 95 | x ⊔ y ≡⟨ cong (x ⊔_) (sym p) ⟩ 96 | x ⊔ x ⊔ z ≡⟨ ⊔-assoc _ _ _ ⟩ 97 | (x ⊔ x) ⊔ z ≡⟨ cong (_⊔ z) (⊔-idem _) ⟩ 98 | x ⊔ z ≡⟨ p ⟩ 99 | y ∎ 100 | -------------------------------------------------------------------------------- 101 | -- (A , ≤) is an order-theoretic lattice 102 | 103 | ⊔-sup₁ : (x y : A) → x ≤ x ⊔ y 104 | ⊔-sup₁ x y = 105 | x ⊔ x ⊔ y ≡⟨ ⊔-assoc _ _ _ ⟩ 106 | (x ⊔ x) ⊔ y ≡⟨ cong (_⊔ y) (⊔-idem _) ⟩ 107 | x ⊔ y ∎ 108 | 109 | ⊔-sup₂ : (x y z : A) → x ≤ z → y ≤ z → x ⊔ y ≤ z 110 | ⊔-sup₂ x y z x≤z y≤z = 111 | (x ⊔ y) ⊔ z ≡⟨ sym (⊔-assoc _ _ _) ⟩ 112 | x ⊔ y ⊔ z ≡⟨ cong (x ⊔_) y≤z ⟩ 113 | x ⊔ z ≡⟨ x≤z ⟩ 114 | z ∎ 115 | 116 | ⊔-∅-injective : (x y : A) → x ⊔ ⊥ ≡ y ⊔ ⊥ → x ≡ y 117 | ⊔-∅-injective x y p = 118 | x ≡⟨ sym (⊔-identityʳ _) ⟩ 119 | x ⊔ ⊥ ≡⟨ p ⟩ 120 | y ⊔ ⊥ ≡⟨ ⊔-identityʳ _ ⟩ 121 | y ∎ 122 | 123 | ⊔-conicalˡ : (x y : A) → x ⊔ y ≡ ⊥ → x ≡ ⊥ 124 | ⊔-conicalˡ x y x⊔y≡⊥ = 125 | ≤-antisym (≤-isAlgOrder _ _ y x⊔y≡⊥) (⊔-identityˡ x) 126 | 127 | ⊔-conicalʳ : (x y : A) → x ⊔ y ≡ ⊥ → y ≡ ⊥ 128 | ⊔-conicalʳ x y x⊔y≡⊥ = ⊔-conicalˡ y x (subst (λ z → z ≡ ⊥) (⊔-comm x y) x⊔y≡⊥) 129 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # FiniteSets 2 | Fintie Sets in Cubical Agda 3 | 4 | An implementation of the type of finite sets in Cubical Agda based on the [Coq formalisation](https://dl.acm.org/citation.cfm?doid=3176245.3167085). 5 | -------------------------------------------------------------------------------- /finiteSets.agda-lib: -------------------------------------------------------------------------------- 1 | name: finite-sets 2 | depend: 3 | cubical 4 | include: . 5 | --------------------------------------------------------------------------------