commit 0d85fceb1f1a6ce5a417d669df7f43ac562d88cd Author: cat Date: Mon Mar 11 20:20:28 2024 +0200 push diff --git a/.agda-lib b/.agda-lib new file mode 100644 index 0000000..84329d8 --- /dev/null +++ b/.agda-lib @@ -0,0 +1,4 @@ +name: agda-symmetries +depend: + . + cubical diff --git a/Cubical/Structures/Arity.agda b/Cubical/Structures/Arity.agda new file mode 100644 index 0000000..6c5a9c0 --- /dev/null +++ b/Cubical/Structures/Arity.agda @@ -0,0 +1,91 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Arity where + +import Cubical.Data.Empty as ⊥ +open import Cubical.Foundations.Everything +open import Cubical.Data.Fin public renaming (Fin to Arity) +open import Cubical.Data.Fin.Recursive public using (rec) +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.List +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + A B : Type ℓ + k : ℕ + +ftwo : Arity (suc (suc (suc k))) +ftwo = fsuc fone + +lookup : ∀ (xs : List A) -> Arity (length xs) -> A +lookup [] num = ⊥.rec (¬Fin0 num) +lookup (x ∷ xs) (zero , prf) = x +lookup (x ∷ xs) (suc num , prf) = lookup xs (num , pred-≤-pred prf) + +tabulate : ∀ n -> (Arity n -> A) -> List A +tabulate zero ^a = [] +tabulate (suc n) ^a = ^a fzero ∷ tabulate n (^a ∘ fsuc) + +length-tabulate : ∀ n -> (^a : Arity n → A) -> length (tabulate n ^a) ≡ n +length-tabulate zero ^a = refl +length-tabulate (suc n) ^a = cong suc (length-tabulate n (^a ∘ fsuc)) + +tabulate-lookup : ∀ (xs : List A) -> tabulate (length xs) (lookup xs) ≡ xs +tabulate-lookup [] = refl +tabulate-lookup (x ∷ xs) = + _ ≡⟨ cong (λ z -> x ∷ tabulate (length xs) z) (funExt λ _ -> cong (lookup xs) (Σ≡Prop (λ _ -> isProp≤) refl)) ⟩ + _ ≡⟨ cong (x ∷_) (tabulate-lookup xs) ⟩ + _ ∎ + +arity-n≡m : ∀ {n m} -> (a : Arity n) -> (b : Arity m) -> (p : n ≡ m) -> a .fst ≡ b .fst -> PathP (λ i -> Arity (p i)) a b +arity-n≡m (v , a) (w , b) p q = ΣPathP (q , toPathP (isProp≤ _ _)) + +lookup-tabulate : ∀ n (^a : Arity n -> A) -> PathP (λ i -> (Arity (length-tabulate n ^a i) -> A)) (lookup (tabulate n ^a)) ^a +lookup-tabulate zero ^a = funExt (⊥.rec ∘ ¬Fin0) +lookup-tabulate (suc n) ^a = toPathP (funExt lemma) i (zero , p) = ^a (0 , toPathP {A = λ j -> 0 < suc (length-tabulate n (^a ∘ fsuc) (~ j))} {! !} i) toPathP (sym (transport-filler _ _) ∙ cong ^a (Σ≡Prop (λ _ -> isProp≤) refl)) i suc-≤-suc (zero-≤ {n = n}) toPathP {x = suc-≤-suc (zero-≤ {n = n})} {! !} i + where + lemma : _ + lemma (zero , p) = sym (transport-filler _ _) ∙ cong ^a (Σ≡Prop (λ _ -> isProp≤) refl) + TODO: Cleanup this mess + lemma (suc w , p) = + _ ≡⟨ sym (transport-filler _ _) ⟩ + _ ≡⟨ congP (λ i f -> f (arity-n≡m (w , + pred-≤-pred + (transp + (λ i → suc w < suc (length-tabulate n (λ x → ^a (fsuc x)) (~ i))) + i0 p)) (w , pred-≤-pred p) (length-tabulate n (^a ∘ fsuc)) refl i)) (lookup-tabulate n (^a ∘ fsuc)) ⟩ + _ ≡⟨ cong ^a (Σ≡Prop (λ _ -> isProp≤) refl) ⟩ + _ ∎ + +lookup2≡i : ∀ (i : Arity 2 -> A) -> lookup (i fzero ∷ i fone ∷ []) ≡ i +lookup2≡i i = funExt lemma + where + lemma : _ + lemma (zero , p) = cong i (Σ≡Prop (λ _ -> isProp≤) refl) + lemma (suc zero , p) = cong i (Σ≡Prop (λ _ -> isProp≤) refl) + lemma (suc (suc n) , p) = ⊥.rec (¬m+n A +◼ = ⊥.rec ∘ ¬Fin0 + +infixr 30 _▸_ + +_▸_ : ∀ {n} -> A -> (Arity n -> A) -> Arity (suc n) -> A +_▸_ {n = zero} x f i = x +_▸_ {n = suc n} x f i = lookup (x ∷ tabulate (suc n) f) (subst Arity (congS (suc ∘ suc) (sym (length-tabulate n (f ∘ fsuc)))) i) + +⟪⟫ : Arity 0 -> A +⟪⟫ = lookup [] + +⟪_⟫ : (a : A) -> Arity 1 -> A +⟪ a ⟫ = lookup [ a ] + +⟪_⨾_⟫ : (a b : A) -> Arity 2 -> A +⟪ a ⨾ b ⟫ = lookup (a ∷ b ∷ []) + +⟪_⨾_⨾_⟫ : (a b c : A) -> Arity 3 -> A +⟪ a ⨾ b ⨾ c ⟫ = lookup (a ∷ b ∷ c ∷ []) + \ No newline at end of file diff --git a/Cubical/Structures/Coh.agda b/Cubical/Structures/Coh.agda new file mode 100644 index 0000000..16199d6 --- /dev/null +++ b/Cubical/Structures/Coh.agda @@ -0,0 +1,34 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Coh where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.Fin +open import Cubical.Data.List as L +open import Cubical.Data.Sigma +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +open import Cubical.Structures.Sig +open import Cubical.Structures.Str +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq + +record CohSig (e n : Level) : Type (ℓ-max (ℓ-suc e) (ℓ-suc n)) where + field + name : Type e + free : name -> Type n +open CohSig public + +data EqSeq {a} {A : Type a} : A -> A -> Type a where + nil : ∀ {a} -> EqSeq a a + cons : ∀ {a b c} -> a ≡ b -> EqSeq b c -> EqSeq a c + +evalEqSeq : ∀ {a} {A : Type a} -> {a b : A} -> EqSeq a b -> a ≡ b +evalEqSeq nil = refl +evalEqSeq (cons p e) = p ∙ evalEqSeq e diff --git a/Cubical/Structures/Eq.agda b/Cubical/Structures/Eq.agda new file mode 100644 index 0000000..a1bf776 --- /dev/null +++ b/Cubical/Structures/Eq.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Eq where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.Fin +open import Cubical.Data.List as L +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +open import Cubical.Structures.Sig +open import Cubical.Structures.Str +open import Cubical.Structures.Tree + +record EqSig (e n : Level) : Type (ℓ-max (ℓ-suc e) (ℓ-suc n)) where + field + name : Type e + free : name -> Type n +open EqSig public + +FinEqSig : (e : Level) -> Type (ℓ-max (ℓ-suc e) (ℓ-suc ℓ-zero)) +FinEqSig = FinSig + +finEqSig : {e : Level} -> FinEqSig e -> EqSig e ℓ-zero +name (finEqSig σ) = σ .fst +free (finEqSig σ) = Fin ∘ σ .snd + +emptyEqSig : EqSig ℓ-zero ℓ-zero +name emptyEqSig = ⊥.⊥ +free emptyEqSig = ⊥.rec + +sumEqSig : {e n e' n' : Level} -> EqSig e n -> EqSig e' n' -> EqSig (ℓ-max e e') (ℓ-max n n') +name (sumEqSig σ τ) = (name σ) ⊎ (name τ) +free (sumEqSig {n' = n} σ τ) (inl x) = Lift {j = n} ((free σ) x) +free (sumEqSig {n = n} σ τ) (inr x) = Lift {j = n} ((free τ) x) + +module _ {f a e n : Level} (σ : Sig f a) (τ : EqSig e n) where + TODO: refactor as (Tree σ Unit -> Tree σ X) × (Tree σ Unit -> Tree σ X) ? + seq : Type (ℓ-max (ℓ-max (ℓ-max f a) e) n) + seq = (e : τ .name) -> Tree σ (τ .free e) × Tree σ (τ .free e) + +emptySEq : seq emptySig emptyEqSig +emptySEq n = ⊥.rec n + +module _ {f a e n s : Level} {σ : Sig f a} {τ : EqSig e n} where + type of structure satisfying equations + TODO: refactor as a coequaliser + infix 30 _⊨_ + _⊨_ : struct s σ -> (ε : seq σ τ) -> Type (ℓ-max s (ℓ-max e n)) + 𝔛 ⊨ ε = (eqn : τ .name) (ρ : τ .free eqn -> 𝔛 .car) + -> sharp σ 𝔛 ρ (ε eqn .fst) ≡ sharp σ 𝔛 ρ (ε eqn .snd) + diff --git a/Cubical/Structures/Free.agda b/Cubical/Structures/Free.agda new file mode 100644 index 0000000..97fde55 --- /dev/null +++ b/Cubical/Structures/Free.agda @@ -0,0 +1,192 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Free where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.List as L +open import Cubical.Data.Sigma +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +open import Cubical.Structures.Sig +open import Cubical.Structures.Str +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq + +module Definition {f a e n s : Level} (σ : Sig f a) (τ : EqSig e (ℓ-max n s)) (ε : seq {n = ℓ-max n s} σ τ) where + ns : Level + ns = ℓ-max n s + + record Free (ℓ ℓ' : Level) (h : HLevel) : Type (ℓ-suc (ℓ-max ℓ' (ℓ-max ℓ (ℓ-max f (ℓ-max a (ℓ-max e ns)))))) where + field + F : (X : Type ℓ) -> Type (ℓ-max ℓ ns) + η : {X : Type ℓ} -> X -> F X + α : {X : Type ℓ} -> sig σ (F X) -> F X + sat : {X : Type ℓ} -> < F X , α > ⊨ ε + isFree : {X : Type ℓ} + {𝔜 : struct (ℓ-max ℓ' ns) σ} + (H : isOfHLevel h (𝔜 .car)) (ϕ : 𝔜 ⊨ ε) + -> isEquiv (\(f : structHom {x = ℓ-max ℓ ns} < F X , α > 𝔜) -> f .fst ∘ η) + + ext : {X : Type ℓ} {𝔜 : struct (ℓ-max ℓ' ns) σ} + (H : isOfHLevel h (𝔜 .car)) (ϕ : 𝔜 ⊨ ε) + -> (hom : X -> 𝔜 .car) -> structHom < F X , α > 𝔜 + ext H ϕ = invIsEq (isFree H ϕ) + + ext-β : {X : Type ℓ} {𝔜 : struct (ℓ-max ℓ' ns) σ} + (H : isOfHLevel h (𝔜 .car)) (ϕ : 𝔜 ⊨ ε) (Hom : structHom < F X , α > 𝔜) + -> ext H ϕ (Hom .fst ∘ η) ≡ Hom + ext-β H ϕ = retIsEq (isFree H ϕ) + + ext-η : {X : Type ℓ} {𝔜 : struct (ℓ-max ℓ' ns) σ} + (H : isOfHLevel h (𝔜 .car)) (ϕ : 𝔜 ⊨ ε) (h : X -> 𝔜 .car) + -> (ext H ϕ h .fst) ∘ η ≡ h + ext-η H ϕ = secIsEq (isFree H ϕ) + + hom≡ : {X : Type ℓ} {𝔜 : struct (ℓ-max ℓ' ns) σ} + -> (H : isOfHLevel h (𝔜 .car)) (ϕ : 𝔜 ⊨ ε) + -> (H1 H2 : structHom < F X , α > 𝔜) + -> H1 .fst ∘ η ≡ H2 .fst ∘ η + -> H1 ≡ H2 + hom≡ H ϕ H1 H2 α = sym (ext-β H ϕ H1) ∙ cong (ext H ϕ) α ∙ ext-β H ϕ H2 + + open Free + module _ {ℓ} {A : Type ℓ} (𝔛 : Free ℓ ℓ 2) (𝔜 : Free ℓ ℓ 2) (isSet𝔛 : isSet (𝔛 .F A)) (isSet𝔜 : isSet (𝔜 .F A)) where + private + str𝔛 : struct (ℓ-max (ℓ-max n s) ℓ) σ + str𝔛 = < 𝔛 .F A , 𝔛 .α > + + str𝔜 : struct (ℓ-max (ℓ-max n s) ℓ) σ + str𝔜 = < 𝔜 .F A , 𝔜 .α > + + ϕ1 : structHom str𝔛 str𝔜 + ϕ1 = ext 𝔛 isSet𝔜 (𝔜 .sat) (𝔜 .η) + + ϕ2 : structHom str𝔜 str𝔛 + ϕ2 = ext 𝔜 isSet𝔛 (𝔛 .sat) (𝔛 .η) + + ϕ1∘ϕ2 : structHom str𝔜 str𝔜 + ϕ1∘ϕ2 = structHom∘ str𝔜 str𝔛 str𝔜 ϕ1 ϕ2 + + ϕ2∘ϕ1 : structHom str𝔛 str𝔛 + ϕ2∘ϕ1 = structHom∘ str𝔛 str𝔜 str𝔛 ϕ2 ϕ1 + + ϕ1∘ϕ2≡ : ϕ1∘ϕ2 .fst ∘ 𝔜 .η ≡ idHom str𝔜 .fst ∘ 𝔜 .η + ϕ1∘ϕ2≡ = + ϕ1 .fst ∘ ((ext 𝔜 isSet𝔛 (𝔛 .sat) (𝔛 .η) .fst) ∘ 𝔜 .η) + ≡⟨ congS (ϕ1 .fst ∘_) (ext-η 𝔜 isSet𝔛 (𝔛 .sat) (𝔛 .η)) ⟩ + ext 𝔛 isSet𝔜 (𝔜 .sat) (𝔜 .η) .fst ∘ 𝔛 .η + ≡⟨ ext-η 𝔛 isSet𝔜 (𝔜 .sat) (𝔜 .η) ⟩ + 𝔜 .η ∎ + + ϕ2∘ϕ1≡ : ϕ2∘ϕ1 .fst ∘ 𝔛 .η ≡ idHom str𝔛 .fst ∘ 𝔛 .η + ϕ2∘ϕ1≡ = + ϕ2 .fst ∘ ((ext 𝔛 isSet𝔜 (𝔜 .sat) (𝔜 .η) .fst) ∘ 𝔛 .η) + ≡⟨ congS (ϕ2 .fst ∘_) (ext-η 𝔛 isSet𝔜 (𝔜 .sat) (𝔜 .η)) ⟩ + ext 𝔜 isSet𝔛 (𝔛 .sat) (𝔛 .η) .fst ∘ 𝔜 .η + ≡⟨ ext-η 𝔜 isSet𝔛 (𝔛 .sat) (𝔛 .η) ⟩ + 𝔛 .η ∎ + + freeIso : Iso (𝔛 .F A) (𝔜 .F A) + freeIso = iso (ϕ1 .fst) (ϕ2 .fst) + (λ x -> congS (λ f -> f .fst x) (hom≡ 𝔜 isSet𝔜 (𝔜 .sat) ϕ1∘ϕ2 (idHom str𝔜) ϕ1∘ϕ2≡)) + (λ x -> congS (λ f -> f .fst x) (hom≡ 𝔛 isSet𝔛 (𝔛 .sat) ϕ2∘ϕ1 (idHom str𝔛) ϕ2∘ϕ1≡)) + + record FreeAux (ℓ ℓ' : Level) (h : HLevel) (F : (X : Type ℓ) -> Type (ℓ-max ℓ ns)) : Type (ℓ-suc (ℓ-max ℓ' (ℓ-max ℓ (ℓ-max f (ℓ-max a (ℓ-max e ns)))))) where + field + η : {X : Type ℓ} -> X -> F X + α : {X : Type ℓ} -> sig σ (F X) -> F X + sat : {X : Type ℓ} -> < F X , α > ⊨ ε + isFree : {X : Type ℓ} + {𝔜 : struct (ℓ-max ℓ' ns) σ} + (H : isOfHLevel h (𝔜 .car)) (ϕ : 𝔜 ⊨ ε) + -> isEquiv (\(f : structHom {x = ℓ-max ℓ ns} < F X , α > 𝔜) -> f .fst ∘ η) + + isoAux : {ℓ ℓ' : Level} {h : HLevel} -> + Iso (Σ[ F ∈ ((X : Type ℓ) -> Type (ℓ-max ℓ ns)) ] FreeAux ℓ ℓ' h F) (Free ℓ ℓ' h) + isoAux {ℓ = ℓ} {ℓ' = ℓ'} {h = h} = iso to from (λ _ -> refl) (λ _ -> refl) + where + to : Σ[ F ∈ ((X : Type ℓ) -> Type (ℓ-max ℓ ns)) ] FreeAux ℓ ℓ' h F -> Free ℓ ℓ' h + Free.F (to (F , aux)) = F + Free.η (to (F , aux)) = FreeAux.η aux + Free.α (to (F , aux)) = FreeAux.α aux + Free.sat (to (F , aux)) = FreeAux.sat aux + Free.isFree (to (F , aux)) = FreeAux.isFree aux + + from : Free ℓ ℓ' h -> Σ[ F ∈ ((X : Type ℓ) -> Type (ℓ-max ℓ ns)) ] FreeAux ℓ ℓ' h F + fst (from free) = Free.F free + FreeAux.η (snd (from free)) = Free.η free + FreeAux.α (snd (from free)) = Free.α free + FreeAux.sat (snd (from free)) = Free.sat free + FreeAux.isFree (snd (from free)) = Free.isFree free + + +constructions of a free structure on a signature and equations +TODO: generalise the universe levels!! +using a HIT + module Construction {f a e n : Level} (σ : Sig f a) (τ : EqSig e n) (ε : seq σ τ) where + + data Free (X : Type n) : Type (ℓ-suc (ℓ-max f (ℓ-max a (ℓ-max e n)))) where + η : X -> Free X + α : sig σ (Free X) -> Free X + sat : mkStruct (Free X) α ⊨ ε + + freeStruct : (X : Type) -> struct σ + car (freeStruct X) = Free X + alg (freeStruct _) = α + + module _ (X : Type) (𝔜 : struct σ) (ϕ : 𝔜 ⊨ ε) where + + private + Y = 𝔜 .fst + β = 𝔜 .snd + + ext : (h : X -> Y) -> Free X -> Y + ext h (η x) = h x + ext h (α (f , o)) = β (f , (ext h ∘ o)) + ext h (sat e ρ i) = ϕ e (ext h ∘ ρ) {!i!} + + module _ where + postulate + Fr : Type (ℓ-max (ℓ-max f a) n) + Fα : sig σ Fr -> Fr + Fs : sat σ τ ε (Fr , Fα) + + module _ (Y : Type (ℓ-max (ℓ-max f a) n)) (α : sig σ Y -> Y) where + postulate + Frec : sat σ τ ε (Y , α) -> Fr -> Y + Fhom : (p : sat σ τ ε (Y , α)) -> walgIsH σ (Fr , Fα) (Y , α) (Frec p) + Feta : (p : sat σ τ ε (Y , α)) (h : Fr -> Y) -> walgIsH σ (Fr , Fα) (Y , α) h -> Frec p ≡ h + +module Construction2 (σ : Sig ℓ-zero ℓ-zero) (τ : EqSig ℓ-zero ℓ-zero) (ε : seq σ τ) where +data _≈_ {X : Type} : Tree σ X -> Tree σ X -> Type (ℓ-suc ℓ-zero) where + ≈-refl : ∀ t -> t ≈ t + ≈-sym : ∀ t s -> t ≈ s -> s ≈ t + ≈-trans : ∀ t s r -> t ≈ s -> s ≈ r -> t ≈ r + ≈-cong : (f : σ .symbol) -> {t s : σ .arity f -> Tree σ X} + -> ((a : σ .arity f) -> t a ≈ s a) + -> node (f , t) ≈ node (f , s) + ≈-eqs : (𝔜 : struct {ℓ-zero} {ℓ-zero} {ℓ-zero} σ) (ϕ : 𝔜 ⊨ ε) + -> (e : τ .name) (ρ : X -> 𝔜 .car) + -> ∀ t s -> sharp σ {𝔜 = 𝔜} ρ t ≡ sharp σ {𝔜 = 𝔜} ρ s + -> t ≈ s + +Free : Type -> Type₁ +Free X = Tree σ X / _≈_ + +freeAlg : (X : Type) -> sig σ (Free X) -> Free X +freeAlg X (f , i) = Q.[ node (f , {!!}) ] needs choice? + +freeStruct : (X : Type) -> struct σ +freeStruct X = Free X , freeAlg X + +module _ (X : Type) (𝔜 : struct σ) (ϕ : 𝔜 ⊨ ε) where + + private + Y = 𝔜 .fst + β = 𝔜 .snd diff --git a/Cubical/Structures/Gpd/Mon/Free.agda b/Cubical/Structures/Gpd/Mon/Free.agda new file mode 100644 index 0000000..e7390b3 --- /dev/null +++ b/Cubical/Structures/Gpd/Mon/Free.agda @@ -0,0 +1,4 @@ +TODO: Define free monoidal groupoids as a HIT + +{-# OPTIONS --cubical --safe --exact-split #-} +module Cubical.Structures.Gpd.Mon.Free where diff --git a/Cubical/Structures/Gpd/Mon/List.agda b/Cubical/Structures/Gpd/Mon/List.agda new file mode 100644 index 0000000..ec1f1ba --- /dev/null +++ b/Cubical/Structures/Gpd/Mon/List.agda @@ -0,0 +1,4 @@ +TODO: Show that List A has all the monoidal coherences + +{-# OPTIONS --cubical --safe --exact-split #-} +module Cubical.Structures.Gpd.Mon.List where \ No newline at end of file diff --git a/Cubical/Structures/Gpd/SMon/Free.agda b/Cubical/Structures/Gpd/SMon/Free.agda new file mode 100644 index 0000000..62c455c --- /dev/null +++ b/Cubical/Structures/Gpd/SMon/Free.agda @@ -0,0 +1,4 @@ +TODO: Define free symmetric monoidal groupoids as a HIT + +{-# OPTIONS --cubical --safe --exact-split #-} +module Cubical.Structures.Gpd.SMon.Free where diff --git a/Cubical/Structures/Gpd/SMon/SList.agda b/Cubical/Structures/Gpd/SMon/SList.agda new file mode 100644 index 0000000..09f60f5 --- /dev/null +++ b/Cubical/Structures/Gpd/SMon/SList.agda @@ -0,0 +1,115 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Gpd.SMon.SList where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma + +infixr 20 _∷_ + +data SList {a} (A : Type a) : Type a where + [] : SList A + _∷_ : (a : A) -> (as : SList A) -> SList A + swap : (a b : A) (cs : SList A) + -> a ∷ b ∷ cs ≡ b ∷ a ∷ cs + swap⁻¹ : (a b : A) (cs : SList A) + -> swap a b cs ≡ sym (swap b a cs) + hexagon– : (a b c : A) (cs : SList A) + -> a ∷ b ∷ c ∷ cs ≡ c ∷ b ∷ a ∷ cs + hexagon↑ : (a b c : A) (cs : SList A) + -> Square (\i -> b ∷ swap a c cs i) (hexagon– a b c cs) + (swap b a (c ∷ cs)) (swap b c (a ∷ cs)) + hexagon↓ : (a b c : A) (cs : SList A) + -> Square (hexagon– a b c cs) (swap a c (b ∷ cs)) + (\i -> a ∷ swap b c cs i) (\i -> c ∷ swap b a cs i) + isGpdSList : isGroupoid (SList A) + +pattern [_] a = a ∷ [] + +private + variable + ℓ : Level + A : Type ℓ + +swap² : (a b : A) (cs : SList A) + -> swap a b cs ∙ swap b a cs ≡ refl +swap² a b cs = cong (_∙ swap b a cs) (swap⁻¹ a b cs) ∙ lCancel (swap b a cs) + +hexagon : (a b c : A) (cs : SList A) + -> (swap a b (c ∷ cs) ∙ cong (b ∷_) (swap a c cs) ∙ swap b c (a ∷ cs)) + ≡ cong (a ∷_) (swap b c cs) ∙ swap a c (b ∷ cs) ∙ cong (c ∷_) (swap a b cs) +hexagon a b c cs = + let hex-up = PathP→compPathL (hexagon↑ a b c cs) + hex-dn = PathP→compPathR (hexagon↓ a b c cs) + in + swap a b (c ∷ cs) ∙ cong (b ∷_) (swap a c cs) ∙ swap b c (a ∷ cs) + ≡⟨ cong (_∙ cong (b ∷_) (swap a c cs) ∙ swap b c (a ∷ cs)) (swap⁻¹ a b (c ∷ cs)) ⟩ + sym (swap b a (c ∷ cs)) ∙ cong (b ∷_) (swap a c cs) ∙ swap b c (a ∷ cs) + ≡⟨ hex-up ⟩ + hexagon– a b c cs + ≡⟨ hex-dn ⟩ + cong (a ∷_) (swap b c cs) ∙ swap a c (b ∷ cs) ∙ cong (c ∷_) (sym (swap b a cs)) + ≡⟨ assoc (cong (a ∷_) (swap b c cs)) (swap a c (b ∷ cs)) (cong (c ∷_) (sym (swap b a cs))) ⟩ + (cong (a ∷_) (swap b c cs) ∙ swap a c (b ∷ cs)) ∙ cong (c ∷_) (sym (swap b a cs)) + ≡⟨ cong ((cong (a ∷_) (swap b c cs) ∙ swap a c (b ∷ cs)) ∙_) (cong (cong (c ∷_)) (sym (swap⁻¹ a b cs))) ⟩ + (cong (a ∷_) (swap b c cs) ∙ swap a c (b ∷ cs)) ∙ cong (c ∷_) (swap a b cs) + ≡⟨ sym (assoc (cong (a ∷_) (swap b c cs)) (swap a c (b ∷ cs)) (cong (c ∷_) (swap a b cs))) ⟩ + cong (a ∷_) (swap b c cs) ∙ swap a c (b ∷ cs) ∙ cong (c ∷_) (swap a b cs) + ∎ + +_++_ : SList A -> SList A -> SList A +[] ++ bs = bs +(a ∷ as) ++ bs = a ∷ (as ++ bs) +swap a b as i ++ bs = swap a b (as ++ bs) i +swap⁻¹ a b as i j ++ bs = swap⁻¹ a b (as ++ bs) i j +hexagon– a b c as i ++ bs = hexagon– a b c (as ++ bs) i +hexagon↑ a b c as i j ++ bs = hexagon↑ a b c (as ++ bs) i j +hexagon↓ a b c as i j ++ bs = hexagon↓ a b c (as ++ bs) i j +isGpdSList as cs p q u v i j k ++ bs = + isGpdSList (as ++ bs) (cs ++ bs) (cong (_++ bs) p) (cong (_++ bs) q) + (cong (cong (_++ bs)) u) (cong (cong (_++ bs)) v) i j k + +++-unitl : (as : SList A) -> [] ++ as ≡ as +++-unitl as = refl + +++-unitr : (as : SList A) -> as ++ [] ≡ as +++-unitr [] = refl +++-unitr (a ∷ as) = cong (a ∷_) (++-unitr as) +++-unitr (swap a b as i) j = swap a b (++-unitr as j) i +++-unitr (swap⁻¹ a b as i j) k = swap⁻¹ a b (++-unitr as k) i j +++-unitr (hexagon– a b c as i) j = hexagon– a b c (++-unitr as j) i +++-unitr (hexagon↑ a b c as i j) k = hexagon↑ a b c (++-unitr as k) i j +++-unitr (hexagon↓ a b c as i j) k = hexagon↓ a b c (++-unitr as k) i j +++-unitr (isGpdSList as cs p q u v i j k) l = + isGpdSList (++-unitr as l) (++-unitr cs l) (cong (\z -> ++-unitr z l) p) (cong (\z -> ++-unitr z l) q) + (cong (cong (\z -> ++-unitr z l)) u) (cong (cong (\z -> ++-unitr z l)) v) i j k + +++-assocr : (as bs cs : SList A) -> (as ++ bs) ++ cs ≡ as ++ (bs ++ cs) +++-assocr [] bs cs = refl +++-assocr (a ∷ as) bs cs = cong (a ∷_) (++-assocr as bs cs) +++-assocr (swap a b as i) bs cs j = swap a b (++-assocr as bs cs j) i +++-assocr (swap⁻¹ a b as i j) bs cs k = swap⁻¹ a b (++-assocr as bs cs k) i j +++-assocr (hexagon– a b c as i) bs cs j = hexagon– a b c (++-assocr as bs cs j) i +++-assocr (hexagon↑ a b c as i j) bs cs k = hexagon↑ a b c (++-assocr as bs cs k) i j +++-assocr (hexagon↓ a b c as i j) bs cs k = hexagon↓ a b c (++-assocr as bs cs k) i j +++-assocr (isGpdSList as ds p q u v i j k) bs cs l = + isGpdSList (++-assocr as bs cs l) (++-assocr ds bs cs l) (cong (\z -> ++-assocr z bs cs l) p) (cong (\z -> ++-assocr z bs cs l) q) + (cong (cong (\z -> ++-assocr z bs cs l)) u) (cong (cong (\z -> ++-assocr z bs cs l)) v) i j k + +TODO: Define commutativity for SList directly or after truncation + To do this directly will probably need coherence for swap + ++-∷ : (a : A) (as : SList A) -> a ∷ as ≡ as ++ [ a ] + ++-∷ a [] = refl + ++-∷ a (b ∷ as) = swap a b as ∙ cong (b ∷_) (++-∷ a as) + ++-∷ a (swap b c as i) = + {!!} + + ++-comm : (as bs : SList A) -> as ++ bs ≡ bs ++ as + ++-comm [] bs = sym (++-unitr bs) + ++-comm (a ∷ as) bs = cong (a ∷_) (++-comm as bs) + ∙ cong (_++ as) (++-∷ a bs) + ∙ ++-assocr bs [ a ] as +++-comm (swap a b as i) bs = {!!} + +TODO: Prove SList is free symmetric monoidal on groupoids +-- diff --git a/Cubical/Structures/Set/CMon/Bag.agda b/Cubical/Structures/Set/CMon/Bag.agda new file mode 100644 index 0000000..5eebfac --- /dev/null +++ b/Cubical/Structures/Set/CMon/Bag.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Bag where + +open import Cubical.Structures.Set.CMon.Bag.Base public +open import Cubical.Structures.Set.CMon.Bag.Free public +open import Cubical.Structures.Set.CMon.Bag.ToCList public + diff --git a/Cubical/Structures/Set/CMon/Bag/Base.agda b/Cubical/Structures/Set/CMon/Bag/Base.agda new file mode 100644 index 0000000..39717e3 --- /dev/null +++ b/Cubical/Structures/Set/CMon/Bag/Base.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Bag.Base where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.List renaming (_∷_ to _∷ₗ_) +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Sigma +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Set.Mon.Array as A +open import Cubical.Structures.Set.Mon.List as L +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity hiding (_/_) +open import Cubical.HITs.SetQuotients as Q +open import Cubical.Structures.Set.CMon.QFreeMon +open import Cubical.Relation.Nullary + +open Iso + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + +_≈_ : ∀ {A : Type ℓ} -> Array A -> Array A -> Type ℓ +_≈_ (n , v) (m , w) = Σ[ σ ∈ Iso (Fin n) (Fin m) ] v ≡ w ∘ σ .fun + +Bag : Type ℓ -> Type ℓ +Bag A = Array A Q./ _≈_ diff --git a/Cubical/Structures/Set/CMon/Bag/FinExcept.agda b/Cubical/Structures/Set/CMon/Bag/FinExcept.agda new file mode 100644 index 0000000..845a8db --- /dev/null +++ b/Cubical/Structures/Set/CMon/Bag/FinExcept.agda @@ -0,0 +1,311 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Bag.FinExcept where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Data.List renaming (_∷_ to _∷ₗ_) +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Sigma +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Set.Mon.Array as A +open import Cubical.Structures.Set.Mon.List as L +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity hiding (_/_) +open import Cubical.Structures.Set.CMon.QFreeMon +open import Cubical.Structures.Set.CMon.Bag.Base +open import Cubical.Structures.Set.CMon.Bag.Free +open import Cubical.Relation.Nullary +open import Cubical.Data.Fin.LehmerCode as E + +private + variable + ℓ : Level + A : Type ℓ + n : ℕ + +abstract + <-not-dense : ∀ {a b c} -> a < b -> b < suc c -> a < c + <-not-dense {a} {b} {c} p q with a ≟ c + ... | lt r = r + ... | eq r = ⊥.rec (¬nb→a≠0 : ∀ {a b} -> a > b -> ¬ a ≡ 0 + a>b→a≠0 {a} {b} (k , p) a≡0 = snotz (sym (+-suc k b) ∙ p ∙ a≡0) + + predℕa a < suc b -> e < a -> predℕ a < b + predℕab→a≠0 q)) (predℕ-≤-predℕ p) + + _ ¬ a ≡ b -> (a < b) ⊎ (b < a) + a (p : ¬ a ≡ b) -> (q : a < b) -> a (p : ¬ a ≡ b) -> (q : a > b) -> a (p : a < b) -> a ≤? b ≡ inl p + ≤?-beta-inl a b p with a ≤? b + ... | inl q = congS inl (isProp≤ _ _) + ... | inr q = ⊥.rec (<-asym p q) + + ≤?-beta-inr : ∀ (a b : ℕ) -> (p : b ≤ a) -> a ≤? b ≡ inr p + ≤?-beta-inr a b p with a ≤? b + ... | inl q = ⊥.rec (<-asym q p) + ... | inr q = congS inr (isProp≤ _ _) + + FinExcept≡ : ∀ {n k} -> {a b : FinExcept {n = n} k} -> a .fst .fst ≡ b .fst .fst -> a ≡ b + FinExcept≡ p = Σ≡Prop (λ _ -> isProp¬ _) (Fin-fst-≡ p) + +pOut : ∀ {n} -> (k : Fin (suc n)) -> FinExcept k -> Fin n +pOut {n} (k , p) ((j , q) , r) = + ⊎.rec + (λ j j , <-not-dense j predℕ j , predℕa Fin n -> FinExcept k +pIn (k , p) (j , q) = + ⊎.rec + (λ j (j , ≤-suc q) , <→≢ j fsuc (j , q) , λ r -> ¬m (Fin n -> A) -> FinExcept k -> A +f ∼ = f ∘ toFinExc + +1+_ : {k : Fin n} -> FinExcept k -> FinExcept (fsuc k) +1+_ {k = k} (j , p) = (fsuc j) , fsuc≠ + where + abstract + fsuc≠ : ¬ fsuc k ≡ fsuc j + fsuc≠ r = p (fsuc-inj r) + +pIn∘Out : (k : Fin (suc n)) -> ∀ j -> pIn k (pOut k j) ≡ j +pIn∘Out (k , p) ((j , q) , r) = + ⊎.rec + (λ j + pIn (k , p) (pOut (k , p) ((j , q) , r)) + ≡⟨⟩ + pIn (k , p) (⊎.rec (λ j j , _) _ (j (j , _) , _) _ (j ≤? k) + ≡⟨ congS (⊎.rec (λ j (j , _) , _) _) (≤?-beta-inl j k j + pIn (k , p) (pOut (k , p) ((j , q) , r)) + ≡⟨⟩ + pIn (k , p) (⊎.rec _ (λ k predℕ j , predℕa fsuc (predℕ j , _) , _) (predℕ j ≤? k) + ≡⟨ congS (⊎.rec _ _) (≤?-beta-inr (predℕ j) k (predℕ-≤-predℕ kb→a≠0 k ∀ j -> pOut k (pIn k j) ≡ j +pOut∘In (k , p) (j , q) = + ⊎.rec + (λ j + pOut (k , p) (pIn (k , p) (j , q)) + ≡⟨⟩ + pOut (k , p) (⊎.rec (λ j (j , ≤-suc q) , _) _ (j ≤? k)) + ≡⟨ congS (pOut (k , p) ∘ ⊎.rec _ _) (≤?-beta-inl j k j j , <-not-dense j + pOut (k , p) (pIn (k , p) (j , q)) + ≡⟨⟩ + pOut (k , p) (⊎.rec _ (λ k≤j -> fsuc (j , q) , _) (j ≤? k)) + ≡⟨ congS (pOut (k , p) ∘ ⊎.rec _ _) (≤?-beta-inr j k k≤j) ⟩ + pOut (k , p) (fsuc (j , q) , _) + ≡⟨⟩ + ⊎.rec _ (λ k predℕ (suc j) , _) (suc j fst (pIn fzero k) ≡ fsuc k +pInZ≡fsuc k = congS (fst ∘ ⊎.rec _ _) (≤?-beta-inr (fst k) 0 zero-≤) + +pIn-fsuc-nat : {k : Fin (suc n)} -> 1+_ ∘ pIn k ≡ pIn (fsuc k) ∘ fsuc +pIn-fsuc-nat {n = zero} {k = k} = funExt \j -> ⊥.rec (¬Fin0 j) +pIn-fsuc-nat {n = suc n} {k = k} = funExt (pIn-fsuc-nat-htpy k) + where + pIn-fsuc-nat-htpy : (k : Fin (suc (suc n))) (j : Fin (suc n)) + -> 1+ (pIn k j) ≡ pIn (fsuc k) (fsuc j) + pIn-fsuc-nat-htpy (k , p) (j , q) = + ⊎.rec + (λ j + let + r = + 1+ pIn (k , p) (j , q) + ≡⟨⟩ + 1+ (⊎.rec (λ j (j , ≤-suc q) , _) _ (j ≤? k)) + ≡⟨ congS (1+_ ∘ ⊎.rec _ _) (≤?-beta-inl j k j fsuc (j , _) , _) _ (suc j ≤? suc k) + ≡⟨ congS (⊎.rec _ _) (≤?-beta-inl (suc j) (suc k) (suc-≤-suc j + let + r = + 1+ pIn (k , p) (j , q) + ≡⟨⟩ + 1+ (⊎.rec _ (λ k≤j -> fsuc (j , q) , _) (j ≤? k)) + ≡⟨ congS (1+_ ∘ ⊎.rec _ _) (≤?-beta-inr j k k≤j) ⟩ + 1+ (fsuc (j , q) , _) + ≡⟨⟩ + fsuc (fsuc (j , q)) , _ ∎ + s = + pIn (fsuc (k , p)) (fsuc (j , q)) + ≡⟨⟩ + ⊎.rec _ (λ sk≤sj -> fsuc (fsuc (j , q)) , _) (suc j ≤? suc k) + ≡⟨ congS (⊎.rec _ _) (≤?-beta-inr (suc j) (suc k) (suc-≤-suc k≤j)) ⟩ + fsuc (fsuc (j , q)) , _ + ≡⟨ FinExcept≡ refl ⟩ + fsuc (fsuc (j , q)) , _ ∎ + in r ∙ (sym s) + ) + (j ≤? k) + + +Aut : (A : Type ℓ) -> Type ℓ +Aut A = Iso A A + +projectIso : {i : Fin n} -> Iso (Unit ⊎ FinExcept i) (Fin n) +fun (projectIso {i = i}) (inl _) = i +fun (projectIso {i = i}) (inr m) = fst m +inv (projectIso {i = i}) m = + decRec (λ i≡m -> inl tt) (λ i≠m -> inr (m , i≠m)) (discreteFin i m) +rightInv (projectIso {i = i}) m with discreteFin i m +... | yes i≡m = i≡m +... | no ¬i≡m = refl +leftInv (projectIso {i = i}) (inl _) with discreteFin i i +... | yes i≡m = refl +... | no ¬i≡m = ⊥.rec (¬i≡m refl) +leftInv (projectIso {i = i}) (inr m) with discreteFin i (fst m) +... | yes i≡m = ⊥.rec (m .snd i≡m) +... | no ¬i≡m = congS inr (FinExcept≡ refl) + +module _ {n : ℕ} where + equivIn : (σ : Aut (Fin (suc n))) -> Iso (FinExcept fzero) (FinExcept (σ .fun fzero)) + Iso.fun (equivIn σ) (k , p) = σ .fun k , p ∘ isoFunInjective σ _ _ + Iso.inv (equivIn σ) (k , p) = σ .inv k , p ∘ isoInvInjective σ _ _ ∘ σ .leftInv fzero ∙_ + Iso.rightInv (equivIn σ) (k , p) = FinExcept≡ (congS fst (σ .rightInv k)) + Iso.leftInv (equivIn σ) (k , p) = FinExcept≡ (congS fst (σ .leftInv k)) + + equivOut : ∀ {k} -> Iso (FinExcept fzero) (FinExcept k) -> Aut (Fin (suc n)) + equivOut {k = k} σ = + Fin (suc n) Iso⟨ invIso projectIso ⟩ + Unit ⊎ FinExcept fzero Iso⟨ ⊎Iso idIso σ ⟩ + Unit ⊎ FinExcept k Iso⟨ projectIso ⟩ + Fin (suc n) ∎Iso + + equivOut-beta-α : ∀ {k} {σ} (x : FinExcept fzero) -> (equivOut {k = k} σ) .fun (fst x) ≡ fst (σ .fun x) + equivOut-beta-α {σ = σ} x with discreteFin fzero (fst x) + ... | yes p = ⊥.rec (x .snd p) + ... | no ¬p = congS (fst ∘ σ .fun) (FinExcept≡ refl) + + equivOut-beta-β : ∀ {k} {σ} (x : FinExcept (equivOut {k = k} σ .fun fzero)) -> equivOut σ .inv (fst x) ≡ σ .inv x .fst + equivOut-beta-β {k = k} {σ = σ} x with discreteFin k (fst x) + ... | yes p = ⊥.rec (x .snd p) + ... | no ¬p = congS (fst ∘ σ .inv) (FinExcept≡ refl) + + equivIn∘Out : ∀ {k} -> (τ : Iso (FinExcept fzero) (FinExcept k)) -> equivIn (equivOut τ) ≡ τ + equivIn∘Out τ = + Iso≡Set isSetFinExcept isSetFinExcept _ _ + (FinExcept≡ ∘ congS fst ∘ equivOut-beta-α) + (FinExcept≡ ∘ congS fst ∘ equivOut-beta-β) + + equivOut∘In : (σ : Aut (Fin (suc n))) -> equivOut (equivIn σ) ≡ σ + equivOut∘In σ = Iso≡Set isSetFin isSetFin _ _ lemma-α lemma-β + where + lemma-α : (x : Fin (suc n)) -> equivOut (equivIn σ) .fun x ≡ σ .fun x + lemma-α x with discreteFin fzero x + ... | yes p = congS (σ .fun) p + ... | no ¬p = refl + lemma-β : (x : Fin (suc n)) -> equivOut (equivIn σ) .inv x ≡ σ .inv x + lemma-β x with discreteFin (σ .fun fzero) x + ... | yes p = sym (σ .leftInv fzero) ∙ congS (σ .inv) p + ... | no ¬p = refl + + G : Iso (Aut (Fin (suc n))) (Σ[ k ∈ Fin (suc n) ] (Iso (FinExcept (fzero {k = n})) (FinExcept k))) + fun G σ = σ .fun fzero , equivIn σ + inv G (k , τ) = equivOut τ + rightInv G (k , τ) = ΣPathP (refl , equivIn∘Out τ) + leftInv G = equivOut∘In + + punch-σ : (σ : Aut (Fin (suc n))) -> Aut (Fin n) + punch-σ σ = + Fin n Iso⟨ invIso pIso ⟩ + FinExcept (fzero {k = n}) Iso⟨ (G .fun σ) .snd ⟩ + FinExcept (σ .fun fzero) Iso⟨ pIso ⟩ + Fin n ∎Iso + + fill-σ : ∀ k -> Aut (Fin (suc n)) + fill-σ k = equivOut $ + FinExcept fzero Iso⟨ pIso ⟩ + Fin n Iso⟨ invIso pIso ⟩ + FinExcept k ∎Iso diff --git a/Cubical/Structures/Set/CMon/Bag/Free.agda b/Cubical/Structures/Set/CMon/Bag/Free.agda new file mode 100644 index 0000000..3093d0b --- /dev/null +++ b/Cubical/Structures/Set/CMon/Bag/Free.agda @@ -0,0 +1,390 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Bag.Free where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.List renaming (_∷_ to _∷ₗ_) +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Sigma +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Set.Mon.Array as A +open import Cubical.Structures.Set.Mon.List as L +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity hiding (_/_) +open import Cubical.Structures.Set.CMon.QFreeMon +open import Cubical.Structures.Set.CMon.Bag.Base +open import Cubical.Relation.Nullary + +open Iso + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + +Bag≈ = _≈_ + +refl≈ : {as : Array A} -> as ≈ as +refl≈ {as = as} = idIso , refl + +sym≈ : {as bs : Array A} -> as ≈ bs -> bs ≈ as +sym≈ {as = (n , f)} {bs = (m , g)} (σ , p) = + invIso σ , congS (g ∘_) (sym (funExt (σ .rightInv))) + ∙ congS (_∘ σ .inv) (sym p) + +trans≈ : {as bs cs : Array A} -> as ≈ bs -> bs ≈ cs -> as ≈ cs +trans≈ {as = (n , f)} {bs = (m , g)} {cs = (o , h)} (σ , p) (τ , q) = + compIso σ τ , sym + ((h ∘ τ .fun) ∘ σ .fun ≡⟨ congS (_∘ σ .fun) (sym q) ⟩ + g ∘ σ .fun ≡⟨ sym p ⟩ + f ∎) + +Fin+-cong : {n m n' m' : ℕ} -> Iso (Fin n) (Fin n') -> Iso (Fin m) (Fin m') -> Iso (Fin (n + m)) (Fin (n' + m')) +Fin+-cong {n} {m} {n'} {m'} σ τ = + compIso (Fin≅Fin+Fin n m) (compIso (⊎Iso σ τ) (invIso (Fin≅Fin+Fin n' m'))) + +⊎Iso-eta : {A B A' B' : Type ℓ} {C : Type ℓ'} (f : A' -> C) (g : B' -> C) + -> (σ : Iso A A') (τ : Iso B B') + -> ⊎.rec (f ∘ σ .fun) (g ∘ τ .fun) ≡ ⊎.rec f g ∘ ⊎Iso σ τ .fun +⊎Iso-eta f g σ τ = ⊎-eta (⊎.rec f g ∘ ⊎Iso σ τ .fun) refl refl + +⊎Swap-eta : {A B : Type ℓ} {C : Type ℓ'} (f : A -> C) (g : B -> C) + -> ⊎.rec f g ≡ ⊎.rec g f ∘ ⊎-swap-Iso .fun +⊎Swap-eta f g = ⊎-eta (⊎.rec g f ∘ ⊎-swap-Iso .fun) refl refl + +cong≈ : {as bs cs ds : Array A} -> as ≈ bs -> cs ≈ ds -> (as ⊕ cs) ≈ (bs ⊕ ds) +cong≈ {as = n , f} {bs = n' , f'} {m , g} {m' , g'} (σ , p) (τ , q) = + Fin+-cong σ τ , + ( + combine n m f g + ≡⟨ cong₂ (combine n m) p q ⟩ + combine n m (f' ∘ σ .fun) (g' ∘ τ .fun) + ≡⟨⟩ + ⊎.rec (f' ∘ σ .fun) (g' ∘ τ .fun) ∘ finSplit n m + ≡⟨ congS (_∘ finSplit n m) (⊎Iso-eta f' g' σ τ) ⟩ + ⊎.rec f' g' ∘ ⊎Iso σ τ .fun ∘ finSplit n m + ≡⟨⟩ + ⊎.rec f' g' ∘ idfun _ ∘ ⊎Iso σ τ .fun ∘ finSplit n m + ≡⟨ congS (\h -> ⊎.rec f' g' ∘ h ∘ ⊎Iso σ τ .fun ∘ finSplit n m) (sym (funExt (Fin≅Fin+Fin n' m' .rightInv))) ⟩ + ⊎.rec f' g' ∘ (Fin≅Fin+Fin n' m' .fun ∘ Fin≅Fin+Fin n' m' .inv) ∘ ⊎Iso σ τ .fun ∘ finSplit n m + ≡⟨⟩ + (⊎.rec f' g' ∘ Fin≅Fin+Fin n' m' .fun) ∘ (Fin≅Fin+Fin n' m' .inv ∘ ⊎Iso σ τ .fun ∘ finSplit n m) + ≡⟨⟩ + combine n' m' f' g' ∘ Fin+-cong σ τ .fun + ∎) + +Fin+-comm : (n m : ℕ) -> Iso (Fin (n + m)) (Fin (m + n)) +Fin+-comm n m = compIso (Fin≅Fin+Fin n m) (compIso ⊎-swap-Iso (invIso (Fin≅Fin+Fin m n))) + +comm≈ : {as bs : Array A} -> (as ⊕ bs) ≈ (bs ⊕ as) +comm≈ {as = n , f} {bs = m , g} = + Fin+-comm n m , sym + ( + ⊎.rec g f ∘ finSplit m n ∘ Fin≅Fin+Fin m n .inv ∘ ⊎-swap-Iso .fun ∘ Fin≅Fin+Fin n m .fun + ≡⟨⟩ + ⊎.rec g f ∘ (Fin≅Fin+Fin m n .fun ∘ Fin≅Fin+Fin m n .inv) ∘ ⊎-swap-Iso .fun ∘ Fin≅Fin+Fin n m .fun + ≡⟨ congS (λ h -> ⊎.rec g f ∘ h ∘ ⊎-swap-Iso .fun ∘ Fin≅Fin+Fin n m .fun) (funExt (Fin≅Fin+Fin m n .rightInv)) ⟩ + ⊎.rec g f ∘ ⊎-swap-Iso .fun ∘ Fin≅Fin+Fin n m .fun + ≡⟨ congS (_∘ Fin≅Fin+Fin n m .fun) (sym (⊎Swap-eta f g)) ⟩ + ⊎.rec f g ∘ Fin≅Fin+Fin n m .fun + ∎) + +fpred' : ∀ {n} -> (x : Fin (suc n)) -> ¬ x ≡ fzero -> Fin n +fpred' (zero , p) q = ⊥.rec (q (Fin-fst-≡ refl)) +fpred' (suc w , p) q = w , pred-≤-pred p + +fpred : ∀ {n} -> Fin (suc (suc n)) -> Fin (suc n) +fpred (zero , p) = fzero +fpred (suc w , p) = w , pred-≤-pred p + +fsuc∘fpred : ∀ {n} -> (x : Fin (suc (suc n))) -> ¬ x ≡ fzero -> fsuc (fpred x) ≡ x +fsuc∘fpred (zero , p) q = ⊥.rec (q (Fin-fst-≡ refl)) +fsuc∘fpred (suc k , p) q = Fin-fst-≡ refl + +fpred∘fsuc : ∀ {n} -> (x : Fin (suc n)) -> fpred (fsuc x) ≡ x +fpred∘fsuc (k , p) = Fin-fst-≡ refl + +isoFunInv : ∀ {A B : Type ℓ} {x y} -> (σ : Iso A B) -> σ .fun x ≡ y -> σ .inv y ≡ x +isoFunInv σ p = congS (σ .inv) (sym p) ∙ σ .leftInv _ + +isoFunInvContra : ∀ {A B : Type ℓ} {x y z} -> (σ : Iso A B) -> σ .fun x ≡ y -> ¬ (z ≡ y) -> ¬ (σ .inv z ≡ x) +isoFunInvContra σ p z≠y q = z≠y (sym (σ .rightInv _) ∙ congS (σ .fun) q ∙ p) + +autInvIs0 : ∀ {n} -> (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) + -> σ .fun fzero ≡ fzero -> σ .inv fzero ≡ fzero +autInvIs0 = isoFunInv + +autSucNot0 : ∀ {n} -> (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) + -> (x : Fin (suc n)) -> σ .fun fzero ≡ fzero -> ¬ σ .fun (fsuc x) ≡ fzero +autSucNot0 σ x p = isoFunInvContra (invIso σ) (isoFunInv σ p) (snotz ∘ congS fst) + +punchOutZero : ∀ {n} (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) -> σ .fun fzero ≡ fzero + -> Iso (Fin (suc n)) (Fin (suc n)) +punchOutZero {n = n} σ p = + iso (punch σ) (punch (invIso σ)) (punch∘punch σ p) (punch∘punch (invIso σ) (autInvIs0 σ p)) + where + punch : Iso (Fin (suc (suc n))) (Fin (suc (suc n))) -> Fin (suc n) -> Fin (suc n) + punch σ = fpred ∘ σ .fun ∘ fsuc + punch∘punch : (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) + -> σ .fun fzero ≡ fzero + -> (x : Fin (suc n)) + -> punch σ (punch (invIso σ) x) ≡ x + punch∘punch σ p x = + punch σ (punch (invIso σ) x) + ≡⟨⟩ + fpred (σ .fun ((fsuc ∘ fpred) (σ .inv (fsuc x)))) + ≡⟨ congS (fpred ∘ σ .fun) (fsuc∘fpred (σ .inv (fsuc x)) (autSucNot0 (invIso σ) x (autInvIs0 σ p))) ⟩ + fpred (σ .fun (σ .inv (fsuc x))) + ≡⟨ congS fpred (σ .rightInv (fsuc x)) ⟩ + fpred (fsuc x) + ≡⟨ fpred∘fsuc x ⟩ + x ∎ + +punchOutZero≡fsuc : ∀ {n} (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) -> (σ-0≡0 : σ .fun fzero ≡ fzero) + -> (w : Fin (suc n)) -> σ .fun (fsuc w) ≡ fsuc (punchOutZero σ σ-0≡0 .fun w) +punchOutZero≡fsuc σ σ-0≡0 w = sym (fsuc∘fpred _ (autSucNot0 σ w σ-0≡0)) + +finSubst : ∀ {n m} -> n ≡ m -> Fin n -> Fin m +finSubst {n = n} {m = m} p (k , q) = k , (subst (k <_) p q) + +Fin≅ : ∀ {n m} -> n ≡ m -> Iso (Fin n) (Fin m) +Fin≅ {n = n} {m = m} p = iso + (finSubst p) + (finSubst (sym p)) + (λ (k , q) -> Fin-fst-≡ refl) + (λ (k , q) -> Fin-fst-≡ refl) + +Fin≅-inj : {n m : ℕ} -> Iso (Fin n) (Fin m) -> n ≡ m +Fin≅-inj {n = n} {m = m} σ = Fin-inj n m (isoToPath σ) + +≈-length : {n m : ℕ} -> Iso (Fin n) (Fin m) -> n ≡ m +≈-length {n = n} {m = m} σ = Fin-inj n m (isoToPath σ) + +module _ {n} (σ : Iso (Fin (suc n)) (Fin (suc n))) where + private + m : ℕ + m = suc n + + cutoff : ℕ + cutoff = (σ .inv fzero) .fst + + cutoff< : cutoff < m + cutoff< = (σ .inv fzero) .snd + + cutoff+- : cutoff + (m ∸ cutoff) ≡ m + cutoff+- = ∸-lemma (<-weaken cutoff<) + + 00 m cutoff cutoff< + + swapAut : Iso (Fin (suc n)) (Fin (suc n)) + swapAut = compIso (Fin≅ (sym cutoff+- ∙ +-comm cutoff _)) (compIso (Fin+-comm (m ∸ cutoff) cutoff) (compIso (Fin≅ cutoff+-) σ)) + + swapAut0≡0 : swapAut .fun fzero ≡ fzero + swapAut0≡0 = + σ .fun (finSubst cutoff+- (⊎.rec finCombine-inl finCombine-inr (fun ⊎-swap-Iso (finSplit (m ∸ cutoff) cutoff (0 , _))))) + ≡⟨ congS (λ z -> σ .fun (finSubst cutoff+- (⊎.rec (finCombine-inl {m = cutoff}) (finCombine-inr {m = cutoff}) (fun ⊎-swap-Iso z)))) (finSplit-beta-inl 0 0 𝔜 .car) where + module 𝔜 = M.CMonSEq 𝔜 𝔜-cmon + + f♯-hom = ArrayDef.Free.ext arrayDef isSet𝔜 (M.cmonSatMon 𝔜-cmon) f + + f♯ : Array A -> 𝔜 .car + f♯ = f♯-hom .fst + + f♯-η : (a : A) -> f♯ (η a) ≡ f a + f♯-η a i = ArrayDef.Free.ext-η arrayDef isSet𝔜 (M.cmonSatMon 𝔜-cmon) f i a + + f♯-hom-⊕ : (as bs : Array A) -> f♯ (as ⊕ bs) ≡ f♯ as 𝔜.⊕ f♯ bs + f♯-hom-⊕ as bs = + f♯ (as ⊕ bs) ≡⟨ sym ((f♯-hom .snd) M.`⊕ ⟪ as ⨾ bs ⟫) ⟩ + 𝔜 .alg (M.`⊕ , (λ w -> f♯ (⟪ as ⨾ bs ⟫ w))) ≡⟨ 𝔜.⊕-eta ⟪ as ⨾ bs ⟫ f♯ ⟩ + f♯ as 𝔜.⊕ f♯ bs ∎ + + f♯-comm : (as bs : Array A) -> f♯ (as ⊕ bs) ≡ f♯ (bs ⊕ as) + f♯-comm as bs = + f♯ (as ⊕ bs) ≡⟨ f♯-hom-⊕ as bs ⟩ + f♯ as 𝔜.⊕ f♯ bs ≡⟨ 𝔜.comm (f♯ as) (f♯ bs) ⟩ + f♯ bs 𝔜.⊕ f♯ as ≡⟨ sym (f♯-hom-⊕ bs as) ⟩ + f♯ (bs ⊕ as) ∎ + + swapAutToAut : ∀ {n} (zs : Fin (suc (suc n)) -> A) (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) + -> f♯ (suc (suc n) , zs ∘ swapAut σ .fun) ≡ f♯ (suc (suc n) , zs ∘ σ .fun) + swapAutToAut {n = n} zs σ = + f♯ (m , zs ∘ swapAut σ .fun) + ≡⟨ congS f♯ lemma-α ⟩ + f♯ (((m ∸ cutoff) , (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff _ ∘ inr)) + ⊕ (cutoff , (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff _ ∘ inl))) + ≡⟨ f♯-comm ((m ∸ cutoff) , (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff _ ∘ inr)) _ ⟩ + f♯ ((cutoff , (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff _ ∘ inl)) + ⊕ ((m ∸ cutoff) , (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff _ ∘ inr))) + ≡⟨ congS f♯ lemma-β ⟩ + f♯ (m , zs ∘ σ .fun) ∎ + where + m : ℕ + m = suc (suc n) + + cutoff : ℕ + cutoff = (σ .inv fzero) .fst + + cutoff< : cutoff < m + cutoff< = (σ .inv fzero) .snd + + cutoff+- : cutoff + (m ∸ cutoff) ≡ m + cutoff+- = ∸-lemma (<-weaken cutoff<) + + lemma-α : Path (Array A) (m , zs ∘ swapAut σ .fun) ((m ∸ cutoff) + cutoff , _) + lemma-α = Array≡ (sym cutoff+- ∙ +-comm cutoff _) λ k k ⊎.rec + (λ k + zs (σ .fun (finSubst cutoff+- (⊎.rec finCombine-inl finCombine-inr (fun ⊎-swap-Iso (finSplit (m ∸ cutoff) cutoff (k , _)))))) + ≡⟨ congS (λ z -> zs (σ .fun (finSubst cutoff+- (⊎.rec (finCombine-inl {m = cutoff}) finCombine-inr (fun ⊎-swap-Iso z))))) (finSplit-beta-inl k k + zs (σ .fun (finSubst cutoff+- (⊎.rec finCombine-inl finCombine-inr (fun ⊎-swap-Iso (finSplit (m ∸ cutoff) cutoff (k , _)))))) + ≡⟨ congS (λ z -> zs (σ .fun (finSubst cutoff+- (⊎.rec (finCombine-inl {m = cutoff}) finCombine-inr (fun ⊎-swap-Iso z))))) (finSplit-beta-inr k _ m∸cutoff≤k (∸-<-lemma (m ∸ cutoff) cutoff k k ⊎.rec + (λ k + ⊎.rec + (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff (m ∸ cutoff) ∘ inl) + (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff (m ∸ cutoff) ∘ inr) + (finSplit cutoff (m ∸ cutoff) (k , _)) + ≡⟨ congS (⊎.rec _ _) (finSplit-beta-inl k k + ⊎.rec + (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff (m ∸ cutoff) ∘ inl) + (zs ∘ σ .fun ∘ finSubst cutoff+- ∘ finCombine cutoff (m ∸ cutoff) ∘ inr) + (finSplit cutoff (m ∸ cutoff) (k , _)) + ≡⟨ congS (⊎.rec _ _) (finSplit-beta-inr k _ cutoff≤k (<-∸-< k m cutoff k A) (σ : Iso (Fin n) (Fin n)) -> f♯ (n , zs ∘ σ .fun) ≡ f♯ (n , zs) + permuteInvariant zero zs σ = + congS f♯ (Array≡ {f = zs ∘ σ .fun} {g = zs} refl \k k<0 -> ⊥.rec (¬-<-zero k<0)) + permuteInvariant (suc zero) zs σ = + congS f♯ (Array≡ {f = zs ∘ σ .fun} {g = zs} refl \k k<1 -> congS zs (isContr→isProp isContrFin1 _ _)) + permuteInvariant (suc (suc n)) zs σ = + let τ = swapAut σ ; τ-0≡0 = swapAut0≡0 σ + IH = permuteInvariant (suc n) (zs ∘ fsuc) (punchOutZero τ τ-0≡0) + in + f♯ (suc (suc n) , zs ∘ σ .fun) + ≡⟨ sym (swapAutToAut zs σ) ⟩ + f♯ (suc (suc n) , zs ∘ τ .fun) + ≡⟨ permuteInvariantOnZero τ τ-0≡0 IH ⟩ + f♯ (suc (suc n) , zs) ∎ + where + permuteInvariantOnZero : (τ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) (τ-0≡0 : τ .fun fzero ≡ fzero) -> (IH : _) + -> f♯ (suc (suc n) , zs ∘ τ .fun) ≡ f♯ (suc (suc n) , zs) + permuteInvariantOnZero τ τ-0≡0 IH = + f♯ (suc (suc n) , zs ∘ τ .fun) + ≡⟨⟩ + f (zs (τ .fun fzero)) 𝔜.⊕ f♯ (suc n , zs ∘ τ .fun ∘ fsuc) + ≡⟨ congS (\z -> f (zs z) 𝔜.⊕ f♯ (suc n , zs ∘ τ .fun ∘ fsuc)) (Fin-fst-≡ (congS fst τ-0≡0)) ⟩ + f (zs fzero) 𝔜.⊕ f♯ (suc n , zs ∘ τ .fun ∘ fsuc) + ≡⟨ congS (\z -> f (zs fzero) 𝔜.⊕ f♯ z) + (Array≡ {f = zs ∘ τ .fun ∘ fsuc} refl \k k≤n -> + congS (zs ∘ τ .fun ∘ fsuc) (Fin-fst-≡ refl) ∙ congS zs (punchOutZero≡fsuc τ τ-0≡0 (k , k≤n))) ⟩ + f (zs fzero) 𝔜.⊕ f♯ (suc n , zs ∘ fsuc ∘ punchOutZero τ τ-0≡0 .fun) + ≡⟨ cong₂ 𝔜._⊕_ (sym (f♯-η (zs fzero))) IH ⟩ + f♯ (η (zs fzero)) 𝔜.⊕ f♯ (suc n , zs ∘ fsuc) + ≡⟨ sym (f♯-hom-⊕ (η (zs fzero)) (suc n , zs ∘ fsuc)) ⟩ + f♯ (η (zs fzero) ⊕ (suc n , zs ∘ fsuc)) + ≡⟨ congS f♯ (η+fsuc zs) ⟩ + f♯ (suc (suc n) , zs) + ∎ + + ≈-resp-♯ : {as bs : Array A} -> as ≈ bs -> f♯ as ≡ f♯ bs + ≈-resp-♯ {as = n , g} {bs = m , h} (σ , p) = + f♯ (n , g) + ≡⟨ congS (λ z -> f♯ (n , z)) p ⟩ + f♯ (n , h ∘ σ .fun) + ≡⟨ congS f♯ (Array≡ n≡m λ _ _ -> refl) ⟩ + f♯ (m , h ∘ σ .fun ∘ (Fin≅ (sym n≡m)) .fun) + ≡⟨⟩ + f♯ (m , h ∘ (compIso (Fin≅ (sym n≡m)) σ) .fun) + ≡⟨ permuteInvariant m h (compIso (Fin≅ (sym n≡m)) σ) ⟩ + f♯ (m , h) ∎ + where + n≡m : n ≡ m + n≡m = Fin≅-inj σ + +module _ {ℓ} (A : Type ℓ) where + open import Cubical.Relation.Binary + module P = BinaryRelation {A = Array A} _≈_ + module R = isPermRel + + isPermRelPerm : isPermRel arrayDef (_≈_ {A = A}) + P.isEquivRel.reflexive (R.isEquivRel isPermRelPerm) _ = refl≈ + P.isEquivRel.symmetric (R.isEquivRel isPermRelPerm) _ _ = sym≈ + P.isEquivRel.transitive (R.isEquivRel isPermRelPerm) _ _ cs = trans≈ {cs = cs} + R.isCongruence isPermRelPerm {as} {bs} {cs} {ds} p q = cong≈ p q + R.isCommutative isPermRelPerm = comm≈ + R.resp-♯ isPermRelPerm {isSet𝔜 = isSet𝔜} 𝔜-cmon f p = ≈-resp-♯ isSet𝔜 𝔜-cmon f p + + PermRel : PermRelation arrayDef A + PermRel = _≈_ , isPermRelPerm + +module BagDef = F.Definition M.MonSig M.CMonEqSig M.CMonSEq + +bagFreeDef : ∀ {ℓ} -> BagDef.Free ℓ ℓ 2 +bagFreeDef = qFreeMonDef (PermRel _) diff --git a/Cubical/Structures/Set/CMon/Bag/ToCList.agda b/Cubical/Structures/Set/CMon/Bag/ToCList.agda new file mode 100644 index 0000000..fc48462 --- /dev/null +++ b/Cubical/Structures/Set/CMon/Bag/ToCList.agda @@ -0,0 +1,313 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Bag.ToCList where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.List renaming (_∷_ to _∷ₗ_) +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Sigma +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Set.Mon.Array as A +open import Cubical.Structures.Set.Mon.List as L +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity hiding (_/_) +open import Cubical.Structures.Set.CMon.QFreeMon +open import Cubical.Structures.Set.CMon.CList +open import Cubical.Structures.Set.CMon.Bag.Base +open import Cubical.Structures.Set.CMon.Bag.Free +open import Cubical.Relation.Nullary +open import Cubical.Data.Fin.LehmerCode hiding (LehmerCode; _∷_; []) + +open import Cubical.Structures.Set.CMon.Bag.FinExcept + +open Iso + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + n : ℕ + +module IsoToCList {ℓ} (A : Type ℓ) where + open import Cubical.Structures.Set.CMon.CList as CL + open import Cubical.HITs.SetQuotients as Q + open BagDef.Free + + module 𝔄 = M.MonSEq < Array A , array-α > array-sat + module 𝔅 = M.CMonSEq < Bag A , bagFreeDef .α > (bagFreeDef .sat) + module ℭ = M.CMonSEq < CList A , clist-α > clist-sat + + abstract needed so Agda wouldn't get stuck + fromCListHom : structHom < CList A , clist-α > < Bag A , bagFreeDef .α > + fromCListHom = ext clistDef squash/ (bagFreeDef .sat) (BagDef.Free.η bagFreeDef) + + fromCList : CList A -> Bag A + fromCList = fromCListHom .fst + + fromCListIsHom : structIsHom < CList A , clist-α > < Bag A , bagFreeDef .α > fromCList + fromCListIsHom = fromCListHom .snd + + fromCList-e : fromCList [] ≡ 𝔅.e + fromCList-e = refl + + fromCList-++ : ∀ xs ys -> fromCList (xs ℭ.⊕ ys) ≡ fromCList xs 𝔅.⊕ fromCList ys + fromCList-++ xs ys = + fromCList (xs ℭ.⊕ ys) ≡⟨ sym (fromCListIsHom M.`⊕ ⟪ xs ⨾ ys ⟫) ⟩ + _ ≡⟨ 𝔅.⊕-eta ⟪ xs ⨾ ys ⟫ fromCList ⟩ + _ ∎ + + fromCList-η : ∀ x -> fromCList (CL.[ x ]) ≡ Q.[ A.η x ] + fromCList-η x = congS (λ f -> f x) + (ext-η clistDef squash/ (bagFreeDef .sat) (BagDef.Free.η bagFreeDef)) + + ListToCListHom : structHom < List A , list-α > < CList A , clist-α > + ListToCListHom = ListDef.Free.ext listDef isSetCList (M.cmonSatMon clist-sat) CL.[_] + + ListToCList : List A -> CList A + ListToCList = ListToCListHom .fst + + ArrayToCListHom : structHom < Array A , array-α > < CList A , clist-α > + ArrayToCListHom = structHom∘ < Array A , array-α > < List A , list-α > < CList A , clist-α > + ListToCListHom ((arrayIsoToList .fun) , arrayIsoToListHom) + + ArrayToCList : Array A -> CList A + ArrayToCList = ArrayToCListHom .fst + + tab : ∀ n -> (Fin n -> A) -> CList A + tab = curry ArrayToCList + + isContr≅ : ∀ {ℓ} {A : Type ℓ} -> isContr A -> isContr (Iso A A) + isContr≅ ϕ = inhProp→isContr idIso \σ1 σ2 -> + SetsIso≡ (isContr→isOfHLevel 2 ϕ) (isContr→isOfHLevel 2 ϕ) + (funExt \x -> isContr→isProp ϕ (σ1 .fun x) (σ2 .fun x)) + (funExt \x -> isContr→isProp ϕ (σ1 .inv x) (σ2 .inv x)) + + isContrFin1≅ : isContr (Iso (Fin 1) (Fin 1)) + isContrFin1≅ = isContr≅ isContrFin1 + + fsuc∘punchOutZero≡ : ∀ {n} + -> (f g : Fin (suc (suc n)) -> A) + -> (σ : Iso (Fin (suc (suc n))) (Fin (suc (suc n)))) (p : f ≡ g ∘ σ .fun) + -> (q : σ .fun fzero ≡ fzero) + -> f ∘ fsuc ≡ g ∘ fsuc ∘ punchOutZero σ q .fun + fsuc∘punchOutZero≡ f g σ p q = + f ∘ fsuc ≡⟨ congS (_∘ fsuc) p ⟩ + g ∘ σ .fun ∘ fsuc ≡⟨ congS (g ∘_) (funExt (punchOutZero≡fsuc σ q)) ⟩ + g ∘ fsuc ∘ punchOutZero σ q .fun ∎ + + toCList-eq : ∀ n + -> (f : Fin n -> A) (g : Fin n -> A) (σ : Iso (Fin n) (Fin n)) (p : f ≡ g ∘ σ .fun) + -> tab n f ≡ tab n g + toCList-eq zero f g σ p = + refl + toCList-eq (suc zero) f g σ p = + let q : σ ≡ idIso + q = isContr→isProp isContrFin1≅ σ idIso + in congS (tab 1) (p ∙ congS (g ∘_) (congS Iso.fun q)) + toCList-eq (suc (suc n)) f g σ p = + ⊎.rec + (λ 0≡σ0 -> + let IH = toCList-eq (suc n) (f ∘ fsuc) (g ∘ fsuc) (punchOutZero σ (sym 0≡σ0)) (fsuc∘punchOutZero≡ f g σ p (sym 0≡σ0)) + in case1 IH (sym 0≡σ0) + ) + (λ (k , Sk≡σ0) -> + case2 k (sym Sk≡σ0) + (toCList-eq (suc n) (f ∘ fsuc) ((g ∼) ∘ pIn (fsuc k)) (punch-σ σ) (sym (IH1-lemma k Sk≡σ0))) + (toCList-eq (suc n) (g-σ k) (g ∘ fsuc) (fill-σ k) (sym (funExt (IH2-lemma k Sk≡σ0)))) + ) + (fsplit (σ .fun fzero)) + where + g-σ : ∀ k -> Fin (suc n) -> A + g-σ k (zero , p) = g (σ .fun fzero) + g-σ k (suc j , p) = (g ∼) (1+ (pIn k (j , predℕ-≤-predℕ p))) + + IH1-lemma : ∀ k -> fsuc k ≡ σ .fun fzero -> (g ∼) ∘ pIn (fsuc k) ∘ punch-σ σ .fun ≡ f ∘ fsuc + IH1-lemma k Sk≡σ0 = + (g ∼) ∘ pIn (fsuc k) ∘ punch-σ σ .fun + ≡⟨ congS (λ z -> (g ∼) ∘ pIn z ∘ punch-σ σ .fun) Sk≡σ0 ⟩ + (g ∼) ∘ pIn (σ .fun fzero) ∘ punch-σ σ .fun + ≡⟨⟩ + (g ∼) ∘ pIn (σ .fun fzero) ∘ pOut (σ .fun fzero) ∘ ((G .fun σ) .snd) .fun ∘ pIn fzero + ≡⟨ congS (λ h -> (g ∼) ∘ h ∘ ((G .fun σ) .snd) .fun ∘ (invIso pIso) .fun) (funExt (pIn∘Out (σ .fun fzero))) ⟩ + (g ∼) ∘ equivIn σ .fun ∘ pIn fzero + ≡⟨⟩ + g ∘ σ .fun ∘ fst ∘ pIn fzero + ≡⟨ congS (_∘ fst ∘ pIn fzero) (sym p) ⟩ + f ∘ fst ∘ pIn fzero + ≡⟨ congS (f ∘_) (funExt pInZ≡fsuc) ⟩ + f ∘ fsuc ∎ + + IH2-lemma : ∀ k -> fsuc k ≡ σ .fun fzero -> (j : Fin (suc n)) -> g (fsuc (fill-σ k .fun j)) ≡ (g-σ k) j + IH2-lemma k Sk≡σ0 (zero , r) = congS g Sk≡σ0 + IH2-lemma k Sk≡σ0 (suc j , r) = + g (fsuc (equivOut {k = k} (compIso pIso (invIso pIso)) .fun (suc j , r))) + ≡⟨⟩ + g (fsuc (equivOut {k = k} (compIso pIso (invIso pIso)) .fun (j' .fst))) + ≡⟨ congS (g ∘ fsuc) (equivOut-beta-α {σ = compIso pIso (invIso pIso)} j') ⟩ + g (fsuc (fst (pIn k (pOut fzero j')))) + ≡⟨⟩ + g (fsuc (fst (pIn k (⊎.rec _ (λ k predℕ (suc j) , _) (suc j σ .fun fzero ≡ fzero + -> tab (suc (suc n)) f ≡ tab (suc (suc n)) g + case1 IH σ0≡0 = + tab (suc (suc n)) f ≡⟨⟩ + f fzero ∷ tab (suc n) (f ∘ fsuc) ≡⟨ congS (_∷ tab (suc n) (f ∘ fsuc)) (funExt⁻ p fzero) ⟩ + g (σ .fun fzero) ∷ tab (suc n) (f ∘ fsuc) ≡⟨ congS (\k -> g k ∷ tab (suc n) (f ∘ fsuc)) σ0≡0 ⟩ + g fzero ∷ tab (suc n) (f ∘ fsuc) ≡⟨ congS (g fzero ∷_) IH ⟩ + g fzero ∷ tab (suc n) (g ∘ fsuc) ≡⟨⟩ + tab (suc (suc n)) g ∎ + case2 : (k : Fin (suc n)) + -> σ .fun fzero ≡ fsuc k + -> tab (suc n) (f ∘ fsuc) ≡ tab (suc n) ((g ∼) ∘ pIn (fsuc k)) + -> tab (suc n) (g-σ k) ≡ tab (suc n) (g ∘ fsuc) + -> tab (suc (suc n)) f ≡ tab (suc (suc n)) g + case2 k σ0≡Sk IH1 IH2 = + comm (f fzero) (g fzero) (tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc)) + (sym (eqn1 IH1)) (sym (eqn2 IH2)) + where + eqn1 : tab (suc n) (f ∘ fsuc) ≡ tab (suc n) ((g ∼) ∘ pIn (fsuc k)) + -> g fzero ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc) ≡ tab (suc n) (f ∘ fsuc) + eqn1 IH = + g fzero ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc) + ≡⟨ congS (λ z -> g z ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc)) (Fin-fst-≡ refl) ⟩ + ((g ∼) ∘ pIn (fsuc k)) fzero ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc) + ≡⟨⟩ + tab (suc n) ((g ∼) ∘ pIn (fsuc k)) + ≡⟨ sym IH ⟩ + tab (suc n) (f ∘ fsuc) ∎ + + g-σ≡ : tab (suc n) (g-σ k) ≡ g (σ .fun fzero) ∷ tab n ((g ∼) ∘ 1+_ ∘ pIn k) + g-σ≡ = congS (λ z -> g (σ .fun fzero) ∷ tab n z) + (funExt λ z -> congS {x = (fst z , pred-≤-pred (snd (fsuc z)))} ((g ∼) ∘ 1+_ ∘ pIn k) (Fin-fst-≡ refl)) + + eqn2 : tab (suc n) (g-σ k) ≡ tab (suc n) (g ∘ fsuc) + -> f fzero ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc) ≡ tab (suc n) (g ∘ fsuc) + eqn2 IH2 = + f fzero ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc) + ≡⟨ congS (λ h -> h fzero ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc)) p ⟩ + g (σ .fun fzero) ∷ tab n ((g ∼) ∘ pIn (fsuc k) ∘ fsuc) + ≡⟨ congS (λ h -> g (σ .fun fzero) ∷ tab n ((g ∼) ∘ h)) (sym pIn-fsuc-nat) ⟩ + g (σ .fun fzero) ∷ tab n ((g ∼) ∘ 1+_ ∘ pIn k) + ≡⟨ sym g-σ≡ ⟩ + tab (suc n) (g-σ k) + ≡⟨ IH2 ⟩ + tab (suc n) (g ∘ fsuc) ∎ + + abstract + toCList-eq' : ∀ n m f g -> (r : (n , f) ≈ (m , g)) -> tab n f ≡ tab m g + toCList-eq' n m f g (σ , p) = + tab n f ≡⟨ toCList-eq n f (g ∘ (finSubst n≡m)) (compIso σ (Fin≅ (sym n≡m))) (sym lemma-α) ⟩ + (uncurry tab) (n , g ∘ finSubst n≡m) ≡⟨ congS (uncurry tab) (Array≡ n≡m λ _ _ -> congS g (Fin-fst-≡ refl)) ⟩ + (uncurry tab) (m , g) ∎ + where + n≡m : n ≡ m + n≡m = ≈-length σ + lemma-α : g ∘ finSubst n≡m ∘ (Fin≅ (sym n≡m)) .fun ∘ σ .fun ≡ f + lemma-α = + g ∘ finSubst n≡m ∘ finSubst (sym n≡m) ∘ σ .fun + ≡⟨ congS {x = finSubst n≡m ∘ finSubst (sym n≡m)} {y = idfun _} (λ h -> g ∘ h ∘ σ .fun) (funExt (λ x -> Fin-fst-≡ refl)) ⟩ + g ∘ σ .fun + ≡⟨ sym p ⟩ + f ∎ + + abstract + toCList : Bag A -> CList A + toCList Q.[ (n , f) ] = tab n f + toCList (eq/ (n , f) (m , g) r i) = toCList-eq' n m f g r i + toCList (squash/ xs ys p q i j) = + isSetCList (toCList xs) (toCList ys) (congS toCList p) (congS toCList q) i j + + toCList-η : (xs : Array A) -> toCList Q.[ xs ] ≡ ArrayToCList xs + toCList-η xs = refl + + toCList-e : toCList 𝔅.e ≡ CL.[] + toCList-e = refl + + toCList-++ : ∀ xs ys -> toCList (xs 𝔅.⊕ ys) ≡ toCList xs ℭ.⊕ toCList ys + toCList-++ = + elimProp (λ _ -> isPropΠ (λ _ -> isSetCList _ _)) λ xs -> + elimProp (λ _ -> isSetCList _ _) λ ys -> + sym (ArrayToCListHom .snd M.`⊕ ⟪ xs ⨾ ys ⟫) + + toCList∘fromCList-η : ∀ x -> toCList (fromCList CL.[ x ]) ≡ CL.[ x ] + toCList∘fromCList-η x = refl + + fromCList∘toCList-η : ∀ x -> fromCList (toCList Q.[ A.η x ]) ≡ Q.[ A.η x ] + fromCList∘toCList-η x = fromCList-η x + + toCList-fromCList : ∀ xs -> toCList (fromCList xs) ≡ xs + toCList-fromCList = + elimCListProp.f _ + (congS toCList fromCList-e ∙ toCList-e) + (λ x {xs} p -> + toCList (fromCList (x ∷ xs)) ≡⟨ congS toCList (fromCList-++ CL.[ x ] xs) ⟩ + toCList (fromCList CL.[ x ] 𝔅.⊕ fromCList xs) ≡⟨ toCList-++ (fromCList CL.[ x ]) (fromCList xs) ⟩ + toCList (fromCList CL.[ x ]) ℭ.⊕ toCList (fromCList xs) ≡⟨ congS (toCList (fromCList CL.[ x ]) ℭ.⊕_) p ⟩ + toCList (fromCList CL.[ x ]) ℭ.⊕ xs ≡⟨ congS {x = toCList (fromCList CL.[ x ])} {y = CL.[ x ]} (ℭ._⊕ xs) (toCList∘fromCList-η x) ⟩ + CL.[ x ] ℭ.⊕ xs + ∎) + (isSetCList _ _) + + fromList-toCList : ∀ xs -> fromCList (toCList xs) ≡ xs + fromList-toCList = elimProp (λ _ -> squash/ _ _) (uncurry lemma) + where + lemma : (n : ℕ) (f : Fin n -> A) -> fromCList (toCList Q.[ n , f ]) ≡ Q.[ n , f ] + lemma zero f = + fromCList (toCList Q.[ zero , f ]) ≡⟨ congS fromCList (toCList-η (zero , f)) ⟩ + fromCList [] ≡⟨ fromCList-e ⟩ + 𝔅.e ≡⟨ congS Q.[_] (e-eta _ (zero , f) refl refl) ⟩ + Q.[ zero , f ] ∎ + lemma (suc n) f = + fromCList (toCList Q.[ suc n , f ]) + ≡⟨ congS fromCList (toCList-η (suc n , f)) ⟩ + fromCList (ArrayToCList (suc n , f)) + ≡⟨ congS (fromCList ∘ ArrayToCList) (sym (η+fsuc f)) ⟩ + fromCList (ArrayToCList (A.η (f fzero) ⊕ (n , f ∘ fsuc))) + ≡⟨ congS fromCList $ sym (ArrayToCListHom .snd M.`⊕ ⟪ A.η (f fzero) ⨾ (n , f ∘ fsuc) ⟫) ⟩ + fromCList (f fzero ∷ ArrayToCList (n , f ∘ fsuc)) + ≡⟨ fromCList-++ CL.[ f fzero ] (ArrayToCList (n , f ∘ fsuc)) ⟩ + fromCList CL.[ f fzero ] 𝔅.⊕ fromCList (ArrayToCList (n , f ∘ fsuc)) + ≡⟨ congS (𝔅._⊕ fromCList (ArrayToCList (n , f ∘ fsuc))) (fromCList-η (f fzero)) ⟩ + Q.[ A.η (f fzero) ] 𝔅.⊕ fromCList (ArrayToCList (n , f ∘ fsuc)) + ≡⟨ congS (λ zs -> Q.[ A.η (f fzero) ] 𝔅.⊕ fromCList zs) (sym (toCList-η (n , f ∘ fsuc))) ⟩ + Q.[ A.η (f fzero) ] 𝔅.⊕ fromCList (toCList Q.[ n , f ∘ fsuc ]) + ≡⟨ congS (Q.[ A.η (f fzero) ] 𝔅.⊕_) (lemma n (f ∘ fsuc)) ⟩ + Q.[ A.η (f fzero) ] 𝔅.⊕ Q.[ n , f ∘ fsuc ] + ≡⟨ QFreeMon.[ A ]-isMonHom (PermRel A) .snd M.`⊕ ⟪ _ ⨾ _ ⟫ ⟩ + Q.[ A.η (f fzero) 𝔄.⊕ (n , f ∘ fsuc) ] + ≡⟨ congS Q.[_] (η+fsuc f) ⟩ + Q.[ suc n , f ] ∎ + + BagToCList : Iso (Bag A) (CList A) + BagToCList = iso toCList fromCList toCList-fromCList fromList-toCList + +bagDef' : ∀ {ℓ ℓ'} -> BagDef.Free ℓ ℓ' 2 +bagDef' {ℓ = ℓ} {ℓ' = ℓ'} = BagDef.isoAux .fun (Bag , bagFreeAux) + where + clistFreeAux : BagDef.FreeAux ℓ ℓ' 2 CList + clistFreeAux = (inv BagDef.isoAux clistDef) .snd + + bagFreeAux : BagDef.FreeAux ℓ ℓ' 2 Bag + bagFreeAux = subst (BagDef.FreeAux ℓ ℓ' 2) + (funExt λ X -> isoToPath $ invIso (IsoToCList.BagToCList X)) clistFreeAux diff --git a/Cubical/Structures/Set/CMon/CList.agda b/Cubical/Structures/Set/CMon/CList.agda new file mode 100644 index 0000000..7de6c5f --- /dev/null +++ b/Cubical/Structures/Set/CMon/CList.agda @@ -0,0 +1,178 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.CList where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +import Cubical.Data.List as L + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity + +infixr 20 _∷_ + +data CList {a} (A : Type a) : Type a where + [] : CList A + _∷_ : (a : A) -> (as : CList A) -> CList A + comm : (a b : A) + -> {as bs : CList A} (cs : CList A) + -> (p : as ≡ b ∷ cs) (q : bs ≡ a ∷ cs) + -> a ∷ as ≡ b ∷ bs + isSetCList : isSet (CList A) + +pattern [_] a = a ∷ [] + +module elimCListSet {ℓ p : Level} {A : Type ℓ} (P : CList A -> Type p) + ([]* : P []) + (_∷*_ : (x : A) {xs : CList A} -> P xs -> P (x ∷ xs)) + (comm* : (a b : A) + -> {as bs : CList A} {cs : CList A} + -> {as* : P as} + -> {bs* : P bs} + -> (cs* : P cs) + -> {p : as ≡ b ∷ cs} {q : bs ≡ a ∷ cs} + -> (bp : PathP (λ i -> P (p i)) as* (b ∷* cs*)) + -> (bq : PathP (λ i -> P (q i)) bs* (a ∷* cs*)) + -> PathP (λ i -> P (comm a b cs p q i)) (a ∷* as*) (b ∷* bs*) + ) + (isSetCList* : {xs : CList A} -> isSet (P xs)) + where + f : (xs : CList A) -> P xs + f [] = []* + f (a ∷ as) = a ∷* f as + f (comm a b {as} {bs} cs p q i) = + comm* a b (f cs) (cong f p) (cong f q) i + f (isSetCList xs ys p q i j) = + isOfHLevel→isOfHLevelDep 2 (λ xs -> isSetCList* {xs = xs}) (f xs) (f ys) (cong f p) (cong f q) (isSetCList xs ys p q) i j + +module elimCListProp {ℓ p : Level} {A : Type ℓ} (P : CList A -> Type p) + ([]* : P []) + (_∷*_ : (x : A) {xs : CList A} -> P xs -> P (x ∷ xs)) + (isSetCList* : {xs : CList A} -> isProp (P xs)) + where + f : (xs : CList A) -> P xs + f = elimCListSet.f P []* _∷*_ + (λ a b {as} {bs} {cs} {as*} {bs*} cs* bp bq -> toPathP (isSetCList* _ (b ∷* bs*))) + (isProp→isSet isSetCList*) + +private + variable + ℓ : Level + A : Type ℓ + +_++_ : CList A -> CList A -> CList A +[] ++ bs = bs +(a ∷ as) ++ bs = a ∷ (as ++ bs) +comm a b cs p q i ++ bs = comm a b (cs ++ bs) (cong (_++ bs) p) (cong (_++ bs) q) i +isSetCList a b p q i j ++ bs = isSetCList (a ++ bs) (b ++ bs) (cong (_++ bs) p) (cong (_++ bs) q) i j + +++-unitl : (as : CList A) -> [] ++ as ≡ as +++-unitl as = refl + +++-unitr : (as : CList A) -> as ++ [] ≡ as +++-unitr = elimCListProp.f _ refl (λ a p -> cong (a ∷_) p) (isSetCList _ _) + +++-assocr : (as bs cs : CList A) -> (as ++ bs) ++ cs ≡ as ++ (bs ++ cs) +++-assocr = elimCListProp.f _ + (λ _ _ -> refl) + (λ x p bs cs -> cong (x ∷_) (p bs cs)) + (isPropΠ λ _ -> isPropΠ λ _ -> isSetCList _ _) + +swap : (a b : A) (cs : CList A) -> a ∷ b ∷ cs ≡ b ∷ a ∷ cs +swap a b cs = comm a b cs refl refl + +++-∷ : (a : A) (as : CList A) -> a ∷ as ≡ as ++ [ a ] +++-∷ a = elimCListProp.f (λ as -> a ∷ as ≡ as ++ [ a ]) + refl + (λ b {as} p -> swap a b as ∙ cong (b ∷_) p) + (isSetCList _ _) + +++-comm : (as bs : CList A) -> as ++ bs ≡ bs ++ as +++-comm = elimCListProp.f _ + (sym ∘ ++-unitr) + (λ a {as} p bs -> cong (a ∷_) (p bs) ∙ cong (_++ as) (++-∷ a bs) ∙ ++-assocr bs [ a ] as) + (isPropΠ λ _ -> isSetCList _ _) + +clist-α : ∀ {n : Level} {X : Type n} -> sig M.MonSig (CList X) -> CList X +clist-α (M.`e , i) = [] +clist-α (M.`⊕ , i) = i fzero ++ i fone + +module Free {x y : Level} {A : Type x} {𝔜 : struct y M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-cmon : 𝔜 ⊨ M.CMonSEq) where + module 𝔜 = M.CMonSEq 𝔜 𝔜-cmon + + 𝔛 : M.CMonStruct + 𝔛 = < CList A , clist-α > + + module _ (f : A -> 𝔜 .car) where + _♯ : CList A -> 𝔜 .car + _♯ = elimCListSet.f _ + 𝔜.e + (λ a p -> f a 𝔜.⊕ p) + (λ a b {ab} {bs} {cs} {as*} {bs*} cs* bp bq -> + f a 𝔜.⊕ as* ≡⟨ cong (f a 𝔜.⊕_) bp ⟩ + f a 𝔜.⊕ f b 𝔜.⊕ cs* ≡⟨ sym (𝔜.assocr _ _ _) ⟩ + (f a 𝔜.⊕ f b) 𝔜.⊕ cs* ≡⟨ cong (𝔜._⊕ cs*) (𝔜.comm _ _) ⟩ + (f b 𝔜.⊕ f a) 𝔜.⊕ cs* ≡⟨ 𝔜.assocr _ _ _ ⟩ + f b 𝔜.⊕ (f a 𝔜.⊕ cs*) ≡⟨ cong (f b 𝔜.⊕_) (sym bq) ⟩ + f b 𝔜.⊕ bs* ∎ + ) + isSet𝔜 + + private + ♯-++ : ∀ xs ys -> (xs ++ ys) ♯ ≡ (xs ♯) 𝔜.⊕ (ys ♯) + ♯-++ = elimCListProp.f _ + (λ ys -> sym (𝔜.unitl (ys ♯))) + (λ a {xs} p ys -> + f a 𝔜.⊕ ((xs ++ ys) ♯) ≡⟨ cong (f a 𝔜.⊕_) (p ys) ⟩ + f a 𝔜.⊕ ((xs ♯) 𝔜.⊕ (ys ♯)) ≡⟨ sym (𝔜.assocr (f a) (xs ♯) (ys ♯)) ⟩ + _ ∎ + ) + (isPropΠ λ _ -> isSet𝔜 _ _) + + ♯-isMonHom : structHom 𝔛 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ ∙ sym (♯-++ (i fzero) (i fone)) + + private + clistEquivLemma : (g : structHom 𝔛 𝔜) -> (x : CList A) -> g .fst x ≡ ((g .fst ∘ [_]) ♯) x + clistEquivLemma (g , homMonWit) = elimCListProp.f _ + (sym (homMonWit M.`e (lookup L.[])) ∙ 𝔜.e-eta) + (λ x {xs} p -> + g (x ∷ xs) ≡⟨ sym (homMonWit M.`⊕ (lookup ([ x ] L.∷ xs L.∷ L.[]))) ⟩ + _ ≡⟨ 𝔜.⊕-eta (lookup ([ x ] L.∷ xs L.∷ L.[])) g ⟩ + _ ≡⟨ cong (g [ x ] 𝔜.⊕_) p ⟩ + _ ∎ + ) + (isSet𝔜 _ _) + + clistEquivLemma-β : (g : structHom 𝔛 𝔜) -> g ≡ ♯-isMonHom (g .fst ∘ [_]) + clistEquivLemma-β g = structHom≡ 𝔛 𝔜 g (♯-isMonHom (g .fst ∘ [_])) isSet𝔜 (funExt (clistEquivLemma g)) + + clistMonEquiv : structHom 𝔛 𝔜 ≃ (A -> 𝔜 .car) + clistMonEquiv = + isoToEquiv (iso (λ g -> g .fst ∘ [_]) ♯-isMonHom (λ g -> funExt (𝔜.unitr ∘ g)) (sym ∘ clistEquivLemma-β)) + +module CListDef = F.Definition M.MonSig M.CMonEqSig M.CMonSEq + +clist-sat : ∀ {n} {X : Type n} -> < CList X , clist-α > ⊨ M.CMonSEq +clist-sat (M.`mon M.`unitl) ρ = ++-unitl (ρ fzero) +clist-sat (M.`mon M.`unitr) ρ = ++-unitr (ρ fzero) +clist-sat (M.`mon M.`assocr) ρ = ++-assocr (ρ fzero) (ρ fone) (ρ ftwo) +clist-sat M.`comm ρ = ++-comm (ρ fzero) (ρ fone) + +clistDef : ∀ {ℓ ℓ'} -> CListDef.Free ℓ ℓ' 2 +F.Definition.Free.F clistDef = CList +F.Definition.Free.η clistDef = [_] +F.Definition.Free.α clistDef = clist-α +F.Definition.Free.sat clistDef = clist-sat +F.Definition.Free.isFree clistDef isSet𝔜 satMon = (Free.clistMonEquiv isSet𝔜 satMon) .snd diff --git a/Cubical/Structures/Set/CMon/Desc.agda b/Cubical/Structures/Set/CMon/Desc.agda new file mode 100644 index 0000000..614685f --- /dev/null +++ b/Cubical/Structures/Set/CMon/Desc.agda @@ -0,0 +1,103 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Desc where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.List +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Functions.Logic as L +open import Cubical.Structures.Arity as F public + +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq + +import Cubical.Structures.Set.Mon.Desc as M + +open M.MonSym + +CMonSym = M.MonSym +CMonAr = M.MonAr +CMonFinSig = M.MonFinSig +CMonSig = M.MonSig + +CMonStruct : ∀ {n} -> Type (ℓ-suc n) +CMonStruct {n} = struct n CMonSig + +CMon→Mon : ∀ {ℓ} -> CMonStruct {ℓ} -> M.MonStruct {ℓ} +car (CMon→Mon 𝔛) = 𝔛 .car +alg (CMon→Mon 𝔛) = 𝔛 .alg + +module CMonStruct {ℓ} (𝔛 : CMonStruct {ℓ}) where + open M.MonStruct (CMon→Mon 𝔛) public + +data CMonEq : Type where + `mon : M.MonEq -> CMonEq + `comm : CMonEq + +CMonEqFree : CMonEq -> ℕ +CMonEqFree (`mon eqn) = M.MonEqFree eqn +CMonEqFree `comm = 2 + +CMonEqSig : EqSig ℓ-zero ℓ-zero +CMonEqSig = finEqSig (CMonEq , CMonEqFree) + +cmonEqLhs : (eq : CMonEq) -> FinTree CMonFinSig (CMonEqFree eq) +cmonEqLhs (`mon eqn) = M.monEqLhs eqn +cmonEqLhs `comm = node (`⊕ , lookup (leaf fzero ∷ leaf fone ∷ [])) + +cmonEqRhs : (eq : CMonEq) -> FinTree CMonFinSig (CMonEqFree eq) +cmonEqRhs (`mon eqn) = M.monEqRhs eqn +cmonEqRhs `comm = node (`⊕ , lookup (leaf fone ∷ leaf fzero ∷ [])) + +CMonSEq : seq CMonSig CMonEqSig +CMonSEq n = cmonEqLhs n , cmonEqRhs n + +cmonSatMon : ∀ {s} {str : struct s CMonSig} -> str ⊨ CMonSEq -> str ⊨ M.MonSEq +cmonSatMon {_} {str} cmonSat eqn ρ = cmonSat (`mon eqn) ρ + +module CMonSEq {ℓ} (𝔛 : CMonStruct {ℓ}) (ϕ : 𝔛 ⊨ CMonSEq) where + open M.MonSEq (CMon→Mon 𝔛) (cmonSatMon ϕ) public + + comm : ∀ m n -> m ⊕ n ≡ n ⊕ m + comm m n = + m ⊕ n + ≡⟨⟩ + 𝔛 .alg (`⊕ , lookup (m ∷ n ∷ [])) + ≡⟨ cong (λ z -> 𝔛 .alg (`⊕ , z)) (funExt lemma1) ⟩ + 𝔛 .alg (`⊕ , (λ x -> sharp CMonSig 𝔛 (lookup (m ∷ n ∷ [])) (lookup (leaf fzero ∷ leaf fone ∷ []) x))) + ≡⟨ ϕ `comm (lookup (m ∷ n ∷ [])) ⟩ + 𝔛 .alg (`⊕ , (λ x -> sharp CMonSig 𝔛 (lookup (m ∷ n ∷ [])) (lookup (leaf fone ∷ leaf fzero ∷ []) x))) + ≡⟨ cong (λ z -> 𝔛 .alg (`⊕ , z)) (sym (funExt lemma2)) ⟩ + 𝔛 .alg (`⊕ , lookup (n ∷ m ∷ [])) + ≡⟨⟩ + n ⊕ m ∎ + where + lemma1 : (w : CMonSig .arity `⊕) -> lookup (m ∷ n ∷ []) w ≡ sharp CMonSig 𝔛 (lookup (m ∷ n ∷ [])) (lookup (leaf fzero ∷ leaf fone ∷ []) w) + lemma1 (zero , p) = refl + lemma1 (suc zero , p) = refl + lemma1 (suc (suc n) , p) = ⊥.rec (¬m+n lookup (n ∷ m ∷ []) w ≡ sharp CMonSig 𝔛 (lookup (m ∷ n ∷ [])) (lookup (leaf fone ∷ leaf fzero ∷ []) w) + lemma2 (zero , p) = refl + lemma2 (suc zero , p) = refl + lemma2 (suc (suc n) , p) = ⊥.rec (¬m+n M.⊔-MonStr ℓ ⊨ CMonSEq +⊔-MonStr-CMonSEq ℓ (`mon eqn) ρ = M.⊔-MonStr-MonSEq ℓ eqn ρ +⊔-MonStr-CMonSEq ℓ `comm ρ = ⊔-comm (ρ fzero) (ρ fone) + +⊓-MonStr-CMonSEq : (ℓ : Level) -> M.⊓-MonStr ℓ ⊨ CMonSEq +⊓-MonStr-CMonSEq ℓ (`mon eqn) ρ = M.⊓-MonStr-MonSEq ℓ eqn ρ +⊓-MonStr-CMonSEq ℓ `comm ρ = ⊓-comm (ρ fzero) (ρ fone) \ No newline at end of file diff --git a/Cubical/Structures/Set/CMon/Free.agda b/Cubical/Structures/Set/CMon/Free.agda new file mode 100644 index 0000000..f131770 --- /dev/null +++ b/Cubical/Structures/Set/CMon/Free.agda @@ -0,0 +1,142 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.Free where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.List +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.Mon.Free as FM +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity + +data FreeCMon {ℓ : Level} (A : Type ℓ) : Type ℓ where + η : (a : A) -> FreeCMon A + e : FreeCMon A + _⊕_ : FreeCMon A -> FreeCMon A -> FreeCMon A + unitl : ∀ m -> e ⊕ m ≡ m + unitr : ∀ m -> m ⊕ e ≡ m + assocr : ∀ m n o -> (m ⊕ n) ⊕ o ≡ m ⊕ (n ⊕ o) + comm : ∀ m n -> m ⊕ n ≡ n ⊕ m + trunc : isSet (FreeCMon A) + +module elimFreeCMonSet {p n : Level} {A : Type n} (P : FreeCMon A -> Type p) + (η* : (a : A) -> P (η a)) + (e* : P e) + (_⊕*_ : {m n : FreeCMon A} -> P m -> P n -> P (m ⊕ n)) + (unitl* : {m : FreeCMon A} (m* : P m) -> PathP (λ i → P (unitl m i)) (e* ⊕* m*) m*) + (unitr* : {m : FreeCMon A} (m* : P m) -> PathP (λ i → P (unitr m i)) (m* ⊕* e*) m*) + (assocr* : {m n o : FreeCMon A} + (m* : P m) -> + (n* : P n) -> + (o* : P o) -> + PathP (λ i → P (assocr m n o i)) ((m* ⊕* n*) ⊕* o*) (m* ⊕* (n* ⊕* o*))) + (comm* : {m n : FreeCMon A} + (m* : P m) -> + (n* : P n) -> + PathP (λ i → P (comm m n i)) (m* ⊕* n*) (n* ⊕* m*)) + (trunc* : {xs : FreeCMon A} -> isSet (P xs)) + where + f : (x : FreeCMon A) -> P x + f (η a) = η* a + f e = e* + f (x ⊕ y) = f x ⊕* f y + f (unitl x i) = unitl* (f x) i + f (unitr x i) = unitr* (f x) i + f (assocr x y z i) = assocr* (f x) (f y) (f z) i + f (comm x y i) = comm* (f x) (f y) i + f (trunc xs ys p q i j) = + isOfHLevel→isOfHLevelDep 2 (\xs -> trunc* {xs = xs}) (f xs) (f ys) (cong f p) (cong f q) (trunc xs ys p q) i j + +module elimFreeCMonProp {p n : Level} {A : Type n} (P : FreeCMon A -> Type p) + (η* : (a : A) -> P (η a)) + (e* : P e) + (_⊕*_ : {m n : FreeCMon A} -> P m -> P n -> P (m ⊕ n)) + (trunc* : {xs : FreeCMon A} -> isProp (P xs)) + where + f : (x : FreeCMon A) -> P x + f = elimFreeCMonSet.f P η* e* _⊕*_ unitl* unitr* assocr* comm* (isProp→isSet trunc*) + where + abstract + unitl* : {m : FreeCMon A} (m* : P m) -> PathP (λ i → P (unitl m i)) (e* ⊕* m*) m* + unitl* {m} m* = toPathP (trunc* (transp (λ i -> P (unitl m i)) i0 (e* ⊕* m*)) m*) + unitr* : {m : FreeCMon A} (m* : P m) -> PathP (λ i → P (unitr m i)) (m* ⊕* e*) m* + unitr* {m} m* = toPathP (trunc* (transp (λ i -> P (unitr m i)) i0 (m* ⊕* e*)) m*) + assocr* : {m n o : FreeCMon A} + (m* : P m) -> + (n* : P n) -> + (o* : P o) -> PathP (λ i → P (assocr m n o i)) ((m* ⊕* n*) ⊕* o*) (m* ⊕* (n* ⊕* o*)) + assocr* {m} {n} {o} m* n* o* = + toPathP (trunc* (transp (λ i -> P (assocr m n o i)) i0 ((m* ⊕* n*) ⊕* o*)) (m* ⊕* (n* ⊕* o*))) + comm* : {m n : FreeCMon A} (m* : P m) (n* : P n) -> PathP (λ i → P (comm m n i)) (m* ⊕* n*) (n* ⊕* m*) + comm* {m} {n} m* n* = toPathP (trunc* (transp (λ i -> P (comm m n i)) i0 (m* ⊕* n*)) (n* ⊕* m*)) + +freeCMon-α : ∀ {ℓ} {X : Type ℓ} -> sig M.MonSig (FreeCMon X) -> FreeCMon X +freeCMon-α (M.`e , _) = e +freeCMon-α (M.`⊕ , i) = i fzero ⊕ i fone + +module Free {x y : Level} {A : Type x} {𝔜 : struct y M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-cmon : 𝔜 ⊨ M.CMonSEq) where + module 𝔜 = M.CMonSEq 𝔜 𝔜-cmon + + 𝔉 : struct x M.MonSig + 𝔉 = < FreeCMon A , freeCMon-α > + + module _ (f : A -> 𝔜 .car) where + _♯ : FreeCMon A -> 𝔜 .car + _♯ (η a) = f a + _♯ e = 𝔜.e + _♯ (m ⊕ n) = (m ♯) 𝔜.⊕ (n ♯) + _♯ (unitl m i) = 𝔜.unitl (m ♯) i + _♯ (unitr m i) = 𝔜.unitr (m ♯) i + _♯ (assocr m n o i) = 𝔜.assocr (m ♯) (n ♯) (o ♯) i + comm m n i ♯ = 𝔜.comm (m ♯) (n ♯) i + (trunc m n p q i j) ♯ = isSet𝔜 (m ♯) (n ♯) (cong _♯ p) (cong _♯ q) i j + + ♯-isMonHom : structHom 𝔉 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ + + private + freeCMonEquivLemma : (g : structHom 𝔉 𝔜) -> (x : FreeCMon A) -> g .fst x ≡ ((g .fst ∘ η) ♯) x + freeCMonEquivLemma (g , homMonWit) = elimFreeCMonProp.f (λ x -> g x ≡ ((g ∘ η) ♯) x) + (λ _ -> refl) + (sym (homMonWit M.`e (lookup [])) ∙ 𝔜.e-eta) + (λ {m} {n} p q -> + g (m ⊕ n) ≡⟨ sym (homMonWit M.`⊕ (lookup (m ∷ n ∷ []))) ⟩ + 𝔜 .alg (M.`⊕ , (λ w -> g (lookup (m ∷ n ∷ []) w))) ≡⟨ 𝔜.⊕-eta (lookup (m ∷ n ∷ [])) g ⟩ + g m 𝔜.⊕ g n ≡⟨ cong₂ 𝔜._⊕_ p q ⟩ + _ ∎ + ) + (isSet𝔜 _ _) + + freeCMonEquivLemma-β : (g : structHom 𝔉 𝔜) -> g ≡ ♯-isMonHom (g .fst ∘ η) + freeCMonEquivLemma-β g = structHom≡ 𝔉 𝔜 g (♯-isMonHom (g .fst ∘ η)) isSet𝔜 (funExt (freeCMonEquivLemma g)) + + freeCMonEquiv : structHom 𝔉 𝔜 ≃ (A -> 𝔜 .car) + freeCMonEquiv = + isoToEquiv (iso (λ g -> g .fst ∘ η) ♯-isMonHom (λ _ -> refl) (sym ∘ freeCMonEquivLemma-β)) + +module FreeCMonDef = F.Definition M.MonSig M.CMonEqSig M.CMonSEq + +freeCMon-sat : ∀ {n} {X : Type n} -> < FreeCMon X , freeCMon-α > ⊨ M.CMonSEq +freeCMon-sat (M.`mon M.`unitl) ρ = unitl (ρ fzero) +freeCMon-sat (M.`mon M.`unitr) ρ = unitr (ρ fzero) +freeCMon-sat (M.`mon M.`assocr) ρ = assocr (ρ fzero) (ρ fone) (ρ ftwo) +freeCMon-sat M.`comm ρ = comm (ρ fzero) (ρ fone) + +freeMonDef : ∀ {ℓ ℓ'} -> FreeCMonDef.Free ℓ ℓ' 2 +F.Definition.Free.F freeMonDef = FreeCMon +F.Definition.Free.η freeMonDef = η +F.Definition.Free.α freeMonDef = freeCMon-α +F.Definition.Free.sat freeMonDef = freeCMon-sat +F.Definition.Free.isFree freeMonDef isSet𝔜 satMon = (Free.freeCMonEquiv isSet𝔜 satMon) .snd diff --git a/Cubical/Structures/Set/CMon/PList.agda b/Cubical/Structures/Set/CMon/PList.agda new file mode 100644 index 0000000..f17a339 --- /dev/null +++ b/Cubical/Structures/Set/CMon/PList.agda @@ -0,0 +1,115 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +Definition taken from https://drops.dagstuhl.de/opus/volltexte/2023/18395/pdf/LIPIcs-ITP-2023-20.pdf +module Cubical.Structures.Set.CMon.PList where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Data.List as L + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.Mon.List as LM +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity hiding (_/_) +open import Cubical.Structures.Set.CMon.QFreeMon + +data Perm {ℓ : Level} {A : Type ℓ} : List A -> List A -> Type ℓ where + perm-refl : ∀ {xs} -> Perm xs xs + perm-swap : ∀ {x y xs ys zs} -> Perm (xs ++ x ∷ y ∷ ys) zs -> Perm (xs ++ y ∷ x ∷ ys) zs + +private + variable + ℓ ℓ₁ ℓ₂ : Level + A B : Type ℓ + +infixr 30 _∙ₚ_ +_∙ₚ_ : ∀ {xs ys zs} -> Perm xs ys -> Perm ys zs -> Perm {A = A} xs zs +perm-refl ∙ₚ q = q +(perm-swap p) ∙ₚ q = perm-swap (p ∙ₚ q) + +perm-sym : ∀ {xs ys} -> Perm xs ys -> Perm {A = A} ys xs +perm-sym perm-refl = perm-refl +perm-sym (perm-swap p) = perm-sym p ∙ₚ perm-swap perm-refl + +perm-subst : ∀ {xs ys} -> xs ≡ ys -> Perm {A = A} xs ys +perm-subst {xs = xs} p = subst (Perm xs) p perm-refl + +perm-∷ : ∀ {x xs ys} -> Perm xs ys -> Perm {A = A} (x ∷ xs) (x ∷ ys) +perm-∷ perm-refl = perm-refl +perm-∷ {x = x} (perm-swap {xs = xs} p) = perm-swap {xs = x ∷ xs} (perm-∷ p) + +perm-prepend : (xs : List A) {ys zs : List A} -> Perm ys zs -> Perm (xs ++ ys) (xs ++ zs) +perm-prepend [] p = p +perm-prepend (x ∷ xs) p = perm-∷ (perm-prepend xs p) + +perm-append : ∀ {xs ys} -> Perm xs ys -> (zs : List A) -> Perm (xs ++ zs) (ys ++ zs) +perm-append perm-refl _ = perm-refl +perm-append (perm-swap {xs = xs} p) _ = + perm-subst (++-assoc xs _ _) ∙ₚ perm-swap (perm-subst (sym (++-assoc xs _ _)) ∙ₚ perm-append p _) + +perm-movehead : (x : A) (xs : List A) {ys : List A} -> Perm (x ∷ xs ++ ys) (xs ++ x ∷ ys) +perm-movehead x [] = perm-refl +perm-movehead x (y ∷ xs) = perm-swap {xs = []} (perm-∷ (perm-movehead x xs)) + +⊕-commₚ : (xs ys : List A) -> Perm (xs ++ ys) (ys ++ xs) +⊕-commₚ xs [] = perm-subst (++-unit-r xs) +⊕-commₚ xs (y ∷ ys) = perm-sym (perm-movehead y xs {ys = ys}) ∙ₚ perm-∷ (⊕-commₚ xs ys) + +module _ {ℓA ℓB} {A : Type ℓA} {𝔜 : struct ℓB M.MonSig} {isSet𝔜 : isSet (𝔜 .car)} (𝔜-cmon : 𝔜 ⊨ M.CMonSEq) (f : A -> 𝔜 .car) where + module 𝔜 = M.CMonSEq 𝔜 𝔜-cmon + + f♯-hom = LM.Free.♯-isMonHom isSet𝔜 (M.cmonSatMon 𝔜-cmon) f + + f♯ : List A -> 𝔜 .car + f♯ = f♯-hom .fst + + f♯-++ : ∀ xs ys -> f♯ (xs ++ ys) ≡ f♯ xs 𝔜.⊕ f♯ ys + f♯-++ xs ys = + f♯ (xs ++ ys) ≡⟨ sym ((f♯-hom .snd) M.`⊕ (lookup (xs ∷ ys ∷ []))) ⟩ + 𝔜 .alg (M.`⊕ , (λ w -> f♯ (lookup (xs ∷ ys ∷ []) w))) ≡⟨ 𝔜.⊕-eta (lookup (xs ∷ ys ∷ [])) f♯ ⟩ + _ ∎ + + f♯-swap : ∀ {x y : A} (xs ys : List A) -> f♯ (xs ++ x ∷ y ∷ ys) ≡ f♯ (xs ++ y ∷ x ∷ ys) + f♯-swap {x} {y} [] ys = + f♯ ((L.[ x ] ++ L.[ y ]) ++ ys) ≡⟨ f♯-++ (L.[ x ] ++ L.[ y ]) ys ⟩ + f♯ (L.[ x ] ++ L.[ y ]) 𝔜.⊕ f♯ ys ≡⟨ cong (𝔜._⊕ f♯ ys) (f♯-++ L.[ x ] L.[ y ]) ⟩ + (f♯ L.[ x ] 𝔜.⊕ f♯ L.[ y ]) 𝔜.⊕ f♯ ys ≡⟨ cong (𝔜._⊕ f♯ ys) (𝔜.comm _ _) ⟩ + (f♯ L.[ y ] 𝔜.⊕ f♯ L.[ x ]) 𝔜.⊕ f♯ ys ≡⟨ cong (𝔜._⊕ f♯ ys) (sym (f♯-++ L.[ y ] L.[ x ])) ⟩ + f♯ (L.[ y ] ++ L.[ x ]) 𝔜.⊕ f♯ ys ≡⟨ sym (f♯-++ (L.[ y ] ++ L.[ x ]) ys) ⟩ + f♯ ((L.[ y ] ++ L.[ x ]) ++ ys) ∎ + f♯-swap {x} {y} (a ∷ as) ys = + f♯ (L.[ a ] ++ (as ++ x ∷ y ∷ ys)) ≡⟨ f♯-++ L.[ a ] (as ++ x ∷ y ∷ ys) ⟩ + f♯ L.[ a ] 𝔜.⊕ f♯ (as ++ x ∷ y ∷ ys) ≡⟨ cong (f♯ L.[ a ] 𝔜.⊕_) (f♯-swap as ys) ⟩ + f♯ L.[ a ] 𝔜.⊕ f♯ (as ++ y ∷ x ∷ ys) ≡⟨ sym (f♯-++ L.[ a ] (as ++ y ∷ x ∷ ys)) ⟩ + f♯ (L.[ a ] ++ (as ++ y ∷ x ∷ ys)) ≡⟨⟩ + f♯ ((a ∷ as) ++ y ∷ x ∷ ys) ∎ + + perm-resp-f♯ : {a b : List A} -> Perm a b -> f♯ a ≡ f♯ b + perm-resp-f♯ perm-refl = refl + perm-resp-f♯ (perm-swap {xs = xs} {ys = ys} p) = f♯-swap xs ys ∙ perm-resp-f♯ p + +module _ {ℓ} (A : Type ℓ) where + open import Cubical.Relation.Binary + module P = BinaryRelation {A = List A} Perm + open isPermRel + + isPermRelPerm : isPermRel LM.listDef (Perm {A = A}) + P.isEquivRel.reflexive (isEquivRel isPermRelPerm) _ = perm-refl + P.isEquivRel.symmetric (isEquivRel isPermRelPerm) _ _ = perm-sym + P.isEquivRel.transitive (isEquivRel isPermRelPerm) _ _ _ = _∙ₚ_ + isCongruence isPermRelPerm {a} {b} {c} {d} p q = perm-prepend a q ∙ₚ perm-append p d + isCommutative isPermRelPerm {a} {b} = ⊕-commₚ a b + resp-♯ isPermRelPerm {isSet𝔜 = isSet𝔜} 𝔜-cmon f p = perm-resp-f♯ {isSet𝔜 = isSet𝔜} 𝔜-cmon f p + + PermRel : PermRelation LM.listDef A + PermRel = Perm , isPermRelPerm + +module PListDef = F.Definition M.MonSig M.CMonEqSig M.CMonSEq + +plistFreeDef : ∀ {ℓ} -> PListDef.Free ℓ ℓ 2 +plistFreeDef = qFreeMonDef (PermRel _) diff --git a/Cubical/Structures/Set/CMon/QFreeMon.agda b/Cubical/Structures/Set/CMon/QFreeMon.agda new file mode 100644 index 0000000..9670aa3 --- /dev/null +++ b/Cubical/Structures/Set/CMon/QFreeMon.agda @@ -0,0 +1,195 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.QFreeMon where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.HITs.SetQuotients as Q +open import Cubical.Data.List as L +open import Cubical.Relation.Binary + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity hiding (_/_) +open import Cubical.Relation.Nullary hiding (⟪_⟫) + +open F.Definition M.MonSig M.MonEqSig M.MonSEq +open F.Definition.Free + +module _ {ℓ ℓ' : Level} (freeMon : Free ℓ ℓ' 2) where + + private + ℱ : Type ℓ -> Type ℓ + ℱ A = freeMon .F A + + PRel : Type ℓ -> Type (ℓ-max ℓ (ℓ-suc ℓ')) + PRel A = Rel (ℱ A) (ℱ A) ℓ' + + record isPermRel {A : Type ℓ} (R : PRel A) : Type (ℓ-max ℓ (ℓ-suc ℓ')) where + constructor permRel + infix 3 _≈_ + _≈_ : Rel (ℱ A) (ℱ A) ℓ' ; _≈_ = R + TODO: Add ⊕ as a helper in Mon.Desc + infixr 10 _⊕_ + _⊕_ : ℱ A -> ℱ A -> ℱ A + a ⊕ b = freeMon .α (M.`⊕ , ⟪ a ⨾ b ⟫) + module R = BinaryRelation R + field + isEquivRel : R.isEquivRel + isCongruence : {a b c d : ℱ A} + -> a ≈ b -> c ≈ d + -> a ⊕ c ≈ b ⊕ d + isCommutative : {a b : ℱ A} + -> a ⊕ b ≈ b ⊕ a + resp-♯ : {a b : ℱ A} {𝔜 : struct ℓ' M.MonSig} {isSet𝔜 : isSet (𝔜 .car)} (𝔜-cmon : 𝔜 ⊨ M.CMonSEq) + -> (f : A -> 𝔜 .car) + -> a ≈ b + -> let f♯ = ext freeMon isSet𝔜 (M.cmonSatMon 𝔜-cmon) f .fst in f♯ a ≡ f♯ b + + refl≈ = isEquivRel .R.isEquivRel.reflexive + trans≈ = isEquivRel .R.isEquivRel.transitive + cong≈ = isCongruence + comm≈ = isCommutative + subst≈-left : {a b c : ℱ A} -> a ≡ b -> a ≈ c -> b ≈ c + subst≈-left {c = c} = subst (_≈ c) + subst≈-right : {a b c : ℱ A} -> b ≡ c -> a ≈ b -> a ≈ c + subst≈-right {a = a} = subst (a ≈_) + subst≈ : {a b c d : ℱ A} -> a ≡ b -> c ≡ d -> a ≈ c -> b ≈ d + subst≈ {a} {b} {c} {d} p q r = trans≈ b c d (subst≈-left p r) (subst≈-right q (refl≈ c)) + + PermRelation : Type ℓ -> Type (ℓ-max ℓ (ℓ-suc ℓ')) + PermRelation A = Σ (PRel A) isPermRel + +module QFreeMon {ℓr ℓB} {freeMon : Free ℓr ℓB 2} (A : Type ℓr) ((R , isPermRelR) : PermRelation freeMon A) where + private + ℱ = freeMon .F A + open isPermRel isPermRelR + + 𝒬 : Type (ℓ-max ℓr ℓB) + 𝒬 = ℱ / _≈_ + + 𝔉 : M.MonStruct + 𝔉 = < ℱ , freeMon .α > + + module 𝔉 = M.MonSEq 𝔉 (freeMon .sat) + + e : ℱ + e = 𝔉.e + + e/ : 𝒬 + e/ = Q.[ e ] + + η/ : A -> 𝒬 + η/ x = Q.[ freeMon .η x ] + + _⊕/_ : 𝒬 -> 𝒬 -> 𝒬 + _⊕/_ = Q.rec2 squash/ + (\a b -> Q.[ a ⊕ b ]) + (\a b c r -> eq/ (a ⊕ c) (b ⊕ c) (cong≈ r (refl≈ c))) + (\a b c r -> eq/ (a ⊕ b) (a ⊕ c) (cong≈ (refl≈ a) r)) + + ⊕-unitl : (a : 𝒬) -> e/ ⊕/ a ≡ a + ⊕-unitl = Q.elimProp + (\_ -> squash/ _ _) + (\a -> eq/ (e ⊕ a) a (subst≈-right (𝔉.unitl a) (refl≈ (e ⊕ a)))) + + ⊕-unitr : (a : 𝒬) -> a ⊕/ e/ ≡ a + ⊕-unitr = Q.elimProp + (\_ -> squash/ _ _) + (\a -> eq/ (a ⊕ e) a (subst≈-right (𝔉.unitr a) (refl≈ (a ⊕ e)))) + + ⊕-assocr : (a b c : 𝒬) -> (a ⊕/ b) ⊕/ c ≡ a ⊕/ (b ⊕/ c) + ⊕-assocr = Q.elimProp + (\_ -> isPropΠ (\_ -> isPropΠ (\_ -> squash/ _ _))) + (\a -> elimProp + (\_ -> isPropΠ (\_ -> squash/ _ _)) + (\b -> elimProp + (\_ -> squash/ _ _) + (\c -> eq/ ((a ⊕ b) ⊕ c) (a ⊕ (b ⊕ c)) (subst≈-right (𝔉.assocr a b c) (refl≈ ((a ⊕ b) ⊕ c)))))) + + ⊕-comm : (a b : 𝒬) -> a ⊕/ b ≡ b ⊕/ a + ⊕-comm = elimProp + (\_ -> isPropΠ (\_ -> squash/ _ _)) + (\a -> elimProp + (\_ -> squash/ _ _) + (\b -> eq/ (a ⊕ b) (b ⊕ a) comm≈)) + + qFreeMon-α : sig M.MonSig 𝒬 -> 𝒬 + qFreeMon-α (M.`e , i) = Q.[ e ] + qFreeMon-α (M.`⊕ , i) = i fzero ⊕/ i fone + + qFreeMon-sat : < 𝒬 , qFreeMon-α > ⊨ M.CMonSEq + qFreeMon-sat (M.`mon M.`unitl) ρ = ⊕-unitl (ρ fzero) + qFreeMon-sat (M.`mon M.`unitr) ρ = ⊕-unitr (ρ fzero) + qFreeMon-sat (M.`mon M.`assocr) ρ = ⊕-assocr (ρ fzero) (ρ fone) (ρ ftwo) + qFreeMon-sat M.`comm ρ = ⊕-comm (ρ fzero) (ρ fone) + + private + 𝔛 : M.CMonStruct + 𝔛 = < 𝒬 , qFreeMon-α > + + module 𝔛 = M.CMonSEq 𝔛 qFreeMon-sat + + [_]-isMonHom : structHom 𝔉 𝔛 + fst [_]-isMonHom = Q.[_] + snd [_]-isMonHom M.`e i = cong _/_.[_] 𝔉.e-eta + snd [_]-isMonHom M.`⊕ i = + 𝔛 .alg (M.`⊕ , (λ x -> Q.[ i x ])) ≡⟨ 𝔛.⊕-eta i Q.[_] ⟩ + Q.[ freeMon .α (M.`⊕ , _) ] ≡⟨ cong (λ z -> Q.[_] {R = _≈_} (freeMon .α (M.`⊕ , z))) (lookup2≡i i) ⟩ + Q.[ freeMon .α (M.`⊕ , i) ] ∎ + + module IsFree {𝔜 : struct ℓB M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-cmon : 𝔜 ⊨ M.CMonSEq) where + module 𝔜 = M.CMonSEq 𝔜 𝔜-cmon + + module _ (f : A -> 𝔜 .car) where + f♯ : structHom 𝔉 𝔜 + f♯ = ext (freeMon) isSet𝔜 (M.cmonSatMon 𝔜-cmon) f + + _♯ : 𝒬 -> 𝔜 .car + _♯ = Q.rec isSet𝔜 (f♯ .fst) (\_ _ -> resp-♯ 𝔜-cmon f) + + private + ♯-++ : ∀ xs ys -> (xs ⊕/ ys) ♯ ≡ (xs ♯) 𝔜.⊕ (ys ♯) + ♯-++ = + elimProp (λ _ -> isPropΠ λ _ -> isSet𝔜 _ _) λ xs -> + elimProp (λ _ -> isSet𝔜 _ _) λ ys -> + f♯ .fst (xs ⊕ ys) ≡⟨ sym (f♯ .snd M.`⊕ (lookup (xs ∷ ys ∷ []))) ⟩ + _ ≡⟨ 𝔜.⊕-eta (lookup (xs ∷ ys ∷ [])) (f♯ .fst) ⟩ + _ ∎ + + ♯-isMonHom : structHom 𝔛 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta ∙ f♯ .snd M.`e (lookup []) + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ ∙ sym (♯-++ (i fzero) (i fone)) + + private + qFreeMonEquivLemma : (g : structHom 𝔛 𝔜) (x : 𝔛 .car) -> g .fst x ≡ ((g .fst ∘ η/) ♯) x + qFreeMonEquivLemma g = elimProp (λ _ -> isSet𝔜 _ _) λ x i -> lemma (~ i) x + where + lemma : (f♯ (((g .fst) ∘ Q.[_]) ∘ freeMon .η)) .fst ≡ (g .fst) ∘ Q.[_] + lemma = cong fst (ext-β (freeMon) isSet𝔜 (M.cmonSatMon 𝔜-cmon) (structHom∘ 𝔉 𝔛 𝔜 g [_]-isMonHom)) + + qFreeMonEquiv : structHom 𝔛 𝔜 ≃ (A -> 𝔜 .car) + qFreeMonEquiv = + isoToEquiv + ( iso + (λ g -> g .fst ∘ η/) + ♯-isMonHom + (ext-η (freeMon) isSet𝔜 (M.cmonSatMon 𝔜-cmon)) + (λ g -> sym (structHom≡ 𝔛 𝔜 g (♯-isMonHom (g .fst ∘ η/)) isSet𝔜 (funExt (qFreeMonEquivLemma g)))) + ) + +module QFreeMonDef = F.Definition M.MonSig M.CMonEqSig M.CMonSEq + +qFreeMonDef : ∀ {ℓ : Level} {freeMon : Free ℓ ℓ 2} (R : {A : Type ℓ} -> PermRelation freeMon A) -> QFreeMonDef.Free ℓ ℓ 2 +F (qFreeMonDef R) A = QFreeMon.𝒬 A R +η (qFreeMonDef R) = QFreeMon.η/ _ R +α (qFreeMonDef R) = QFreeMon.qFreeMon-α _ R +sat (qFreeMonDef R) = QFreeMon.qFreeMon-sat _ R +isFree (qFreeMonDef R) isSet𝔜 𝔜-cmon = (QFreeMon.IsFree.qFreeMonEquiv _ R isSet𝔜 𝔜-cmon) .snd diff --git a/Cubical/Structures/Set/CMon/SList.agda b/Cubical/Structures/Set/CMon/SList.agda new file mode 100644 index 0000000..1df4c1d --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.SList where + +open import Cubical.Structures.Set.CMon.SList.Base public +open import Cubical.Structures.Set.CMon.SList.Length public +open import Cubical.Structures.Set.CMon.SList.Membership public diff --git a/Cubical/Structures/Set/CMon/SList/Base.agda b/Cubical/Structures/Set/CMon/SList/Base.agda new file mode 100644 index 0000000..c169444 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Base.agda @@ -0,0 +1,98 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.SList.Base where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +import Cubical.Data.List as L + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.HITs.FiniteMultiset public renaming (FMSet to SList; comm to swap) + +pattern [_] a = a ∷ [] + +private + variable + ℓ : Level + A : Type ℓ + +slist-α : ∀ {n : Level} {X : Type n} -> sig M.MonSig (SList X) -> SList X +slist-α (M.`e , i) = [] +slist-α (M.`⊕ , i) = i fzero ++ i fone + +module Free {x y : Level} {A : Type x} {𝔜 : struct y M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-cmon : 𝔜 ⊨ M.CMonSEq) where + module 𝔜 = M.CMonSEq 𝔜 𝔜-cmon + + 𝔛 : M.CMonStruct + 𝔛 = < SList A , slist-α > + + module _ (f : A -> 𝔜 .car) where + _♯ : SList A -> 𝔜 .car + _♯ = Elim.f 𝔜.e + (λ x xs -> (f x) 𝔜.⊕ xs) + (λ a b xs i -> + ( sym (𝔜.assocr (f a) (f b) xs) + ∙ cong (𝔜._⊕ xs) (𝔜.comm (f a) (f b)) + ∙ 𝔜.assocr (f b) (f a) xs + ) i + ) + (λ _ -> isSet𝔜) + + private + ♯-++ : ∀ xs ys -> (xs ++ ys) ♯ ≡ (xs ♯) 𝔜.⊕ (ys ♯) + ♯-++ = ElimProp.f (isPropΠ λ _ -> isSet𝔜 _ _) + (λ ys -> sym (𝔜.unitl (ys ♯))) + (λ a {xs} p ys -> + f a 𝔜.⊕ ((xs ++ ys) ♯) ≡⟨ cong (f a 𝔜.⊕_) (p ys) ⟩ + f a 𝔜.⊕ ((xs ♯) 𝔜.⊕ (ys ♯)) ≡⟨ sym (𝔜.assocr (f a) (xs ♯) (ys ♯)) ⟩ + _ + ∎) + + ♯-isMonHom : structHom 𝔛 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ ∙ sym (♯-++ (i fzero) (i fone)) + + private + slistEquivLemma : (g : structHom 𝔛 𝔜) -> (x : SList A) -> g .fst x ≡ ((g .fst ∘ [_]) ♯) x + slistEquivLemma (g , homMonWit) = ElimProp.f (isSet𝔜 _ _) + (sym (homMonWit M.`e (lookup L.[])) ∙ 𝔜.e-eta) + (λ x {xs} p -> + g (x ∷ xs) ≡⟨ sym (homMonWit M.`⊕ (lookup ([ x ] L.∷ xs L.∷ L.[]))) ⟩ + _ ≡⟨ 𝔜.⊕-eta (lookup ([ x ] L.∷ xs L.∷ L.[])) g ⟩ + _ ≡⟨ cong (g [ x ] 𝔜.⊕_) p ⟩ + _ ∎ + ) + + slistEquivLemma-β : (g : structHom 𝔛 𝔜) -> g ≡ ♯-isMonHom (g .fst ∘ [_]) + slistEquivLemma-β g = structHom≡ 𝔛 𝔜 g (♯-isMonHom (g .fst ∘ [_])) isSet𝔜 (funExt (slistEquivLemma g)) + + slistMonEquiv : structHom 𝔛 𝔜 ≃ (A -> 𝔜 .car) + slistMonEquiv = + isoToEquiv (iso (λ g -> g .fst ∘ [_]) ♯-isMonHom (λ g -> funExt (𝔜.unitr ∘ g)) (sym ∘ slistEquivLemma-β)) + +module SListDef = F.Definition M.MonSig M.CMonEqSig M.CMonSEq + +slist-sat : ∀ {n} {X : Type n} -> < SList X , slist-α > ⊨ M.CMonSEq +slist-sat (M.`mon M.`unitl) ρ = unitl-++ (ρ fzero) +slist-sat (M.`mon M.`unitr) ρ = unitr-++ (ρ fzero) +slist-sat (M.`mon M.`assocr) ρ = sym (assoc-++ (ρ fzero) (ρ fone) (ρ ftwo)) +slist-sat M.`comm ρ = comm-++ (ρ fzero) (ρ fone) + +slistDef : ∀ {ℓ ℓ'} -> SListDef.Free ℓ ℓ' 2 +F.Definition.Free.F slistDef = SList +F.Definition.Free.η slistDef = [_] +F.Definition.Free.α slistDef = slist-α +F.Definition.Free.sat slistDef = slist-sat +F.Definition.Free.isFree slistDef isSet𝔜 satMon = (Free.slistMonEquiv isSet𝔜 satMon) .snd diff --git a/Cubical/Structures/Set/CMon/SList/Length.agda b/Cubical/Structures/Set/CMon/SList/Length.agda new file mode 100644 index 0000000..5941b6f --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Length.agda @@ -0,0 +1,148 @@ +{-# OPTIONS --cubical --exact-split --safe #-} + +module Cubical.Structures.Set.CMon.SList.Length where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Data.Sum +open import Cubical.Data.Prod +open import Cubical.Data.Empty as E +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Functions.Embedding + +open import Cubical.Structures.Set.CMon.SList.Base as SList + +private + variable + ℓ : Level + A B : Type ℓ + ϕ : isSet A + a x : A + xs ys : SList A + +disj-nil-cons : x ∷ xs ≡ [] -> ⊥ +disj-nil-cons p = transport (λ i → fst (code (p i))) tt + where code : SList A -> hProp ℓ-zero + code = Elim.f (⊥ , isProp⊥) + (λ _ _ -> Unit , isPropUnit) + (λ _ _ _ _ -> Unit , isPropUnit) + (λ _ -> isSetHProp) + +disj-cons-nil : [] ≡ x ∷ xs -> ⊥ +disj-cons-nil p = disj-nil-cons (sym p) + +length : SList A → ℕ +length = Elim.f 0 (λ _ n -> suc n) (λ _ _ _ -> refl) (λ _ -> isSetℕ) + +length-++ : (x y : SList A) → length (x ++ y) ≡ length x + length y +length-++ x y = ElimProp.f {B = λ xs → length (xs ++ y) ≡ length xs + length y} + (isSetℕ _ _) refl (λ a p -> cong suc p) x + +lenZero-out : (xs : SList A) → length xs ≡ 0 → [] ≡ xs +lenZero-out = ElimProp.f (isPropΠ λ _ → trunc _ _) + (λ _ -> refl) + (λ x p q -> E.rec (snotz q)) + +lenZero-eqv : (xs : SList A) → (length xs ≡ 0) ≃ ([] ≡ xs) +lenZero-eqv xs = propBiimpl→Equiv (isSetℕ _ _) (trunc _ _) (lenZero-out xs) (λ p i → length (p (~ i))) + +is-sing-pred : SList A → A → Type _ +is-sing-pred xs a = [ a ] ≡ xs + +is-sing-pred-prop : (xs : SList A) (z : A) → isProp (is-sing-pred xs z) +is-sing-pred-prop xs a = trunc [ a ] xs + +is-sing : SList A → Type _ +is-sing xs = Σ _ (is-sing-pred xs) + +Sing : Type ℓ → Type ℓ +Sing A = Σ (SList A) is-sing + +[_]-is-sing : (a : A) → is-sing [ a ] +[ a ]-is-sing = a , refl + +sing=isProp : ∀ {x y : A} → isProp ([ x ] ≡ [ y ]) +sing=isProp {x = x} {y = y} = trunc [ x ] [ y ] + +sing=-in : ∀ {x y : A} → x ≡ y → [ x ] ≡ [ y ] +sing=-in p i = [ p i ] + +sing=isContr : ∀ {x y : A} → x ≡ y → isContr ([ x ] ≡ [ y ]) +sing=isContr p = sing=-in p , sing=isProp (sing=-in p) + +lenOne-down : (x : A) (xs : SList A) → length (x ∷ xs) ≡ 1 → [] ≡ xs +lenOne-down x xs p = lenZero-out xs (injSuc p) + +lenOne-cons : (x : A) (xs : SList A) → length (x ∷ xs) ≡ 1 → is-sing (x ∷ xs) +lenOne-cons x xs p = + let q = lenOne-down x xs p + in transport (λ i → is-sing (x ∷ q i)) [ x ]-is-sing + +lenTwo-⊥ : (x y : A) (xs : SList A) → (length (y ∷ x ∷ xs) ≡ 1) ≃ ⊥ +lenTwo-⊥ x y xs = (λ p → E.rec (snotz (injSuc p))) , record { equiv-proof = E.elim } + +isContrlenTwo-⊥→X : (X : Type ℓ) (x y : A) (xs : SList A) → isContr (length (y ∷ x ∷ xs) ≡ 1 → X) +isContrlenTwo-⊥→X X x y xs = subst (λ Z → isContr (Z → X)) (sym (ua (lenTwo-⊥ x y xs))) (isContr⊥→A {A = X}) + +lenOne-swap : {A : Type ℓ} (x y : A) (xs : SList A) + → PathP (λ i → length (SList.swap x y xs i) ≡ 1 → is-sing (SList.swap x y xs i)) + (λ p → lenOne-cons x (y ∷ xs) p) + (λ p → lenOne-cons y (x ∷ xs) p) +lenOne-swap {A = A} x y xs = + transport (sym (PathP≡Path (λ i → length (SList.swap x y xs i) ≡ 1 → is-sing (SList.swap x y xs i)) _ _)) + (isContr→isProp (isContrlenTwo-⊥→X _ x y xs) _ _) + +lenOne : Type ℓ → Type _ +lenOne A = Σ (SList A) (λ xs → length xs ≡ 1) + +module _ {ϕ : isSet A} where + + is-sing-set : (xs : SList A) → isSet (is-sing xs) + is-sing-set xs = isSetΣ ϕ λ x → isProp→isSet (is-sing-pred-prop xs x) + + lenOne-out : (xs : SList A) → length xs ≡ 1 → is-sing xs + lenOne-out = Elim.f (λ p → E.rec (znots p)) + (λ x {xs} f p → lenOne-cons x xs p) + (λ x y {xs} f → lenOne-swap x y xs) + λ xs → isSetΠ (λ p → is-sing-set xs) + + head : lenOne A → A + head (xs , p) = lenOne-out xs p .fst + + head-β : (a : A) → head ([ a ] , refl) ≡ a + head-β a i = transp (λ _ → A) i a + + lenOnePath : (a b : A) → Type _ + lenOnePath a b = Path (lenOne A) ([ a ] , refl) ([ b ] , refl) + + lenOnePath-in : {a b : A} → [ a ] ≡ [ b ] → lenOnePath a b + lenOnePath-in = Σ≡Prop (λ xs → isSetℕ _ _) + + [-]-inj : {a b : A} → [ a ] ≡ [ b ] → a ≡ b + [-]-inj {a = a} {b = b} p = aux (lenOnePath-in p) + where aux : lenOnePath a b → a ≡ b + aux p = sym (head-β a) ∙ cong head p ∙ head-β b + + is-sing-prop : (xs : SList A) → isProp (is-sing xs) + is-sing-prop xs (a , ψ) (b , ξ) = Σ≡Prop (is-sing-pred-prop xs) ([-]-inj (ψ ∙ sym ξ)) + + lenOne-eqv : (xs : SList A) → (length xs ≡ 1) ≃ (is-sing xs) + lenOne-eqv xs = propBiimpl→Equiv (isSetℕ _ _) (is-sing-prop xs) (lenOne-out xs) (λ χ → cong length (sym (χ .snd))) + + lenOne-set-eqv : lenOne A ≃ A + lenOne-set-eqv = isoToEquiv (iso head g head-β g-f) + where g : A → lenOne A + g a = [ a ] , refl + g-f : (as : lenOne A) → g (head as) ≡ as + g-f (as , ϕ) = + let (a , ψ) = lenOne-out as ϕ + in Σ≡Prop (λ xs → isSetℕ _ _) {u = ([ a ] , refl)} {v = (as , ϕ)} ψ + + sing-set-eqv : Sing A ≃ A + sing-set-eqv = isoToEquiv (iso f (λ a → [ a ] , [ a ]-is-sing) (λ _ → refl) ret) + where f : Sing A → A + f s = s .snd .fst + ret : (s : Sing A) → ([ f s ] , f s , refl) ≡ s + ret (s , ψ) = Σ≡Prop is-sing-prop (ψ .snd) diff --git a/Cubical/Structures/Set/CMon/SList/Membership.agda b/Cubical/Structures/Set/CMon/SList/Membership.agda new file mode 100644 index 0000000..1fdc1c6 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Membership.agda @@ -0,0 +1,82 @@ +{-# OPTIONS --cubical --exact-split --safe #-} + +module Cubical.Structures.Set.CMon.SList.Membership where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +import Cubical.Data.List as L +open import Cubical.Functions.Logic as L +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Sum as ⊎ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.Structures.Set.Mon.List + +open import Cubical.Structures.Set.CMon.SList.Base as SList + +private + variable + ℓ : Level + A : Type ℓ + +list→slist-Hom : structHom < L.List A , list-α > < SList A , slist-α > +list→slist-Hom = ListDef.Free.ext listDef trunc (M.cmonSatMon slist-sat) [_] + +list→slist : L.List A -> SList A +list→slist = list→slist-Hom .fst + +module Membership* {ℓ} {A : Type ℓ} (isSetA : isSet A) where + open SList.Free {A = A} isSetHProp (M.⊔-MonStr-CMonSEq ℓ) + + ∈*Prop : A -> SList A -> hProp ℓ + ∈*Prop x = (λ y -> (x ≡ y) , isSetA x y) ♯ + + _∈*_ : A -> SList A -> Type ℓ + x ∈* xs = ∈*Prop x xs .fst + + isProp-∈* : (x : A) -> (xs : SList A) -> isProp (x ∈* xs) + isProp-∈* x xs = (∈*Prop x xs) .snd + + x∈*xs : ∀ x xs -> x ∈* (x ∷ xs) + x∈*xs x xs = L.inl refl + + x∈*[x] : ∀ x -> x ∈* [ x ] + x∈*[x] x = x∈*xs x [] + + ∈*-∷ : ∀ x y xs -> x ∈* xs -> x ∈* (y ∷ xs) + ∈*-∷ x y xs p = L.inr p + + ∈*-++ : ∀ x xs ys -> x ∈* ys -> x ∈* (xs ++ ys) + ∈*-++ x xs ys p = + ElimProp.f {B = λ zs -> x ∈* (zs ++ ys)} (λ {zs} -> isProp-∈* x (zs ++ ys)) p + (λ a {zs} q -> ∈*-∷ x a (zs ++ ys) q) + xs + + ¬∈*[] : ∀ x -> (x ∈* []) -> ⊥.⊥ + ¬∈*[] x = ⊥.rec* + + x∈*[y]→x≡y : ∀ x y -> x ∈* [ y ] -> x ≡ y + x∈*[y]→x≡y x y = P.rec (isSetA x y) $ ⊎.rec (idfun _) ⊥.rec* + + open Membership isSetA + + ∈→∈* : ∀ x xs -> x ∈ xs -> x ∈* (list→slist xs) + ∈→∈* x (y L.∷ ys) = P.rec + (isProp-∈* x (list→slist (y L.∷ ys))) + (⊎.rec L.inl (L.inr ∘ ∈→∈* x ys)) + + ∈*→∈ : ∀ x xs -> x ∈* (list→slist xs) -> x ∈ xs + ∈*→∈ x (y L.∷ ys) = P.rec + (isProp-∈ x (y L.∷ ys)) + (⊎.rec L.inl (L.inr ∘ ∈*→∈ x ys)) \ No newline at end of file diff --git a/Cubical/Structures/Set/CMon/SList/Seely.agda b/Cubical/Structures/Set/CMon/SList/Seely.agda new file mode 100644 index 0000000..c77e446 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Seely.agda @@ -0,0 +1,141 @@ +{-# OPTIONS --cubical --exact-split --safe #-} + +module Cubical.Structures.Set.CMon.SList.Seely where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +import Cubical.Data.List as L +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Sum as ⊎ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.Structures.Set.Mon.List + +open import Cubical.Structures.Set.CMon.SList.Base as SList + +module _ {ℓ} {A B : Type ℓ} where + + open SListDef.Free + + isSetSList× : isSet (SList A × SList B) + isSetSList× = isSet× trunc trunc + + slist×-α : sig M.MonSig (SList A × SList B) -> (SList A × SList B) + slist×-α (M.`e , i) = [] , [] + slist×-α (M.`⊕ , i) = (i fzero) .fst ++ (i fone) .fst , (i fzero) .snd ++ (i fone) .snd + + slist×-sat : < (SList A × SList B) , slist×-α > ⊨ M.CMonSEq + slist×-sat (M.`mon M.`unitl) ρ = refl + slist×-sat (M.`mon M.`unitr) ρ = ≡-× (slist-sat (M.`mon M.`unitr) (fst ∘ ρ)) ((slist-sat (M.`mon M.`unitr) (snd ∘ ρ))) + slist×-sat (M.`mon M.`assocr) ρ = ≡-× (slist-sat (M.`mon M.`assocr) (fst ∘ ρ)) ((slist-sat (M.`mon M.`assocr) (snd ∘ ρ))) + slist×-sat M.`comm ρ = ≡-× (slist-sat M.`comm (fst ∘ ρ)) ((slist-sat M.`comm (snd ∘ ρ))) + + f-η : A ⊎ B -> SList A × SList B + f-η (inl x) = [ x ] , [] + f-η (inr x) = [] , [ x ] + + f-hom : structHom < SList (A ⊎ B) , slist-α > < (SList A × SList B) , slist×-α > + f-hom = ext slistDef isSetSList× slist×-sat f-η + + f : SList (A ⊎ B) -> SList A × SList B + f = f-hom .fst + + mmap : ∀ {X Y : Type ℓ} -> (X -> Y) -> SList X -> SList Y + mmap f = ext slistDef trunc slist-sat ([_] ∘ f) .fst + + mmap-++ : ∀ {X Y : Type ℓ} -> ∀ f xs ys -> mmap {X = X} {Y = Y} f (xs ++ ys) ≡ mmap f xs ++ mmap f ys + mmap-++ f xs ys = sym (ext slistDef trunc slist-sat ([_] ∘ f) .snd M.`⊕ ⟪ xs ⨾ ys ⟫) + + mmap-∷ : ∀ {X Y : Type ℓ} -> ∀ f x xs -> mmap {X = X} {Y = Y} f (x ∷ xs) ≡ f x ∷ mmap f xs + mmap-∷ f x xs = mmap-++ f [ x ] xs + + g : SList A × SList B -> SList (A ⊎ B) + g (as , bs) = mmap inl as ++ mmap inr bs + + g-++ : ∀ xs ys -> g xs ++ g ys ≡ g (xs .fst ++ ys .fst , xs .snd ++ ys .snd) + g-++ (as , bs) (cs , ds) = sym $ + g (as ++ cs , bs ++ ds) + ≡⟨ cong (_++ mmap inr (bs ++ ds)) (mmap-++ inl as cs) ⟩ + (mmap inl as ++ mmap inl cs) ++ (mmap inr (bs ++ ds)) + ≡⟨ cong ((mmap inl as ++ mmap inl cs) ++_) (mmap-++ inr bs ds) ⟩ + (mmap inl as ++ mmap inl cs) ++ (mmap inr bs ++ mmap inr ds) + ≡⟨ assoc-++ (mmap inl as ++ mmap inl cs) (mmap inr bs) (mmap inr ds) ⟩ + ((mmap inl as ++ mmap inl cs) ++ mmap inr bs) ++ mmap inr ds + ≡⟨ cong (_++ mmap inr ds) (sym (assoc-++ (mmap inl as) (mmap inl cs) (mmap inr bs))) ⟩ + (mmap inl as ++ (mmap inl cs ++ mmap inr bs)) ++ mmap inr ds + ≡⟨ cong (λ z → (mmap inl as ++ z) ++ mmap inr ds) (comm-++ (mmap inl cs) (mmap inr bs)) ⟩ + (mmap inl as ++ (mmap inr bs ++ mmap inl cs)) ++ mmap inr ds + ≡⟨ cong (_++ mmap inr ds) (assoc-++ (mmap inl as) (mmap inr bs) (mmap inl cs)) ⟩ + ((mmap inl as ++ mmap inr bs) ++ mmap inl cs) ++ mmap inr ds + ≡⟨ sym (assoc-++ (mmap inl as ++ mmap inr bs) (mmap inl cs) (mmap inr ds)) ⟩ + (mmap inl as ++ mmap inr bs) ++ (mmap inl cs ++ mmap inr ds) + ≡⟨⟩ + g (as , bs) ++ g (cs , ds) + ∎ + + g-hom : structHom < (SList A × SList B) , slist×-α > < SList (A ⊎ B) , slist-α > + g-hom = g , g-is-hom + where + g-is-hom : structIsHom < SList A × SList B , slist×-α > < SList (A ⊎ B) , slist-α > g + g-is-hom M.`e i = refl + g-is-hom M.`⊕ i = g-++ (i fzero) (i fone) + + module _ {ℓ} {X : Type ℓ} (h : structHom < SList X , slist-α > < SList X , slist-α >) (h-η : ∀ x -> h .fst [ x ] ≡ [ x ]) where + univ-htpy : ∀ xs -> h .fst xs ≡ xs + univ-htpy xs = h~η♯ xs ∙ η♯~id xs + where + h~η♯ : ∀ xs -> h .fst xs ≡ ext slistDef trunc slist-sat [_] .fst xs + h~η♯ = ElimProp.f (trunc _ _) (sym (h .snd M.`e ⟪⟫)) λ x {xs} p -> + h .fst (x ∷ xs) ≡⟨ sym (h .snd M.`⊕ ⟪ [ x ] ⨾ xs ⟫) ⟩ + h .fst [ x ] ++ h .fst xs ≡⟨ congS (_++ h .fst xs) (h-η x) ⟩ + x ∷ h .fst xs ≡⟨ congS (x ∷_) p ⟩ + x ∷ ext slistDef trunc slist-sat [_] .fst xs ∎ + η♯~id : ∀ xs -> ext slistDef trunc slist-sat [_] .fst xs ≡ xs + η♯~id xs = congS (λ h -> h .fst xs) (ext-β slistDef trunc slist-sat (idHom < SList X , slist-α >)) + + g-f : ∀ xs -> g (f xs) ≡ xs + g-f = univ-htpy (structHom∘ _ _ < SList (A ⊎ B) , slist-α > g-hom f-hom) lemma + where + lemma : ∀ x -> g (f [ x ]) ≡ [ x ] + lemma (inl x) = refl + lemma (inr x) = refl + + f-mmap-inl : ∀ as -> f (mmap inl as) ≡ (as , []) + f-mmap-inl = ElimProp.f (isSetSList× _ _) refl λ x {xs} p -> + f (mmap inl (x ∷ xs)) ≡⟨ cong f (mmap-∷ inl x xs) ⟩ + f (inl x ∷ mmap inl xs) ≡⟨ sym (f-hom .snd M.`⊕ ⟪ [ inl x ] ⨾ mmap inl xs ⟫) ⟩ + x ∷ f (mmap inl xs) .fst , f (mmap inl xs) .snd ≡⟨ congS (λ z -> x ∷ z .fst , z .snd) p ⟩ + x ∷ xs , [] ∎ + + f-mmap-inr : ∀ as -> f (mmap inr as) ≡ ([] , as) + f-mmap-inr = ElimProp.f (isSetSList× _ _) refl λ x {xs} p -> + f (mmap inr (x ∷ xs)) ≡⟨ cong f (mmap-∷ inr x xs) ⟩ + f (inr x ∷ mmap inr xs) ≡⟨ sym (f-hom .snd M.`⊕ ⟪ [ inr x ] ⨾ mmap inr xs ⟫) ⟩ + f (mmap inr xs) .fst , x ∷ f (mmap inr xs) .snd ≡⟨ congS (λ z -> z .fst , x ∷ z .snd) p ⟩ + [] , x ∷ xs ∎ + + f-g : ∀ xs -> f (g xs) ≡ xs + f-g (as , bs) = + f (g (as , bs)) + ≡⟨ sym (f-hom .snd M.`⊕ ⟪ mmap inl as ⨾ mmap inr bs ⟫) ⟩ + (f (mmap inl as) .fst ++ f (mmap inr bs) .fst) , (f (mmap inl as) .snd ++ f (mmap inr bs) .snd) + ≡⟨ congS (λ z -> (z .fst ++ f (mmap inr bs) .fst) , (z .snd ++ f (mmap inr bs) .snd)) (f-mmap-inl as) ⟩ + as ++ f (mmap inr bs) .fst , [] ++ f (mmap inr bs) .snd + ≡⟨ congS (λ z -> as ++ z .fst , [] ++ z .snd) (f-mmap-inr bs) ⟩ + as ++ [] , bs + ≡⟨ congS (λ zs -> zs , bs) (unitr-++ as) ⟩ + as , bs ∎ + + seely : SList (A ⊎ B) ≃ SList A × SList B + seely = isoToEquiv (iso f g f-g g-f) \ No newline at end of file diff --git a/Cubical/Structures/Set/CMon/SList/Sort.agda b/Cubical/Structures/Set/CMon/SList/Sort.agda new file mode 100644 index 0000000..c44b539 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Sort.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.CMon.SList.Sort where + +open import Cubical.Structures.Set.CMon.SList.Sort.Base public +open import Cubical.Structures.Set.CMon.SList.Sort.Order public +open import Cubical.Structures.Set.CMon.SList.Sort.Sort public +open import Cubical.Structures.Set.CMon.SList.Sort.Equiv public diff --git a/Cubical/Structures/Set/CMon/SList/Sort/Base.agda b/Cubical/Structures/Set/CMon/SList/Sort/Base.agda new file mode 100644 index 0000000..8bb7f5e --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Sort/Base.agda @@ -0,0 +1,165 @@ +{-# OPTIONS --cubical --safe --exact-split -WnoUnsupportedIndexedMatch #-} + +module Cubical.Structures.Set.CMon.SList.Sort.Base where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order renaming (_≤_ to _≤ℕ_; _<_ to _<ℕ_) +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Order +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.HLevels +open import Cubical.Data.List +open import Cubical.HITs.PropositionalTruncation as P +import Cubical.Data.List as L +open import Cubical.Functions.Logic as L hiding (¬_; ⊥) + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.Structures.Set.Mon.List +open import Cubical.Structures.Set.CMon.SList.Base renaming (_∷_ to _∷*_; [] to []*; [_] to [_]*; _++_ to _++*_) +import Cubical.Structures.Set.CMon.SList.Base as S +open import Cubical.Structures.Set.CMon.SList.Length as S +open import Cubical.Structures.Set.CMon.SList.Membership as S + +open Iso + +private + variable + ℓ : Level + A : Type ℓ + +head-maybe : List A -> Maybe A +head-maybe [] = nothing +head-maybe (x ∷ xs) = just x + +module Sort {A : Type ℓ} (isSetA : isSet A) (sort : SList A -> List A) where + open Membership isSetA + + is-sorted : List A -> Type _ + is-sorted list = ∥ fiber sort list ∥₁ + + is-section : Type _ + is-section = ∀ xs -> list→slist (sort xs) ≡ xs + + isProp-is-section : isProp is-section + isProp-is-section = isPropΠ (λ _ -> trunc _ _) + + is-head-least : Type _ + is-head-least = ∀ x y xs -> is-sorted (x ∷ xs) -> y ∈ (x ∷ xs) -> is-sorted (x ∷ y ∷ []) + + is-tail-sort : Type _ + is-tail-sort = ∀ x xs -> is-sorted (x ∷ xs) -> is-sorted xs + + is-sort : Type _ + is-sort = is-head-least × is-tail-sort + + isProp-is-head-least : isProp is-head-least + isProp-is-head-least = isPropΠ5 λ _ _ _ _ _ -> squash₁ + + isProp-is-tail-sort : isProp is-tail-sort + isProp-is-tail-sort = isPropΠ3 (λ _ _ _ -> squash₁) + + isProp-is-sort : isProp is-sort + isProp-is-sort = isProp× isProp-is-head-least isProp-is-tail-sort + + is-sort-section : Type _ + is-sort-section = is-section × is-sort + + isProp-is-sort-section : isProp is-sort-section + isProp-is-sort-section = isOfHLevelΣ 1 isProp-is-section (λ _ -> isProp-is-sort) + + module Section (sort≡ : is-section) where + open Membership* isSetA + + list→slist-η : ∀ xs -> (x : A) -> list→slist xs ≡ [ x ]* -> xs ≡ [ x ] + list→slist-η [] x p = ⊥.rec (znots (congS S.length p)) + list→slist-η (x ∷ []) y p = congS [_] ([-]-inj {ϕ = isSetA} p) + list→slist-η (x ∷ y ∷ xs) z p = ⊥.rec (snotz (injSuc (congS S.length p))) + + sort-length≡-α : ∀ (xs : List A) -> L.length xs ≡ S.length (list→slist xs) + sort-length≡-α [] = refl + sort-length≡-α (x ∷ xs) = congS suc (sort-length≡-α xs) + + sort-length≡ : ∀ xs -> L.length (sort xs) ≡ S.length xs + sort-length≡ xs = sort-length≡-α (sort xs) ∙ congS S.length (sort≡ xs) + + length-0 : ∀ (xs : List A) -> L.length xs ≡ 0 -> xs ≡ [] + length-0 [] p = refl + length-0 (x ∷ xs) p = ⊥.rec (snotz p) + + sort-[] : ∀ xs -> sort xs ≡ [] -> xs ≡ []* + sort-[] xs p = sym (sort≡ xs) ∙ congS list→slist p + + sort-[]' : sort []* ≡ [] + sort-[]' = length-0 (sort []*) (sort-length≡ []*) + + sort-[-] : ∀ x -> sort [ x ]* ≡ [ x ] + sort-[-] x = list→slist-η (sort [ x ]*) x (sort≡ [ x ]*) + + sort-∈ : ∀ x xs -> x ∈* xs -> x ∈ sort xs + sort-∈ x xs p = ∈*→∈ x (sort xs) (subst (x ∈*_) (sym (sort≡ xs)) p) + + sort-∈* : ∀ x xs -> x ∈ sort xs -> x ∈* xs + sort-∈* x xs p = subst (x ∈*_) (sort≡ xs) (∈→∈* x (sort xs) p) + + sort-unique : ∀ xs -> is-sorted xs -> sort (list→slist xs) ≡ xs + sort-unique xs = P.rec (isOfHLevelList 0 isSetA _ xs) λ (ys , p) -> + sym (congS sort (sym (sort≡ ys) ∙ congS list→slist p)) ∙ p + + sort-choice-lemma : ∀ x -> sort (x ∷* x ∷* []*) ≡ x ∷ x ∷ [] + sort-choice-lemma x with sort (x ∷* x ∷* []*) | inspect sort (x ∷* x ∷* []*) + ... | [] | [ p ]ᵢ = ⊥.rec (snotz (sym (sort-length≡ (x ∷* x ∷* []*)) ∙ congS L.length p)) + ... | x₁ ∷ [] | [ p ]ᵢ = ⊥.rec (snotz (injSuc (sym (sort-length≡ (x ∷* x ∷* []*)) ∙ congS L.length p))) + ... | x₁ ∷ x₂ ∷ x₃ ∷ xs | [ p ]ᵢ = ⊥.rec (znots (injSuc (injSuc (sym (sort-length≡ (x ∷* x ∷* []*)) ∙ congS L.length p)))) + ... | a ∷ b ∷ [] | [ p ]ᵢ = + P.rec (isOfHLevelList 0 isSetA _ _) + (⊎.rec lemma1 (lemma1 ∘ x∈[y]→x≡y a x)) + (sort-∈* a (x ∷* x ∷* []*) (subst (a ∈_) (sym p) (x∈xs a [ b ]))) + where + lemma2 : a ≡ x -> b ≡ x -> a ∷ b ∷ [] ≡ x ∷ x ∷ [] + lemma2 q r = cong₂ (λ u v -> u ∷ v ∷ []) q r + lemma1 : a ≡ x -> a ∷ b ∷ [] ≡ x ∷ x ∷ [] + lemma1 q = + P.rec (isOfHLevelList 0 isSetA _ _) + (⊎.rec (lemma2 q) (lemma2 q ∘ x∈[y]→x≡y b x)) + (sort-∈* b (x ∷* x ∷* []*) (subst (b ∈_) (sym p) (L.inr (L.inl refl)))) + + sort-choice : ∀ x y -> (sort (x ∷* y ∷* []*) ≡ x ∷ y ∷ []) ⊔′ (sort (x ∷* y ∷* []*) ≡ y ∷ x ∷ []) + sort-choice x y with sort (x ∷* y ∷* []*) | inspect sort (x ∷* y ∷* []*) + ... | [] | [ p ]ᵢ = ⊥.rec (snotz (sym (sort-length≡ (x ∷* y ∷* []*)) ∙ congS L.length p)) + ... | x₁ ∷ [] | [ p ]ᵢ = ⊥.rec (snotz (injSuc (sym (sort-length≡ (x ∷* y ∷* []*)) ∙ congS L.length p))) + ... | x₁ ∷ x₂ ∷ x₃ ∷ xs | [ p ]ᵢ = ⊥.rec (znots (injSuc (injSuc (sym (sort-length≡ (x ∷* y ∷* []*)) ∙ congS L.length p)))) + ... | a ∷ b ∷ [] | [ p ]ᵢ = + P.rec squash₁ + (⊎.rec + (λ x≡a -> P.rec squash₁ + (⊎.rec + (λ y≡a -> L.inl (sym p ∙ subst (λ u -> sort (x ∷* [ u ]*) ≡ x ∷ u ∷ []) (x≡a ∙ sym y≡a) (sort-choice-lemma x))) + (λ y∈[b] -> L.inl (cong₂ (λ u v → u ∷ v ∷ []) (sym x≡a) (sym (x∈[y]→x≡y y b y∈[b])))) + ) + (subst (y ∈_) p (sort-∈ y (x ∷* y ∷* []*) (L.inr (L.inl refl)))) + ) + (λ x∈[b] -> P.rec squash₁ + (⊎.rec + (λ y≡a -> L.inr (cong₂ (λ u v → u ∷ v ∷ []) (sym y≡a) (sym (x∈[y]→x≡y x b x∈[b])))) + (λ y∈[b] -> + let x≡y = (x∈[y]→x≡y x b x∈[b]) ∙ sym (x∈[y]→x≡y y b y∈[b]) + in L.inl (sym p ∙ subst (λ u -> sort (x ∷* [ u ]*) ≡ x ∷ u ∷ []) x≡y (sort-choice-lemma x)) + ) + ) + (subst (y ∈_) p (sort-∈ y (x ∷* y ∷* []*) (L.inr (L.inl refl)))) + ) + ) + (subst (x ∈_) p (sort-∈ x (x ∷* y ∷* []*) (L.inl refl))) diff --git a/Cubical/Structures/Set/CMon/SList/Sort/Equiv.agda b/Cubical/Structures/Set/CMon/SList/Sort/Equiv.agda new file mode 100644 index 0000000..dc05969 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Sort/Equiv.agda @@ -0,0 +1,172 @@ +{-# OPTIONS --cubical --exact-split -WnoUnsupportedIndexedMatch --safe #-} + +module Cubical.Structures.Set.CMon.SList.Sort.Equiv where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order renaming (_≤_ to _≤ℕ_; _<_ to _<ℕ_) +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Order +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.HLevels +open import Cubical.Data.List +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Functions.Embedding +import Cubical.Data.List as L +open import Cubical.Functions.Logic as L hiding (¬_; ⊥) + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.Structures.Set.Mon.List +open import Cubical.Structures.Set.CMon.SList.Base renaming (_∷_ to _∷*_; [] to []*; [_] to [_]*; _++_ to _++*_) +import Cubical.Structures.Set.CMon.SList.Base as S +open import Cubical.Structures.Set.CMon.SList.Length as S +open import Cubical.Structures.Set.CMon.SList.Membership as S +open import Cubical.Structures.Set.CMon.SList.Sort.Base +open import Cubical.Structures.Set.CMon.SList.Sort.Sort +open import Cubical.Structures.Set.CMon.SList.Sort.Order + +open Iso + +private + variable + ℓ : Level + A : Type ℓ + +module Sort↔Order {ℓ : Level} {A : Type ℓ} (isSetA : isSet A) where + open Sort isSetA + open Sort.Section isSetA + open Sort→Order isSetA + open Order→Sort + open IsToset + + IsDecOrder : (A -> A -> Type ℓ) -> Type _ + IsDecOrder _≤_ = IsToset _≤_ × (∀ x y -> Dec (x ≤ y)) + + HasDecOrder : Type _ + HasDecOrder = Σ _ IsDecOrder + + IsHeadLeastSection : (SList A -> List A) -> Type _ + IsHeadLeastSection q = is-section q × is-head-least q + + IsSortSection : (SList A -> List A) -> Type _ + IsSortSection q = is-section q × is-head-least q × is-tail-sort q + + HasHeadLeastSectionAndIsDiscrete : Type _ + HasHeadLeastSectionAndIsDiscrete = (Σ _ IsHeadLeastSection) × (Discrete A) + + HasSortSectionAndIsDiscrete : Type _ + HasSortSectionAndIsDiscrete = (Σ _ IsSortSection) × (Discrete A) + + IsSortSection→IsHeadLeastSection : ∀ q -> IsSortSection q -> IsHeadLeastSection q + IsSortSection→IsHeadLeastSection q (section , head-least , _) = section , head-least + + order→sort : HasDecOrder -> HasSortSectionAndIsDiscrete + order→sort (_≤_ , isToset , isDec) = + (sort _≤_ isToset isDec , subst (λ isSetA' -> Sort.is-sort-section isSetA' (sort _≤_ isToset isDec)) (isPropIsSet _ _) (Order→Sort.sort-is-sort-section _≤_ isToset isDec)) , isDiscreteA _≤_ isToset isDec + + order→head-least : HasDecOrder -> HasHeadLeastSectionAndIsDiscrete + order→head-least p = let ((q , q-is-sort) , r) = order→sort p in (q , IsSortSection→IsHeadLeastSection q q-is-sort) , r + + head-least→order : HasHeadLeastSectionAndIsDiscrete -> HasDecOrder + head-least→order ((s , s-is-section , s-is-sort) , discA) = + _≤_ s s-is-section , ≤-isToset s s-is-section s-is-sort , dec-≤ s s-is-section discA + + sort→order : HasSortSectionAndIsDiscrete -> HasDecOrder + sort→order ((q , q-is-sort) , r) = head-least→order ((q , IsSortSection→IsHeadLeastSection q q-is-sort) , r) + + order→head-least→order : ∀ x -> head-least→order (order→head-least x) ≡ x + order→head-least→order (_≤_ , isToset , isDec) = + Σ≡Prop (λ _≤'_ -> isOfHLevelΣ 1 (isPropIsToset _) (λ p -> isPropΠ2 λ x y -> isPropDec (is-prop-valued p x y))) (sym ≤-≡) + where + _≤*_ : A -> A -> Type _ + _≤*_ = head-least→order (order→head-least (_≤_ , isToset , isDec)) .fst + + ≤*-isToset : IsToset _≤*_ + ≤*-isToset = head-least→order (order→head-least (_≤_ , isToset , isDec)) .snd .fst + + iso-to : ∀ x y -> x ≤ y -> x ≤* y + iso-to x y x≤y with isDec x y + ... | yes p = refl + ... | no ¬p = ⊥.rec (¬p x≤y) + + iso-from : ∀ x y -> x ≤* y -> x ≤ y + iso-from x y x≤y with isDec x y + ... | yes p = p + ... | no ¬p = ⊥.rec (¬p (subst (_≤ y) (just-inj y x x≤y) (is-refl isToset y))) + + iso-≤ : ∀ x y -> Iso (x ≤ y) (x ≤* y) + iso-≤ x y = iso (iso-to x y) (iso-from x y) (λ p -> is-prop-valued ≤*-isToset x y _ p) (λ p -> is-prop-valued isToset x y _ p) + ≤-≡ : _≤_ ≡ _≤*_ + ≤-≡ = funExt λ x -> funExt λ y -> isoToPath (iso-≤ x y) + + sort→order→sort : ∀ x -> order→sort (sort→order x) ≡ x + sort→order→sort ((s , s-is-section , s-is-sort) , discA) = + Σ≡Prop (λ _ -> isPropDiscrete) + $ Σ≡Prop isProp-is-sort-section + $ sym + $ funExt λ xs -> unique-sort' _≤*_ ≤*-isToset ≤*-dec s xs (s-is-section , ∣_∣₁ ∘ s-is-sort') + where + s' : SList A -> List A + s' = order→sort (sort→order ((s , s-is-section , s-is-sort) , discA)) .fst .fst + + s'-is-sort-section : is-sort-section s' + s'-is-sort-section = order→sort (sort→order ((s , s-is-section , s-is-sort) , discA)) .fst .snd + + _≤*_ : A -> A -> Type _ + _≤*_ = _≤_ s s-is-section + + ≤*-isToset : IsToset _≤*_ + ≤*-isToset = ≤-isToset s s-is-section (s-is-sort .fst) + + ≤*-dec : ∀ x y -> Dec (x ≤* y) + ≤*-dec = dec-≤ s s-is-section discA + + Sorted* : List A -> Type _ + Sorted* = Sorted _≤*_ ≤*-isToset ≤*-dec + + s-is-sort'' : ∀ xs n -> n ≡ L.length (s xs) -> Sorted* (s xs) + s-is-sort'' xs n p with n | s xs | inspect s xs + ... | zero | [] | _ = sorted-nil + ... | zero | y ∷ ys | _ = ⊥.rec (znots p) + ... | suc _ | [] | _ = ⊥.rec (snotz p) + ... | suc _ | y ∷ [] | _ = sorted-one y + ... | suc m | y ∷ z ∷ zs | [ q ]ᵢ = induction (s-is-sort'' (list→slist (z ∷ zs)) m (injSuc p ∙ wit)) + where + wit : suc (L.length zs) ≡ L.length (s (list→slist (z ∷ zs))) + wit = sym $ + L.length (s (list→slist (z ∷ zs))) ≡⟨ sort-length≡ s s-is-section (list→slist (z ∷ zs)) ⟩ + S.length (list→slist (z ∷ zs)) ≡⟨ sym (sort-length≡-α s s-is-section (z ∷ zs)) ⟩ + L.length (z ∷ zs) ∎ + + z∷zs-sorted : s (list→slist (z ∷ zs)) ≡ z ∷ zs + z∷zs-sorted = sort-unique s s-is-section (z ∷ zs) (s-is-sort .snd y (z ∷ zs) ∣ _ , q ∣₁) + + induction : Sorted* (s (list→slist (z ∷ zs))) -> Sorted* (y ∷ z ∷ zs) + induction IH = + sorted-cons y z zs + (is-sorted→≤ s s-is-section y z (s-is-sort .fst y z _ ∣ _ , q ∣₁ (L.inr (L.inl refl)))) + (subst Sorted* z∷zs-sorted IH) + + s-is-sort' : ∀ xs -> Sorted* (s xs) + s-is-sort' xs = s-is-sort'' xs (L.length (s xs)) refl helps with termination checking + + without tail sort, we cannot construct a full equivalence + order⊂head-least : isEmbedding order→head-least + order⊂head-least = injEmbedding (isSet× (isSetΣ (isSetΠ (λ _ -> isOfHLevelList 0 isSetA)) λ q -> isProp→isSet (isProp× (isProp-is-section q) (isProp-is-head-least q))) (isProp→isSet isPropDiscrete)) + (λ p -> sym (order→head-least→order _) ∙ congS head-least→order p ∙ order→head-least→order _) + + with tail sort, we can construct a full equivalence + sort↔order : Iso HasDecOrder HasSortSectionAndIsDiscrete + sort↔order = iso order→sort sort→order sort→order→sort order→head-least→order diff --git a/Cubical/Structures/Set/CMon/SList/Sort/Order.agda b/Cubical/Structures/Set/CMon/SList/Sort/Order.agda new file mode 100644 index 0000000..dd6ad06 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Sort/Order.agda @@ -0,0 +1,410 @@ +{-# OPTIONS --cubical --safe --exact-split -WnoUnsupportedIndexedMatch #-} + +module Cubical.Structures.Set.CMon.SList.Sort.Order where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order renaming (_≤_ to _≤ℕ_; _<_ to _<ℕ_) +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Order +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.HLevels +open import Cubical.Data.List +open import Cubical.HITs.PropositionalTruncation as P +import Cubical.Data.List as L +open import Cubical.Functions.Logic as L hiding (¬_; ⊥) + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.Structures.Set.Mon.List +open import Cubical.Structures.Set.CMon.SList.Base renaming (_∷_ to _∷*_; [] to []*; [_] to [_]*; _++_ to _++*_) +import Cubical.Structures.Set.CMon.SList.Base as S +open import Cubical.Structures.Set.CMon.SList.Length as S +open import Cubical.Structures.Set.CMon.SList.Membership as S +open import Cubical.Structures.Set.CMon.SList.Sort.Base + +open Iso + +private + variable + ℓ : Level + A : Type ℓ + +module Order→Sort {A : Type ℓ} (_≤_ : A -> A -> Type ℓ) (≤-isToset : IsToset _≤_) (_≤?_ : ∀ x y -> Dec (x ≤ y)) where + open IsToset ≤-isToset + open Membership is-set + open Membership* is-set + + isDiscreteA : Discrete A + isDiscreteA x y with x ≤? y | y ≤? x + ... | yes p | yes q = yes (is-antisym x y p q) + ... | yes p | no ¬q = no λ r -> ¬q (subst (_≤ x) r (is-refl x)) + ... | no ¬p | yes q = no λ r -> ¬p (subst (x ≤_) r (is-refl x)) + ... | no ¬p | no ¬q = ⊥.rec $ P.rec isProp⊥ (⊎.rec ¬p ¬q) (is-strongly-connected x y) + + A-is-loset : Loset _ _ + A-is-loset = Toset→Loset (A , tosetstr _≤_ ≤-isToset) isDiscreteA + + _<_ : A -> A -> Type ℓ + _<_ = LosetStr._<_ (A-is-loset .snd) + + <-isLoset : IsLoset _<_ + <-isLoset = LosetStr.isLoset (A-is-loset .snd) + + insert : A -> List A -> List A + insert x [] = [ x ] + insert x (y ∷ ys) with x ≤? y + ... | yes p = x ∷ y ∷ ys + ... | no ¬p = y ∷ insert x ys + + insert-β-1 : ∀ x y ys -> x ≤ y -> insert x (y ∷ ys) ≡ x ∷ y ∷ ys + insert-β-1 x y ys p with x ≤? y + ... | yes _ = refl + ... | no ¬p = ⊥.rec (¬p p) + + insert-β-2 : ∀ x y ys -> ¬ (x ≤ y) -> insert x (y ∷ ys) ≡ y ∷ insert x ys + insert-β-2 x y ys ¬p with x ≤? y + ... | yes p = ⊥.rec (¬p p) + ... | no ¬_ = refl + + insert-≤ : ∀ x y xs ys -> insert x (insert y xs) ≡ x ∷ y ∷ ys -> x ≤ y + insert-≤ x y [] ys p with x ≤? y + ... | yes x≤y = x≤y + ... | no ¬x≤y = subst (_≤ y) (cons-inj₁ p) (is-refl y) + insert-≤ x y (z ∷ zs) ys p with y ≤? z + ... | yes y≤z = lemma + where + lemma : x ≤ y + lemma with x ≤? y + ... | yes x≤y = x≤y + ... | no ¬x≤y = subst (_≤ y) (cons-inj₁ p) (is-refl y) + ... | no ¬y≤z = lemma + where + lemma : x ≤ y + lemma with x ≤? z + ... | yes x≤z = subst (x ≤_) (cons-inj₁ (cons-inj₂ p)) x≤z + ... | no ¬x≤z = ⊥.rec (¬x≤z (subst (_≤ z) (cons-inj₁ p) (is-refl z))) + + insert-cons : ∀ x xs ys -> x ∷ xs ≡ insert x ys -> xs ≡ ys + insert-cons x xs [] p = cons-inj₂ p + insert-cons x xs (y ∷ ys) p with x ≤? y + ... | yes x≤y = cons-inj₂ p + ... | no ¬x≤y = ⊥.rec (¬x≤y (subst (x ≤_) (cons-inj₁ p) (is-refl x))) + + private + not-total-impossible : ∀ {x y} -> ¬(x ≤ y) -> ¬(y ≤ x) -> ⊥ + not-total-impossible {x} {y} p q = + P.rec isProp⊥ (⊎.rec p q) (is-strongly-connected x y) + + insert-insert : ∀ x y xs -> insert x (insert y xs) ≡ insert y (insert x xs) + insert-insert x y [] with x ≤? y | y ≤? x + ... | yes x≤y | no ¬y≤x = refl + ... | no ¬x≤y | yes y≤x = refl + ... | no ¬x≤y | no ¬y≤x = ⊥.rec (not-total-impossible ¬x≤y ¬y≤x) + ... | yes x≤y | yes y≤x = let x≡y = is-antisym x y x≤y y≤x in cong₂ (λ u v -> u ∷ v ∷ []) x≡y (sym x≡y) + insert-insert x y (z ∷ zs) with y ≤? z | x ≤? z + ... | yes y≤z | yes x≤z = case1 y≤z x≤z + where + case1 : y ≤ z -> x ≤ z -> insert x (y ∷ z ∷ zs) ≡ insert y (x ∷ z ∷ zs) + case1 y≤z x≤z with x ≤? y | y ≤? x + ... | yes x≤y | yes y≤x = let x≡y = is-antisym x y x≤y y≤x in cong₂ (λ u v -> (u ∷ v ∷ z ∷ zs)) x≡y (sym x≡y) + ... | yes x≤y | no ¬y≤x = sym (congS (x ∷_) (insert-β-1 y z zs y≤z)) + ... | no ¬x≤y | yes y≤x = congS (y ∷_) (insert-β-1 x z zs x≤z) + ... | no ¬x≤y | no ¬y≤x = ⊥.rec (not-total-impossible ¬x≤y ¬y≤x) + ... | yes y≤z | no ¬x≤z = case2 y≤z ¬x≤z + where + case2 : y ≤ z -> ¬(x ≤ z) -> insert x (y ∷ z ∷ zs) ≡ insert y (z ∷ insert x zs) + case2 y≤z ¬x≤z with x ≤? y + ... | yes x≤y = ⊥.rec (¬x≤z (is-trans x y z x≤y y≤z)) + ... | no ¬x≤y = congS (y ∷_) (insert-β-2 x z zs ¬x≤z) ∙ sym (insert-β-1 y z _ y≤z) + ... | no ¬y≤z | yes x≤z = case3 ¬y≤z x≤z + where + case3 : ¬(y ≤ z) -> x ≤ z -> insert x (z ∷ insert y zs) ≡ insert y (x ∷ z ∷ zs) + case3 ¬y≤z x≤z with y ≤? x + ... | yes y≤x = ⊥.rec (¬y≤z (is-trans y x z y≤x x≤z)) + ... | no ¬y≤x = insert-β-1 x z _ x≤z ∙ congS (x ∷_) (sym (insert-β-2 y z zs ¬y≤z)) + ... | no ¬y≤z | no ¬x≤z = + insert x (z ∷ insert y zs) ≡⟨ insert-β-2 x z _ ¬x≤z ⟩ + z ∷ insert x (insert y zs) ≡⟨ congS (z ∷_) (insert-insert x y zs) ⟩ + z ∷ insert y (insert x zs) ≡⟨ sym (insert-β-2 y z _ ¬y≤z) ⟩ + insert y (z ∷ insert x zs) ∎ + + sort : SList A -> List A + sort = Elim.f [] + (λ x xs -> insert x xs) + (λ x y xs -> insert-insert x y xs) + (λ _ -> isOfHLevelList 0 is-set) + + insert-is-permute : ∀ x xs -> list→slist (x ∷ xs) ≡ list→slist (insert x xs) + insert-is-permute x [] = refl + insert-is-permute x (y ∷ ys) with x ≤? y + ... | yes p = refl + ... | no ¬p = + x ∷* y ∷* list→slist ys ≡⟨ swap x y _ ⟩ + y ∷* x ∷* list→slist ys ≡⟨ congS (y ∷*_) (insert-is-permute x ys) ⟩ + y ∷* list→slist (insert x ys) ∎ + + open Sort is-set + + sort-is-permute : is-section sort + sort-is-permute = ElimProp.f (trunc _ _) refl lemma + where + lemma : ∀ x {xs} p -> list→slist (sort (x ∷* xs)) ≡ x ∷* xs + lemma x {xs} p = sym $ + x ∷* xs ≡⟨ congS (x ∷*_) (sym p) ⟩ + list→slist (x ∷ sort xs) ≡⟨ insert-is-permute x (sort xs) ⟩ + list→slist (insert x (sort xs)) ≡⟨⟩ + list→slist (sort (x ∷* xs)) ∎ + + tail-is-sorted : ∀ x xs -> is-sorted sort (x ∷ xs) -> is-sorted sort xs + tail-is-sorted x xs = P.rec squash₁ (uncurry lemma) + where + lemma : ∀ ys p -> is-sorted sort xs + lemma ys p = ∣ (list→slist xs) , sym (insert-cons x _ _ sort-proof) ∣₁ + where + ys-proof : ys ≡ x ∷* list→slist xs + ys-proof = sym (sort-is-permute ys) ∙ (congS list→slist p) + sort-proof : x ∷ xs ≡ insert x (sort (list→slist xs)) + sort-proof = sym p ∙ congS sort ys-proof + + data Sorted : List A -> Type ℓ where + sorted-nil : Sorted [] + sorted-one : ∀ x -> Sorted [ x ] + sorted-cons : ∀ x y zs -> x ≤ y -> Sorted (y ∷ zs) -> Sorted (x ∷ y ∷ zs) + + open Sort.Section is-set + + is-sort' : (SList A -> List A) -> Type _ + is-sort' f = (∀ xs -> list→slist (f xs) ≡ xs) × (∀ xs -> ∥ Sorted (f xs) ∥₁) + + tail-sorted' : ∀ {x xs} -> Sorted (x ∷ xs) -> Sorted xs + tail-sorted' (sorted-one ._) = sorted-nil + tail-sorted' (sorted-cons ._ _ _ _ p) = p + + ≤-tail : ∀ {x y ys} -> y ∈ (x ∷ ys) -> Sorted (x ∷ ys) -> x ≤ y + ≤-tail {y = y} p (sorted-one x) = subst (_≤ y) (x∈[y]→x≡y y x p) (is-refl y) + ≤-tail {x} {y = z} p (sorted-cons x y zs q r) = + P.rec (is-prop-valued x z) + (⊎.rec + (λ z≡x -> subst (x ≤_) (sym z≡x) (is-refl x)) + (λ z∈ys -> is-trans x y z q (≤-tail z∈ys r)) + ) p + + smallest-sorted : ∀ x xs -> (∀ y -> y ∈ xs -> x ≤ y) -> Sorted xs -> Sorted (x ∷ xs) + smallest-sorted x .[] p sorted-nil = + sorted-one x + smallest-sorted x .([ y ]) p (sorted-one y) = + sorted-cons x y [] (p y (x∈[x] y)) (sorted-one y) + smallest-sorted x .(_ ∷ _ ∷ _) p (sorted-cons y z zs y≤z r) = + sorted-cons x y (z ∷ zs) (p y (x∈xs y (z ∷ zs))) (smallest-sorted y (z ∷ zs) lemma r) + where + lemma : ∀ a -> a ∈ (z ∷ zs) -> y ≤ a + lemma a = P.rec (is-prop-valued y a) $ ⊎.rec + (λ a≡z -> subst (y ≤_) (sym a≡z) y≤z) + (λ a∈zs -> is-trans y z a y≤z (≤-tail (L.inr a∈zs) r)) + + insert-∈ : ∀ {x} {y} ys -> x ∈ insert y ys -> x ∈ (y ∷ ys) + insert-∈ {x} {y} ys p = ∈*→∈ x (y ∷ ys) + (subst (x ∈*_) (sym (insert-is-permute y ys)) (∈→∈* x (insert y ys) p)) + + insert-is-sorted : ∀ x xs -> Sorted xs -> Sorted (insert x xs) + insert-is-sorted x [] p = sorted-one x + insert-is-sorted x (y ∷ ys) p with x ≤? y + ... | yes q = sorted-cons x y ys q p + ... | no ¬q = smallest-sorted y (insert x ys) (lemma ys p) IH + where + IH : Sorted (insert x ys) + IH = insert-is-sorted x ys (tail-sorted' p) + y≤x : y ≤ x + y≤x = P.rec (is-prop-valued y x) (⊎.rec (idfun _) (⊥.rec ∘ ¬q)) (is-strongly-connected y x) + lemma : ∀ zs -> Sorted (y ∷ zs) -> ∀ z -> z ∈ insert x zs → y ≤ z + lemma zs p z r = P.rec (is-prop-valued y z) + (⊎.rec (λ z≡x -> subst (y ≤_) (sym z≡x) y≤x) (λ z∈zs -> ≤-tail (L.inr z∈zs) p)) (insert-∈ zs r) + + sort-is-sorted' : ∀ xs -> ∥ Sorted (sort xs) ∥₁ + sort-is-sorted' = ElimProp.f squash₁ ∣ sorted-nil ∣₁ + λ x -> P.rec squash₁ λ p -> ∣ (insert-is-sorted x _ p) ∣₁ + + sort-is-sorted'' : ∀ xs -> is-sorted sort xs -> ∥ Sorted xs ∥₁ + sort-is-sorted'' xs = P.rec squash₁ λ (ys , p) -> + P.map (subst Sorted p) (sort-is-sorted' ys) + + Step 1. show both sort xs and sort->order->sort xs give sorted list + Step 2. apply this lemma + Step 3. get sort->order->sort = sort + unique-sorted-xs : ∀ xs ys -> list→slist xs ≡ list→slist ys -> Sorted xs -> Sorted ys -> xs ≡ ys + unique-sorted-xs [] [] p xs-sorted ys-sorted = refl + unique-sorted-xs [] (y ∷ ys) p xs-sorted ys-sorted = ⊥.rec (znots (congS S.length p)) + unique-sorted-xs (x ∷ xs) [] p xs-sorted ys-sorted = ⊥.rec (snotz (congS S.length p)) + unique-sorted-xs (x ∷ xs) (y ∷ ys) p xs-sorted ys-sorted = + cong₂ _∷_ x≡y (unique-sorted-xs xs ys xs≡ys (tail-sorted' xs-sorted) (tail-sorted' ys-sorted)) + where + x≤y : x ≤ y + x≤y = ≤-tail (∈*→∈ y (x ∷ xs) (subst (y ∈*_) (sym p) (L.inl refl))) xs-sorted + y≤x : y ≤ x + y≤x = ≤-tail (∈*→∈ x (y ∷ ys) (subst (x ∈*_) p (L.inl refl))) ys-sorted + x≡y : x ≡ y + x≡y = is-antisym x y x≤y y≤x + xs≡ys : list→slist xs ≡ list→slist ys + xs≡ys = + list→slist xs ≡⟨ remove1-≡-lemma isDiscreteA (list→slist xs) refl ⟩ + remove1 isDiscreteA x (list→slist (x ∷ xs)) ≡⟨ congS (remove1 isDiscreteA x) p ⟩ + remove1 isDiscreteA x (list→slist (y ∷ ys)) ≡⟨ sym (remove1-≡-lemma isDiscreteA (list→slist ys) x≡y) ⟩ + list→slist ys ∎ + + unique-sort : ∀ f -> is-sort' f -> f ≡ sort + unique-sort f (f-is-permute , f-is-sorted) = funExt λ xs -> + P.rec2 (isOfHLevelList 0 is-set _ _) + (unique-sorted-xs (f xs) (sort xs) (f-is-permute xs ∙ sym (sort-is-permute xs))) + (f-is-sorted xs) + (sort-is-sorted' xs) + + unique-sort' : ∀ f xs -> is-sort' f -> f xs ≡ sort xs + unique-sort' f xs p = congS (λ g -> g xs) (unique-sort f p) + + private + sort→order-lemma : ∀ x y xs -> is-sorted sort (x ∷ y ∷ xs) -> x ≤ y + sort→order-lemma x y xs = P.rec (is-prop-valued x y) (uncurry lemma) + where + lemma : ∀ ys p -> x ≤ y + lemma ys p = insert-≤ x y (sort tail) xs (congS sort tail-proof ∙ p) + where + tail : SList A + tail = list→slist xs + tail-proof : x ∷* y ∷* tail ≡ ys + tail-proof = sym (congS list→slist p) ∙ sort-is-permute ys + + sort→order : ∀ x y xs -> is-sorted sort (x ∷ xs) -> y ∈ (x ∷ xs) -> x ≤ y + sort→order x y [] p y∈xs = subst (_≤ y) (x∈[y]→x≡y y x y∈xs) (is-refl y) + sort→order x y (z ∷ zs) p y∈x∷z∷zs with isDiscreteA x y + ... | yes x≡y = subst (x ≤_) x≡y (is-refl x) + ... | no ¬x≡y = + P.rec (is-prop-valued x y) (⊎.rec (⊥.rec ∘ ¬x≡y ∘ sym) lemma) y∈x∷z∷zs + where + lemma : y ∈ (z ∷ zs) -> x ≤ y + lemma y∈z∷zs = is-trans x z y + (sort→order-lemma x z zs p) + (sort→order z y zs (tail-is-sorted x (z ∷ zs) p) y∈z∷zs) + + sort-is-head-least : is-head-least sort + sort-is-head-least x y xs p y∈xs = ∣ (x ∷* y ∷* []* , insert-β-1 x y [] (sort→order x y xs p y∈xs)) ∣₁ + + sort-is-sort-section : is-sort-section sort + sort-is-sort-section = sort-is-permute , sort-is-head-least , tail-is-sorted + + + swap1-2 : List A -> List A + swap1-2 [] = [] + swap1-2 (x ∷ []) = x ∷ [] + swap1-2 (x ∷ y ∷ xs) = y ∷ x ∷ xs + + swap1-2-id : ∀ xs -> swap1-2 (swap1-2 xs) ≡ xs + swap1-2-id [] = refl + swap1-2-id (x ∷ []) = refl + swap1-2-id (x ∷ y ∷ xs) = refl + + swap2-3 : List A -> List A + swap2-3 [] = [] + swap2-3 (x ∷ xs) = x ∷ swap1-2 xs + + swap2-3-id : ∀ xs -> swap2-3 (swap2-3 xs) ≡ xs + swap2-3-id [] = refl + swap2-3-id (x ∷ xs) = congS (x ∷_) (swap1-2-id xs) + + swap2-3-is-permute : ∀ f -> is-section f -> is-section (swap2-3 ∘ f) + swap2-3-is-permute f f-is-permute xs with f xs | inspect f xs + ... | [] | [ p ]ᵢ = sym (congS list→slist p) ∙ f-is-permute xs + ... | x ∷ [] | [ p ]ᵢ = sym (congS list→slist p) ∙ f-is-permute xs + ... | x ∷ y ∷ [] | [ p ]ᵢ = sym (congS list→slist p) ∙ f-is-permute xs + ... | x ∷ y ∷ z ∷ ys | [ p ]ᵢ = + x ∷* z ∷* y ∷* list→slist ys ≡⟨ congS (x ∷*_) (swap z y (list→slist ys)) ⟩ + x ∷* y ∷* z ∷* list→slist ys ≡⟨ sym (congS list→slist p) ⟩ + list→slist (f xs) ≡⟨ f-is-permute xs ⟩ + xs ∎ + + head-least-only : SList A -> List A + head-least-only = swap2-3 ∘ sort + + head-least-only-is-permute : is-section head-least-only + head-least-only-is-permute = swap2-3-is-permute sort sort-is-permute + + private + head-least-is-same-for-2 : ∀ u v -> is-sorted sort (u ∷ v ∷ []) -> is-sorted head-least-only (u ∷ v ∷ []) + head-least-is-same-for-2 u v = P.map λ (ys , q) -> ys , congS swap2-3 q + + head-least→sorted : ∀ x xs -> is-sorted head-least-only (x ∷ xs) -> is-sorted sort (x ∷ swap1-2 xs) + head-least→sorted x [] = + P.map λ (ys , q) -> ys , sym (swap2-3-id (sort ys)) ∙ congS swap2-3 q + head-least→sorted x (y ∷ []) = + P.map λ (ys , q) -> ys , sym (swap2-3-id (sort ys)) ∙ congS swap2-3 q + head-least→sorted x (y ∷ z ∷ xs) = + P.map λ (ys , q) -> ys , sym (swap2-3-id (sort ys)) ∙ congS swap2-3 q + + swap1-2-∈ : ∀ x xs -> x ∈ xs -> x ∈ swap1-2 xs + swap1-2-∈ x [] x∈ = x∈ + swap1-2-∈ x (y ∷ []) x∈ = x∈ + swap1-2-∈ x (y ∷ z ∷ xs) = P.rec squash₁ + (⊎.rec (L.inr ∘ L.inl) (P.rec squash₁ (⊎.rec L.inl (L.inr ∘ L.inr)))) + + swap2-3-∈ : ∀ x xs -> x ∈ xs -> x ∈ swap2-3 xs + swap2-3-∈ x (y ∷ xs) = P.rec squash₁ (⊎.rec L.inl (L.inr ∘ swap1-2-∈ x xs)) + + is-sorted→≤ : ∀ {x y} -> is-sorted sort (x ∷ y ∷ []) -> x ≤ y + is-sorted→≤ {x} {y} p = P.rec (is-prop-valued x y) + (≤-tail {ys = y ∷ []} (L.inr (L.inl refl))) + (sort-is-sorted'' (x ∷ y ∷ []) p) + + head-least-tail-sort→x≡y : ∀ x y -> is-tail-sort head-least-only -> x ≤ y -> x ≡ y + head-least-tail-sort→x≡y x y h-tail-sort x≤y = + is-antisym x y x≤y (is-sorted→≤ (lemma1 ∣ x ∷* x ∷* y ∷* []* , lemma2 ∣₁)) + where + [1, 2, 3] sorted by sort -> [1, 3, 2] sorted by head-least -> [3, 2] sorted by head-least -> [3, 2] sorted by sort + lemma1 : is-sorted sort (x ∷ x ∷ y ∷ []) -> is-sorted sort (y ∷ x ∷ []) + lemma1 = P.rec squash₁ λ (ys , q) -> head-least→sorted y [ x ] (h-tail-sort x _ ∣ ys , congS swap2-3 q ∣₁) + lemma2 : sort (x ∷* x ∷* [ y ]*) ≡ x ∷ x ∷ y ∷ [] + lemma2 = + insert x (insert x [ y ]) ≡⟨ congS (insert x) (insert-β-1 x y [] x≤y) ⟩ + insert x (x ∷ y ∷ []) ≡⟨ insert-β-1 x x [ y ] (is-refl x) ⟩ + x ∷ x ∷ [ y ] ∎ + + head-least-only-is-head-least : is-head-least head-least-only + head-least-only-is-head-least x y xs p y∈xs = head-least-is-same-for-2 x y + (sort-is-head-least x y (swap1-2 xs) (head-least→sorted x xs p) (swap2-3-∈ y (x ∷ xs) y∈xs)) + + head-least-tail-sort→isProp-A : is-tail-sort head-least-only -> isProp A + head-least-tail-sort→isProp-A h-tail-sort x y = P.rec (is-set _ _) + (⊎.rec (head-least-tail-sort→x≡y x y h-tail-sort) (sym ∘ (head-least-tail-sort→x≡y y x h-tail-sort))) + (is-strongly-connected x y) + +module Order→Sort-Example where + + ≤ℕ-isToset : IsToset _≤ℕ_ + ≤ℕ-isToset = istoset isSetℕ + (λ _ _ -> isProp≤) + (λ _ -> ≤-refl) + (λ _ _ _ -> ≤-trans) + (λ _ _ -> ≤-antisym) + lemma + where + <→≤ : ∀ {n m} -> n <ℕ m -> n ≤ℕ m + <→≤ (k , p) = suc k , sym (+-suc k _) ∙ p + lemma : BinaryRelation.isStronglyConnected _≤ℕ_ + lemma x y = ∣ ⊎.rec ⊎.inl (_⊎_.inr ∘ <→≤) (splitℕ-≤ x y) ∣₁ + + open Order→Sort _≤ℕ_ ≤ℕ-isToset ≤Dec + + _ : sort (4 ∷* 6 ∷* 1 ∷* 2 ∷* []*) ≡ (1 ∷ 2 ∷ 4 ∷ 6 ∷ []) + _ = refl + + \ No newline at end of file diff --git a/Cubical/Structures/Set/CMon/SList/Sort/Sort.agda b/Cubical/Structures/Set/CMon/SList/Sort/Sort.agda new file mode 100644 index 0000000..3dfbb86 --- /dev/null +++ b/Cubical/Structures/Set/CMon/SList/Sort/Sort.agda @@ -0,0 +1,178 @@ +{-# OPTIONS --cubical --safe --exact-split -WnoUnsupportedIndexedMatch #-} + +module Cubical.Structures.Set.CMon.SList.Sort.Sort where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order renaming (_≤_ to _≤ℕ_; _<_ to _<ℕ_) +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Empty as ⊥ +open import Cubical.Induction.WellFounded +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Order +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.HLevels +open import Cubical.Data.List +open import Cubical.HITs.PropositionalTruncation as P +import Cubical.Data.List as L +open import Cubical.Functions.Logic as L hiding (¬_; ⊥) + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.Structures.Set.Mon.List +open import Cubical.Structures.Set.CMon.SList.Base renaming (_∷_ to _∷*_; [] to []*; [_] to [_]*; _++_ to _++*_) +import Cubical.Structures.Set.CMon.SList.Base as S +open import Cubical.Structures.Set.CMon.SList.Length as S +open import Cubical.Structures.Set.CMon.SList.Membership as S +open import Cubical.Structures.Set.CMon.SList.Sort.Base + +open Iso + +private + variable + ℓ : Level + A : Type ℓ + +module Sort→Order (isSetA : isSet A) (sort : SList A -> List A) (sort≡ : ∀ xs -> list→slist (sort xs) ≡ xs) where + + isSetMaybeA : isSet (Maybe A) + isSetMaybeA = isOfHLevelMaybe 0 isSetA + + isSetListA : isSet (List A) + isSetListA = isOfHLevelList 0 isSetA + + private + module 𝔖 = M.CMonSEq < SList A , slist-α > slist-sat + + open Membership isSetA + open Membership* isSetA + open Sort isSetA sort + open Sort.Section isSetA sort sort≡ + + least : SList A -> Maybe A + least xs = head-maybe (sort xs) + + least-nothing : ∀ xs -> least xs ≡ nothing -> xs ≡ []* + least-nothing xs p with sort xs | inspect sort xs + ... | [] | [ q ]ᵢ = sort-[] xs q + ... | y ∷ ys | [ q ]ᵢ = ⊥.rec (¬just≡nothing p) + + least-Σ : ∀ x xs -> least xs ≡ just x -> Σ[ ys ∈ List A ] (sort xs ≡ x ∷ ys) + least-Σ x xs p with sort xs + ... | [] = ⊥.rec (¬nothing≡just p) + ... | y ∷ ys = ys , congS (_∷ ys) (just-inj y x p) + + least-in : ∀ x xs -> least xs ≡ just x -> x ∈* xs + least-in x xs p = + let (ys , q) = least-Σ x xs p + x∷ys≡xs = congS list→slist (sym q) ∙ sort≡ xs + in subst (x ∈*_) x∷ys≡xs (x∈*xs x (list→slist ys)) + + least-choice : ∀ x y -> (least (x ∷* [ y ]*) ≡ just x) ⊔′ (least (x ∷* [ y ]*) ≡ just y) + least-choice x y = P.rec squash₁ + (⊎.rec + (L.inl ∘ congS head-maybe) + (L.inr ∘ congS head-maybe) + ) + (sort-choice x y) + + _≤_ : A -> A -> Type _ + x ≤ y = least (x ∷* y ∷* []*) ≡ just x + + isProp-≤ : ∀ {a} {b} -> isProp (a ≤ b) + isProp-≤ = isSetMaybeA _ _ + + ≤-Prop : ∀ x y -> hProp _ + ≤-Prop x y = (x ≤ y) , isProp-≤ + + refl-≤ : ∀ x -> x ≤ x + refl-≤ x = P.rec isProp-≤ (⊎.rec (idfun _) (idfun _)) (least-choice x x) + + antisym-≤ : ∀ x y -> x ≤ y -> y ≤ x -> x ≡ y + antisym-≤ x y p q = P.rec (isSetA x y) + (⊎.rec + (λ xy -> just-inj x y $ + just x ≡⟨ sym xy ⟩ + least (x ∷* y ∷* []*) ≡⟨ congS least (swap x y []*) ⟩ + least (y ∷* x ∷* []*) ≡⟨ q ⟩ + just y + ∎) + (λ yx -> just-inj x y $ + just x ≡⟨ sym p ⟩ + least (x ∷* [ y ]*) ≡⟨ yx ⟩ + just y + ∎)) + (least-choice x y) + + total-≤ : ∀ x y -> (x ≤ y) ⊔′ (y ≤ x) + total-≤ x y = P.rec squash₁ + (⊎.rec + L.inl + (λ p -> L.inr $ + least (y ∷* [ x ]*) ≡⟨ congS least (swap y x []*) ⟩ + least (x ∷* [ y ]*) ≡⟨ p ⟩ + just y + ∎)) + (least-choice x y) + + dec-≤ : (discA : Discrete A) -> ∀ x y -> Dec (x ≤ y) + dec-≤ discA x y = discreteMaybe discA _ _ + + is-sorted→≤ : ∀ x y -> is-sorted (x ∷ y ∷ []) -> x ≤ y + is-sorted→≤ x y = P.rec (isSetMaybeA _ _) λ (xs , p) -> + congS head-maybe (congS sort (sym (sym (sort≡ xs) ∙ congS list→slist p)) ∙ p) + + module _ (sort-is-sort : is-head-least) where + trans-≤ : ∀ x y z -> x ≤ y -> y ≤ z -> x ≤ z + trans-≤ x y z x≤y y≤z with least (x ∷* y ∷* z ∷* []*) | inspect least (x ∷* y ∷* z ∷* []*) + ... | nothing | [ p ]ᵢ = ⊥.rec (snotz (congS S.length (least-nothing _ p))) + ... | just u | [ p ]ᵢ = + P.rec (isSetMaybeA _ _) + (⊎.rec case1 + (P.rec (isSetMaybeA _ _) + (⊎.rec case2 (case3 ∘ x∈[y]→x≡y _ _)) + ) + ) + (least-in u (x ∷* y ∷* z ∷* []*) p) + where + tail' : Σ[ xs ∈ List A ] u ∷ xs ≡ sort (x ∷* y ∷* z ∷* []*) + tail' with sort (x ∷* y ∷* z ∷* []*) + ... | [] = ⊥.rec (¬nothing≡just p) + ... | v ∷ xs = xs , congS (_∷ xs) (just-inj _ _ (sym p)) + tail : List A + tail = tail' .fst + tail-proof : u ∷ tail ≡ sort (x ∷* y ∷* z ∷* []*) + tail-proof = tail' .snd + u∷tail-is-sorted : is-sorted (u ∷ tail) + u∷tail-is-sorted = ∣ ((x ∷* y ∷* z ∷* []*) , sym tail-proof) ∣₁ + u-is-smallest : ∀ v -> v ∈* (x ∷* y ∷* z ∷* []*) -> u ≤ v + u-is-smallest v q = + is-sorted→≤ u v (sort-is-sort u v tail u∷tail-is-sorted (subst (v ∈_) (sym tail-proof) (sort-∈ v _ q))) + case1 : u ≡ x -> x ≤ z + case1 u≡x = subst (_≤ z) u≡x (u-is-smallest z (L.inr (L.inr (L.inl refl)))) + case2 : u ≡ y -> x ≤ z + case2 u≡y = subst (_≤ z) (antisym-≤ y x y≤x x≤y) y≤z + where + y≤x : y ≤ x + y≤x = subst (_≤ x) u≡y (u-is-smallest x (L.inl refl)) + case3 : u ≡ z -> x ≤ z + case3 u≡z = subst (x ≤_) (antisym-≤ y z y≤z z≤y) x≤y + where + z≤y : z ≤ y + z≤y = subst (_≤ y) u≡z (u-is-smallest y (L.inr (L.inl refl))) + + ≤-isToset : IsToset _≤_ + IsToset.is-set ≤-isToset = isSetA + IsToset.is-prop-valued ≤-isToset x y = isOfHLevelMaybe 0 isSetA _ _ + IsToset.is-refl ≤-isToset = refl-≤ + IsToset.is-trans ≤-isToset = trans-≤ + IsToset.is-antisym ≤-isToset = antisym-≤ + IsToset.is-strongly-connected ≤-isToset = total-≤ diff --git a/Cubical/Structures/Set/Empty.agda b/Cubical/Structures/Set/Empty.agda new file mode 100644 index 0000000..6ae37e0 --- /dev/null +++ b/Cubical/Structures/Set/Empty.agda @@ -0,0 +1,65 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.Empty where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +import Cubical.Data.List as L + +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity + +private + variable + ℓ ℓ' : Level + A : Type ℓ + +empty-α : ∀ (A : Type ℓ) -> sig emptySig A -> A +empty-α _ (x , _) = ⊥.rec x + +emptyHomDegen : (𝔜 : struct ℓ' emptySig) -> structHom < A , empty-α A > 𝔜 ≃ (A -> 𝔜 .car) +emptyHomDegen _ = Σ-contractSnd λ _ -> isContrΠ⊥ + +module EmptyDef = F.Definition emptySig emptyEqSig emptySEq + +empty-sat : ∀ (A : Type ℓ) -> < A , empty-α A > ⊨ emptySEq +empty-sat _ eqn ρ = ⊥.rec eqn + +treeEmpty≃ : Tree emptySig A ≃ A +treeEmpty≃ = isoToEquiv (iso from leaf (λ _ -> refl) leaf∘from) + where + from : Tree emptySig A -> A + from (leaf x) = x + + leaf∘from : retract from leaf + leaf∘from (leaf x) = refl + +treeDef : ∀ {ℓ ℓ'} -> EmptyDef.Free ℓ ℓ' 2 +F.Definition.Free.F treeDef = Tree emptySig +F.Definition.Free.η treeDef = leaf +F.Definition.Free.α treeDef = empty-α (Tree emptySig _) +F.Definition.Free.sat treeDef = empty-sat (Tree emptySig _) +F.Definition.Free.isFree (treeDef {ℓ = ℓ}) {X = A} {𝔜 = 𝔜} H ϕ = lemma .snd + where + 𝔗 : struct ℓ emptySig + 𝔗 = < Tree emptySig A , empty-α (Tree emptySig A) > + + lemma : structHom 𝔗 𝔜 ≃ (A -> 𝔜 .car) + lemma = + structHom 𝔗 𝔜 ≃⟨ emptyHomDegen 𝔜 ⟩ + (𝔗 .car -> 𝔜 .car) ≃⟨ equiv→ treeEmpty≃ (idEquiv (𝔜 .car)) ⟩ + (A -> 𝔜 .car) ■ + +anyDef : ∀ {ℓ ℓ'} -> EmptyDef.Free ℓ ℓ' 2 +F.Definition.Free.F anyDef A = A +F.Definition.Free.η anyDef a = a +F.Definition.Free.α anyDef = empty-α _ +F.Definition.Free.sat anyDef = empty-sat _ +F.Definition.Free.isFree anyDef {𝔜 = 𝔜} _ _ = emptyHomDegen 𝔜 .snd \ No newline at end of file diff --git a/Cubical/Structures/Set/Mon/Array.agda b/Cubical/Structures/Set/Mon/Array.agda new file mode 100644 index 0000000..7b32b17 --- /dev/null +++ b/Cubical/Structures/Set/Mon/Array.agda @@ -0,0 +1,493 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.Mon.Array where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.List renaming (_∷_ to _∷ₗ_) +open import Cubical.Data.Fin +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Sum as ⊎ +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity + +open Iso + +private + variable + ℓ ℓ₁ ℓ₂ : Level + A B C : Type ℓ + n m : ℕ + +Array : Type ℓ -> Type ℓ +Array A = Σ[ n ∈ ℕ ] (Fin n -> A) + +isSetArray : isSet A -> isSet (Array A) +isSetArray isSetA = isOfHLevelΣ 2 isSetℕ λ _ -> isOfHLevelΠ 2 λ _ -> isSetA + +tptLemma : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (P : A -> Type ℓ'') {a b : A} (p : a ≡ b) (f : P a -> B) (k : P b) + -> transport (\i -> P (p i) -> B) f k ≡ f (transport (\i -> P (p (~ i))) k) +tptLemma A B P {a} p f k = sym (transport-filler (λ _ -> B) (f (transport (\i -> P (p (~ i))) k))) + +Array≡ : ∀ {ℓ} {A : Type ℓ} {n m} {f : Fin n -> A} {g : Fin m -> A} + -> (n=m : n ≡ m) + -> (∀ (k : ℕ) (k f (k , subst (k <_) (sym n=m) k Path (Array A) (n , f) (m , g) +Array≡ {A = A} {n = n} {m} {f} {g} p h = ΣPathP (p , toPathP (funExt lemma)) + where + lemma : (k : Fin m) -> transport (\i -> Fin (p i) -> A) f k ≡ g k + lemma k = tptLemma ℕ A Fin p f k ∙ h (k .fst) (k .snd) + +ℕ≡→Fin̄≅ : ∀ {n m} -> n ≡ m -> Fin n ≃ Fin m +ℕ≡→Fin̄≅ {n = n} {m = m} p = univalence .fst (cong Fin p) + +⊎-inl-beta : (B : Type ℓ) (f : A -> C) (g : B -> C) -> (a : A) -> ⊎.rec f g (inl a) ≡ f a +⊎-inl-beta B f g a = refl + +⊎-inr-beta : (A : Type ℓ) (f : A -> C) (g : B -> C) -> (b : B) -> ⊎.rec f g (inr b) ≡ g b +⊎-inr-beta A f g b = refl + +⊎-eta : {f : A -> C} {g : B -> C} -> (h : A ⊎ B -> C) (h1 : f ≡ h ∘ inl) (h2 : g ≡ h ∘ inr) -> ⊎.rec f g ≡ h +⊎-eta h h1 h2 i (inl a) = h1 i a +⊎-eta h h1 h2 i (inr b) = h2 i b + +∸-<-lemma⁻ : (m n o : ℕ) -> m ≤ o -> o ∸ m < n -> o < m + n +∸-<-lemma⁻ zero n o m≤o o∸m k < m + n -> (k < m) ⊎ (m ≤ k) -> Fin m ⊎ Fin n +finSplitAux m n k k Fin (m + n) -> Fin m ⊎ Fin n +finSplit m n (k , k finSplit m n (k , o inl (k , p)) (isProp≤ k finSplit m n (k , k finSplit m n (k , ϕ)) (isProp≤ k finSplit m n (k , ∸-<-lemma⁻ m n k m≤k k∸m inr (k ∸ m , p)) (isProp≤ (∸-<-lemma m n k (∸-<-lemma⁻ m n k m≤k! k∸m n + m ∸ n ≡ m +n+m∸n=m n m = congS (_∸ n) (+-comm n m) ∙ m+n∸n=m n m + +finSplit-beta-inr : ∀ {m n} (k : ℕ) (k finSplit m n (k , k finSplit m n (k , ϕ)) (isProp≤ k finSplit m n (m + k , <-k+ {k = m} k isProp≤) (n+m∸n=m m k)) ⟩ + inr (k , k Fin (m + n) +finCombine-inl {m = m} {n = n} (k , k Fin (m + n) +finCombine-inr {n = n} {m = m} (k , k Fin m ⊎ Fin n -> Fin (m + n) +finCombine m n = ⊎.rec finCombine-inl finCombine-inr + +finSplit∘finCombine : ∀ m n x -> (finSplit m n ∘ finCombine m n) x ≡ x +finSplit∘finCombine m n = + ⊎.elim (\(k , k + finSplit m n (finCombine m n (inl (k , k + finSplit m n (finCombine m n (inr (k , k (finCombine m n ∘ finSplit m n) x ≡ x +finCombine∘finSplit m n (k , k + finCombine m n (finSplit m n (k , k isProp≤) refl ⟩ + (k , k + finCombine m n (finSplit m n (k , k isProp≤) (≤-∸-k m≤k ∙ n+m∸n=m m k) ⟩ + (k , k Iso (Fin (m + n)) (Fin m ⊎ Fin n) +Fin≅Fin+Fin m n = iso (finSplit m n) (finCombine m n) (finSplit∘finCombine m n) (finCombine∘finSplit m n) + +combine : ∀ n m -> (Fin n -> A) -> (Fin m -> A) -> (Fin (n + m) -> A) +combine n m as bs w = ⊎.rec as bs (finSplit n m w) + +_⊕_ : Array A -> Array A -> Array A +(n , as) ⊕ (m , bs) = n + m , combine n m as bs + +e-fun : Fin 0 -> A +e-fun = ⊥.rec ∘ ¬Fin0 + +e : Array A +e = 0 , e-fun + +e-eta : ∀ (xs ys : Array A) -> xs .fst ≡ 0 -> ys .fst ≡ 0 -> xs ≡ ys +e-eta (n , xs) (m , ys) p q = ΣPathP (p ∙ sym q , toPathP (funExt lemma)) + where + lemma : _ + lemma x = ⊥.rec (¬Fin0 (subst Fin q x)) + +η : A -> Array A +η x = 1 , λ _ -> x + +zero-+ : ∀ m → 0 + m ≡ m +zero-+ m = refl + +⊕-unitl : ∀ {ℓ} {A : Type ℓ} -> (xs : Array A) -> e ⊕ xs ≡ xs +⊕-unitl (n , f) = Array≡ (zero-+ n) \k k + ⊎.rec e-fun f (finSplit 0 n (k , subst (k <_) (sym (zero-+ n)) k (xs : Array A) -> xs ⊕ e ≡ xs +⊕-unitr (n , f) = Array≡ (+-zero n) \k k + ⊎.rec f e-fun (finSplit n 0 (k , subst (k <_) (sym (+-zero n)) k k < n + (m + o) -> n + m ≤ k -> k ∸ (n + m) < o +assocr-then-∸ k n m o p q = ∸-<-lemma (n + m) o k (subst (k <_) (+-assoc n m o) p) q + +⊕-assocr-left-beta : ∀ {ℓ} {A : Type ℓ} {m n o : ℕ} (as : Fin m -> A) (bs : Fin n -> A) (cs : Fin o -> A) + -> (k : ℕ) (k ⊎.rec as (⊎.rec bs cs ∘ finSplit n o) (finSplit m (n + o) (k , k + ⊎.rec as (⊎.rec bs cs ∘ finSplit n o) (finSplit m (n + o) (k , k + ⊎.rec as (⊎.rec bs cs ∘ finSplit n o) (finSplit m (n + o) (k , k m + n ≤ k -> m ≤ k +m+n≤k→m≤k n m k (o , p) = (o + n) , sym (+-assoc o n m) ∙ congS (o +_) (+-comm n m) ∙ p + +n+m≤k→m≤k∸n : ∀ n m k -> n + m ≤ k -> m ≤ k ∸ n +n+m≤k→m≤k∸n n m k p = subst (_≤ k ∸ n) (∸+ m n) (≤-∸-≤ (n + m) k n p) + +⊕-assocr : ∀ {ℓ} {A : Type ℓ} (m n o : Array A) -> (m ⊕ n) ⊕ o ≡ m ⊕ (n ⊕ o) +⊕-assocr (n , as) (m , bs) (o , cs) = Array≡ (sym (+-assoc n m o)) \k k + ⊎.rec (\k + ⊎.rec (⊎.rec as bs ∘ finSplit n m) cs (finSplit (n + m) o (k , subst (k <_) (+-assoc n m o) k + ⊎.rec (⊎.rec as bs ∘ finSplit n m) cs (finSplit (n + m) o (k , subst (k <_) (+-assoc n m o) k isProp≤) (sym (∸-+-assoc k n m))) ⟩ + cs (k ∸ n ∸ m , subst (_< o) (sym (∸-+-assoc k n m)) (assocr-then-∸ k n m o k A) -> η (xs fzero) ⊕ (n , xs ∘ fsuc) ≡ (suc n , xs) +η+fsuc {n = n} xs = Array≡ refl λ k k ⊎.rec + (λ k<1 -> + ⊎.rec (λ _ -> xs fzero) _ (finSplit 1 n (k , _)) + ≡⟨ congS (⊎.rec _ _) (finSplit-beta-inl k k<1 _) ⟩ + xs fzero + ≡⟨ congS xs (Σ≡Prop (λ _ -> isProp≤) (sym (≤0→≡0 (pred-≤-pred k<1)))) ⟩ + xs (k , k + ⊎.rec _ (xs ∘ fsuc) (finSplit 1 n (k , _)) + ≡⟨ congS (⊎.rec _ _) (finSplit-beta-inr k _ 1≤k (<-∸-< k (suc n) 1 k n < m -> m < suc n -> ⊥.⊥ +¬n y + suc z) (sym p) ⟩ + y + (suc x + suc n) ≡⟨ +-assoc y (suc x) (suc n) ⟩ + (y + suc x) + suc n ∎ + lemma-β : 0 ≡ suc (y + x) + lemma-β = (sym (n∸n (suc n))) ∙ cong (_∸ suc n) lemma-α ∙ +∸ (y + suc x) (suc n) ∙ +-suc y x + +⊕-split : ∀ n m (xs : Fin (suc n) -> A) (ys : Fin m -> A) -> + (n + m , (λ w -> combine (suc n) m xs ys (fsuc w))) + ≡ ((n , (λ w -> xs (fsuc w))) ⊕ (m , ys)) +⊕-split n m xs ys = Array≡ refl λ k k ⊎.rec + (λ sk sym $ + ⊎.rec (xs ∘ fsuc) ys (finSplit n m (k , k + let + k∸n Array A +array-α (M.`e , i) = e +array-α (M.`⊕ , i) = i fzero ⊕ i fone + +module Free {x y : Level} {A : Type x} {𝔜 : struct y M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-monoid : 𝔜 ⊨ M.MonSEq) where + module 𝔜 = M.MonSEq 𝔜 𝔜-monoid + + 𝔄 : M.MonStruct + 𝔄 = < Array A , array-α > + + module _ (f : A -> 𝔜 .car) where + ♯^ : (n : ℕ) -> (Fin n -> A) -> 𝔜 .car + ♯^ zero _ = 𝔜.e + ♯^ (suc n) xs = f (xs fzero) 𝔜.⊕ ♯^ n (xs ∘ fsuc) + + _♯ : Array A -> 𝔜 .car + _♯ = uncurry ♯^ + + ♯-η∘ : ∀ n (xs : Fin (suc n) -> A) + -> (η (xs fzero) ♯) 𝔜.⊕ ((n , xs ∘ fsuc) ♯) + ≡ ((η (xs fzero) ⊕ (n , xs ∘ fsuc)) ♯) + ♯-η∘ n xs = + (η (xs fzero) ♯) 𝔜.⊕ ((n , xs ∘ fsuc) ♯) ≡⟨ cong (𝔜._⊕ ((n , xs ∘ fsuc) ♯)) (𝔜.unitr _) ⟩ + f (xs fzero) 𝔜.⊕ ((n , xs ∘ fsuc) ♯) ≡⟨⟩ + (suc n , xs) ♯ ≡⟨ cong _♯ (sym (η+fsuc xs)) ⟩ + ((η (xs fzero) ⊕ (n , xs ∘ fsuc)) ♯) ∎ + + ♯-++^ : ∀ n xs m ys -> ((n , xs) ⊕ (m , ys)) ♯ ≡ ((n , xs) ♯) 𝔜.⊕ ((m , ys) ♯) + ♯-++^ zero xs m ys = + ((zero , xs) ⊕ (m , ys)) ♯ ≡⟨ cong (λ z -> (z ⊕ (m , ys)) ♯) (e-eta (zero , xs) e refl refl) ⟩ + (e ⊕ (m , ys)) ♯ ≡⟨ cong _♯ (⊕-unitl (m , ys)) ⟩ + (m , ys) ♯ ≡⟨ sym (𝔜.unitl _) ⟩ + 𝔜.e 𝔜.⊕ ((m , ys) ♯) ∎ + ♯-++^ (suc n) xs m ys = + f (xs fzero) 𝔜.⊕ ((n + m , _) ♯) + ≡⟨ cong (λ z -> f (xs fzero) 𝔜.⊕ (z ♯)) (⊕-split n m xs ys) ⟩ + f (xs fzero) 𝔜.⊕ (((n , xs ∘ fsuc) ⊕ (m , ys)) ♯) + ≡⟨ cong (f (xs fzero) 𝔜.⊕_) (♯-++^ n _ m _) ⟩ + f (xs fzero) 𝔜.⊕ ((n , xs ∘ fsuc) ♯) 𝔜.⊕ ((m , ys) ♯) + ≡⟨ sym (𝔜.assocr _ _ _) ⟩ + (f (xs fzero) 𝔜.⊕ ((n , xs ∘ fsuc) ♯)) 𝔜.⊕ ((m , ys) ♯) + ≡⟨ cong (λ z -> (z 𝔜.⊕ ((n , xs ∘ fsuc) ♯)) 𝔜.⊕ ((m , ys) ♯) ) (sym (𝔜.unitr _)) ⟩ + ((η (xs fzero) ♯) 𝔜.⊕ ((n , xs ∘ fsuc) ♯)) 𝔜.⊕ ((m , ys) ♯) + ≡⟨ cong (𝔜._⊕ ((m , ys) ♯)) (♯-η∘ n xs) ⟩ + ((η (xs fzero) ⊕ (n , xs ∘ fsuc)) ♯) 𝔜.⊕ ((m , ys) ♯) + ≡⟨ cong (λ z -> (z ♯) 𝔜.⊕ ((m , ys) ♯)) (η+fsuc xs) ⟩ + ((suc n , xs) ♯) 𝔜.⊕ ((m , ys) ♯) ∎ + + ♯-++ : ∀ xs ys -> (xs ⊕ ys) ♯ ≡ (xs ♯) 𝔜.⊕ (ys ♯) + ♯-++ (n , xs) (m , ys) = ♯-++^ n xs m ys + + ♯-isMonHom : structHom 𝔄 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ ∙ sym (♯-++ (i fzero) (i fone)) + + private + arrayEquivLemma : (g : structHom 𝔄 𝔜) (n : ℕ) (xs : Fin n -> A) -> g .fst (n , xs) ≡ ((g .fst ∘ η) ♯) (n , xs) + arrayEquivLemma (g , homMonWit) zero xs = + g (0 , xs) ≡⟨ cong g (e-eta _ _ refl refl) ⟩ + g e ≡⟨ sym (homMonWit M.`e (lookup [])) ∙ 𝔜.e-eta ⟩ + 𝔜.e ≡⟨⟩ + ((g ∘ η) ♯) (zero , xs) ∎ + arrayEquivLemma (g , homMonWit) (suc n) xs = + g (suc n , xs) ≡⟨ cong g (sym (η+fsuc xs)) ⟩ + g (η (xs fzero) ⊕ (n , xs ∘ fsuc)) ≡⟨ sym (homMonWit M.`⊕ (lookup (η (xs fzero) ∷ₗ (n , xs ∘ fsuc) ∷ₗ []))) ⟩ + _ ≡⟨ 𝔜.⊕-eta (lookup ((η (xs fzero)) ∷ₗ (n , xs ∘ fsuc) ∷ₗ [])) g ⟩ + g (η (xs fzero)) 𝔜.⊕ g (n , xs ∘ fsuc) ≡⟨ cong (g (η (xs fzero)) 𝔜.⊕_) (arrayEquivLemma (g , homMonWit) n (xs ∘ fsuc)) ⟩ + g (η (xs fzero)) 𝔜.⊕ ((g ∘ η) ♯) (n , xs ∘ fsuc) ∎ + + arrayEquivLemma-β : (g : structHom 𝔄 𝔜) -> g ≡ ♯-isMonHom (g .fst ∘ η) + arrayEquivLemma-β g = structHom≡ 𝔄 𝔜 g (♯-isMonHom (g .fst ∘ η)) isSet𝔜 (funExt λ (n , p) -> arrayEquivLemma g n p) + + arrayEquiv : structHom 𝔄 𝔜 ≃ (A -> 𝔜 .car) + arrayEquiv = + isoToEquiv (iso (λ g -> g .fst ∘ η) ♯-isMonHom (λ g -> funExt (𝔜.unitr ∘ g)) (sym ∘ arrayEquivLemma-β)) + +module ArrayDef = F.Definition M.MonSig M.MonEqSig M.MonSEq + +array-str : ∀ {n} (A : Type n) -> struct n M.MonSig +array-str A = < Array A , array-α > + +array-sat : ∀ {n} {X : Type n} -> array-str X ⊨ M.MonSEq +array-sat M.`unitl ρ = ⊕-unitl (ρ fzero) +array-sat M.`unitr ρ = ⊕-unitr (ρ fzero) +array-sat M.`assocr ρ = ⊕-assocr (ρ fzero) (ρ fone) (ρ ftwo) + +arrayDef : ∀ {ℓ ℓ'} -> ArrayDef.Free ℓ ℓ' 2 +F.Definition.Free.F arrayDef = Array +F.Definition.Free.η arrayDef = η +F.Definition.Free.α arrayDef = array-α +F.Definition.Free.sat arrayDef = array-sat +F.Definition.Free.isFree arrayDef isSet𝔜 satMon = (Free.arrayEquiv isSet𝔜 satMon) .snd + +direct proof of isomorphism between Array and List +without using the universal property of Array as a free monoid +arrayIsoToList : ∀ {ℓ} {A : Type ℓ} -> Iso (Array A) (List A) +arrayIsoToList {A = A} = iso (uncurry tabulate) from tabulate-lookup from∘to + where + from : List A -> Array A + from xs = length xs , lookup xs + + from∘to : ∀ xs -> from (uncurry tabulate xs) ≡ xs + from∘to (n , xs) = ΣPathP (length-tabulate n xs , lookup-tabulate n xs) + +array≡List : ∀ {ℓ} -> Array {ℓ = ℓ} ≡ List +array≡List = funExt λ _ -> isoToPath arrayIsoToList + +import Cubical.Structures.Set.Mon.List as LM + +arrayDef' : ∀ {ℓ ℓ'} -> ArrayDef.Free ℓ ℓ' 2 +arrayDef' {ℓ = ℓ} {ℓ' = ℓ'} = fun ArrayDef.isoAux (Array , arrayFreeAux) + where + listFreeAux : ArrayDef.FreeAux ℓ ℓ' 2 List + listFreeAux = (inv ArrayDef.isoAux (LM.listDef {ℓ = ℓ} {ℓ' = ℓ'})) .snd + + arrayFreeAux : ArrayDef.FreeAux ℓ ℓ' 2 Array + arrayFreeAux = subst (ArrayDef.FreeAux ℓ ℓ' 2) (sym array≡List) listFreeAux + +private + arrayIsoToList+η : ∀ {ℓ} {A : Type ℓ} -> (x : A) (ys : Array A) + -> arrayIsoToList .fun (η x ⊕ ys) ≡ arrayIsoToList .fun (η x) ++ arrayIsoToList .fun ys + arrayIsoToList+η x ys = + congS (λ z -> x ∷ₗ (uncurry tabulate) z) $ Array≡ refl $ λ k k + congS (⊎.rec _ _) (finSplit-beta-inr (suc k) (suc-≤-suc _) (suc-≤-suc zero-≤) k (f : Fin n -> A) (ys : Array A) + -> arrayIsoToList .fun (n , f) ++ arrayIsoToList .fun ys ≡ arrayIsoToList .fun ((n , f) ⊕ ys) + arrayIsoToList++ zero f ys = congS (arrayIsoToList .fun) $ sym $ + (zero , f) ⊕ ys ≡⟨ congS (_⊕ ys) (e-eta (zero , f) e refl refl) ⟩ + e ⊕ ys ≡⟨ ⊕-unitl ys ⟩ + ys ∎ + arrayIsoToList++ (suc n) f ys = + arrayIsoToList .fun (suc n , f) ++ arrayIsoToList .fun ys + ≡⟨ congS (λ z -> arrayIsoToList .fun z ++ arrayIsoToList .fun ys) $ sym (η+fsuc f) ⟩ + arrayIsoToList .fun (η (f fzero) ⊕ (n , f ∘ fsuc)) ++ arrayIsoToList .fun ys + ≡⟨ congS (_++ arrayIsoToList .fun ys) (arrayIsoToList+η (f fzero) (n , f ∘ fsuc)) ⟩ + (arrayIsoToList .fun (η (f fzero)) ++ arrayIsoToList .fun (n , f ∘ fsuc)) ++ arrayIsoToList .fun ys + ≡⟨ ++-assoc (arrayIsoToList .fun (η (f fzero))) (arrayIsoToList .fun (n , f ∘ fsuc)) _ ⟩ + arrayIsoToList .fun (η (f fzero)) ++ (arrayIsoToList .fun (n , f ∘ fsuc) ++ arrayIsoToList .fun ys) + ≡⟨ congS (arrayIsoToList .fun (η (f fzero)) ++_) (arrayIsoToList++ n (f ∘ fsuc) ys) ⟩ + arrayIsoToList .fun (η (f fzero)) ++ arrayIsoToList .fun ((n , f ∘ fsuc) ⊕ ys) + ≡⟨ sym (arrayIsoToList+η (f fzero) ((n , f ∘ fsuc) ⊕ ys)) ⟩ + arrayIsoToList .fun (η (f fzero) ⊕ ((n , f ∘ fsuc) ⊕ ys)) + ≡⟨ congS (arrayIsoToList .fun) (sym (⊕-assocr (η (f fzero)) (n , f ∘ fsuc) ys)) ⟩ + arrayIsoToList .fun ((η (f fzero) ⊕ (n , f ∘ fsuc)) ⊕ ys) + ≡⟨ congS (λ zs -> arrayIsoToList .fun (zs ⊕ ys)) (η+fsuc f) ⟩ + arrayIsoToList .fun ((suc n , f) ⊕ ys) ∎ + +module _ {ℓ} {A : Type ℓ} where + open ArrayDef.Free + private + module 𝔄 = M.MonSEq < Array A , array-α > array-sat + + arrayIsoToListHom : structIsHom < Array A , array-α > < List A , LM.list-α > (arrayIsoToList .fun) + arrayIsoToListHom M.`e i = refl + arrayIsoToListHom M.`⊕ i = + arrayIsoToList .fun (i fzero) ++ arrayIsoToList .fun (i fone) + ≡⟨ arrayIsoToList++ (fst (i fzero)) (snd (i fzero)) (i fone) ⟩ + arrayIsoToList .fun (i fzero ⊕ i fone) + ≡⟨ congS (arrayIsoToList .fun) (sym (𝔄.⊕-eta i (idfun _))) ⟩ + arrayIsoToList .fun (i fzero ⊕ i fone) ∎ \ No newline at end of file diff --git a/Cubical/Structures/Set/Mon/Desc.agda b/Cubical/Structures/Set/Mon/Desc.agda new file mode 100644 index 0000000..1da1e32 --- /dev/null +++ b/Cubical/Structures/Set/Mon/Desc.agda @@ -0,0 +1,187 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.Mon.Desc where + +open import Cubical.Foundations.Everything +open import Cubical.Functions.Logic as L +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.List +open import Cubical.Data.Sigma +open import Cubical.Data.Sum +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree as T +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity as F + +data MonSym : Type where + `e : MonSym + `⊕ : MonSym + +MonAr : MonSym -> ℕ +MonAr `e = 0 +MonAr `⊕ = 2 + +MonFinSig : FinSig ℓ-zero +MonFinSig = MonSym , MonAr + +MonSig : Sig ℓ-zero ℓ-zero +MonSig = finSig MonFinSig + +MonStruct : ∀ {n} -> Type (ℓ-suc n) +MonStruct {n} = struct n MonSig + +module MonStruct {ℓ} (𝔛 : MonStruct {ℓ}) where + e : 𝔛 .car + e = 𝔛 .alg (`e , lookup []) + + e-eta : {i j : Arity 0 -> 𝔛 .car} -> 𝔛 .alg (`e , i) ≡ 𝔛 .alg (`e , j) + e-eta {i} = cong (\j -> 𝔛 .alg (`e , j)) (funExt λ z -> lookup [] z) + + infixr 40 _⊕_ + _⊕_ : 𝔛 .car -> 𝔛 .car -> 𝔛 .car + _⊕_ x y = 𝔛 .alg (`⊕ , lookup (x ∷ y ∷ [])) + + ⊕-eta : ∀ {ℓ} {A : Type ℓ} (i : Arity 2 -> A) (_♯ : A -> 𝔛 .car) -> 𝔛 .alg (`⊕ , (λ w -> i w ♯)) ≡ (i fzero ♯) ⊕ (i fone ♯) + ⊕-eta i _♯ = cong (λ z -> 𝔛 .alg (`⊕ , z)) (funExt lemma) + where + lemma : (x : Arity 2) -> (i x ♯) ≡ lookup ((i fzero ♯) ∷ (i fone ♯) ∷ []) x + lemma (zero , p) = cong (_♯ ∘ i) (Σ≡Prop (λ _ -> isProp≤) refl) + lemma (suc zero , p) = cong (_♯ ∘ i) (Σ≡Prop (λ _ -> isProp≤) refl) + lemma (suc (suc n) , p) = ⊥.rec (¬m+n ℕ +MonEqFree `unitl = 1 +MonEqFree `unitr = 1 +MonEqFree `assocr = 3 + +MonEqSig : EqSig ℓ-zero ℓ-zero +MonEqSig = finEqSig (MonEq , MonEqFree) + +monEqLhs : (eq : MonEq) -> FinTree MonFinSig (MonEqFree eq) +monEqLhs `unitl = node (`⊕ , lookup (node (`e , lookup []) ∷ leaf fzero ∷ [])) +monEqLhs `unitr = node (`⊕ , lookup (leaf fzero ∷ node (`e , lookup []) ∷ [])) +monEqLhs `assocr = node (`⊕ , lookup (node (`⊕ , lookup (leaf fzero ∷ leaf fone ∷ [])) ∷ leaf ftwo ∷ [])) + +monEqRhs : (eq : MonEq) -> FinTree MonFinSig (MonEqFree eq) +monEqRhs `unitl = leaf fzero +monEqRhs `unitr = leaf fzero +monEqRhs `assocr = node (`⊕ , lookup (leaf fzero ∷ node (`⊕ , lookup (leaf fone ∷ leaf ftwo ∷ [])) ∷ [])) + +MonSEq : seq MonSig MonEqSig +MonSEq n = monEqLhs n , monEqRhs n + +module MonSEq {ℓ} (𝔛 : MonStruct {ℓ}) (ϕ : 𝔛 ⊨ MonSEq) where + open MonStruct 𝔛 public + + unitl : ∀ m -> e ⊕ m ≡ m + unitl m = + e ⊕ m + ≡⟨⟩ + 𝔛 .alg (`⊕ , lookup (𝔛 .alg (`e , _) ∷ m ∷ [])) + ≡⟨ cong (\w -> 𝔛 .alg (`⊕ , w)) (funExt lemma) ⟩ + 𝔛 .alg (`⊕ , (λ x -> sharp (finSig (MonSym , MonAr)) 𝔛 (lookup (m ∷ [])) (lookup (node (`e , _) ∷ leaf fzero ∷ []) x))) + ≡⟨ ϕ `unitl (lookup [ m ]) ⟩ + m ∎ + where + lemma : (w : MonSig .arity `⊕) -> lookup (𝔛 .alg (`e , (λ num → ⊥.rec (¬Fin0 num))) ∷ m ∷ []) w ≡ sharp (finSig (MonSym , MonAr)) 𝔛 (lookup (m ∷ [])) (lookup (node (`e , (λ num → ⊥.rec (¬Fin0 num))) ∷ leaf fzero ∷ []) w) + lemma (zero , p) = sym e-eta + lemma (suc zero , p) = refl + lemma (suc (suc n) , p) = ⊥.rec (¬m+n m ⊕ e ≡ m + unitr m = + m ⊕ e + ≡⟨⟩ + 𝔛 .alg (`⊕ , lookup (m ∷ 𝔛 .alg (`e , _) ∷ [])) + ≡⟨ cong (\w -> 𝔛 .alg (`⊕ , w)) (funExt lemma) ⟩ + 𝔛 .alg (`⊕ , (λ x -> sharp MonSig 𝔛 (lookup [ m ]) (lookup (leaf fzero ∷ node (`e , (λ num → ⊥.rec (¬Fin0 num))) ∷ []) x))) + ≡⟨ ϕ `unitr (lookup [ m ]) ⟩ + m ∎ + where + lemma : (x : MonSig .arity `⊕) -> lookup (m ∷ 𝔛 .alg (`e , (λ num → ⊥.rec (¬Fin0 num))) ∷ []) x ≡ sharp MonSig 𝔛 (lookup [ m ]) (lookup (leaf fzero ∷ node (`e , (λ num → ⊥.rec (¬Fin0 num))) ∷ []) x) + lemma (zero , p) = refl + lemma (suc zero , p) = sym e-eta + lemma (suc (suc n) , p) = ⊥.rec (¬m+n (m ⊕ n) ⊕ o ≡ m ⊕ (n ⊕ o) + assocr m n o = + (m ⊕ n) ⊕ o + ≡⟨⟩ + 𝔛 .alg (`⊕ , lookup (𝔛 .alg (`⊕ , lookup (m ∷ n ∷ [])) ∷ o ∷ [])) + ≡⟨ cong (\w -> 𝔛 .alg (`⊕ , w)) (funExt lemma1) ⟩ + 𝔛 .alg (`⊕ , (\w -> sharp MonSig 𝔛 (lookup (m ∷ n ∷ o ∷ [])) (lookup (node (`⊕ , lookup (leaf fzero ∷ leaf fone ∷ [])) ∷ leaf ftwo ∷ []) w))) + ≡⟨ ϕ `assocr (lookup (m ∷ n ∷ o ∷ [])) ⟩ + 𝔛 .alg (`⊕ , (λ w -> sharp MonSig 𝔛 (lookup (m ∷ n ∷ o ∷ [])) (lookup (leaf fzero ∷ node (`⊕ , lookup (leaf fone ∷ leaf ftwo ∷ [])) ∷ []) w))) + ≡⟨ cong (\w -> 𝔛 .alg (`⊕ , w)) (sym (funExt lemma3)) ⟩ + 𝔛 .alg (`⊕ , lookup (m ∷ 𝔛 .alg (`⊕ , lookup (n ∷ o ∷ [])) ∷ [])) + ≡⟨⟩ + m ⊕ (n ⊕ o) ∎ + where + lemma1 : (w : MonSig .arity `⊕) -> lookup (𝔛 .alg (`⊕ , lookup (m ∷ n ∷ [])) ∷ o ∷ []) w ≡ sharp MonSig 𝔛 (lookup (m ∷ n ∷ o ∷ [])) (lookup (node (`⊕ , lookup (leaf fzero ∷ leaf fone ∷ [])) ∷ leaf ftwo ∷ []) w) + lemma2 : (w : MonSig .arity `⊕) -> lookup (m ∷ n ∷ []) w ≡ sharp MonSig 𝔛 (lookup (m ∷ n ∷ o ∷ [])) (lookup (leaf fzero ∷ leaf fone ∷ []) w) + + lemma1 (zero , p) = cong (λ o → 𝔛 .alg (`⊕ , o)) (funExt lemma2) + lemma1 (suc zero , p) = refl + lemma1 (suc (suc n) , p) = ⊥.rec (¬m+n lookup (m ∷ 𝔛 .alg (`⊕ , lookup (n ∷ o ∷ [])) ∷ []) w ≡ sharp MonSig 𝔛 (lookup (m ∷ n ∷ o ∷ [])) (lookup (leaf fzero ∷ node (`⊕ , lookup (leaf fone ∷ leaf ftwo ∷ [])) ∷ []) w) + lemma4 : (w : MonSig .arity `⊕) -> lookup (n ∷ o ∷ []) w ≡ sharp MonSig 𝔛 (lookup (m ∷ n ∷ o ∷ [])) (lookup (leaf fone ∷ leaf ftwo ∷ []) w) + + lemma3 (zero , p) = refl + lemma3 (suc zero , p) = cong (λ w → 𝔛 .alg (`⊕ , w)) (funExt lemma4) + lemma3 (suc (suc n) , p) = ⊥.rec (¬m+n lookup (x ∷ xs) (fsuc a) ≡ sharp MonSig 𝔛 {!!} (lookup {!!} a) + lemma f = {!!} + +ℕ-MonStr : MonStruct +car ℕ-MonStr = ℕ +alg ℕ-MonStr (`e , _) = 0 +alg ℕ-MonStr (`⊕ , i) = i fzero + i fone + +ℕ-MonStr-MonSEq : ℕ-MonStr ⊨ MonSEq +ℕ-MonStr-MonSEq `unitl ρ = refl +ℕ-MonStr-MonSEq `unitr ρ = +-zero (ρ fzero) +ℕ-MonStr-MonSEq `assocr ρ = sym (+-assoc (ρ fzero) (ρ fone) (ρ ftwo)) + +⊔-MonStr : (ℓ : Level) -> MonStruct {n = ℓ-suc ℓ} +car (⊔-MonStr ℓ) = hProp ℓ +alg (⊔-MonStr ℓ) (`e , i) = ⊥* , isProp⊥* +alg (⊔-MonStr ℓ) (`⊕ , i) = i fzero ⊔ i fone + +⊔-MonStr-MonSEq : (ℓ : Level) -> ⊔-MonStr ℓ ⊨ MonSEq +⊔-MonStr-MonSEq ℓ `unitl ρ = + ⇒∶ ⊔-elim (⊥* , isProp⊥*) (ρ fzero) (λ _ -> ρ fzero) ⊥.rec* (idfun _) + ⇐∶ L.inr +⊔-MonStr-MonSEq ℓ `unitr ρ = + ⇔toPath (⊔-elim (ρ fzero) (⊥* , isProp⊥*) (λ _ -> ρ fzero) (idfun _) ⊥.rec*) L.inl +⊔-MonStr-MonSEq ℓ `assocr ρ = + sym (⊔-assoc (ρ fzero) (ρ fone) (ρ ftwo)) + +⊓-MonStr : (ℓ : Level) -> MonStruct {n = ℓ-suc ℓ} +car (⊓-MonStr ℓ) = hProp ℓ +alg (⊓-MonStr ℓ) (`e , i) = L.⊤ +alg (⊓-MonStr ℓ) (`⊕ , i) = i fzero ⊓ i fone + +⊓-MonStr-MonSEq : (ℓ : Level) -> ⊓-MonStr ℓ ⊨ MonSEq +⊓-MonStr-MonSEq ℓ `unitl ρ = ⊓-identityˡ (ρ fzero) +⊓-MonStr-MonSEq ℓ `unitr ρ = ⊓-identityʳ (ρ fzero) +⊓-MonStr-MonSEq ℓ `assocr ρ = sym (⊓-assoc (ρ fzero) (ρ fone) (ρ ftwo)) \ No newline at end of file diff --git a/Cubical/Structures/Set/Mon/Free.agda b/Cubical/Structures/Set/Mon/Free.agda new file mode 100644 index 0000000..08924b4 --- /dev/null +++ b/Cubical/Structures/Set/Mon/Free.agda @@ -0,0 +1,144 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.Mon.Free where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.List +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity + +data FreeMon {ℓ : Level} (A : Type ℓ) : Type ℓ where + η : (a : A) -> FreeMon A + e : FreeMon A + _⊕_ : FreeMon A -> FreeMon A -> FreeMon A + unitl : ∀ m -> e ⊕ m ≡ m + unitr : ∀ m -> m ⊕ e ≡ m + assocr : ∀ m n o -> (m ⊕ n) ⊕ o ≡ m ⊕ (n ⊕ o) + trunc : isSet (FreeMon A) + +module elimFreeMonSet {p n : Level} {A : Type n} (P : FreeMon A -> Type p) + (η* : (a : A) -> P (η a)) + (e* : P e) + (_⊕*_ : {m n : FreeMon A} -> P m -> P n -> P (m ⊕ n)) + (unitl* : {m : FreeMon A} (m* : P m) -> PathP (λ i → P (unitl m i)) (e* ⊕* m*) m*) + (unitr* : {m : FreeMon A} (m* : P m) -> PathP (λ i → P (unitr m i)) (m* ⊕* e*) m*) + (assocr* : {m n o : FreeMon A} + (m* : P m) -> + (n* : P n) -> + (o* : P o) -> PathP (λ i → P (assocr m n o i)) ((m* ⊕* n*) ⊕* o*) (m* ⊕* (n* ⊕* o*))) + (trunc* : {xs : FreeMon A} -> isSet (P xs)) + where + f : (x : FreeMon A) -> P x + f (η a) = η* a + f e = e* + f (x ⊕ y) = f x ⊕* f y + f (unitl x i) = unitl* (f x) i + f (unitr x i) = unitr* (f x) i + f (assocr x y z i) = assocr* (f x) (f y) (f z) i + f (trunc xs ys p q i j) = + isOfHLevel→isOfHLevelDep 2 (\xs -> trunc* {xs = xs}) (f xs) (f ys) (cong f p) (cong f q) (trunc xs ys p q) i j + +module recFreeMonSet {p n : Level} {A : Type n} (P : Type p) + (η* : (a : A) -> P) + (e* : P) + (_⊕*_ : P -> P -> P) + (unitl* : (m* : P) -> PathP (λ i → P) (e* ⊕* m*) m*) + (unitr* : (m* : P) -> PathP (λ i → P) (m* ⊕* e*) m*) + (assocr* : (m* : P) -> + (n* : P) -> + (o* : P) -> PathP (λ i → P) ((m* ⊕* n*) ⊕* o*) (m* ⊕* (n* ⊕* o*))) + (trunc* : isSet P) + where + f : (x : FreeMon A) -> P + f = elimFreeMonSet.f (\_ -> P) η* e* _⊕*_ unitl* unitr* assocr* trunc* + +module elimFreeMonProp {p n : Level} {A : Type n} (P : FreeMon A -> Type p) + (η* : (a : A) -> P (η a)) + (e* : P e) + (_⊕*_ : {m n : FreeMon A} -> P m -> P n -> P (m ⊕ n)) + (trunc* : {xs : FreeMon A} -> isProp (P xs)) + where + f : (x : FreeMon A) -> P x + f = elimFreeMonSet.f P η* e* _⊕*_ unitl* unitr* assocr* (isProp→isSet trunc*) + where + abstract + unitl* : {m : FreeMon A} (m* : P m) -> PathP (λ i → P (unitl m i)) (e* ⊕* m*) m* + unitl* {m} m* = toPathP (trunc* (transp (λ i -> P (unitl m i)) i0 (e* ⊕* m*)) m*) + unitr* : {m : FreeMon A} (m* : P m) -> PathP (λ i → P (unitr m i)) (m* ⊕* e*) m* + unitr* {m} m* = toPathP (trunc* (transp (λ i -> P (unitr m i)) i0 (m* ⊕* e*)) m*) + assocr* : {m n o : FreeMon A} + (m* : P m) -> + (n* : P n) -> + (o* : P o) -> PathP (λ i → P (assocr m n o i)) ((m* ⊕* n*) ⊕* o*) (m* ⊕* (n* ⊕* o*)) + assocr* {m} {n} {o} m* n* o* = + toPathP (trunc* (transp (λ i -> P (assocr m n o i)) i0 ((m* ⊕* n*) ⊕* o*)) (m* ⊕* (n* ⊕* o*))) + +freeMon-α : ∀ {n : Level} {X : Type n} -> sig M.MonSig (FreeMon X) -> FreeMon X +freeMon-α (M.`e , i) = e +freeMon-α (M.`⊕ , i) = i fzero ⊕ i fone + +module Free {x y : Level} {A : Type x} {𝔜 : struct y M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-monoid : 𝔜 ⊨ M.MonSEq) where + module 𝔜 = M.MonSEq 𝔜 𝔜-monoid + + 𝔉 : struct x M.MonSig + 𝔉 = < FreeMon A , freeMon-α > + + module _ (f : A -> 𝔜 .car) where + _♯ : FreeMon A -> 𝔜 .car + (η a) ♯ = f a + e ♯ = 𝔜.e + (m ⊕ n) ♯ = (m ♯) 𝔜.⊕ (n ♯) + (unitl m i) ♯ = 𝔜.unitl (m ♯) i + (unitr m i) ♯ = 𝔜.unitr (m ♯) i + (assocr m n o i) ♯ = 𝔜.assocr (m ♯) (n ♯) (o ♯) i + (trunc m n p q i j) ♯ = isSet𝔜 (m ♯) (n ♯) (cong _♯ p) (cong _♯ q) i j + + ♯-isMonHom : structHom 𝔉 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ + + private + freeMonEquivLemma : (g : structHom 𝔉 𝔜) -> (x : FreeMon A) -> g .fst x ≡ ((g .fst ∘ η) ♯) x + freeMonEquivLemma (g , homMonWit) = elimFreeMonProp.f (λ x -> g x ≡ ((g ∘ η) ♯) x) + (λ _ -> refl) + (sym (homMonWit M.`e (lookup [])) ∙ 𝔜.e-eta) + (λ {m} {n} p q -> + g (m ⊕ n) ≡⟨ sym (homMonWit M.`⊕ (lookup (m ∷ n ∷ []))) ⟩ + 𝔜 .alg (M.`⊕ , (λ w -> g (lookup (m ∷ n ∷ []) w))) ≡⟨ 𝔜.⊕-eta (lookup (m ∷ n ∷ [])) g ⟩ + g m 𝔜.⊕ g n ≡⟨ cong₂ 𝔜._⊕_ p q ⟩ + _ ∎ + ) + (isSet𝔜 _ _) + + freeMonEquivLemma-β : (g : structHom 𝔉 𝔜) -> g ≡ ♯-isMonHom (g .fst ∘ η) + freeMonEquivLemma-β g = structHom≡ 𝔉 𝔜 g (♯-isMonHom (g .fst ∘ η)) isSet𝔜 (funExt (freeMonEquivLemma g)) + + freeMonEquiv : structHom 𝔉 𝔜 ≃ (A -> 𝔜 .car) + freeMonEquiv = + isoToEquiv (iso (λ g -> g .fst ∘ η) ♯-isMonHom (λ _ -> refl) (sym ∘ freeMonEquivLemma-β)) + +module FreeMonDef = F.Definition M.MonSig M.MonEqSig M.MonSEq + +freeMon-sat : ∀ {n} {X : Type n} -> < FreeMon X , freeMon-α > ⊨ M.MonSEq +freeMon-sat M.`unitl ρ = unitl (ρ fzero) +freeMon-sat M.`unitr ρ = unitr (ρ fzero) +freeMon-sat M.`assocr ρ = assocr (ρ fzero) (ρ fone) (ρ ftwo) + +freeMonDef : ∀ {ℓ ℓ'} -> FreeMonDef.Free ℓ ℓ' 2 +F.Definition.Free.F freeMonDef = FreeMon +F.Definition.Free.η freeMonDef = η +F.Definition.Free.α freeMonDef = freeMon-α +F.Definition.Free.sat freeMonDef = freeMon-sat +F.Definition.Free.isFree freeMonDef isSet𝔜 satMon = (Free.freeMonEquiv isSet𝔜 satMon) .snd + \ No newline at end of file diff --git a/Cubical/Structures/Set/Mon/List.agda b/Cubical/Structures/Set/Mon/List.agda new file mode 100644 index 0000000..426a8f4 --- /dev/null +++ b/Cubical/Structures/Set/Mon/List.agda @@ -0,0 +1,120 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Set.Mon.List where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.List +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +import Cubical.Data.Empty as ⊥ +open import Cubical.Functions.Logic as L + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Sum as ⊎ + +private + variable + ℓ : Level + A B : Type ℓ + +list-α : sig M.MonSig (List A) -> List A +list-α (M.`e , i) = [] +list-α (M.`⊕ , i) = i fzero ++ i fone + +module Free {x y : Level} {A : Type x} {𝔜 : struct y M.MonSig} (isSet𝔜 : isSet (𝔜 .car)) (𝔜-monoid : 𝔜 ⊨ M.MonSEq) where + module 𝔜 = M.MonSEq 𝔜 𝔜-monoid + + 𝔏 : M.MonStruct + 𝔏 = < List A , list-α > + + module _ (f : A -> 𝔜 .car) where + _♯ : List A -> 𝔜 .car + [] ♯ = 𝔜.e + (x ∷ xs) ♯ = f x 𝔜.⊕ (xs ♯) + + private + ♯-++ : ∀ xs ys -> (xs ++ ys) ♯ ≡ (xs ♯) 𝔜.⊕ (ys ♯) + ♯-++ [] ys = sym (𝔜.unitl (ys ♯)) + ♯-++ (x ∷ xs) ys = cong (f x 𝔜.⊕_) (♯-++ xs ys) ∙ sym (𝔜.assocr (f x) (xs ♯) (ys ♯)) + + ♯-isMonHom : structHom 𝔏 𝔜 + fst ♯-isMonHom = _♯ + snd ♯-isMonHom M.`e i = 𝔜.e-eta + snd ♯-isMonHom M.`⊕ i = 𝔜.⊕-eta i _♯ ∙ sym (♯-++ (i fzero) (i fone)) + + private + listEquivLemma : (g : structHom 𝔏 𝔜) -> (x : List A) -> g .fst x ≡ ((g .fst ∘ [_]) ♯) x + listEquivLemma (g , homMonWit) [] = sym (homMonWit M.`e (lookup [])) ∙ 𝔜.e-eta + listEquivLemma (g , homMonWit) (x ∷ xs) = + g (x ∷ xs) ≡⟨ sym (homMonWit M.`⊕ (lookup ([ x ] ∷ xs ∷ []))) ⟩ + 𝔜 .alg (M.`⊕ , (λ w -> g (lookup ((x ∷ []) ∷ xs ∷ []) w))) ≡⟨ 𝔜.⊕-eta (lookup ([ x ] ∷ xs ∷ [])) g ⟩ + g [ x ] 𝔜.⊕ g xs ≡⟨ cong (g [ x ] 𝔜.⊕_) (listEquivLemma (g , homMonWit) xs) ⟩ + _ ∎ + + listEquivLemma-β : (g : structHom 𝔏 𝔜) -> g ≡ ♯-isMonHom (g .fst ∘ [_]) + listEquivLemma-β g = structHom≡ 𝔏 𝔜 g (♯-isMonHom (g .fst ∘ [_])) isSet𝔜 (funExt (listEquivLemma g)) + + listEquiv : structHom 𝔏 𝔜 ≃ (A -> 𝔜 .car) + listEquiv = + isoToEquiv (iso (λ g -> g .fst ∘ [_]) ♯-isMonHom (λ g -> funExt (𝔜.unitr ∘ g)) (sym ∘ listEquivLemma-β)) + +module ListDef = F.Definition M.MonSig M.MonEqSig M.MonSEq + +list-sat : ∀ {n} {X : Type n} -> < List X , list-α > ⊨ M.MonSEq +list-sat M.`unitl ρ = refl +list-sat M.`unitr ρ = ++-unit-r (ρ fzero) +list-sat M.`assocr ρ = ++-assoc (ρ fzero) (ρ fone) (ρ ftwo) + +listDef : ∀ {ℓ ℓ'} -> ListDef.Free ℓ ℓ' 2 +F.Definition.Free.F listDef = List +F.Definition.Free.η listDef = [_] +F.Definition.Free.α listDef = list-α +F.Definition.Free.sat listDef = list-sat +F.Definition.Free.isFree listDef isSet𝔜 satMon = (Free.listEquiv isSet𝔜 satMon) .snd + +list-⊥ : (List ⊥.⊥) ≃ Unit +list-⊥ = isoToEquiv (iso (λ _ -> tt) (λ _ -> []) (λ _ -> isPropUnit _ _) lemma) + where + lemma : ∀ xs -> [] ≡ xs + lemma [] = refl + lemma (x ∷ xs) = ⊥.rec x + +module Membership {ℓ} {A : Type ℓ} (isSetA : isSet A) where + open Free {A = A} isSetHProp (M.⊔-MonStr-MonSEq ℓ) + + ∈Prop : A -> List A -> hProp ℓ + ∈Prop x = (λ y -> (x ≡ y) , isSetA x y) ♯ + + _∈_ : A -> List A -> Type ℓ + x ∈ xs = ∈Prop x xs .fst + + isProp-∈ : (x : A) -> (xs : List A) -> isProp (x ∈ xs) + isProp-∈ x xs = (∈Prop x xs) .snd + + x∈xs : ∀ x xs -> x ∈ (x ∷ xs) + x∈xs x xs = L.inl refl + + x∈[x] : ∀ x -> x ∈ [ x ] + x∈[x] x = x∈xs x [] + + ∈-∷ : ∀ x y xs -> x ∈ xs -> x ∈ (y ∷ xs) + ∈-∷ x y xs p = L.inr p + + ∈-++ : ∀ x xs ys -> x ∈ ys -> x ∈ (xs ++ ys) + ∈-++ x [] ys p = p + ∈-++ x (a ∷ as) ys p = ∈-∷ x a (as ++ ys) (∈-++ x as ys p) + + ¬∈[] : ∀ x -> (x ∈ []) -> ⊥.⊥ + ¬∈[] x = ⊥.rec* + + x∈[y]→x≡y : ∀ x y -> x ∈ [ y ] -> x ≡ y + x∈[y]→x≡y x y = P.rec (isSetA x y) $ ⊎.rec (idfun _) ⊥.rec* diff --git a/Cubical/Structures/Sig.agda b/Cubical/Structures/Sig.agda new file mode 100644 index 0000000..5ad9f06 --- /dev/null +++ b/Cubical/Structures/Sig.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Sig where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.Fin +open import Cubical.Data.List as L +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +record Sig (f a : Level) : Type ((ℓ-suc f) ⊔ (ℓ-suc a)) where + field + symbol : Type f + arity : symbol -> Type a +open Sig public + +FinSig : (f : Level) -> Type (ℓ-suc f) +FinSig f = Σ (Type f) \sym -> sym -> ℕ + +finSig : {f : Level} -> FinSig f -> Sig f ℓ-zero +symbol (finSig σ) = σ .fst +arity (finSig σ) = Fin ∘ σ .snd + +emptySig : Sig ℓ-zero ℓ-zero +symbol emptySig = ⊥.⊥ +arity emptySig = ⊥.rec + +sumSig : {f a g b : Level} -> Sig f a -> Sig g b -> Sig (ℓ-max f g) (ℓ-max a b) +symbol (sumSig σ τ) = symbol σ ⊎ symbol τ +arity (sumSig {b = b} σ τ) (inl x) = Lift {j = b} ((arity σ) x) +arity (sumSig {a = a} σ τ) (inr x) = Lift {j = a} ((arity τ) x) + +signature functor +module _ {f a n : Level} (σ : Sig f a) where + sig : Type n -> Type (ℓ-max f (ℓ-max a n)) + sig X = Σ (σ .symbol) \f -> σ .arity f -> X + + sig≡ : {X : Type n} (f : σ .symbol) {i j : σ .arity f -> X} + -> ((a : σ .arity f) -> i a ≡ j a) + -> Path (sig X) (f , i) (f , j) + sig≡ f H = ΣPathP (refl , funExt H) + + sigF : {X Y : Type n} -> (X -> Y) -> sig X -> sig Y + sigF h (f , i) = f , h ∘ i + +module _ {f n : Level} (σ : FinSig f) where + sigFin : (X : Type n) -> Type (ℓ-max f n) + sigFin X = sig (finSig σ) X + + sigFin≡ : {X : Type n} (f : σ .fst) {i j : finSig σ .arity f -> X} + -> ((a : finSig σ .arity f) -> i a ≡ j a) + -> Path (sigFin X) (f , i) (f , j) + sigFin≡ = sig≡ (finSig σ) diff --git a/Cubical/Structures/Str.agda b/Cubical/Structures/Str.agda new file mode 100644 index 0000000..895bb43 --- /dev/null +++ b/Cubical/Structures/Str.agda @@ -0,0 +1,49 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Str where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.List as L +open import Cubical.Data.Sigma +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +open import Cubical.Structures.Sig + +record struct {f a : Level} (n : Level) (σ : Sig f a) : Type (ℓ-max f (ℓ-max a (ℓ-suc n))) where + constructor <_,_> + field + car : Type n + alg : sig σ car -> car +open struct public + +module _ {f a x y : Level} {σ : Sig f a} (𝔛 : struct x σ) (𝔜 : struct y σ) where + structIsHom : (h : 𝔛 .car -> 𝔜 .car) -> Type (ℓ-max f (ℓ-max a (ℓ-max x y))) + structIsHom h = + ((f : σ .symbol) -> (i : σ .arity f -> 𝔛 .car) -> 𝔜 .alg (f , h ∘ i) ≡ h (𝔛 .alg (f , i))) + + structHom : Type (ℓ-max f (ℓ-max a (ℓ-max x y))) + structHom = Σ[ h ∈ (𝔛 .car -> 𝔜 .car) ] structIsHom h + + structHom≡ : (g h : structHom) -> isSet (𝔜 .car) -> g .fst ≡ h .fst -> g ≡ h + structHom≡ (g-f , g-hom) (h-f , h-hom) isSetY = + Σ≡Prop (\fun -> isPropΠ \f -> isPropΠ \o -> isSetY (𝔜 .alg (f , fun ∘ o)) (fun (𝔛 .alg (f , o)))) + +module _ {f a x : Level} {σ : Sig f a} (𝔛 : struct x σ) where + idHom : structHom 𝔛 𝔛 + idHom = idfun _ , \f i -> refl + +module _ {f a x y z : Level} {σ : Sig f a} (𝔛 : struct x σ) (𝔜 : struct y σ) (ℨ : struct z σ) where + structHom∘ : (g : structHom 𝔜 ℨ) -> (h : structHom 𝔛 𝔜) -> structHom 𝔛 ℨ + structHom∘ (g-f , g-hom) (h-f , h-hom) = g-f ∘ h-f , lemma-α + where + lemma-α : structIsHom 𝔛 ℨ (g-f ∘ h-f) + lemma-α eqn i = + ℨ .alg (eqn , g-f ∘ h-f ∘ i) ≡⟨ g-hom eqn (h-f ∘ i) ⟩ + g-f (𝔜 .alg (eqn , h-f ∘ i)) ≡⟨ cong g-f (h-hom eqn i) ⟩ + _ ∎ diff --git a/Cubical/Structures/Tree.agda b/Cubical/Structures/Tree.agda new file mode 100644 index 0000000..a38a746 --- /dev/null +++ b/Cubical/Structures/Tree.agda @@ -0,0 +1,134 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Cubical.Structures.Tree where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.Fin +open import Cubical.Structures.Sig +open import Cubical.Structures.Str +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as S +open import Cubical.Foundations.Isomorphism renaming (Iso to _≅_) + +module _ {f a n : Level} (σ : Sig f a) where + data Tree (V : Type n) : Type (ℓ-max (ℓ-max f a) n) where + leaf : V -> Tree V + node : sig σ (Tree V) -> Tree V + open Tree + + +module _ {f a n y : Level} (σ : Sig f a) {V : Type n} where + open import Cubical.Data.W.Indexed + + S : Type n -> Type (ℓ-max f n) + S _ = V ⊎ (σ .symbol) + + P : (V : Type n) -> S V -> Type a + P V (inl v) = ⊥* + P V (inr f) = σ .arity f + + inX : ∀ V (s : S V) -> P V s -> Type n + inX V s p = V + + RepTree : Type (ℓ-max f (ℓ-max a (ℓ-suc n))) + RepTree = IW S P inX V + + Tree→RepTree : Tree σ V -> RepTree + Tree→RepTree (leaf v) = node (inl v) ⊥.rec* + Tree→RepTree (node (f , i)) = node (inr f) (Tree→RepTree ∘ i) + + RepTree→Tree : RepTree -> Tree σ V + RepTree→Tree (node (inl v) subtree) = leaf v + RepTree→Tree (node (inr f) subtree) = node (f , RepTree→Tree ∘ subtree) + + Tree→RepTree→Tree : ∀ t -> RepTree→Tree (Tree→RepTree t) ≡ t + Tree→RepTree→Tree (leaf v) = refl + Tree→RepTree→Tree (node (f , i)) j = node (f , \a -> Tree→RepTree→Tree (i a) j) + + RepTree→Tree→RepTree : ∀ r -> Tree→RepTree (RepTree→Tree r) ≡ r + RepTree→Tree→RepTree (node (inl v) subtree) = congS (node (inl v)) (funExt (⊥.rec ∘ lower)) + RepTree→Tree→RepTree (node (inr f) subtree) = congS (node (inr f)) (funExt (RepTree→Tree→RepTree ∘ subtree)) + + isoRepTree : Tree σ V ≅ RepTree + Iso.fun isoRepTree = Tree→RepTree + Iso.inv isoRepTree = RepTree→Tree + Iso.rightInv isoRepTree = RepTree→Tree→RepTree + Iso.leftInv isoRepTree = Tree→RepTree→Tree + + isOfHLevelMax : ∀ {ℓ} {n m : HLevel} {A : Type ℓ} + -> isOfHLevel n A + -> isOfHLevel (max n m) A + isOfHLevelMax {n = n} {m = m} {A = A} p = + let + (k , o) = left-≤-max {m = n} {n = m} + in + subst (λ l -> isOfHLevel l A) o (isOfHLevelPlus k p) + + isOfHLevelS : (h h' : HLevel) + (p : isOfHLevel (2 + h) V) (q : isOfHLevel (2 + h') (σ .symbol)) + -> isOfHLevel (max (2 + h) (2 + h')) (S V) + isOfHLevelS h h' p q = + isOfHLevel⊎ _ + (isOfHLevelMax {n = 2 + h} {m = 2 + h'} p) + (subst (λ h'' -> isOfHLevel h'' (σ .symbol)) (maxComm (2 + h') (2 + h)) (isOfHLevelMax {n = 2 + h'} {m = 2 + h} q)) + + isOfHLevelRepTree : ∀ {h h' : HLevel} + -> isOfHLevel (2 + h) V + -> isOfHLevel (2 + h') (σ .symbol) + -> isOfHLevel (max (2 + h) (2 + h')) RepTree + isOfHLevelRepTree {h} {h'} p q = + isOfHLevelSuc-IW (max (suc h) (suc h')) (λ _ -> isOfHLevelPath' _ (isOfHLevelS _ _ p q)) V + + isOfHLevelTree : ∀ {h h' : HLevel} + -> isOfHLevel (2 + h) V + -> isOfHLevel (2 + h') (σ .symbol) + -> isOfHLevel (max (2 + h) (2 + h')) (Tree σ V) + isOfHLevelTree {h} {h'} p q = + isOfHLevelRetract (max (2 + h) (2 + h')) + Tree→RepTree + RepTree→Tree + Tree→RepTree→Tree + (isOfHLevelRepTree p q) + +module _ {f : Level} (σ : FinSig f) where + FinTree : (k : ℕ) -> Type f + FinTree k = Tree (finSig σ) (Fin k) + +module _ {f a : Level} (σ : Sig f a) where + algTr : ∀ {x} (X : Type x) -> struct (ℓ-max f (ℓ-max a x)) σ + car (algTr X) = Tree σ X + alg (algTr X) = node + +module _ {f a : Level} (σ : Sig f a) {x y} {X : Type x} (𝔜 : struct y σ) where + private + 𝔛 : struct (ℓ-max f (ℓ-max a x)) σ + 𝔛 = algTr σ X + + sharp : (X -> 𝔜 .car) -> Tree σ X -> 𝔜 .car + sharp ρ (leaf v) = ρ v + sharp ρ (node (f , o)) = 𝔜 .alg (f , sharp ρ ∘ o) + + eval : (X -> 𝔜 .car) -> structHom 𝔛 𝔜 + eval h = sharp h , λ _ _ -> refl + + sharp-eta : (g : structHom 𝔛 𝔜) -> (tr : Tree σ X) -> g .fst tr ≡ sharp (g .fst ∘ leaf) tr + sharp-eta g (leaf x) = refl + sharp-eta (g-f , g-hom) (node x) = + g-f (node x) ≡⟨ sym (g-hom (x .fst) (x .snd)) ⟩ + 𝔜 .alg (x .fst , (λ y → g-f (x .snd y))) ≡⟨ cong (λ z → 𝔜 .alg (x .fst , z)) (funExt λ y -> sharp-eta ((g-f , g-hom)) (x .snd y)) ⟩ + 𝔜 .alg (x .fst , (λ y → sharp (g-f ∘ leaf) (x .snd y))) + ∎ + + sharp-hom-eta : isSet (𝔜 .car) -> (g : structHom 𝔛 𝔜) -> g ≡ eval (g .fst ∘ leaf) + sharp-hom-eta p g = structHom≡ 𝔛 𝔜 g (eval (g .fst ∘ leaf)) p (funExt (sharp-eta g)) + + trEquiv : isSet (𝔜 .car) -> structHom 𝔛 𝔜 ≃ (X -> 𝔜 .car) + trEquiv isSetY = isoToEquiv (iso (\g -> g .fst ∘ leaf) eval (\_ -> refl) (sym ∘ sharp-hom-eta isSetY)) + + trIsEquiv : isSet (𝔜 .car) -> isEquiv (\g -> g .fst ∘ leaf) + trIsEquiv = snd ∘ trEquiv diff --git a/Everything.agda b/Everything.agda new file mode 100644 index 0000000..2a10c8a --- /dev/null +++ b/Everything.agda @@ -0,0 +1,3 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module Everything where diff --git a/Experiments/Fin.agda b/Experiments/Fin.agda new file mode 100644 index 0000000..4a4af8e --- /dev/null +++ b/Experiments/Fin.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --cubical --allow-unsolved-metas #-} + +module Experiments.Fin where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Fin +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Sum as ⊎ +import Cubical.Data.Empty as ⊥ + +open import Cubical.Structures.Set.Mon.Array +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity +import Cubical.Structures.Set.Mon.List as LM + +open Iso + +open import Agda.Builtin.Reflection hiding (Type) +open import Agda.Builtin.String +open import Cubical.Reflection.Base +open import Cubical.Tactics.FunctorSolver.Solver +open import Cubical.Tactics.Reflection + +private + variable + ℓ : Level + A : Type ℓ + + +infixr 5 _∷_ + +data FMSet (A : Type ℓ) : Type ℓ where + [] : FMSet A + _∷_ : (x : A) → (xs : FMSet A) → FMSet A + comm : ∀ x y xs → x ∷ y ∷ xs ≡ y ∷ x ∷ xs + trunc : isSet (FMSet A) + +_++_ : ∀ (xs ys : FMSet A) → FMSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys +comm x y xs i ++ ys = comm x y (xs ++ ys) i +trunc xs zs p q i j ++ ys = {! trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) !} + trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j diff --git a/Experiments/FreeStr.agda b/Experiments/FreeStr.agda new file mode 100644 index 0000000..9dc2dde --- /dev/null +++ b/Experiments/FreeStr.agda @@ -0,0 +1,141 @@ +{-# OPTIONS --cubical #-} + +module Experiments.FreeStr where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.FinData as F +open import Cubical.Data.List as L +open import Cubical.Data.List.FinData as F +open import Cubical.Data.Sigma +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +module _ {f a e n : Level} (σ : Sig f a) (τ : EqSig e n) (ε : seq σ τ) where + srel : (X : Type n) -> Tree σ X -> Tree σ X -> Type {!!} + srel X l r = sat σ τ ε {!!} + + Fr : (X : Type n) -> walg σ + Fr X = (Tree σ X / {!!}) , {!!} + +module _ {f a e n : Level} (σ : Sig f a) (τ : EqSig e n) where + + evalEq : {X : Type n} -> (n : τ .name) -> (ρ : X -> τ .free n) -> Tree σ X -> Tree σ (τ .free n) + evalEq {X = X} n ρ = _♯ σ {X = X} {Y = TreeStr σ (τ .free n)} (leaf ∘ ρ) + + eqs : {X : Type n} -> Tree σ X -> Tree σ X -> Type (ℓ-max (ℓ-max f a) (ℓ-max e n)) + eqs {X = X} lhs rhs = (n : τ .name) (ρ : X -> τ .free n) -> evalEq n ρ lhs ≡ evalEq n ρ rhs + + Free : Type n -> Type (ℓ-max (ℓ-max f a) (ℓ-max e n)) + Free X = Tree σ X / eqs + +module _ {f a e n : Level} (σ : Sig f a) (τ : EqSig e n) where + + t2f : {X : Type n} -> Tree σ X -> Free σ τ X + t2f = Q.[_] + + f2t : {X : Type n} -> Free σ τ X -> ∥ Tree σ X ∥₁ + f2t F = let p = ([]surjective F) in P.map fst p + + FreeStr : (X : Type n) -> Str (ℓ-max (ℓ-max (ℓ-max f a) e) n) σ + car (FreeStr X) = Free σ τ X + ops (FreeStr X) f o = {!!} + isSetStr (FreeStr X) = {!!} + +ℕ-MonStrEq : Eqs ℓ-zero MonSig ℕ-MonStr +Eqs.name ℕ-MonStrEq = MonEq +Eqs.free ℕ-MonStrEq = MonEqFree +Eqs.lhs ℕ-MonStrEq = MonEqLhs +Eqs.rhs ℕ-MonStrEq = MonEqRhs +Eqs.equ ℕ-MonStrEq unitl = refl +Eqs.equ ℕ-MonStrEq unitr = funExt \v -> +-zero (v zero) +Eqs.equ ℕ-MonStrEq assocr = funExt \v -> sym (+-assoc (v zero) (v one) (v two)) + +Free-MonStr : (X : Type) -> Str ℓ-zero MonSig +Str.car (Free-MonStr X) = Free MonSig X +Str.ops (Free-MonStr X) e f = op e f +Str.ops (Free-MonStr X) ⊕ f = op ⊕ f +Str.isSetStr (Free-MonStr X) = isSetStr + + +Sig : (f a : Level) -> Type (ℓ-max (ℓ-suc f) (ℓ-suc a)) +Sig f a = Σ[ F ∈ Type f ] (F -> Type a) + +Op : {a x : Level} -> Type a -> Type x -> Type (ℓ-max a x) +Op A X = (A -> X) -> X + +Str : {f a : Level} (x : Level) -> Sig f a -> Type (ℓ-max (ℓ-max f a) (ℓ-suc x)) +Str {f} {a} x (F , ar) = Σ[ X ∈ Type x ] ((f : F) -> Op (ar f) X) + +MonSig : Sig ℓ-zero ℓ-zero +MonSig = MonDesc , MonAr + +MonStr = Str ℓ-zero MonSig + +ℕ-MonStr : MonStr +ℕ-MonStr = ℕ , \{ e f -> 0 + ; ⊕ f -> f zero + f (suc zero) } +-- + +module FreeStr {f a : Level} (σ @ (F , ar) : Sig f a) where + + data Free {x} (X : Type x) : Type {!!} where + η : X -> Free X + op : (f : F) -> Op (ar f) (Free X) -> Free X + isSetStr : isSet (Free X) + +Op : (a : Level) -> ℕ -> Type a -> Type a +Op a n A = (Fin n -> A) -> A + +Str : {f : Level} (a : Level) -> Sig f -> Type {!!} +Str {f} a (F , ar) = Σ[ A ∈ Type a ] ((f : F) -> Op a (ar f) A) + +MonSig : Sig ℓ-zero +MonSig = Fin 2 , lookup (0 ∷ 2 ∷ []) + +ℕ-MonStr : Str ℓ-zero MonSig +ℕ-MonStr = ℕ , F.elim {!!} (\_ -> 0) {!!} + +F.elim (\{k} f -> Op ℓ-zero {!!} ℕ) {!!} {!!} +(f : Fin 2) → (Fin (lookup (0 ∷ 2 ∷ []) f) → ℕ) → ℕ +(f : fst MonSig) → Op ℓ-zero (snd MonSig f) ℕ + +F.elim (\f -> Op ℓ-zero (rec 0 2 f) ℕ) (\_ -> 0) (\_ -> {!!}) + +record FreeS (ℓ : Level) (S : Type ℓ → Type ℓ') : Type (ℓ-max (ℓ-suc ℓ) ℓ') where + field + F : Type ℓ -> Type ℓ + F-S : (A : Type ℓ) -> S (F A) + η : (A : Type ℓ) -> A -> F A + +freeS : (S : Type ℓ -> Type ℓ') -> FreeS ℓ S +monad T, T-alg + +monads on hSet + +record Monad {ℓ : Level} : Type (ℓ-suc ℓ) where + field + T : Type ℓ -> Type ℓ + map : ∀ {A B} -> (A -> B) -> T A -> T B + η : ∀ {A} -> A -> T A + μ : ∀ {A} -> T (T A) -> T A + +record T-Alg {ℓ} (T : Monad {ℓ}) : Type (ℓ-suc ℓ) where + module T = Monad T + field + el : Type ℓ + alg : T.T el -> el + unit : alg ∘ T.η ≡ idfun _ + action : alg ∘ T.map alg ≡ alg ∘ T.μ + +record T-AlgHom {ℓ} {T : Monad {ℓ}} (α β : T-Alg T) : Type (ℓ-suc ℓ) where + module T = Monad T + module α = T-Alg α + module β = T-Alg β + field + f : α.el -> β.el + f-coh : β.alg ∘ T.map f ≡ f ∘ α.alg diff --git a/Experiments/FreeStrFin.agda b/Experiments/FreeStrFin.agda new file mode 100644 index 0000000..19e3f49 --- /dev/null +++ b/Experiments/FreeStrFin.agda @@ -0,0 +1,113 @@ +{-# OPTIONS --cubical --type-in-type #-} + +module _ where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Image +open import Cubical.HITs.PropositionalTruncation as P +open import Cubical.Data.Nat +open import Cubical.Data.FinData as F +open import Cubical.Data.List as L +open import Cubical.Data.List.FinData as F +open import Cubical.Data.Sigma +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.SetQuotients as Q +open import Agda.Primitive + +record Sig : Type₁ where + field + symbol : Type + arity : symbol -> ℕ +open Sig public + +module _ (σ : Sig) where + sig : Type -> Type + sig X = Σ (σ .symbol) \f -> Fin (σ .arity f) -> X + + sigF : {X Y : Type} -> (X -> Y) -> sig X -> sig Y + sigF h (f , i) = f , h ∘ i + +module _ (σ : Sig) where + struct : Type₁ + struct = Σ Type (\X -> sig σ X -> X) + + structIsHom : (𝔛 𝔜 : struct) (h : 𝔛 .fst -> 𝔜 .fst) -> Type + structIsHom (X , α) (Y , β) h = + ((f , i) : sig σ X) -> β (f , h ∘ i) ≡ h (α (f , i)) + + structHom : struct -> struct -> Type + structHom 𝔛 𝔜 = Σ[ h ∈ (𝔛 .fst -> 𝔜 .fst) ] structIsHom 𝔛 𝔜 h + +record EqSig : Type₁ where + field + name : Type + free : name -> Type +open EqSig public + +data Tree (σ : Sig) (V : Type) : Type where + leaf : V -> Tree σ V + node : sig σ (Tree σ V) -> Tree σ V + +module _ {σ : Sig} (𝔛 : struct σ) where + ext : {V : Type} -> (V -> 𝔛 .fst) -> Tree σ V -> 𝔛 .fst + ext h (leaf v) = h v + ext h (node (f , i)) = 𝔛 .snd (f , (ext h ∘ i)) + + extHom : {V : Type} -> (V -> 𝔛 .fst) -> structHom σ (Tree σ V , node) 𝔛 + extHom h = ext h , \_ → refl + +module _ (σ : Sig) (τ : EqSig) where + eqs : Type + eqs = (e : τ .name) -> Tree σ (τ .free e) × Tree σ (τ .free e) + +module _ {σ : Sig} {τ : EqSig} where + infix 30 _⊨_ + _⊨_ : struct σ -> eqs σ τ -> Type + 𝔛 ⊨ ε = (e : τ .name) (ρ : τ .free e -> 𝔛 .fst) + -> ext 𝔛 ρ (ε e .fst) ≡ ext 𝔛 ρ (ε e .snd) + +module Free1 (σ : Sig) (τ : EqSig) (ε : eqs σ τ) where + + congruence relation generated by equations + data _≈_ {X : Type} : Tree σ X -> Tree σ X -> Type₁ where + ≈-refl : ∀ t -> t ≈ t + ≈-sym : ∀ t s -> t ≈ s -> s ≈ t + ≈-trans : ∀ t s r -> t ≈ s -> s ≈ r -> t ≈ r + ≈-cong : (f : σ .symbol) -> {t s : Fin (σ .arity f) -> Tree σ X} + -> ((a : Fin (σ .arity f)) -> t a ≈ s a) + -> node (f , t) ≈ node (f , s) + ≈-eqs : (𝔜 : struct σ) (ϕ : 𝔜 ⊨ ε) + -> (e : τ .name) (ρ : X -> 𝔜 .fst) + -> ∀ t s -> ext 𝔜 ρ t ≡ ext 𝔜 ρ t + -> t ≈ s + + Free : Type -> Type₁ + Free X = Tree σ X / _≈_ + + test1 : {X : Type} (n : ℕ) -> (Fin n -> Free X) -> (Fin n -> Tree σ X) + test1 zero i () + test1 (suc n) i zero = {!!} ?? + test1 (suc n) i (suc k) = test1 n (i ∘ weakenFin) k + + freeAlg : (X : Type) -> sig σ (Free X) -> Free X + freeAlg X (f , i) = Q.[ node (f , {!!}) ] needs choice? + + freeStruct : (X : Type) -> struct σ + freeStruct X = Free X , freeAlg X + +module Free2 (σ : Sig) (τ : EqSig) (ε : eqs σ τ) where + + mutual + data Free (X : Type) : Type where + η : X -> Free X + α : sig σ (Free X) -> Free X + sat : (Free X , α) ⊨ ε + + ext 𝔜 (sharp ϕ h ∘ ρ) (ε₁ e .snd) != + sharp ϕ h (ext (Free X , α) ρ (ε₁ e .snd)) + + sharp : {X : Type} {𝔜 : struct σ} (ϕ : 𝔜 ⊨ ε) (h : X -> 𝔜 .fst) -> Free X -> 𝔜 .fst + sharp ϕ h (η x) = h x + sharp {𝔜 = 𝔜} ϕ h (α (f , o)) = 𝔜 .snd (f , (sharp ϕ h ∘ o)) + sharp ϕ h (sat e ρ i) = ϕ e (sharp ϕ h ∘ ρ) {!i!} diff --git a/Experiments/ListArray.agda b/Experiments/ListArray.agda new file mode 100644 index 0000000..6133367 --- /dev/null +++ b/Experiments/ListArray.agda @@ -0,0 +1,63 @@ +{-# OPTIONS --cubical #-} + +module Experiments.ListArray where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.List renaming (_∷_ to _∷ₗ_) +open import Cubical.Data.Fin +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Sum as ⊎ +open import Cubical.Induction.WellFounded +import Cubical.Data.Empty as ⊥ + +import Cubical.Structures.Set.Mon.Desc as M +import Cubical.Structures.Free as F +open import Cubical.Structures.Sig +open import Cubical.Structures.Str public +open import Cubical.Structures.Tree +open import Cubical.Structures.Eq +open import Cubical.Structures.Arity + +open import Cubical.Structures.Set.Mon.Array +open import Cubical.Structures.Set.Mon.List + +open Iso + +private + variable + ℓ : Level + A : Type ℓ + +an-array : Array ℕ +an-array = 3 , lemma + where + lemma : Fin 3 -> ℕ + lemma (0 , p) = 5 + lemma (1 , p) = 9 + lemma (2 , p) = 2 + lemma (suc (suc (suc n)) , p) = ⊥.rec (¬-<-zero (<-k+-cancel p)) + +module MonDef = F.Definition M.MonSig M.MonEqSig M.MonSEq + +want : List ℕ +want = 5 ∷ₗ 9 ∷ₗ 2 ∷ₗ [] + +to-list : structHom < Array ℕ , array-α > < List ℕ , list-α > +to-list = MonDef.Free.ext arrayDef (isOfHLevelList 0 isSetℕ) list-sat [_] + +_ : to-list .fst an-array ≡ want +_ = refl + +to-list' : structHom < Array ℕ , MonDef.Free.α (arrayDef' {ℓ' = ℓ-zero}) > < List ℕ , list-α > +to-list' = MonDef.Free.ext arrayDef' (isOfHLevelList 0 isSetℕ) list-sat [_] + +_ : to-list' .fst an-array ≡ want +_ = refl + +to-list'' : Array ℕ -> List ℕ +to-list'' = MonDef.freeIso arrayDef listDef (isSetArray isSetℕ) (isOfHLevelList 0 isSetℕ) .fun + +_ : to-list'' an-array ≡ want +_ = refl diff --git a/Experiments/ListMon.agda b/Experiments/ListMon.agda new file mode 100644 index 0000000..1711710 --- /dev/null +++ b/Experiments/ListMon.agda @@ -0,0 +1,112 @@ +{-# OPTIONS --cubical #-} + +module Experiments.ListMon where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +import Mon as M + +data List (A : Type) : Type where + [] : List A + _∷_ : (a : A) -> List A -> List A + trunc : isSet (List A) + +module elimListSet {p : Level} {A : Type} (P : List A -> Type p) + ([]* : P []) + (_∷*_ : (x : A) {xs : List A} -> P xs -> P (x ∷ xs)) + (trunc* : {xs : List A} -> isSet (P xs)) + where + f : (xs : List A) -> P xs + f [] = []* + f (x ∷ xs) = x ∷* f xs + f (trunc xs ys p q i j) = + isOfHLevel→isOfHLevelDep 2 (\xs -> trunc* {xs = xs}) (f xs) (f ys) (cong f p) (cong f q) (trunc xs ys p q) i j + +module elimListProp {p : Level} {A : Type} (P : List A -> Type p) + ([]* : P []) + (_∷*_ : (x : A) {xs : List A} -> P xs -> P (x ∷ xs)) + (trunc* : {xs : List A} -> isProp (P xs)) + where + f : (xs : List A) -> P xs + f = elimListSet.f P []* _∷*_ (isProp→isSet trunc*) + +[_] : {A : Type} -> A -> List A +[ a ] = a ∷ [] + +_⊕_ : {A : Type} -> List A -> List A -> List A +[] ⊕ ys = ys +(x ∷ xs) ⊕ ys = x ∷ (xs ⊕ ys) +trunc xs ys p q i j ⊕ zs = trunc (xs ⊕ zs) (ys ⊕ zs) (cong (_⊕ zs) p) (cong (_⊕ zs) q) i j + +unitl : {A : Type} -> ∀ (m : List A) -> [] ⊕ m ≡ m +unitl _ = refl + +unitr : {A : Type} -> ∀ (m : List A) -> m ⊕ [] ≡ m +unitr = + elimListProp.f _ + refl + (\x p -> cong (x ∷_) p) + (trunc _ _) + +assocr : {A : Type} -> ∀ (m n o : List A) -> (m ⊕ n) ⊕ o ≡ m ⊕ (n ⊕ o) +assocr = + elimListProp.f _ + (λ _ _ -> refl) + (λ x p n o -> cong (x ∷_) (p n o)) + (isPropΠ λ _ -> isPropΠ λ _ -> trunc _ _) + +listMon : (A : Type) -> M.MonStruct (List A) +M.MonStruct.e (listMon _) = [] +M.MonStruct._⊕_ (listMon _) = _⊕_ +M.MonStruct.unitl (listMon _) = unitl +M.MonStruct.unitr (listMon _) = unitr +M.MonStruct.assocr (listMon _) = assocr +M.MonStruct.trunc (listMon _) = trunc + +module _ {A B : Type} (M : M.MonStruct B) where + module B = M.MonStruct M + module _ (f : A -> B) where + + _♯ : List A -> B + (_♯) [] = B.e + (_♯) (x ∷ xs) = f x B.⊕ (_♯) xs + (_♯) (trunc m n p q i j) = B.trunc ((_♯) m) ((_♯) n) (cong (_♯) p) (cong (_♯) q) i j + + _♯-isMonHom : M.isMonHom (listMon A) M _♯ + M.isMonHom.f-e _♯-isMonHom = refl + M.isMonHom.f-⊕ _♯-isMonHom = + elimListProp.f _ + (λ b -> sym (B.unitl (b ♯))) + (λ x {xs} p b -> + f x B.⊕ ((xs ⊕ b) ♯) ≡⟨ cong (f x B.⊕_) (p b) ⟩ + f x B.⊕ ((xs ♯) B.⊕ (b ♯)) ≡⟨ sym (B.assocr (f x) _ _) ⟩ + (f x B.⊕ (xs ♯)) B.⊕ (b ♯) + ∎ + ) + (isPropΠ λ _ -> B.trunc _ _) + + private + listMonEquivLemma : (f : List A -> B) -> M.isMonHom (listMon A) M f -> (x : List A) -> f x ≡ ((f ∘ [_]) ♯) x + listMonEquivLemma f homMonWit = elimListProp.f _ + (M.isMonHom.f-e homMonWit) + (λ x {xs} p -> + f ([ x ] ⊕ xs) ≡⟨ M.isMonHom.f-⊕ homMonWit [ x ] xs ⟩ + f [ x ] B.⊕ f xs ≡⟨ cong (f [ x ] B.⊕_) p ⟩ + (f ∘ [_]) x B.⊕ ((f ∘ [_]) ♯) xs + ∎) + (B.trunc _ _) + + listMonEquivLemma-β : (f : List A -> B) -> M.isMonHom (listMon A) M f -> ((f ∘ [_]) ♯) ≡ f + listMonEquivLemma-β f homMonWit i x = listMonEquivLemma f homMonWit x (~ i) + + listMonEquiv : M.MonHom (listMon A) M ≃ (A -> B) + listMonEquiv = isoToEquiv + ( iso + (λ (f , ϕ) -> f ∘ [_]) + (λ f -> (f ♯) , (f ♯-isMonHom)) + (λ f i x -> B.unitr (f x) i) + (λ (f , homMonWit) -> Σ≡Prop M.isMonHom-isProp (listMonEquivLemma-β f homMonWit)) + ) + + listMonIsEquiv : isEquiv {A = M.MonHom (listMon A) M} (\(f , ϕ) -> f ∘ [_]) + listMonIsEquiv = listMonEquiv .snd diff --git a/Experiments/Mon.agda b/Experiments/Mon.agda new file mode 100644 index 0000000..295e512 --- /dev/null +++ b/Experiments/Mon.agda @@ -0,0 +1,44 @@ +{-# OPTIONS --cubical #-} + +module Experiments.Mon where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Sigma +open import Cubical.Functions.FunExtEquiv +open import Agda.Primitive + +record MonStruct (A : Type) : Type where + field + e : A + _⊕_ : A -> A -> A + unitl : ∀ m -> e ⊕ m ≡ m + unitr : ∀ m -> m ⊕ e ≡ m + assocr : ∀ m n o -> (m ⊕ n) ⊕ o ≡ m ⊕ (n ⊕ o) + trunc : isSet A + +record isMonHom {A B : Type} (M : MonStruct A) (N : MonStruct B) (f : A -> B) : Type where + pattern + inductive + constructor monHom + + module M = MonStruct M + module N = MonStruct N + + field + f-e : f (M.`e) ≡ N.e + f-⊕ : ∀ a b -> f (a M.`⊕ b) ≡ f a N.⊕ f b + + f-unitl : ∀ a -> cong f (M.unitl a) ≡ f-⊕ M.`e a ∙ cong (N._⊕ f a) f-e ∙ N.unitl (f a) + f-unitl a = N.trunc _ _ _ _ + +module _ {A B : Type} {M : MonStruct A} {N : MonStruct B} (f : A -> B) where + module M = MonStruct M + module N = MonStruct N + + isMonHom-isProp : isProp (isMonHom M N f) + isMonHom-isProp (monHom x-e x-⊕) (monHom y-e y-⊕) = + cong₂ monHom (N.trunc _ _ x-e y-e) ((isPropΠ λ _ → isPropΠ (λ _ → N.trunc _ _)) x-⊕ y-⊕) + +MonHom : {A B : Type} (M : MonStruct A) (N : MonStruct B) -> Type +MonHom {A} {B} M N = Σ[ f ∈ (A -> B) ] isMonHom M N f + \ No newline at end of file diff --git a/Experiments/Norm.agda b/Experiments/Norm.agda new file mode 100644 index 0000000..a8e50d3 --- /dev/null +++ b/Experiments/Norm.agda @@ -0,0 +1,67 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +open import Agda.Primitive + +TODO: Fix levels in Free so this isn't necessary +module Experiments.Norm {ℓ : Level} {A : Set ℓ} where + +open import Cubical.Foundations.Everything +open import Cubical.Functions.Embedding +open import Cubical.Data.Sigma +open import Cubical.Data.Fin +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Sum as ⊎ +open import Cubical.Induction.WellFounded +import Cubical.Data.Empty as ⊥ +open import Cubical.HITs.SetQuotients as Q + +import Cubical.Structures.Set.CMon.Desc as M +import Cubical.Structures.Free as F + +open import Cubical.Structures.Set.CMon.Free as F +open import Cubical.Structures.Set.CMon.QFreeMon +open import Cubical.Structures.Set.CMon.Bag +open import Cubical.Structures.Set.CMon.SList as S + +module CMonDef = F.Definition M.CMonSig M.CMonEqSig M.CMonSEq + +module FCM = CMonDef.Free + +ηᵇ : A -> Bag A +ηᵇ = FCM.η bagFreeDef + +ηˢ : A -> SList A +ηˢ = FCM.η {ℓ' = ℓ} slistDef + +ηˢ♯ : structHom < Bag A , FCM.α bagFreeDef > < SList A , FCM.α {ℓ' = ℓ} slistDef > +ηˢ♯ = FCM.ext bagFreeDef S.trunc (FCM.sat {ℓ' = ℓ} S.slistDef) ηˢ + +ηᵇ♯ : structHom < SList A , FCM.α {ℓ' = ℓ} slistDef > < Bag A , FCM.α bagFreeDef > +ηᵇ♯ = FCM.ext slistDef Q.squash/ (FCM.sat bagFreeDef) (FCM.η bagFreeDef) + +ηᵇ♯∘ηˢ♯ : structHom < Bag A , FCM.α bagFreeDef > < Bag A , FCM.α bagFreeDef > +ηᵇ♯∘ηˢ♯ = structHom∘ < Bag _ , FCM.α bagFreeDef > < SList _ , FCM.α {ℓ' = ℓ} slistDef > < Bag _ , FCM.α bagFreeDef > ηᵇ♯ ηˢ♯ + +ηᵇ♯∘ηˢ♯-β : ηᵇ♯∘ηˢ♯ .fst ∘ ηᵇ ≡ ηᵇ +ηᵇ♯∘ηˢ♯-β = + ηᵇ♯ .fst ∘ ηˢ♯ .fst ∘ ηᵇ + ≡⟨ cong (ηᵇ♯ .fst ∘_) (FCM.ext-η bagFreeDef S.trunc (FCM.sat {ℓ' = ℓ} slistDef) ηˢ) ⟩ + ηᵇ♯ .fst ∘ ηˢ + ≡⟨ FCM.ext-η slistDef Q.squash/ (FCM.sat bagFreeDef) ηᵇ ⟩ + ηᵇ + ∎ + +ηᵇ♯∘ηˢ♯-η : ηᵇ♯∘ηˢ♯ ≡ idHom < Bag A , FCM.α bagFreeDef > +ηᵇ♯∘ηˢ♯-η = FCM.hom≡ bagFreeDef Q.squash/ (FCM.sat bagFreeDef) ηᵇ♯∘ηˢ♯ (idHom _) ηᵇ♯∘ηˢ♯-β + +𝔫𝔣 : Bag A -> SList A +𝔫𝔣 = ηˢ♯ .fst + +𝔫𝔣-inj : {as bs : Bag A} -> 𝔫𝔣 as ≡ 𝔫𝔣 bs -> as ≡ bs +𝔫𝔣-inj {as} {bs} p = sym (funExt⁻ (cong fst ηᵇ♯∘ηˢ♯-η) as) ∙ cong (ηᵇ♯ .fst) p ∙ funExt⁻ (cong fst ηᵇ♯∘ηˢ♯-η) bs + +norm : isEmbedding 𝔫𝔣 +norm = injEmbedding trunc 𝔫𝔣-inj + +also equivalence diff --git a/Experiments/TList.agda b/Experiments/TList.agda new file mode 100644 index 0000000..ada4120 --- /dev/null +++ b/Experiments/TList.agda @@ -0,0 +1,58 @@ +{-# OPTIONS --cubical #-} + +module Experiments.TList where + +open import Cubical.Foundations.Everything +open import Cubical.Data.List +open import Cubical.HITs.SetTruncation as S + +MList : ∀ {ℓ} -> Type ℓ -> Type ℓ +MList A = ∥ List A ∥₂ + +variable + ℓ : Level + A B C : Type ℓ + +inc : (A -> List B) -> (A -> MList B) +inc = ∣_∣₂ ∘_ + +eta : A -> MList A +eta = inc [_] + +map₂ : (A -> B -> C) -> ∥ A ∥₂ -> ∥ B ∥₂ -> ∥ C ∥₂ +map₂ f ∣ a ∣₂ = S.map (f a) +map₂ f (squash₂ a a' p q i j) = {!!} + +fmap : (A -> B) -> ∥ A ∥₂ -> ∥ B ∥₂ +fmap f ∣ a ∣₂ = ∣ f a ∣₂ +fmap f (squash₂ a b p q i j) = + squash₂ (fmap f a) (fmap f b) (cong (fmap f) p) (cong (fmap f) q) i j + +module _ (P : List A -> Type ℓ) + (h : (xs : List A) -> P xs) + (trunc : {xs : List A} -> isSet (P xs)) + where + f : (xs : MList A) -> P {!!} + f = {!!} + +e : MList A +e = ∣ [] ∣₂ + +_⊕_ : MList A -> MList A -> MList A +_⊕_ = map₂ _++_ + +map-inc : ∀ {f : A -> B} {a} -> S.map f (∣ a ∣₂) ≡ ∣ f a ∣₂ +map-inc = refl + +map₂-inc : ∀ {f : A -> B -> C} {a b} -> map₂ f a ∣ b ∣₂ ≡ S.map (\a -> f a b) a +map₂-inc {a = a} = {!!} + +congmap : ((a b : A) -> f a ≡ g b) + -> (a' b' : ∥ A ∥₂) -> map f a' ≡ map g b' + +unitr : (xs : MList A) -> xs ⊕ e ≡ xs +unitr xs = + map₂ _++_ xs e ≡⟨ {!!} ⟩ + S.map (\xs -> xs ⊕ e) e ≡⟨ {!!} ⟩ + idfun _ xs + ∎ diff --git a/index.agda b/index.agda new file mode 100644 index 0000000..8fc3dca --- /dev/null +++ b/index.agda @@ -0,0 +1,49 @@ +{-# OPTIONS --cubical --safe --exact-split #-} + +module index where + +an organised list of modules: +-- +universal algebra +algebraic theories and 2-theories +import Cubical.Structures.Sig +import Cubical.Structures.Str +import Cubical.Structures.Eq +import Cubical.Structures.Coh + +free algebras +import Cubical.Structures.Free + +free monoids on sets +import Cubical.Structures.Set.Mon.Free +import Cubical.Structures.Set.Mon.List +import Cubical.Structures.Set.Mon.Array + +free commutative monoids on sets +import Cubical.Structures.Set.CMon.Free +import Cubical.Structures.Set.CMon.SList +import Cubical.Structures.Set.CMon.CList +import Cubical.Structures.Set.CMon.PList +import Cubical.Structures.Set.CMon.QFreeMon +import Cubical.Structures.Set.CMon.Bag + +free monoidal groupoids on groupoids +import Cubical.Structures.Gpd.Mon.Free +import Cubical.Structures.Gpd.Mon.List + +free symmetric monoidal groupoids on groupoids +import Cubical.Structures.Gpd.SMon.Free +import Cubical.Structures.Gpd.SMon.SList + +sorting and order equivalence +import Cubical.Structures.Set.CMon.SList.Sort + +combinatorics properties +import Cubical.Structures.Set.CMon.SList.Seely + +useful experiments +import Experiments.Norm + +an exhaustive list of all modules: +-- +import Everything